Theory SEC1v2_0_Test_Vectors

theory SEC1v2_0_Test_Vectors
  imports "FIPS186_4_Curves"
          "FIPS180_4"
begin

text ‹In this theory, we show that test vectors provided by NIST are validated against the
translation of SEC1 v2.0 to HOL/Isabelle, SEC1v2_0.thy.  These test vectors cover key pair 
generation, key validation, the Diffie-Hellman primitive, the MQV primitive, the digital 
signature algorithm (both signing and verifying operations).  The test vectors cover the NIST
curves over prime fields, P-224, P-256, P-382, and P-521.  For certain operations, NIST also
provided test vectors for P-192.  For the digital signature algorithm, a hash is required.  The
test vectors cover SHA-224, -256, -384, and -512 for each of the prime field curves.›

section‹NIST Test Vectors for ECDSA: Curve P-224, Hash SHA-224›
  
subsection‹Test Vector: P224_SHA224_TV01›

subsubsection‹Given Test Vector Values›

definition P224_SHA224_TV01_Msg :: octets where
  "P224_SHA224_TV01_Msg = nat_to_octets_len 0x699325d6fc8fbbb4981a6ded3c3a54ad2e4e3db8a5669201912064c64e700c139248cdc19495df081c3fc60245b9f25fc9e301b845b3d703a694986e4641ae3c7e5a19e6d6edbf1d61e535f49a8fad5f4ac26397cfec682f161a5fcd32c5e780668b0181a91955157635536a22367308036e2070f544ad4fff3d5122c76fad5d 128"

definition P224_SHA224_TV01_d :: nat where
  "P224_SHA224_TV01_d = 0x16797b5c0c7ed5461e2ff1b88e6eafa03c0f46bf072000dfc830d615"

definition P224_SHA224_TV01_Qx :: nat where
  "P224_SHA224_TV01_Qx = 0x605495756e6e88f1d07ae5f98787af9b4da8a641d1a9492a12174eab"

definition P224_SHA224_TV01_Qy :: nat where
  "P224_SHA224_TV01_Qy = 0xf5cc733b17decc806ef1df861a42505d0af9ef7c3df3959b8dfc6669"

definition P224_SHA224_TV01_Q :: "int point" where
  "P224_SHA224_TV01_Q = Point (int P224_SHA224_TV01_Qx) (int P224_SHA224_TV01_Qy)"

definition P224_SHA224_TV01_k :: nat where
  "P224_SHA224_TV01_k = 0xd9a5a7328117f48b4b8dd8c17dae722e756b3ff64bd29a527137eec0"

definition P224_SHA224_TV01_R :: nat where
  "P224_SHA224_TV01_R = 0x2fc2cff8cdd4866b1d74e45b07d333af46b7af0888049d0fdbc7b0d6"

definition P224_SHA224_TV01_S :: nat where
  "P224_SHA224_TV01_S = 0x8d9cc4c8ea93e0fd9d6431b9a1fd99b88f281793396321b11dac41eb"

definition P224_SHA224_TV01_Sig :: "nat × nat" where
  "P224_SHA224_TV01_Sig = (P224_SHA224_TV01_R, P224_SHA224_TV01_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA224_TV01_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV01_d" 
  by eval

lemma P224_SHA224_TV01_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA224_TV01_Q" 
  by eval

lemma P224_SHA224_TV01_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA224_TV01_d = P224_SHA224_TV01_Q" 
  by eval

lemma P224_SHA224_TV01_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA224_TV01_d P224_SHA224_TV01_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA224_TV01_dQvalidPair' P224_SHA224_TV01_d_valid) 

lemma P224_SHA224_TV01_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV01_k" 
  by eval

lemma P224_SHA224_TV01_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA224octets P224_SHA224_TV01_d P224_SHA224_TV01_Msg P224_SHA224_TV01_k = Some P224_SHA224_TV01_Sig" 
  by eval

lemma P224_SHA224_TV01_Verify: 
  "SEC1_P224_ECDSA_Verify SHA224octets P224_SHA224_TV01_Q P224_SHA224_TV01_Msg P224_SHA224_TV01_Sig" 
  by eval

subsection‹Test Vector: P224_SHA224_TV02›

subsubsection‹Given Test Vector Values›

definition P224_SHA224_TV02_Msg :: octets where
  "P224_SHA224_TV02_Msg = nat_to_octets_len 0x7de42b44db0aa8bfdcdac9add227e8f0cc7ad1d94693beb5e1d325e5f3f85b3bd033fc25e9469a89733a65d1fa641f7e67d668e7c71d736233c4cba20eb83c368c506affe77946b5e2ec693798aecd7ff943cd8fab90affddf5ad5b8d1af332e6c5fe4a2df16837700b2781e08821d4fbdd8373517f5b19f9e63b89cfeeeef6f 128"

definition P224_SHA224_TV02_d :: nat where
  "P224_SHA224_TV02_d = 0xcf020a1ff36c28511191482ed1e5259c60d383606c581948c3fbe2c5"

definition P224_SHA224_TV02_Qx :: nat where
  "P224_SHA224_TV02_Qx = 0xfa21f85b99d3dc18c6d53351fbcb1e2d029c00fa7d1663a3dd94695e"

definition P224_SHA224_TV02_Qy :: nat where
  "P224_SHA224_TV02_Qy = 0xe9e79578f8988b168edff1a8b34a5ed9598cc20acd1f0aed36715d88"

definition P224_SHA224_TV02_Q :: "int point" where
  "P224_SHA224_TV02_Q = Point (int P224_SHA224_TV02_Qx) (int P224_SHA224_TV02_Qy)"

definition P224_SHA224_TV02_k :: nat where
  "P224_SHA224_TV02_k = 0xc780d047454824af98677cf310117e5f9e99627d02414f136aed8e83"

definition P224_SHA224_TV02_R :: nat where
  "P224_SHA224_TV02_R = 0x45145f06b566ec9fd0fee1b6c6551a4535c7a3bbfc0fede45f4f5038"

definition P224_SHA224_TV02_S :: nat where
  "P224_SHA224_TV02_S = 0x7302dff12545b069cf27df49b26e4781270585463656f2834917c3ca"

definition P224_SHA224_TV02_Sig :: "nat × nat" where
  "P224_SHA224_TV02_Sig = (P224_SHA224_TV02_R, P224_SHA224_TV02_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA224_TV02_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV02_d" 
  by eval

lemma P224_SHA224_TV02_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA224_TV02_Q" 
  by eval

lemma P224_SHA224_TV02_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA224_TV02_d = P224_SHA224_TV02_Q" 
  by eval

lemma P224_SHA224_TV02_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA224_TV02_d P224_SHA224_TV02_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA224_TV02_dQvalidPair' P224_SHA224_TV02_d_valid) 

lemma P224_SHA224_TV02_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV02_k" 
  by eval

lemma P224_SHA224_TV02_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA224octets P224_SHA224_TV02_d P224_SHA224_TV02_Msg P224_SHA224_TV02_k = Some P224_SHA224_TV02_Sig" 
  by eval

lemma P224_SHA224_TV02_Verify: 
  "SEC1_P224_ECDSA_Verify SHA224octets P224_SHA224_TV02_Q P224_SHA224_TV02_Msg P224_SHA224_TV02_Sig" 
  by eval

subsection‹Test Vector: P224_SHA224_TV03›

subsubsection‹Given Test Vector Values›

definition P224_SHA224_TV03_Msg :: octets where
  "P224_SHA224_TV03_Msg = nat_to_octets_len 0xaf0da3adab82784909e2b3dadcecba21eced3c60d7572023dea171044d9a10e8ba67d31b04904541b87fff32a10ccc6580869055fec6216a00320a28899859a6b61faba58a0bc10c2ba07ea16f214c3ddcc9fc5622ad1253b63fe7e95227ae3c9caa9962cffc8b1c4e8260036469d25ab0c8e3643a820b8b3a4d8d43e4b728f9 128"

definition P224_SHA224_TV03_d :: nat where
  "P224_SHA224_TV03_d = 0xdde6f173fa9f307d206ce46b4f02851ebce9638a989330249fd30b73"

definition P224_SHA224_TV03_Qx :: nat where
  "P224_SHA224_TV03_Qx = 0xfc21a99b060afb0d9dbf3250ea3c4da10be94ce627a65874d8e4a630"

definition P224_SHA224_TV03_Qy :: nat where
  "P224_SHA224_TV03_Qy = 0xe8373ab7190890326aac4aacca3eba89e15d1086a05434dd033fd3f3"

definition P224_SHA224_TV03_Q :: "int point" where
  "P224_SHA224_TV03_Q = Point (int P224_SHA224_TV03_Qx) (int P224_SHA224_TV03_Qy)"

definition P224_SHA224_TV03_k :: nat where
  "P224_SHA224_TV03_k = 0x6629366a156840477df4875cfba4f8faa809e394893e1f5525326d07"

definition P224_SHA224_TV03_R :: nat where
  "P224_SHA224_TV03_R = 0x41f8e2b1ae5add7c24da8725a067585a3ad6d5a9ed9580beb226f23a"

definition P224_SHA224_TV03_S :: nat where
  "P224_SHA224_TV03_S = 0xa5d71bff02dce997305dd337128046f36714398f4ef6647599712fae"

definition P224_SHA224_TV03_Sig :: "nat × nat" where
  "P224_SHA224_TV03_Sig = (P224_SHA224_TV03_R, P224_SHA224_TV03_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA224_TV03_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV03_d" 
  by eval

lemma P224_SHA224_TV03_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA224_TV03_Q" 
  by eval

lemma P224_SHA224_TV03_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA224_TV03_d = P224_SHA224_TV03_Q" 
  by eval

lemma P224_SHA224_TV03_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA224_TV03_d P224_SHA224_TV03_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA224_TV03_dQvalidPair' P224_SHA224_TV03_d_valid) 

lemma P224_SHA224_TV03_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV03_k" 
  by eval

lemma P224_SHA224_TV03_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA224octets P224_SHA224_TV03_d P224_SHA224_TV03_Msg P224_SHA224_TV03_k = Some P224_SHA224_TV03_Sig" 
  by eval

lemma P224_SHA224_TV03_Verify: 
  "SEC1_P224_ECDSA_Verify SHA224octets P224_SHA224_TV03_Q P224_SHA224_TV03_Msg P224_SHA224_TV03_Sig" 
  by eval

subsection‹Test Vector: P224_SHA224_TV04›

subsubsection‹Given Test Vector Values›

definition P224_SHA224_TV04_Msg :: octets where
  "P224_SHA224_TV04_Msg = nat_to_octets_len 0xcfa56ae89727df6b7266f69d6636bf738f9e4f15f49c42a0123edac4b3743f32ea52389f919ceb90575c4184897773b2f2fc5b3fcb354880f15c93383215d3c2551fcc1b4180a1ac0f69c969bbc306acd115ce3976eff518540f43ad4076dbb5fbad9ce9b3234f1148b8f5e059192ff480fc4bcbd00d25f4d9f5ed4ba5693b6c 128"

definition P224_SHA224_TV04_d :: nat where
  "P224_SHA224_TV04_d = 0xaeee9071248f077590ac647794b678ad371f8e0f1e14e9fbff49671e"

definition P224_SHA224_TV04_Qx :: nat where
  "P224_SHA224_TV04_Qx = 0xfad0a34991bbf89982ad9cf89337b4bd2565f84d5bdd004289fc1cc3"

definition P224_SHA224_TV04_Qy :: nat where
  "P224_SHA224_TV04_Qy = 0x5d8b6764f28c8163a12855a5c266efeb9388df4994b85a8b4f1bd3bc"

definition P224_SHA224_TV04_Q :: "int point" where
  "P224_SHA224_TV04_Q = Point (int P224_SHA224_TV04_Qx) (int P224_SHA224_TV04_Qy)"

definition P224_SHA224_TV04_k :: nat where
  "P224_SHA224_TV04_k = 0x1d35d027cd5a569e25c5768c48ed0c2b127c0f99cb4e52ea094fe689"

definition P224_SHA224_TV04_R :: nat where
  "P224_SHA224_TV04_R = 0x2258184ef9f0fa698735379972ce9adf034af76017668bfcdab978de"

definition P224_SHA224_TV04_S :: nat where
  "P224_SHA224_TV04_S = 0x866fb8e505dea6c909c2c9143ec869d1bac2282cf12366130ff2146c"

definition P224_SHA224_TV04_Sig :: "nat × nat" where
  "P224_SHA224_TV04_Sig = (P224_SHA224_TV04_R, P224_SHA224_TV04_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA224_TV04_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV04_d" 
  by eval

lemma P224_SHA224_TV04_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA224_TV04_Q" 
  by eval

lemma P224_SHA224_TV04_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA224_TV04_d = P224_SHA224_TV04_Q" 
  by eval

lemma P224_SHA224_TV04_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA224_TV04_d P224_SHA224_TV04_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA224_TV04_dQvalidPair' P224_SHA224_TV04_d_valid) 

lemma P224_SHA224_TV04_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV04_k" 
  by eval

lemma P224_SHA224_TV04_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA224octets P224_SHA224_TV04_d P224_SHA224_TV04_Msg P224_SHA224_TV04_k = Some P224_SHA224_TV04_Sig" 
  by eval

lemma P224_SHA224_TV04_Verify: 
  "SEC1_P224_ECDSA_Verify SHA224octets P224_SHA224_TV04_Q P224_SHA224_TV04_Msg P224_SHA224_TV04_Sig" 
  by eval

subsection‹Test Vector: P224_SHA224_TV05›

subsubsection‹Given Test Vector Values›

definition P224_SHA224_TV05_Msg :: octets where
  "P224_SHA224_TV05_Msg = nat_to_octets_len 0xc223c8009018321b987a615c3414d2bb15954933569ca989de32d6bf11107bc47a330ab6d88d9b50d106cf5777d1b736b14bc48deda1bc573a9a7dd42cd061860645306dce7a5ba8c60f135a6a21999421ce8c4670fe7287a7e9ea3aa1e0fa82721f33e6e823957fe86e2283c89ef92b13cd0333c4bb70865ae1919bf538ea34 128"

definition P224_SHA224_TV05_d :: nat where
  "P224_SHA224_TV05_d = 0x29c204b2954e1406a015020f9d6b3d7c00658298feb2d17440b2c1a4"

definition P224_SHA224_TV05_Qx :: nat where
  "P224_SHA224_TV05_Qx = 0x0e0fc15e775a75d45f872e5021b554cc0579da19125e1a49299c7630"

definition P224_SHA224_TV05_Qy :: nat where
  "P224_SHA224_TV05_Qy = 0xcb64fe462d025ae2a1394746bdbf8251f7ca5a1d6bb13e0edf6b7b09"

definition P224_SHA224_TV05_Q :: "int point" where
  "P224_SHA224_TV05_Q = Point (int P224_SHA224_TV05_Qx) (int P224_SHA224_TV05_Qy)"

definition P224_SHA224_TV05_k :: nat where
  "P224_SHA224_TV05_k = 0x39547c10bb947d69f6c3af701f2528e011a1e80a6d04cc5a37466c02"

definition P224_SHA224_TV05_R :: nat where
  "P224_SHA224_TV05_R = 0x86622c376d326cdf679bcabf8eb034bf49f0c188f3fc3afd0006325d"

definition P224_SHA224_TV05_S :: nat where
  "P224_SHA224_TV05_S = 0x26613d3b33c70e635d7a998f254a5b15d2a3642bf321e8cff08f1e84"

definition P224_SHA224_TV05_Sig :: "nat × nat" where
  "P224_SHA224_TV05_Sig = (P224_SHA224_TV05_R, P224_SHA224_TV05_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA224_TV05_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV05_d" 
  by eval

lemma P224_SHA224_TV05_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA224_TV05_Q" 
  by eval

lemma P224_SHA224_TV05_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA224_TV05_d = P224_SHA224_TV05_Q" 
  by eval

lemma P224_SHA224_TV05_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA224_TV05_d P224_SHA224_TV05_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA224_TV05_dQvalidPair' P224_SHA224_TV05_d_valid) 

lemma P224_SHA224_TV05_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV05_k" 
  by eval

lemma P224_SHA224_TV05_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA224octets P224_SHA224_TV05_d P224_SHA224_TV05_Msg P224_SHA224_TV05_k = Some P224_SHA224_TV05_Sig" 
  by eval

lemma P224_SHA224_TV05_Verify: 
  "SEC1_P224_ECDSA_Verify SHA224octets P224_SHA224_TV05_Q P224_SHA224_TV05_Msg P224_SHA224_TV05_Sig" 
  by eval

subsection‹Test Vector: P224_SHA224_TV06›

subsubsection‹Given Test Vector Values›

definition P224_SHA224_TV06_Msg :: octets where
  "P224_SHA224_TV06_Msg = nat_to_octets_len 0x1c27273d95182c74c100d85b5c08f4b26874c2abc87f127f304aedbf52ef6540eba16dd664ae1e9e30ea1e66ff9cc9ab5a80b5bcbd19dde88a29ff10b50a6abd73388e8071306c68d0c9f6caa26b7e68de29312be959b9f4a5481f5a2ad2070a396ed3de21096541cf58c4a13308e08867565bf2df9d649357a83cdcf18d2cd9 128"

definition P224_SHA224_TV06_d :: nat where
  "P224_SHA224_TV06_d = 0x8986a97b24be042a1547642f19678de4e281a68f1e794e343dabb131"

definition P224_SHA224_TV06_Qx :: nat where
  "P224_SHA224_TV06_Qx = 0x2c070e68e8478341938f3d5026a1fe01e778cdffbebbdd7a4cd29209"

definition P224_SHA224_TV06_Qy :: nat where
  "P224_SHA224_TV06_Qy = 0xcde21c9c7c6590ba300715a7adac278385a5175b6b4ea749c4b6a681"

definition P224_SHA224_TV06_Q :: "int point" where
  "P224_SHA224_TV06_Q = Point (int P224_SHA224_TV06_Qx) (int P224_SHA224_TV06_Qy)"

definition P224_SHA224_TV06_k :: nat where
  "P224_SHA224_TV06_k = 0x509712f9c0f3370f6a09154159975945f0107dd1cee7327c68eaa90b"

definition P224_SHA224_TV06_R :: nat where
  "P224_SHA224_TV06_R = 0x57afda5139b180de96373c3d649700682e37efd56ae182335f081013"

definition P224_SHA224_TV06_S :: nat where
  "P224_SHA224_TV06_S = 0xeb6cd58650cfb26dfdf21de32fa17464a6efc46830eedc16977342e6"

definition P224_SHA224_TV06_Sig :: "nat × nat" where
  "P224_SHA224_TV06_Sig = (P224_SHA224_TV06_R, P224_SHA224_TV06_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA224_TV06_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV06_d" 
  by eval

lemma P224_SHA224_TV06_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA224_TV06_Q" 
  by eval

lemma P224_SHA224_TV06_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA224_TV06_d = P224_SHA224_TV06_Q" 
  by eval

lemma P224_SHA224_TV06_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA224_TV06_d P224_SHA224_TV06_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA224_TV06_dQvalidPair' P224_SHA224_TV06_d_valid) 

lemma P224_SHA224_TV06_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV06_k" 
  by eval

lemma P224_SHA224_TV06_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA224octets P224_SHA224_TV06_d P224_SHA224_TV06_Msg P224_SHA224_TV06_k = Some P224_SHA224_TV06_Sig" 
  by eval

lemma P224_SHA224_TV06_Verify: 
  "SEC1_P224_ECDSA_Verify SHA224octets P224_SHA224_TV06_Q P224_SHA224_TV06_Msg P224_SHA224_TV06_Sig" 
  by eval

subsection‹Test Vector: P224_SHA224_TV07›

subsubsection‹Given Test Vector Values›

definition P224_SHA224_TV07_Msg :: octets where
  "P224_SHA224_TV07_Msg = nat_to_octets_len 0x069ae374971627f6b8503f3aa63ab52bcf4f3fcae65b98cdbbf917a5b08a10dc760056714db279806a8d43485320e6fee0f1e0562e077ee270ace8d3c478d79bcdff9cf8b92fdea68421d4a276f8e62ae379387ae06b60af9eb3c40bd7a768aeffccdc8a08bc78ca2eca18061058043a0e441209c5c594842838a4d9d778a053 128"

definition P224_SHA224_TV07_d :: nat where
  "P224_SHA224_TV07_d = 0xd9aa95e14cb34980cfddadddfa92bde1310acaff249f73ff5b09a974"

definition P224_SHA224_TV07_Qx :: nat where
  "P224_SHA224_TV07_Qx = 0x3a0d4b8e5fad1ea1abb8d3fb742cd45cd0b76d136e5bbb33206ad120"

definition P224_SHA224_TV07_Qy :: nat where
  "P224_SHA224_TV07_Qy = 0xc90ac83276b2fa3757b0f226cd7360a313bc96fd8329c76a7306cc7d"

definition P224_SHA224_TV07_Q :: "int point" where
  "P224_SHA224_TV07_Q = Point (int P224_SHA224_TV07_Qx) (int P224_SHA224_TV07_Qy)"

definition P224_SHA224_TV07_k :: nat where
  "P224_SHA224_TV07_k = 0x1f1739af68a3cee7c5f09e9e09d6485d9cd64cc4085bc2bc89795aaf"

definition P224_SHA224_TV07_R :: nat where
  "P224_SHA224_TV07_R = 0x09bbdd003532d025d7c3204c00747cd52ecdfbc7ce3dde8ffbea23e1"

definition P224_SHA224_TV07_S :: nat where
  "P224_SHA224_TV07_S = 0x1e745e80948779a5cc8dc5cb193beebb550ec9c2647f4948bf58ba7d"

definition P224_SHA224_TV07_Sig :: "nat × nat" where
  "P224_SHA224_TV07_Sig = (P224_SHA224_TV07_R, P224_SHA224_TV07_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA224_TV07_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV07_d" 
  by eval

lemma P224_SHA224_TV07_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA224_TV07_Q" 
  by eval

lemma P224_SHA224_TV07_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA224_TV07_d = P224_SHA224_TV07_Q" 
  by eval

lemma P224_SHA224_TV07_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA224_TV07_d P224_SHA224_TV07_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA224_TV07_dQvalidPair' P224_SHA224_TV07_d_valid) 

lemma P224_SHA224_TV07_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV07_k" 
  by eval

lemma P224_SHA224_TV07_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA224octets P224_SHA224_TV07_d P224_SHA224_TV07_Msg P224_SHA224_TV07_k = Some P224_SHA224_TV07_Sig" 
  by eval

lemma P224_SHA224_TV07_Verify: 
  "SEC1_P224_ECDSA_Verify SHA224octets P224_SHA224_TV07_Q P224_SHA224_TV07_Msg P224_SHA224_TV07_Sig" 
  by eval

subsection‹Test Vector: P224_SHA224_TV08›

subsubsection‹Given Test Vector Values›

definition P224_SHA224_TV08_Msg :: octets where
  "P224_SHA224_TV08_Msg = nat_to_octets_len 0xd0d5ae3e33600aa21c1606caec449eee678c87cb593594be1fbb048cc7cfd076e5cc7132ebe290c4c014e7a517a0d5972759acfa1438d9d2e5d236d19ac92136f6252b7e5bea7588dcba6522b6b18128f003ecab5cb4908832fb5a375cf820f8f0e9ee870653a73dc2282f2d45622a2f0e85cba05c567baf1b9862b79a4b244e 128"

definition P224_SHA224_TV08_d :: nat where
  "P224_SHA224_TV08_d = 0x380fb6154ad3d2e755a17df1f047f84712d4ec9e47d34d4054ea29a8"

definition P224_SHA224_TV08_Qx :: nat where
  "P224_SHA224_TV08_Qx = 0x4772c27cca3348b1801ae87b01cb564c8cf9b81c23cc74468a907927"

definition P224_SHA224_TV08_Qy :: nat where
  "P224_SHA224_TV08_Qy = 0xde9d253935b09617a1655c42d385bf48504e06fa386f5fa533a21dcb"

definition P224_SHA224_TV08_Q :: "int point" where
  "P224_SHA224_TV08_Q = Point (int P224_SHA224_TV08_Qx) (int P224_SHA224_TV08_Qy)"

definition P224_SHA224_TV08_k :: nat where
  "P224_SHA224_TV08_k = 0x14dbdffa326ba2f3d64f79ff966d9ee6c1aba0d51e9a8e59f5686dc1"

definition P224_SHA224_TV08_R :: nat where
  "P224_SHA224_TV08_R = 0xff6d52a09ca4c3b82da0440864d6717e1be0b50b6dcf5e1d74c0ff56"

definition P224_SHA224_TV08_S :: nat where
  "P224_SHA224_TV08_S = 0x09490be77bc834c1efaa23410dcbf800e6fae40d62a737214c5a4418"

definition P224_SHA224_TV08_Sig :: "nat × nat" where
  "P224_SHA224_TV08_Sig = (P224_SHA224_TV08_R, P224_SHA224_TV08_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA224_TV08_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV08_d" 
  by eval

lemma P224_SHA224_TV08_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA224_TV08_Q" 
  by eval

lemma P224_SHA224_TV08_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA224_TV08_d = P224_SHA224_TV08_Q" 
  by eval

lemma P224_SHA224_TV08_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA224_TV08_d P224_SHA224_TV08_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA224_TV08_dQvalidPair' P224_SHA224_TV08_d_valid) 

lemma P224_SHA224_TV08_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV08_k" 
  by eval

lemma P224_SHA224_TV08_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA224octets P224_SHA224_TV08_d P224_SHA224_TV08_Msg P224_SHA224_TV08_k = Some P224_SHA224_TV08_Sig" 
  by eval

lemma P224_SHA224_TV08_Verify: 
  "SEC1_P224_ECDSA_Verify SHA224octets P224_SHA224_TV08_Q P224_SHA224_TV08_Msg P224_SHA224_TV08_Sig" 
  by eval

subsection‹Test Vector: P224_SHA224_TV09›

subsubsection‹Given Test Vector Values›

definition P224_SHA224_TV09_Msg :: octets where
  "P224_SHA224_TV09_Msg = nat_to_octets_len 0x79b7375ae7a4f2e4adad8765d14c1540cd9979db38076c157c1837c760ca6febbb18fd42152335929b735e1a08041bd38d315cd4c6b7dd2729de8752f531f07fe4ddc4f1899debc0311eef0019170b58e08895b439ddf09fbf0aeb1e2fd35c2ef7ae402308c3637733802601dd218fb14c22f57870835b10818369d57d318405 128"

definition P224_SHA224_TV09_d :: nat where
  "P224_SHA224_TV09_d = 0x6b98ec50d6b7f7ebc3a2183ff9388f75e924243827ddded8721186e2"

definition P224_SHA224_TV09_Qx :: nat where
  "P224_SHA224_TV09_Qx = 0x1f249911b125348e6e0a473479105cc4b8cfb4fa32d897810fc69ffe"

definition P224_SHA224_TV09_Qy :: nat where
  "P224_SHA224_TV09_Qy = 0xa17db03b9877d1b6328329061ea67aec5a38a884362e9e5b7d7642dc"

definition P224_SHA224_TV09_Q :: "int point" where
  "P224_SHA224_TV09_Q = Point (int P224_SHA224_TV09_Qx) (int P224_SHA224_TV09_Qy)"

definition P224_SHA224_TV09_k :: nat where
  "P224_SHA224_TV09_k = 0xab3a41fedc77d1f96f3103cc7dce215bf45054a755cf101735fef503"

definition P224_SHA224_TV09_R :: nat where
  "P224_SHA224_TV09_R = 0x70ccc0824542e296d17a79320d422f1edcf9253840dafe4427033f40"

definition P224_SHA224_TV09_S :: nat where
  "P224_SHA224_TV09_S = 0xe3823699c355b61ab1894be3371765fae2b720405a7ce5e790ca8c00"

definition P224_SHA224_TV09_Sig :: "nat × nat" where
  "P224_SHA224_TV09_Sig = (P224_SHA224_TV09_R, P224_SHA224_TV09_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA224_TV09_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV09_d" 
  by eval

lemma P224_SHA224_TV09_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA224_TV09_Q" 
  by eval

lemma P224_SHA224_TV09_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA224_TV09_d = P224_SHA224_TV09_Q" 
  by eval

lemma P224_SHA224_TV09_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA224_TV09_d P224_SHA224_TV09_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA224_TV09_dQvalidPair' P224_SHA224_TV09_d_valid) 

lemma P224_SHA224_TV09_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV09_k" 
  by eval

lemma P224_SHA224_TV09_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA224octets P224_SHA224_TV09_d P224_SHA224_TV09_Msg P224_SHA224_TV09_k = Some P224_SHA224_TV09_Sig" 
  by eval

lemma P224_SHA224_TV09_Verify: 
  "SEC1_P224_ECDSA_Verify SHA224octets P224_SHA224_TV09_Q P224_SHA224_TV09_Msg P224_SHA224_TV09_Sig" 
  by eval

subsection‹Test Vector: P224_SHA224_TV10›

subsubsection‹Given Test Vector Values›

definition P224_SHA224_TV10_Msg :: octets where
  "P224_SHA224_TV10_Msg = nat_to_octets_len 0x8c7de96e6880d5b6efc19646b9d3d56490775cb3faab342e64db2e388c4bd9e94c4e69a63ccdb7e007a19711e69c06f106b71c983a6d97c4589045666c6ab5ea7b5b6d096ddf6fd35b819f1506a3c37ddd40929504f9f079c8d83820fc8493f97b2298aebe48fdb4ff472b29018fc2b1163a22bfbb1de413e8645e871291a9f6 128"

definition P224_SHA224_TV10_d :: nat where
  "P224_SHA224_TV10_d = 0x8dda0ef4170bf73077d685e7709f6f747ced08eb4cde98ef06ab7bd7"

definition P224_SHA224_TV10_Qx :: nat where
  "P224_SHA224_TV10_Qx = 0x7df67b960ee7a2cb62b22932457360ab1e046c1ec84b91ae65642003"

definition P224_SHA224_TV10_Qy :: nat where
  "P224_SHA224_TV10_Qy = 0xc764ca9fc1b0cc2233fa57bdcfedaab0131fb7b5f557d6ca57f4afe0"

definition P224_SHA224_TV10_Q :: "int point" where
  "P224_SHA224_TV10_Q = Point (int P224_SHA224_TV10_Qx) (int P224_SHA224_TV10_Qy)"

definition P224_SHA224_TV10_k :: nat where
  "P224_SHA224_TV10_k = 0x9ef6ebd178a76402968bc8ec8b257174a04fb5e2d65c1ab34ab039b9"

definition P224_SHA224_TV10_R :: nat where
  "P224_SHA224_TV10_R = 0xeef9e8428105704133e0f19636c89e570485e577786df2b09f99602a"

definition P224_SHA224_TV10_S :: nat where
  "P224_SHA224_TV10_S = 0x8c01f0162891e4b9536243cb86a6e5c177323cca09777366caf2693c"

definition P224_SHA224_TV10_Sig :: "nat × nat" where
  "P224_SHA224_TV10_Sig = (P224_SHA224_TV10_R, P224_SHA224_TV10_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA224_TV10_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV10_d" 
  by eval

lemma P224_SHA224_TV10_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA224_TV10_Q" 
  by eval

lemma P224_SHA224_TV10_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA224_TV10_d = P224_SHA224_TV10_Q" 
  by eval

lemma P224_SHA224_TV10_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA224_TV10_d P224_SHA224_TV10_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA224_TV10_dQvalidPair' P224_SHA224_TV10_d_valid) 

lemma P224_SHA224_TV10_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV10_k" 
  by eval

lemma P224_SHA224_TV10_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA224octets P224_SHA224_TV10_d P224_SHA224_TV10_Msg P224_SHA224_TV10_k = Some P224_SHA224_TV10_Sig" 
  by eval

lemma P224_SHA224_TV10_Verify: 
  "SEC1_P224_ECDSA_Verify SHA224octets P224_SHA224_TV10_Q P224_SHA224_TV10_Msg P224_SHA224_TV10_Sig" 
  by eval

subsection‹Test Vector: P224_SHA224_TV11›

subsubsection‹Given Test Vector Values›

definition P224_SHA224_TV11_Msg :: octets where
  "P224_SHA224_TV11_Msg = nat_to_octets_len 0xc89766374c5a5ccef5823e7a9b54af835ac56afbbb517bd77bfecf3fea876bd0cc9ea486e3d685cfe3fb05f25d9c67992cd7863c80a55c7a263249eb3996c4698ad7381131bf3700b7b24d7ca281a100cf2b750e7f0f933e662a08d9f9e47d779fb03754bd20931262ff381a2fe7d1dc94f4a0520de73fa72020494d3133ecf7 128"

definition P224_SHA224_TV11_d :: nat where
  "P224_SHA224_TV11_d = 0x3dbe18cd88fa49febfcb60f0369a67b2379a466d906ac46a8b8d522b"

definition P224_SHA224_TV11_Qx :: nat where
  "P224_SHA224_TV11_Qx = 0xb10150fd797eb870d377f1dbfa197f7d0f0ad29965af573ec13cc42a"

definition P224_SHA224_TV11_Qy :: nat where
  "P224_SHA224_TV11_Qy = 0x17b63ccefbe27fb2a1139e5757b1082aeaa564f478c23a8f631eed5c"

definition P224_SHA224_TV11_Q :: "int point" where
  "P224_SHA224_TV11_Q = Point (int P224_SHA224_TV11_Qx) (int P224_SHA224_TV11_Qy)"

definition P224_SHA224_TV11_k :: nat where
  "P224_SHA224_TV11_k = 0x385803b262ee2ee875838b3a645a745d2e199ae112ef73a25d68d15f"

definition P224_SHA224_TV11_R :: nat where
  "P224_SHA224_TV11_R = 0x1d293b697f297af77872582eb7f543dc250ec79ad453300d264a3b70"

definition P224_SHA224_TV11_S :: nat where
  "P224_SHA224_TV11_S = 0x517a91b89c4859fcc10834242e710c5f0fed90ac938aa5ccdb7c66de"

definition P224_SHA224_TV11_Sig :: "nat × nat" where
  "P224_SHA224_TV11_Sig = (P224_SHA224_TV11_R, P224_SHA224_TV11_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA224_TV11_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV11_d" 
  by eval

lemma P224_SHA224_TV11_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA224_TV11_Q" 
  by eval

lemma P224_SHA224_TV11_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA224_TV11_d = P224_SHA224_TV11_Q" 
  by eval

lemma P224_SHA224_TV11_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA224_TV11_d P224_SHA224_TV11_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA224_TV11_dQvalidPair' P224_SHA224_TV11_d_valid) 

lemma P224_SHA224_TV11_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV11_k" 
  by eval

lemma P224_SHA224_TV11_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA224octets P224_SHA224_TV11_d P224_SHA224_TV11_Msg P224_SHA224_TV11_k = Some P224_SHA224_TV11_Sig" 
  by eval

lemma P224_SHA224_TV11_Verify: 
  "SEC1_P224_ECDSA_Verify SHA224octets P224_SHA224_TV11_Q P224_SHA224_TV11_Msg P224_SHA224_TV11_Sig" 
  by eval

subsection‹Test Vector: P224_SHA224_TV12›

subsubsection‹Given Test Vector Values›

definition P224_SHA224_TV12_Msg :: octets where
  "P224_SHA224_TV12_Msg = nat_to_octets_len 0x30f0e3b502eec5646929d48fd46aa73991d82079c7bd50a38b38ec0bd84167c8cf5ba39bec26999e70208af9b445046cd9d20c82b7629ca1e51bdd00daddbc35f9eb036a15ac57898642d9db09479a38cc80a2e41e380c8a766b2d623de2de798e1eabc02234b89b85d60154460c3bf12764f3fbf17fcccc82df516a2fbe4ecf 128"

definition P224_SHA224_TV12_d :: nat where
  "P224_SHA224_TV12_d = 0xc906b667f38c5135ea96c95722c713dbd125d61156a546f49ddaadc6"

definition P224_SHA224_TV12_Qx :: nat where
  "P224_SHA224_TV12_Qx = 0x3c9b4ef1748a1925578658d3af51995b989ad760790157b25fe09826"

definition P224_SHA224_TV12_Qy :: nat where
  "P224_SHA224_TV12_Qy = 0x55648f4ff4edfb899e9a13bd8d20f5c24b35dc6a6a4e42ed5983b4a0"

definition P224_SHA224_TV12_Q :: "int point" where
  "P224_SHA224_TV12_Q = Point (int P224_SHA224_TV12_Qx) (int P224_SHA224_TV12_Qy)"

definition P224_SHA224_TV12_k :: nat where
  "P224_SHA224_TV12_k = 0xb04d78d8ac40fefadb99f389a06d93f6b5b72198c1be02dbff6195f0"

definition P224_SHA224_TV12_R :: nat where
  "P224_SHA224_TV12_R = 0x4bdd3c84647bad93dcaffd1b54eb87fc61a5704b19d7e6d756d11ad0"

definition P224_SHA224_TV12_S :: nat where
  "P224_SHA224_TV12_S = 0xfdd81e5dca54158514f44ba2330271eff4c618330328451e2d93b9fb"

definition P224_SHA224_TV12_Sig :: "nat × nat" where
  "P224_SHA224_TV12_Sig = (P224_SHA224_TV12_R, P224_SHA224_TV12_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA224_TV12_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV12_d" 
  by eval

lemma P224_SHA224_TV12_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA224_TV12_Q" 
  by eval

lemma P224_SHA224_TV12_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA224_TV12_d = P224_SHA224_TV12_Q" 
  by eval

lemma P224_SHA224_TV12_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA224_TV12_d P224_SHA224_TV12_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA224_TV12_dQvalidPair' P224_SHA224_TV12_d_valid) 

lemma P224_SHA224_TV12_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV12_k" 
  by eval

lemma P224_SHA224_TV12_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA224octets P224_SHA224_TV12_d P224_SHA224_TV12_Msg P224_SHA224_TV12_k = Some P224_SHA224_TV12_Sig" 
  by eval

lemma P224_SHA224_TV12_Verify: 
  "SEC1_P224_ECDSA_Verify SHA224octets P224_SHA224_TV12_Q P224_SHA224_TV12_Msg P224_SHA224_TV12_Sig" 
  by eval

subsection‹Test Vector: P224_SHA224_TV13›

subsubsection‹Given Test Vector Values›

definition P224_SHA224_TV13_Msg :: octets where
  "P224_SHA224_TV13_Msg = nat_to_octets_len 0x6bbb4bf987c8e5069e47c1a541b48b8a3e6d14bfd9ac6dfaa7503b64ab5e1a55f63e91cf5c3e703ac27ad88756dd7fb2d73b909fc15302d0592b974d47e72e60ed339a40b34d39a49b69ea4a5d26ce86f3ca00a70f1cd416a6a5722e8f39d1f0e966981803d6f46dac34e4c7640204cd0d9f1e53fc3acf30096cd00fa80b3ae9 128"

definition P224_SHA224_TV13_d :: nat where
  "P224_SHA224_TV13_d = 0x3456745fbd51eac9b8095cd687b112f93d1b58352dbe02c66bb9b0cc"

definition P224_SHA224_TV13_Qx :: nat where
  "P224_SHA224_TV13_Qx = 0xf0acdfbc75a748a4a0ac55281754b5c4a364b7d61c5390b334daae10"

definition P224_SHA224_TV13_Qy :: nat where
  "P224_SHA224_TV13_Qy = 0x86587a6768f235bf523fbfc6e062c7401ac2b0242cfe4e5fb34f4057"

definition P224_SHA224_TV13_Q :: "int point" where
  "P224_SHA224_TV13_Q = Point (int P224_SHA224_TV13_Qx) (int P224_SHA224_TV13_Qy)"

definition P224_SHA224_TV13_k :: nat where
  "P224_SHA224_TV13_k = 0x854b20c61bcdf7a89959dbf0985880bb14b628f01c65ef4f6446f1c1"

definition P224_SHA224_TV13_R :: nat where
  "P224_SHA224_TV13_R = 0xa2601fbb9fe89f39814735febb349143baa934170ffb91c6448a7823"

definition P224_SHA224_TV13_S :: nat where
  "P224_SHA224_TV13_S = 0xbf90f9305616020a0e34ef30803fc15fa97dffc0948452bbf6cb5f66"

definition P224_SHA224_TV13_Sig :: "nat × nat" where
  "P224_SHA224_TV13_Sig = (P224_SHA224_TV13_R, P224_SHA224_TV13_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA224_TV13_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV13_d" 
  by eval

lemma P224_SHA224_TV13_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA224_TV13_Q" 
  by eval

lemma P224_SHA224_TV13_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA224_TV13_d = P224_SHA224_TV13_Q" 
  by eval

lemma P224_SHA224_TV13_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA224_TV13_d P224_SHA224_TV13_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA224_TV13_dQvalidPair' P224_SHA224_TV13_d_valid) 

lemma P224_SHA224_TV13_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV13_k" 
  by eval

lemma P224_SHA224_TV13_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA224octets P224_SHA224_TV13_d P224_SHA224_TV13_Msg P224_SHA224_TV13_k = Some P224_SHA224_TV13_Sig" 
  by eval

lemma P224_SHA224_TV13_Verify: 
  "SEC1_P224_ECDSA_Verify SHA224octets P224_SHA224_TV13_Q P224_SHA224_TV13_Msg P224_SHA224_TV13_Sig" 
  by eval

subsection‹Test Vector: P224_SHA224_TV14›

subsubsection‹Given Test Vector Values›

definition P224_SHA224_TV14_Msg :: octets where
  "P224_SHA224_TV14_Msg = nat_to_octets_len 0x05b8f8e56214d4217323f2066f974f638f0b83689fc4ed1201848230efdc1fbca8f70359cecc921050141d3b02c2f17aa306fc2ce5fc06e7d0f4be162fcd985a0b687b4ba09b681cb52ffe890bf5bb4a104cb2e770c04df433013605eb8c72a09902f4246d6c22b8c191ef1b0bece10d5ce2744fc7345307dd1b41b6eff0ca89 128"

definition P224_SHA224_TV14_d :: nat where
  "P224_SHA224_TV14_d = 0x2c522af64baaca7b7a08044312f5e265ec6e09b2272f462cc705e4c3"

definition P224_SHA224_TV14_Qx :: nat where
  "P224_SHA224_TV14_Qx = 0x5fad3c047074b5de1960247d0cc216b4e3fb7f3b9cd960575c8479fc"

definition P224_SHA224_TV14_Qy :: nat where
  "P224_SHA224_TV14_Qy = 0xe4fc9c7f05ff0b040eb171fdd2a1dfe2572c564c2003a08c3179a422"

definition P224_SHA224_TV14_Q :: "int point" where
  "P224_SHA224_TV14_Q = Point (int P224_SHA224_TV14_Qx) (int P224_SHA224_TV14_Qy)"

definition P224_SHA224_TV14_k :: nat where
  "P224_SHA224_TV14_k = 0x9267763383f8db55eed5b1ca8f4937dc2e0ca6175066dc3d4a4586af"

definition P224_SHA224_TV14_R :: nat where
  "P224_SHA224_TV14_R = 0x422e2e9fe535eb62f11f5f8ce87cf2e9ec65e61c06737cf6a0019ae6"

definition P224_SHA224_TV14_S :: nat where
  "P224_SHA224_TV14_S = 0x116cfcf0965b7bc63aecade71d189d7e98a0434b124f2afbe3ccf0a9"

definition P224_SHA224_TV14_Sig :: "nat × nat" where
  "P224_SHA224_TV14_Sig = (P224_SHA224_TV14_R, P224_SHA224_TV14_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA224_TV14_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV14_d" 
  by eval

lemma P224_SHA224_TV14_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA224_TV14_Q" 
  by eval

lemma P224_SHA224_TV14_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA224_TV14_d = P224_SHA224_TV14_Q" 
  by eval

lemma P224_SHA224_TV14_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA224_TV14_d P224_SHA224_TV14_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA224_TV14_dQvalidPair' P224_SHA224_TV14_d_valid) 

lemma P224_SHA224_TV14_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV14_k" 
  by eval

lemma P224_SHA224_TV14_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA224octets P224_SHA224_TV14_d P224_SHA224_TV14_Msg P224_SHA224_TV14_k = Some P224_SHA224_TV14_Sig" 
  by eval

lemma P224_SHA224_TV14_Verify: 
  "SEC1_P224_ECDSA_Verify SHA224octets P224_SHA224_TV14_Q P224_SHA224_TV14_Msg P224_SHA224_TV14_Sig" 
  by eval

subsection‹Test Vector: P224_SHA224_TV15›

subsubsection‹Given Test Vector Values›

definition P224_SHA224_TV15_Msg :: octets where
  "P224_SHA224_TV15_Msg = nat_to_octets_len 0xe5c979f0832242b143077bce6ef146a53bb4c53abfc033473c59f3c4095a68b7a504b609f2ab163b5f88f374f0f3bff8762278b1f1c37323b9ed448e3de33e6443796a9ecaa466aa75175375418186c352018a57ce874e44ae72401d5c0f401b5a51804724c10653fded9066e8994d36a137fdeb9364601daeef09fd174dde4a 128"

definition P224_SHA224_TV15_d :: nat where
  "P224_SHA224_TV15_d = 0x3eff7d07edda14e8beba397accfee060dbe2a41587a703bbe0a0b912"

definition P224_SHA224_TV15_Qx :: nat where
  "P224_SHA224_TV15_Qx = 0x6dd84f4d66f362844e41a7913c40b4aad5fa9ba56bb44c2d2ed9efac"

definition P224_SHA224_TV15_Qy :: nat where
  "P224_SHA224_TV15_Qy = 0x15f65ebcdf2fd9f8035385a330bdabec0f1cd9cc7bc31d2fadbe7cda"

definition P224_SHA224_TV15_Q :: "int point" where
  "P224_SHA224_TV15_Q = Point (int P224_SHA224_TV15_Qx) (int P224_SHA224_TV15_Qy)"

definition P224_SHA224_TV15_k :: nat where
  "P224_SHA224_TV15_k = 0x7bb48839d7717bab1fdde89bf4f7b4509d1c2c12510925e13655dead"

definition P224_SHA224_TV15_R :: nat where
  "P224_SHA224_TV15_R = 0x127051d85326049115f307af2bc426f6c2d08f4774a0b496fb6982b1"

definition P224_SHA224_TV15_S :: nat where
  "P224_SHA224_TV15_S = 0x6857e84418c1d1179333b4e5307e92abade0b74f7521ad78044bf597"

definition P224_SHA224_TV15_Sig :: "nat × nat" where
  "P224_SHA224_TV15_Sig = (P224_SHA224_TV15_R, P224_SHA224_TV15_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA224_TV15_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV15_d" 
  by eval

lemma P224_SHA224_TV15_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA224_TV15_Q" 
  by eval

lemma P224_SHA224_TV15_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA224_TV15_d = P224_SHA224_TV15_Q" 
  by eval

lemma P224_SHA224_TV15_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA224_TV15_d P224_SHA224_TV15_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA224_TV15_dQvalidPair' P224_SHA224_TV15_d_valid) 

lemma P224_SHA224_TV15_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA224_TV15_k" 
  by eval

lemma P224_SHA224_TV15_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA224octets P224_SHA224_TV15_d P224_SHA224_TV15_Msg P224_SHA224_TV15_k = Some P224_SHA224_TV15_Sig" 
  by eval

lemma P224_SHA224_TV15_Verify: 
  "SEC1_P224_ECDSA_Verify SHA224octets P224_SHA224_TV15_Q P224_SHA224_TV15_Msg P224_SHA224_TV15_Sig" 
  by eval

section‹NIST Test Vectors for ECDSA: Curve P-224, Hash SHA-256›
  
subsection‹Test Vector: P224_SHA256_TV01›

subsubsection‹Given Test Vector Values›

definition P224_SHA256_TV01_Msg :: octets where
  "P224_SHA256_TV01_Msg = nat_to_octets_len 0x2b49de971bb0f705a3fb5914eb7638d72884a6c3550667dbfdf301adf26bde02f387fd426a31be6c9ff8bfe8690c8113c88576427f1466508458349fc86036afcfb66448b947707e791e71f558b2bf4e7e7507773aaf4e9af51eda95cbce0a0f752b216f8a54a045d47801ff410ee411a1b66a516f278327df2462fb5619470e 128"

definition P224_SHA256_TV01_d :: nat where
  "P224_SHA256_TV01_d = 0x888fc992893bdd8aa02c80768832605d020b81ae0b25474154ec89aa"

definition P224_SHA256_TV01_Qx :: nat where
  "P224_SHA256_TV01_Qx = 0x4c741e4d20103670b7161ae72271082155838418084335338ac38fa4"

definition P224_SHA256_TV01_Qy :: nat where
  "P224_SHA256_TV01_Qy = 0xdb7919151ac28587b72bad7ab180ec8e95ab9e2c8d81d9b9d7e2e383"

definition P224_SHA256_TV01_Q :: "int point" where
  "P224_SHA256_TV01_Q = Point (int P224_SHA256_TV01_Qx) (int P224_SHA256_TV01_Qy)"

definition P224_SHA256_TV01_k :: nat where
  "P224_SHA256_TV01_k = 0x06f7a56007825433c4c61153df1a135eee2f38ec687b492ed40d9c90"

definition P224_SHA256_TV01_R :: nat where
  "P224_SHA256_TV01_R = 0x0909c9b9cae8d2790e29db6afdb45c04f5b072c4c20410c7dc9b6772"

definition P224_SHA256_TV01_S :: nat where
  "P224_SHA256_TV01_S = 0x298f4fcae1fe271da1e0345d11d07a1fca43f58af4c113b909eedea0"

definition P224_SHA256_TV01_Sig :: "nat × nat" where
  "P224_SHA256_TV01_Sig = (P224_SHA256_TV01_R, P224_SHA256_TV01_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA256_TV01_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV01_d" 
  by eval

lemma P224_SHA256_TV01_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA256_TV01_Q" 
  by eval

lemma P224_SHA256_TV01_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA256_TV01_d = P224_SHA256_TV01_Q" 
  by eval

lemma P224_SHA256_TV01_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA256_TV01_d P224_SHA256_TV01_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA256_TV01_dQvalidPair' P224_SHA256_TV01_d_valid) 

lemma P224_SHA256_TV01_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV01_k" 
  by eval

lemma P224_SHA256_TV01_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA256octets P224_SHA256_TV01_d P224_SHA256_TV01_Msg P224_SHA256_TV01_k = Some P224_SHA256_TV01_Sig" 
  by eval

lemma P224_SHA256_TV01_Verify: 
  "SEC1_P224_ECDSA_Verify SHA256octets P224_SHA256_TV01_Q P224_SHA256_TV01_Msg P224_SHA256_TV01_Sig" 
  by eval

subsection‹Test Vector: P224_SHA256_TV02›

subsubsection‹Given Test Vector Values›

definition P224_SHA256_TV02_Msg :: octets where
  "P224_SHA256_TV02_Msg = nat_to_octets_len 0x1fa7201d96ad4d190415f2656d1387fa886afc38e5cd18b8c60da367acf32c627d2c9ea19ef3f030e559fc2a21695cdbb65ddf6ba36a70af0d3fa292a32de31da6acc6108ab2be8bd37843338f0c37c2d62648d3d49013edeb9e179dadf78bf885f95e712fcdfcc8a172e47c09ab159f3a00ed7b930f628c3c48257e92fc7407 128"

definition P224_SHA256_TV02_d :: nat where
  "P224_SHA256_TV02_d = 0x5b5a3e186e7d5b9b0fbdfc74a05e0a3d85dc4be4c87269190c839972"

definition P224_SHA256_TV02_Qx :: nat where
  "P224_SHA256_TV02_Qx = 0x897089f4ef05b943eeac06589f0e09ccc571a6add3eb1610a2fc830f"

definition P224_SHA256_TV02_Qy :: nat where
  "P224_SHA256_TV02_Qy = 0x62ba3f6b3e6f0f062058b93e6f25b6041246c5be13584a41cae7e244"

definition P224_SHA256_TV02_Q :: "int point" where
  "P224_SHA256_TV02_Q = Point (int P224_SHA256_TV02_Qx) (int P224_SHA256_TV02_Qy)"

definition P224_SHA256_TV02_k :: nat where
  "P224_SHA256_TV02_k = 0x5b6f7eca2bcc5899fce41b8169d48cd57cf0c4a1b66a30a150072676"

definition P224_SHA256_TV02_R :: nat where
  "P224_SHA256_TV02_R = 0xf12c9985d454ffbc899ebbbb6cf43e3debcac7f19029f8f2f35cce31"

definition P224_SHA256_TV02_S :: nat where
  "P224_SHA256_TV02_S = 0x12fcb848adbd8b1b4c72b2b54a04d936e4a5f480ae2a3ea2e3c1baae"

definition P224_SHA256_TV02_Sig :: "nat × nat" where
  "P224_SHA256_TV02_Sig = (P224_SHA256_TV02_R, P224_SHA256_TV02_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA256_TV02_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV02_d" 
  by eval

lemma P224_SHA256_TV02_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA256_TV02_Q" 
  by eval

lemma P224_SHA256_TV02_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA256_TV02_d = P224_SHA256_TV02_Q" 
  by eval

lemma P224_SHA256_TV02_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA256_TV02_d P224_SHA256_TV02_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA256_TV02_dQvalidPair' P224_SHA256_TV02_d_valid) 

lemma P224_SHA256_TV02_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV02_k" 
  by eval

lemma P224_SHA256_TV02_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA256octets P224_SHA256_TV02_d P224_SHA256_TV02_Msg P224_SHA256_TV02_k = Some P224_SHA256_TV02_Sig" 
  by eval

lemma P224_SHA256_TV02_Verify: 
  "SEC1_P224_ECDSA_Verify SHA256octets P224_SHA256_TV02_Q P224_SHA256_TV02_Msg P224_SHA256_TV02_Sig" 
  by eval

subsection‹Test Vector: P224_SHA256_TV03›

subsubsection‹Given Test Vector Values›

definition P224_SHA256_TV03_Msg :: octets where
  "P224_SHA256_TV03_Msg = nat_to_octets_len 0x74715fe10748a5b98b138f390f7ca9629c584c5d6ad268fc455c8de2e800b73fa1ea9aaee85de58baa2ce9ce68d822fc31842c6b153baef3a12bf6b4541f74af65430ae931a64c8b4950ad1c76b31aea8c229b3623390e233c112586aa5907bbe419841f54f0a7d6d19c003b91dc84bbb59b14ec477a1e9d194c137e21c75bbb 128"

definition P224_SHA256_TV03_d :: nat where
  "P224_SHA256_TV03_d = 0xf60b3a4d4e31c7005a3d2d0f91cb096d016a8ddb5ab10ecb2a549170"

definition P224_SHA256_TV03_Qx :: nat where
  "P224_SHA256_TV03_Qx = 0x40a4ab1e6a9f84b4dedb81795e6a7124d1cfdfd7ec64c5d4b9e32666"

definition P224_SHA256_TV03_Qy :: nat where
  "P224_SHA256_TV03_Qy = 0x83aa32a3c2fc068e62626f2dafce5d7f050e826e5c145cd2d13d1b27"

definition P224_SHA256_TV03_Q :: "int point" where
  "P224_SHA256_TV03_Q = Point (int P224_SHA256_TV03_Qx) (int P224_SHA256_TV03_Qy)"

definition P224_SHA256_TV03_k :: nat where
  "P224_SHA256_TV03_k = 0xc31150420dfb38ba8347e29add189ec3e38c14b0c541497fb90bf395"

definition P224_SHA256_TV03_R :: nat where
  "P224_SHA256_TV03_R = 0xbf6c6daa89b21211ea2c9f45192d91603378d46b1a5057962dafaf12"

definition P224_SHA256_TV03_S :: nat where
  "P224_SHA256_TV03_S = 0xcb6b237950e0f0369323055cd1f643528c7a64616f75b11c4ddd63c7"

definition P224_SHA256_TV03_Sig :: "nat × nat" where
  "P224_SHA256_TV03_Sig = (P224_SHA256_TV03_R, P224_SHA256_TV03_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA256_TV03_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV03_d" 
  by eval

lemma P224_SHA256_TV03_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA256_TV03_Q" 
  by eval

lemma P224_SHA256_TV03_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA256_TV03_d = P224_SHA256_TV03_Q" 
  by eval

lemma P224_SHA256_TV03_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA256_TV03_d P224_SHA256_TV03_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA256_TV03_dQvalidPair' P224_SHA256_TV03_d_valid) 

lemma P224_SHA256_TV03_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV03_k" 
  by eval

lemma P224_SHA256_TV03_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA256octets P224_SHA256_TV03_d P224_SHA256_TV03_Msg P224_SHA256_TV03_k = Some P224_SHA256_TV03_Sig" 
  by eval

lemma P224_SHA256_TV03_Verify: 
  "SEC1_P224_ECDSA_Verify SHA256octets P224_SHA256_TV03_Q P224_SHA256_TV03_Msg P224_SHA256_TV03_Sig" 
  by eval

subsection‹Test Vector: P224_SHA256_TV04›

subsubsection‹Given Test Vector Values›

definition P224_SHA256_TV04_Msg :: octets where
  "P224_SHA256_TV04_Msg = nat_to_octets_len 0xd10131982dd1a1d839aba383cd72855bf41061c0cb04dfa1acad3181f240341d744ca6002b52f25fb3c63f16d050c4a4ef2c0ebf5f16ce987558f4b9d4a5ad3c6b81b617de00e04ba32282d8bf223bfedbb325b741dfdc8f56fa85c65d42f05f6a1330d8cc6664ad32050dd7b9e3993f4d6c91e5e12cbd9e82196e009ad22560 128"

definition P224_SHA256_TV04_d :: nat where
  "P224_SHA256_TV04_d = 0xc8fc474d3b1cba5981348de5aef0839e376f9f18e7588f1eed7c8c85"

definition P224_SHA256_TV04_Qx :: nat where
  "P224_SHA256_TV04_Qx = 0x66f49457ed15f67ed4042195856f052fe774077f61cebcb9efddc365"

definition P224_SHA256_TV04_Qy :: nat where
  "P224_SHA256_TV04_Qy = 0x3a6e3f3423eec7308a69eb1b0416d67cc3b84d24f251d7cbdb45c079"

definition P224_SHA256_TV04_Q :: "int point" where
  "P224_SHA256_TV04_Q = Point (int P224_SHA256_TV04_Qx) (int P224_SHA256_TV04_Qy)"

definition P224_SHA256_TV04_k :: nat where
  "P224_SHA256_TV04_k = 0x5e5405ae9ab6164bb476c1bb021ec78480e0488736e4f8222920fbd9"

definition P224_SHA256_TV04_R :: nat where
  "P224_SHA256_TV04_R = 0x7b7beaf9f696ca1a8051527478c4c075ab45aa4768937886dbf38618"

definition P224_SHA256_TV04_S :: nat where
  "P224_SHA256_TV04_S = 0x93d4cf110a37c5a6f15c4e6024822118539e860dee2f60b8c3f462f6"

definition P224_SHA256_TV04_Sig :: "nat × nat" where
  "P224_SHA256_TV04_Sig = (P224_SHA256_TV04_R, P224_SHA256_TV04_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA256_TV04_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV04_d" 
  by eval

lemma P224_SHA256_TV04_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA256_TV04_Q" 
  by eval

lemma P224_SHA256_TV04_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA256_TV04_d = P224_SHA256_TV04_Q" 
  by eval

lemma P224_SHA256_TV04_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA256_TV04_d P224_SHA256_TV04_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA256_TV04_dQvalidPair' P224_SHA256_TV04_d_valid) 

lemma P224_SHA256_TV04_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV04_k" 
  by eval

lemma P224_SHA256_TV04_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA256octets P224_SHA256_TV04_d P224_SHA256_TV04_Msg P224_SHA256_TV04_k = Some P224_SHA256_TV04_Sig" 
  by eval

lemma P224_SHA256_TV04_Verify: 
  "SEC1_P224_ECDSA_Verify SHA256octets P224_SHA256_TV04_Q P224_SHA256_TV04_Msg P224_SHA256_TV04_Sig" 
  by eval

subsection‹Test Vector: P224_SHA256_TV05›

subsubsection‹Given Test Vector Values›

definition P224_SHA256_TV05_Msg :: octets where
  "P224_SHA256_TV05_Msg = nat_to_octets_len 0xef9dbd90ded96ad627a0a987ab90537a3e7acc1fdfa991088e9d999fd726e3ce1e1bd89a7df08d8c2bf51085254c89dc67bc21e8a1a93f33a38c18c0ce3880e958ac3e3dbe8aec49f981821c4ac6812dd29fab3a9ebe7fbd799fb50f12021b48d1d9abca8842547b3b99befa612cc8b4ca5f9412e0352e72ab1344a0ac2913db 128"

definition P224_SHA256_TV05_d :: nat where
  "P224_SHA256_TV05_d = 0x04ef5d2a45341e2ace9af8a6ebd25f6cde45453f55b7a724eb6c21f6"

definition P224_SHA256_TV05_Qx :: nat where
  "P224_SHA256_TV05_Qx = 0x8d642868e4d0f55ee62a2052e6b806b566d2ac79dbde7939fe725773"

definition P224_SHA256_TV05_Qy :: nat where
  "P224_SHA256_TV05_Qy = 0x79505a57cd56904d2523b3e1281e9021167657d38aeb7d42fc8ec849"

definition P224_SHA256_TV05_Q :: "int point" where
  "P224_SHA256_TV05_Q = Point (int P224_SHA256_TV05_Qx) (int P224_SHA256_TV05_Qy)"

definition P224_SHA256_TV05_k :: nat where
  "P224_SHA256_TV05_k = 0xec60ea6f3d6b74d102e5574182566b7e79a69699a307fee70a2d0d22"

definition P224_SHA256_TV05_R :: nat where
  "P224_SHA256_TV05_R = 0x2fd7fcbb7832c97ce325301dd338b279a9e28b8933284d49c6eabcf6"

definition P224_SHA256_TV05_S :: nat where
  "P224_SHA256_TV05_S = 0x550b2f1efc312805a6ed8f252e692d8ee19eaa5bcd5d0cda63a1a3f0"

definition P224_SHA256_TV05_Sig :: "nat × nat" where
  "P224_SHA256_TV05_Sig = (P224_SHA256_TV05_R, P224_SHA256_TV05_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA256_TV05_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV05_d" 
  by eval

lemma P224_SHA256_TV05_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA256_TV05_Q" 
  by eval

lemma P224_SHA256_TV05_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA256_TV05_d = P224_SHA256_TV05_Q" 
  by eval

lemma P224_SHA256_TV05_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA256_TV05_d P224_SHA256_TV05_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA256_TV05_dQvalidPair' P224_SHA256_TV05_d_valid) 

lemma P224_SHA256_TV05_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV05_k" 
  by eval

lemma P224_SHA256_TV05_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA256octets P224_SHA256_TV05_d P224_SHA256_TV05_Msg P224_SHA256_TV05_k = Some P224_SHA256_TV05_Sig" 
  by eval

lemma P224_SHA256_TV05_Verify: 
  "SEC1_P224_ECDSA_Verify SHA256octets P224_SHA256_TV05_Q P224_SHA256_TV05_Msg P224_SHA256_TV05_Sig" 
  by eval

subsection‹Test Vector: P224_SHA256_TV06›

subsubsection‹Given Test Vector Values›

definition P224_SHA256_TV06_Msg :: octets where
  "P224_SHA256_TV06_Msg = nat_to_octets_len 0x4cc91f744ac858d3577e48813219aa3538dd813b186b42d1e6218376f07cc1cc448ddd6b37240e98bf953f49cf54d65c12878b33c0bf6eb1c60254f0b6fa974f847e53abc56773eef6f29885dfc619e6a48fc15a667ca94001a0c945b6357a53221b0f4b266181456b0d2d25e90708777f1a6f85971c00140c631c1991e0fd06 128"

definition P224_SHA256_TV06_d :: nat where
  "P224_SHA256_TV06_d = 0x35d4bbe77d149812339e85c79483cb270bdac56bbf30b5ef3d1f4d39"

definition P224_SHA256_TV06_Qx :: nat where
  "P224_SHA256_TV06_Qx = 0x7924b1d7f5920cce98e25094e40f2eb3eb80d70b17e14b3d36c3671c"

definition P224_SHA256_TV06_Qy :: nat where
  "P224_SHA256_TV06_Qy = 0x26c5af35f71e61858582b7cc2b41790597c53ee514ffdf7a289d108c"

definition P224_SHA256_TV06_Q :: "int point" where
  "P224_SHA256_TV06_Q = Point (int P224_SHA256_TV06_Qx) (int P224_SHA256_TV06_Qy)"

definition P224_SHA256_TV06_k :: nat where
  "P224_SHA256_TV06_k = 0x751869c1d0e79eb30aae8fbfb6d97bfa332123fd6b6c72c9cd3c1796"

definition P224_SHA256_TV06_R :: nat where
  "P224_SHA256_TV06_R = 0x26bb1b92b0f01e94eba5fa429271371db527ce857abba13bd1103f64"

definition P224_SHA256_TV06_S :: nat where
  "P224_SHA256_TV06_S = 0x836aba9c63e1252c2b2d72a21e6a41b82241ebe32647e7f814652bcb"

definition P224_SHA256_TV06_Sig :: "nat × nat" where
  "P224_SHA256_TV06_Sig = (P224_SHA256_TV06_R, P224_SHA256_TV06_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA256_TV06_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV06_d" 
  by eval

lemma P224_SHA256_TV06_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA256_TV06_Q" 
  by eval

lemma P224_SHA256_TV06_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA256_TV06_d = P224_SHA256_TV06_Q" 
  by eval

lemma P224_SHA256_TV06_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA256_TV06_d P224_SHA256_TV06_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA256_TV06_dQvalidPair' P224_SHA256_TV06_d_valid) 

lemma P224_SHA256_TV06_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV06_k" 
  by eval

lemma P224_SHA256_TV06_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA256octets P224_SHA256_TV06_d P224_SHA256_TV06_Msg P224_SHA256_TV06_k = Some P224_SHA256_TV06_Sig" 
  by eval

lemma P224_SHA256_TV06_Verify: 
  "SEC1_P224_ECDSA_Verify SHA256octets P224_SHA256_TV06_Q P224_SHA256_TV06_Msg P224_SHA256_TV06_Sig" 
  by eval

subsection‹Test Vector: P224_SHA256_TV07›

subsubsection‹Given Test Vector Values›

definition P224_SHA256_TV07_Msg :: octets where
  "P224_SHA256_TV07_Msg = nat_to_octets_len 0x58f43cc1924de4bc5867664adbc9d26b4f096a43aca47c27c52851b006dc2a658919ef9ce5b5ac48372703be15ac51631c2bd84b88f479f113b0569a9a09e230ec1e8e573474c6075284d3e57d973829af35325d9e7dab4a5f9b065155bbcaff3642a82ef4c9b9e127d3575c050721653da3b087d3fa394192897a5519527d19 128"

definition P224_SHA256_TV07_d :: nat where
  "P224_SHA256_TV07_d = 0x2c291a393281b75264c9b8817af684fa86a1cdc900822f74039dc5d6"

definition P224_SHA256_TV07_Qx :: nat where
  "P224_SHA256_TV07_Qx = 0x18cb5826ad60e6696bf07655032a3749f6577ca36da3ccd6e66a137c"

definition P224_SHA256_TV07_Qy :: nat where
  "P224_SHA256_TV07_Qy = 0x194e14820fe02d784fd1363ff7a30399518309765bd3f4412d646da2"

definition P224_SHA256_TV07_Q :: "int point" where
  "P224_SHA256_TV07_Q = Point (int P224_SHA256_TV07_Qx) (int P224_SHA256_TV07_Qy)"

definition P224_SHA256_TV07_k :: nat where
  "P224_SHA256_TV07_k = 0xe2a860416229dfd3f5a5cc92344ca015093a543943a0d8f73bf2b2fd"

definition P224_SHA256_TV07_R :: nat where
  "P224_SHA256_TV07_R = 0x00e300c1ef4a8c4ca5da6413856f8981db49de29bdf03f32ffc3ceab"

definition P224_SHA256_TV07_S :: nat where
  "P224_SHA256_TV07_S = 0xf250f18a51ba5f63e1584097841099fa6ae4e98ee458c061d1d5aed7"

definition P224_SHA256_TV07_Sig :: "nat × nat" where
  "P224_SHA256_TV07_Sig = (P224_SHA256_TV07_R, P224_SHA256_TV07_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA256_TV07_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV07_d" 
  by eval

lemma P224_SHA256_TV07_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA256_TV07_Q" 
  by eval

lemma P224_SHA256_TV07_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA256_TV07_d = P224_SHA256_TV07_Q" 
  by eval

lemma P224_SHA256_TV07_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA256_TV07_d P224_SHA256_TV07_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA256_TV07_dQvalidPair' P224_SHA256_TV07_d_valid) 

lemma P224_SHA256_TV07_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV07_k" 
  by eval

lemma P224_SHA256_TV07_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA256octets P224_SHA256_TV07_d P224_SHA256_TV07_Msg P224_SHA256_TV07_k = Some P224_SHA256_TV07_Sig" 
  by eval

lemma P224_SHA256_TV07_Verify: 
  "SEC1_P224_ECDSA_Verify SHA256octets P224_SHA256_TV07_Q P224_SHA256_TV07_Msg P224_SHA256_TV07_Sig" 
  by eval

subsection‹Test Vector: P224_SHA256_TV08›

subsubsection‹Given Test Vector Values›

definition P224_SHA256_TV08_Msg :: octets where
  "P224_SHA256_TV08_Msg = nat_to_octets_len 0x113a2806b052fde683ee09453098e402204155afb3776fd1cad3a9103421d327eab8f9ec0dd050ffcc83f93b34ea707705fabeccfe43ab1a71c95298fd3ec769d99ead1066950eee677d225816e0faad19cf69e1b35d16771689e2092cafe16d7c0dd7b0db73fffb8d0f3eaed83004dd21e753530ec939c89ba25578fa5f785b 128"

definition P224_SHA256_TV08_d :: nat where
  "P224_SHA256_TV08_d = 0x831ea25dbeda33d272a1382c5def0e83929170ab06a629eed6ee244b"

definition P224_SHA256_TV08_Qx :: nat where
  "P224_SHA256_TV08_Qx = 0x076518e393940d42dfd09819409d66966d8c9189c83d554a9cc8a082"

definition P224_SHA256_TV08_Qy :: nat where
  "P224_SHA256_TV08_Qy = 0x44d0ceaf4c0f50e46bea4a52e30423ce3ada19edd363ac5694c65cb8"

definition P224_SHA256_TV08_Q :: "int point" where
  "P224_SHA256_TV08_Q = Point (int P224_SHA256_TV08_Qx) (int P224_SHA256_TV08_Qy)"

definition P224_SHA256_TV08_k :: nat where
  "P224_SHA256_TV08_k = 0x6be6dd9f6a083915ccba54626caf12d246d3aece0a7eda7d8d85599c"

definition P224_SHA256_TV08_R :: nat where
  "P224_SHA256_TV08_R = 0xff1460946e06fb6f5d35e8d2625ca70ffb9b45308e3fabf6ad8351b1"

definition P224_SHA256_TV08_S :: nat where
  "P224_SHA256_TV08_S = 0x6029aa3990918e8cb8a388d53b0772e5cdfff49c3405fe0d3a95933a"

definition P224_SHA256_TV08_Sig :: "nat × nat" where
  "P224_SHA256_TV08_Sig = (P224_SHA256_TV08_R, P224_SHA256_TV08_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA256_TV08_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV08_d" 
  by eval

lemma P224_SHA256_TV08_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA256_TV08_Q" 
  by eval

lemma P224_SHA256_TV08_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA256_TV08_d = P224_SHA256_TV08_Q" 
  by eval

lemma P224_SHA256_TV08_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA256_TV08_d P224_SHA256_TV08_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA256_TV08_dQvalidPair' P224_SHA256_TV08_d_valid) 

lemma P224_SHA256_TV08_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV08_k" 
  by eval

lemma P224_SHA256_TV08_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA256octets P224_SHA256_TV08_d P224_SHA256_TV08_Msg P224_SHA256_TV08_k = Some P224_SHA256_TV08_Sig" 
  by eval

lemma P224_SHA256_TV08_Verify: 
  "SEC1_P224_ECDSA_Verify SHA256octets P224_SHA256_TV08_Q P224_SHA256_TV08_Msg P224_SHA256_TV08_Sig" 
  by eval

subsection‹Test Vector: P224_SHA256_TV09›

subsubsection‹Given Test Vector Values›

definition P224_SHA256_TV09_Msg :: octets where
  "P224_SHA256_TV09_Msg = nat_to_octets_len 0x64cbfc8f2e2149a31b3e8a80c4a552f6c62aaeb7990b6e0ee55500a9d17be04213406578caf315951086dff5c2af3b5ce17d425d185101ef26f86396ba3a129a4f3f8e2dd595f59efb6c0f5c2dcc394569d7268695e9ac7daa84203f1f1895f1f9e4b514a5c9cd23baa63454710144fe735ad9b8f42d8c43267aa434a26d7e5f 128"

definition P224_SHA256_TV09_d :: nat where
  "P224_SHA256_TV09_d = 0x70f74c7324ef137318b610ead8ddc5b964e0eed3750b20612fc2e67b"

definition P224_SHA256_TV09_Qx :: nat where
  "P224_SHA256_TV09_Qx = 0x279649e2a2918e683520cde3fc98b0ae58a7100e8de35e7c9cc797b6"

definition P224_SHA256_TV09_Qy :: nat where
  "P224_SHA256_TV09_Qy = 0xaa4de6be34be61f02880139787b9038f4554a8ef1c994b887c2974b5"

definition P224_SHA256_TV09_Q :: "int point" where
  "P224_SHA256_TV09_Q = Point (int P224_SHA256_TV09_Qx) (int P224_SHA256_TV09_Qy)"

definition P224_SHA256_TV09_k :: nat where
  "P224_SHA256_TV09_k = 0x8e984864f86f7a2a73f3edda17dbccd13fac8fa4b872814abf223b1b"

definition P224_SHA256_TV09_R :: nat where
  "P224_SHA256_TV09_R = 0x3b18736fa11d04e27e2614cda03a63ec11a180f357b0b3192920d09c"

definition P224_SHA256_TV09_S :: nat where
  "P224_SHA256_TV09_S = 0x2f0f3dbd570727b14fbb29155538e62c930dd51c4035275c1365dc60"

definition P224_SHA256_TV09_Sig :: "nat × nat" where
  "P224_SHA256_TV09_Sig = (P224_SHA256_TV09_R, P224_SHA256_TV09_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA256_TV09_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV09_d" 
  by eval

lemma P224_SHA256_TV09_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA256_TV09_Q" 
  by eval

lemma P224_SHA256_TV09_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA256_TV09_d = P224_SHA256_TV09_Q" 
  by eval

lemma P224_SHA256_TV09_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA256_TV09_d P224_SHA256_TV09_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA256_TV09_dQvalidPair' P224_SHA256_TV09_d_valid) 

lemma P224_SHA256_TV09_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV09_k" 
  by eval

lemma P224_SHA256_TV09_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA256octets P224_SHA256_TV09_d P224_SHA256_TV09_Msg P224_SHA256_TV09_k = Some P224_SHA256_TV09_Sig" 
  by eval

lemma P224_SHA256_TV09_Verify: 
  "SEC1_P224_ECDSA_Verify SHA256octets P224_SHA256_TV09_Q P224_SHA256_TV09_Msg P224_SHA256_TV09_Sig" 
  by eval

subsection‹Test Vector: P224_SHA256_TV10›

subsubsection‹Given Test Vector Values›

definition P224_SHA256_TV10_Msg :: octets where
  "P224_SHA256_TV10_Msg = nat_to_octets_len 0xa10a11c8e30fff118d371daf824f16c08200b83ea059436466a4611ccac93b2dea2de8c1006f946196aef7fe9b0c251a391b0340f21797798278b412ff2b53842eec6450728e2bca062f8337a2c204b9ea04ff660cd4d4db559f2f11c4d8ef199021339fcc82396f7a93926cf5f247e37d8067fe50692de54f102bd5ab51925c 128"

definition P224_SHA256_TV10_d :: nat where
  "P224_SHA256_TV10_d = 0x026be5789886d25039c11d7d58a11a6e1d52cb1d5657561f2165b8a8"

definition P224_SHA256_TV10_Qx :: nat where
  "P224_SHA256_TV10_Qx = 0x3fa617c50b177da1a2bdb98b780ad21ad1195c4bd24465f6187de3c9"

definition P224_SHA256_TV10_Qy :: nat where
  "P224_SHA256_TV10_Qy = 0xe3fd8d8876dfd03a4a4e31a1acad3a08d983826d286c250c4e5620c1"

definition P224_SHA256_TV10_Q :: "int point" where
  "P224_SHA256_TV10_Q = Point (int P224_SHA256_TV10_Qx) (int P224_SHA256_TV10_Qy)"

definition P224_SHA256_TV10_k :: nat where
  "P224_SHA256_TV10_k = 0x0128b8e3f50731eb5fcc223517fc0cf6b96cd1d2807eb4524bc46f77"

definition P224_SHA256_TV10_R :: nat where
  "P224_SHA256_TV10_R = 0x3a6b633f96f3d0b6d54f7fb29ac33709e4f0dd8fa0e51606ed9765ca"

definition P224_SHA256_TV10_S :: nat where
  "P224_SHA256_TV10_S = 0x63e8c119dfa51784decd864f6911f2210a80f8f02d472d88df10d119"

definition P224_SHA256_TV10_Sig :: "nat × nat" where
  "P224_SHA256_TV10_Sig = (P224_SHA256_TV10_R, P224_SHA256_TV10_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA256_TV10_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV10_d" 
  by eval

lemma P224_SHA256_TV10_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA256_TV10_Q" 
  by eval

lemma P224_SHA256_TV10_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA256_TV10_d = P224_SHA256_TV10_Q" 
  by eval

lemma P224_SHA256_TV10_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA256_TV10_d P224_SHA256_TV10_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA256_TV10_dQvalidPair' P224_SHA256_TV10_d_valid) 

lemma P224_SHA256_TV10_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV10_k" 
  by eval

lemma P224_SHA256_TV10_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA256octets P224_SHA256_TV10_d P224_SHA256_TV10_Msg P224_SHA256_TV10_k = Some P224_SHA256_TV10_Sig" 
  by eval

lemma P224_SHA256_TV10_Verify: 
  "SEC1_P224_ECDSA_Verify SHA256octets P224_SHA256_TV10_Q P224_SHA256_TV10_Msg P224_SHA256_TV10_Sig" 
  by eval

subsection‹Test Vector: P224_SHA256_TV11›

subsubsection‹Given Test Vector Values›

definition P224_SHA256_TV11_Msg :: octets where
  "P224_SHA256_TV11_Msg = nat_to_octets_len 0xb3f720bf566ffa369259f4361959ae0641d2755ec264a4c4349981df2b02563275b2b9adb5aee47f7a456760a971991ffed6b17809bb9694138d1677fa916123795239353158fc6b22d10f20d26f5d2dcd8c56c44373eea5b93067dba2d7c5318dac2e9e8714873cb1b37f58c011fd14fa1e535554efe05f468bfc8e11cd8b99 128"

definition P224_SHA256_TV11_d :: nat where
  "P224_SHA256_TV11_d = 0xe79c18d935c2839644762867aa793201f96a3cde080c5968412ce784"

definition P224_SHA256_TV11_Qx :: nat where
  "P224_SHA256_TV11_Qx = 0xb7ae1e992b1c7fde1141f40bd913358538ca0f07f62b729f13cea327"

definition P224_SHA256_TV11_Qy :: nat where
  "P224_SHA256_TV11_Qy = 0x811252d12120e04805fc171a439d382c43b68a21e1a0bdf5e4ec1da4"

definition P224_SHA256_TV11_Q :: "int point" where
  "P224_SHA256_TV11_Q = Point (int P224_SHA256_TV11_Qx) (int P224_SHA256_TV11_Qy)"

definition P224_SHA256_TV11_k :: nat where
  "P224_SHA256_TV11_k = 0x7abedab1d36f4f0959a03d968b27dd5708223b66e0fc48594d827361"

definition P224_SHA256_TV11_R :: nat where
  "P224_SHA256_TV11_R = 0xd35047d74e1e7305bb8c1a94e8ae47cb1591c3437a3e185e00afe710"

definition P224_SHA256_TV11_S :: nat where
  "P224_SHA256_TV11_S = 0xd9c425c9d5feb776ac8952e6c4eee0ecd68aef2f0e7bff2e49c9185e"

definition P224_SHA256_TV11_Sig :: "nat × nat" where
  "P224_SHA256_TV11_Sig = (P224_SHA256_TV11_R, P224_SHA256_TV11_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA256_TV11_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV11_d" 
  by eval

lemma P224_SHA256_TV11_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA256_TV11_Q" 
  by eval

lemma P224_SHA256_TV11_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA256_TV11_d = P224_SHA256_TV11_Q" 
  by eval

lemma P224_SHA256_TV11_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA256_TV11_d P224_SHA256_TV11_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA256_TV11_dQvalidPair' P224_SHA256_TV11_d_valid) 

lemma P224_SHA256_TV11_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV11_k" 
  by eval

lemma P224_SHA256_TV11_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA256octets P224_SHA256_TV11_d P224_SHA256_TV11_Msg P224_SHA256_TV11_k = Some P224_SHA256_TV11_Sig" 
  by eval

lemma P224_SHA256_TV11_Verify: 
  "SEC1_P224_ECDSA_Verify SHA256octets P224_SHA256_TV11_Q P224_SHA256_TV11_Msg P224_SHA256_TV11_Sig" 
  by eval

subsection‹Test Vector: P224_SHA256_TV12›

subsubsection‹Given Test Vector Values›

definition P224_SHA256_TV12_Msg :: octets where
  "P224_SHA256_TV12_Msg = nat_to_octets_len 0x0a398a46df7ccc48d1e7833f8bbc67100f1ef77a62dc78bbc115b2a662f9591fbaaa91ad3d788e2fdd1b3164e45293d4f5686c151296901768028ac80ded4bf89c647ad35f0c7c4cb318c0c757c1d83c44d850e5fd4677281b3f13b1ee54de79c8c042813f9d3312dcc6111a68299cb7e829557d7f3d96e702f65aefc6499415 128"

definition P224_SHA256_TV12_d :: nat where
  "P224_SHA256_TV12_d = 0x0d087f9d1f8ae29c9cf791490efc4a5789a9d52038c4b1d22494ad8c"

definition P224_SHA256_TV12_Qx :: nat where
  "P224_SHA256_TV12_Qx = 0xcd95cf8fb1cd21690f40d647f2353672a1076cc6c46bddaad2d0fc56"

definition P224_SHA256_TV12_Qy :: nat where
  "P224_SHA256_TV12_Qy = 0x934262f74d9ee0f8a2754f64cb7415923d64bf00c94a39b52803f577"

definition P224_SHA256_TV12_Q :: "int point" where
  "P224_SHA256_TV12_Q = Point (int P224_SHA256_TV12_Qx) (int P224_SHA256_TV12_Qy)"

definition P224_SHA256_TV12_k :: nat where
  "P224_SHA256_TV12_k = 0x557d0e3995dc6377b3911546dd7aeaeec62a6d8f2af6a274382fc37f"

definition P224_SHA256_TV12_R :: nat where
  "P224_SHA256_TV12_R = 0x56df0ea6afdcc232ceb41729eec00cf906b69b6e28423a36d3c92cc5"

definition P224_SHA256_TV12_S :: nat where
  "P224_SHA256_TV12_S = 0xf4f70fd948c9a147f55317fdea7b8a84c33e721014552d5800d63edc"

definition P224_SHA256_TV12_Sig :: "nat × nat" where
  "P224_SHA256_TV12_Sig = (P224_SHA256_TV12_R, P224_SHA256_TV12_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA256_TV12_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV12_d" 
  by eval

lemma P224_SHA256_TV12_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA256_TV12_Q" 
  by eval

lemma P224_SHA256_TV12_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA256_TV12_d = P224_SHA256_TV12_Q" 
  by eval

lemma P224_SHA256_TV12_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA256_TV12_d P224_SHA256_TV12_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA256_TV12_dQvalidPair' P224_SHA256_TV12_d_valid) 

lemma P224_SHA256_TV12_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV12_k" 
  by eval

lemma P224_SHA256_TV12_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA256octets P224_SHA256_TV12_d P224_SHA256_TV12_Msg P224_SHA256_TV12_k = Some P224_SHA256_TV12_Sig" 
  by eval

lemma P224_SHA256_TV12_Verify: 
  "SEC1_P224_ECDSA_Verify SHA256octets P224_SHA256_TV12_Q P224_SHA256_TV12_Msg P224_SHA256_TV12_Sig" 
  by eval

subsection‹Test Vector: P224_SHA256_TV13›

subsubsection‹Given Test Vector Values›

definition P224_SHA256_TV13_Msg :: octets where
  "P224_SHA256_TV13_Msg = nat_to_octets_len 0x8c33616821a6038b448d8918668977fcf1ef5aa0cf7c341837b39bbcc9bca875a3757f4b392630e9995b9bbe4eb66978b877586adaa02f99d2344dae082a7603351d8ffcfca081ab403cd0acb90d078dd1d0789c2eb3185c62bff2d9f04cd38e509e3b83c12ed0a5c6808fc42f7ba5b06acdc496c8ad9be648ee6a4505f8560f 128"

definition P224_SHA256_TV13_d :: nat where
  "P224_SHA256_TV13_d = 0x0830aebb6577d3a3be3ba54a4501c987b0e0bb593267b9bbadb66583"

definition P224_SHA256_TV13_Qx :: nat where
  "P224_SHA256_TV13_Qx = 0xb88652020e083ccc1c43dc83d1881884dd4c7e3b4e3460b344b1ea64"

definition P224_SHA256_TV13_Qy :: nat where
  "P224_SHA256_TV13_Qy = 0x22b69b517f86d7c26dc37c0f8feb4bb07fe876149fbcc3334fd2805b"

definition P224_SHA256_TV13_Q :: "int point" where
  "P224_SHA256_TV13_Q = Point (int P224_SHA256_TV13_Qx) (int P224_SHA256_TV13_Qy)"

definition P224_SHA256_TV13_k :: nat where
  "P224_SHA256_TV13_k = 0xe4f4a3280574c704c2fde47ca81ec883d27f2c5a961a294db7cda9d2"

definition P224_SHA256_TV13_R :: nat where
  "P224_SHA256_TV13_R = 0xb30b8a0079d9a134b5e1618c2ac63e3fbe0e95866b9dbc5f423f2707"

definition P224_SHA256_TV13_S :: nat where
  "P224_SHA256_TV13_S = 0x3dc36746610271ef66e0aa52cc2ccadc5c9b08dc769e4dc4f6538c11"

definition P224_SHA256_TV13_Sig :: "nat × nat" where
  "P224_SHA256_TV13_Sig = (P224_SHA256_TV13_R, P224_SHA256_TV13_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA256_TV13_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV13_d" 
  by eval

lemma P224_SHA256_TV13_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA256_TV13_Q" 
  by eval

lemma P224_SHA256_TV13_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA256_TV13_d = P224_SHA256_TV13_Q" 
  by eval

lemma P224_SHA256_TV13_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA256_TV13_d P224_SHA256_TV13_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA256_TV13_dQvalidPair' P224_SHA256_TV13_d_valid) 

lemma P224_SHA256_TV13_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV13_k" 
  by eval

lemma P224_SHA256_TV13_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA256octets P224_SHA256_TV13_d P224_SHA256_TV13_Msg P224_SHA256_TV13_k = Some P224_SHA256_TV13_Sig" 
  by eval

lemma P224_SHA256_TV13_Verify: 
  "SEC1_P224_ECDSA_Verify SHA256octets P224_SHA256_TV13_Q P224_SHA256_TV13_Msg P224_SHA256_TV13_Sig" 
  by eval

subsection‹Test Vector: P224_SHA256_TV14›

subsubsection‹Given Test Vector Values›

definition P224_SHA256_TV14_Msg :: octets where
  "P224_SHA256_TV14_Msg = nat_to_octets_len 0x94d56535fd4edfe67a0daa6579f9d53bf6b7b8830ae2aeb62892ff59f18756ddf2811b449c7d20d65d54f8507de4e7c50eaa084830637812aa4b250a4d61ab67845be36e4a41cdc0a70f8d6e3a63d4514f0dc197e6486015046a316153d5f3a3a4a0ae1ed7ea5fa55e12e73d333333685c02e0eb636234ea7e6d4b76b4b76b5a 128"

definition P224_SHA256_TV14_d :: nat where
  "P224_SHA256_TV14_d = 0x2acc9b97e625263e8e4cd164302c7d1e078bfcdd706111a13ccda5b2"

definition P224_SHA256_TV14_Qx :: nat where
  "P224_SHA256_TV14_Qx = 0xce1a06f82df874dded37cca03b56c0648e4e8917ecd40ee73ee61588"

definition P224_SHA256_TV14_Qy :: nat where
  "P224_SHA256_TV14_Qy = 0xceb6177b8f1ac7c5c6e6e1f7737cc3026952ee392badd2cd7af32f9d"

definition P224_SHA256_TV14_Q :: "int point" where
  "P224_SHA256_TV14_Q = Point (int P224_SHA256_TV14_Qx) (int P224_SHA256_TV14_Qy)"

definition P224_SHA256_TV14_k :: nat where
  "P224_SHA256_TV14_k = 0xe401fa80f96480d437ed4f61a783888062ec33d530b188fd48016a6d"

definition P224_SHA256_TV14_R :: nat where
  "P224_SHA256_TV14_R = 0x28674f447c4742e4087bbccfb522fbad4e18b56031d2ce8f532b078a"

definition P224_SHA256_TV14_S :: nat where
  "P224_SHA256_TV14_S = 0xa5a7a13d15b423dd17771f73cea98d89dbffa846cc209b45c0e29b76"

definition P224_SHA256_TV14_Sig :: "nat × nat" where
  "P224_SHA256_TV14_Sig = (P224_SHA256_TV14_R, P224_SHA256_TV14_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA256_TV14_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV14_d" 
  by eval

lemma P224_SHA256_TV14_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA256_TV14_Q" 
  by eval

lemma P224_SHA256_TV14_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA256_TV14_d = P224_SHA256_TV14_Q" 
  by eval

lemma P224_SHA256_TV14_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA256_TV14_d P224_SHA256_TV14_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA256_TV14_dQvalidPair' P224_SHA256_TV14_d_valid) 

lemma P224_SHA256_TV14_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV14_k" 
  by eval

lemma P224_SHA256_TV14_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA256octets P224_SHA256_TV14_d P224_SHA256_TV14_Msg P224_SHA256_TV14_k = Some P224_SHA256_TV14_Sig" 
  by eval

lemma P224_SHA256_TV14_Verify: 
  "SEC1_P224_ECDSA_Verify SHA256octets P224_SHA256_TV14_Q P224_SHA256_TV14_Msg P224_SHA256_TV14_Sig" 
  by eval

subsection‹Test Vector: P224_SHA256_TV15›

subsubsection‹Given Test Vector Values›

definition P224_SHA256_TV15_Msg :: octets where
  "P224_SHA256_TV15_Msg = nat_to_octets_len 0x5d8ebdf9eb28b47bdafaa36bf0b66a9eaf99b6c83959da4f2b1151b4f4ecd28fb115a64c0cb9491093a7e9b9c53ec423e4c72e7765bb9c818da0e8c428667e44474a71db4867130c77c40bfd8544b2d7b9d6464d2b8e6a48482153256a32437c3a747231f51134dd14c703407e31146a6fcde23bededcf16950486e90ca69ac0 128"

definition P224_SHA256_TV15_d :: nat where
  "P224_SHA256_TV15_d = 0xf4e873d4fb944fb52323406f933815092b7672221de4d1c45917f3fc"

definition P224_SHA256_TV15_Qx :: nat where
  "P224_SHA256_TV15_Qx = 0x0dc2cdddb990341adb1de73f02d87fc3822485a659a15145f4251d5f"

definition P224_SHA256_TV15_Qy :: nat where
  "P224_SHA256_TV15_Qy = 0xcf78b2a83c7352eda1af2c74e1804ea04b35f76c04e89d90281dc2bb"

definition P224_SHA256_TV15_Q :: "int point" where
  "P224_SHA256_TV15_Q = Point (int P224_SHA256_TV15_Qx) (int P224_SHA256_TV15_Qy)"

definition P224_SHA256_TV15_k :: nat where
  "P224_SHA256_TV15_k = 0x5d1476c682a64162fd2fdc82696fc8cab1469a86f707ea2757416e40"

definition P224_SHA256_TV15_R :: nat where
  "P224_SHA256_TV15_R = 0x82982b38ed465138df4018d7cfb835edcb591cb57446ca49d163782b"

definition P224_SHA256_TV15_S :: nat where
  "P224_SHA256_TV15_S = 0x8ef1d7b326cabee7f7ab95b7b98d3c27a069c0fd95a1599c0ccb422b"

definition P224_SHA256_TV15_Sig :: "nat × nat" where
  "P224_SHA256_TV15_Sig = (P224_SHA256_TV15_R, P224_SHA256_TV15_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA256_TV15_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV15_d" 
  by eval

lemma P224_SHA256_TV15_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA256_TV15_Q" 
  by eval

lemma P224_SHA256_TV15_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA256_TV15_d = P224_SHA256_TV15_Q" 
  by eval

lemma P224_SHA256_TV15_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA256_TV15_d P224_SHA256_TV15_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA256_TV15_dQvalidPair' P224_SHA256_TV15_d_valid) 

lemma P224_SHA256_TV15_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA256_TV15_k" 
  by eval

lemma P224_SHA256_TV15_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA256octets P224_SHA256_TV15_d P224_SHA256_TV15_Msg P224_SHA256_TV15_k = Some P224_SHA256_TV15_Sig" 
  by eval

lemma P224_SHA256_TV15_Verify: 
  "SEC1_P224_ECDSA_Verify SHA256octets P224_SHA256_TV15_Q P224_SHA256_TV15_Msg P224_SHA256_TV15_Sig" 
  by eval

section‹NIST Test Vectors for ECDSA: Curve P-224, Hash SHA-384›
  
subsection‹Test Vector: P224_SHA384_TV01›

subsubsection‹Given Test Vector Values›

definition P224_SHA384_TV01_Msg :: octets where
  "P224_SHA384_TV01_Msg = nat_to_octets_len 0x25e4416695f77551fdce276355528ccf1ddc2483821c5d22d751d50111ca2fadc6593b52c74f4b5957494f1df25b0b2f86950d0d19229ec6506fee8581d2dd09d48418b146ff16bd84a17ca0dc83b1888eb407376da6c8a88fa1e60b8c2a2471dfde4b3996ef673d5bde3d70c434dc9f2488e9de16ae657d29e5e59ec922a1ec 128"

definition P224_SHA384_TV01_d :: nat where
  "P224_SHA384_TV01_d = 0x62c572ee0d6f81b27e591d788bfc2f42b5105d2663078dfb58069ebd"

definition P224_SHA384_TV01_Qx :: nat where
  "P224_SHA384_TV01_Qx = 0xbd6ba605639b98fa8113a16a3bb004ddfaec901c98a931206165f4a5"

definition P224_SHA384_TV01_Qy :: nat where
  "P224_SHA384_TV01_Qy = 0xa3190b10ef39e88abd60b2293b4707512b45c6c5ed5794cc11454427"

definition P224_SHA384_TV01_Q :: "int point" where
  "P224_SHA384_TV01_Q = Point (int P224_SHA384_TV01_Qx) (int P224_SHA384_TV01_Qy)"

definition P224_SHA384_TV01_k :: nat where
  "P224_SHA384_TV01_k = 0x0f0bb1e428bcdebf4dc62a5278068efc0f8ce75f89e89b3630f102b2"

definition P224_SHA384_TV01_R :: nat where
  "P224_SHA384_TV01_R = 0xaac0ea27e129f544abcc77f110e70bbdd5aa3e425dc39d5e8887025d"

definition P224_SHA384_TV01_S :: nat where
  "P224_SHA384_TV01_S = 0x10e5dd06aee6b8419a04aa33d9d5678b0039c3acc3c4b61fe106bfdc"

definition P224_SHA384_TV01_Sig :: "nat × nat" where
  "P224_SHA384_TV01_Sig = (P224_SHA384_TV01_R, P224_SHA384_TV01_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA384_TV01_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV01_d" 
  by eval

lemma P224_SHA384_TV01_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA384_TV01_Q" 
  by eval

lemma P224_SHA384_TV01_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA384_TV01_d = P224_SHA384_TV01_Q" 
  by eval

lemma P224_SHA384_TV01_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA384_TV01_d P224_SHA384_TV01_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA384_TV01_dQvalidPair' P224_SHA384_TV01_d_valid) 

lemma P224_SHA384_TV01_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV01_k" 
  by eval

lemma P224_SHA384_TV01_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA384octets P224_SHA384_TV01_d P224_SHA384_TV01_Msg P224_SHA384_TV01_k = Some P224_SHA384_TV01_Sig" 
  by eval

lemma P224_SHA384_TV01_Verify: 
  "SEC1_P224_ECDSA_Verify SHA384octets P224_SHA384_TV01_Q P224_SHA384_TV01_Msg P224_SHA384_TV01_Sig" 
  by eval

subsection‹Test Vector: P224_SHA384_TV02›

subsubsection‹Given Test Vector Values›

definition P224_SHA384_TV02_Msg :: octets where
  "P224_SHA384_TV02_Msg = nat_to_octets_len 0x9164d633a553deccf3cbd2effccf1387fa3177cd28c95d94a7d1a3e159c5e5c027758cc26493301b2f4d141d8d07a5fe5fead987ce5f30abeafcb48c302afc6c2309f0e93d9b6818cbb6972d222cb7b01302dfe202ae83b89f53150ae4a0e2b8fc0fd1091f19b4ab2e6ab213ab322d04f2c5f57113bfad3c5675227237abf773 128"

definition P224_SHA384_TV02_d :: nat where
  "P224_SHA384_TV02_d = 0xe2f86bf73ba9336fa023343060f038e9ad41e5fe868e9f80574619a3"

definition P224_SHA384_TV02_Qx :: nat where
  "P224_SHA384_TV02_Qx = 0xf5d5346f17898ea6bbdfff19c216a8757a5dc37b95315f5481628381"

definition P224_SHA384_TV02_Qy :: nat where
  "P224_SHA384_TV02_Qy = 0xae61fd172ac8b7a4f13870a932dece465834cbd4f50bbcfb802c824e"

definition P224_SHA384_TV02_Q :: "int point" where
  "P224_SHA384_TV02_Q = Point (int P224_SHA384_TV02_Qx) (int P224_SHA384_TV02_Qy)"

definition P224_SHA384_TV02_k :: nat where
  "P224_SHA384_TV02_k = 0x35724ac043e3b44b73b5a7919cf675190306d26aa67c27c28c873534"

definition P224_SHA384_TV02_R :: nat where
  "P224_SHA384_TV02_R = 0x535147c265af138eec50c7fb570bcc8d2e6f675597b0fcc034e536bc"

definition P224_SHA384_TV02_S :: nat where
  "P224_SHA384_TV02_S = 0x743812c188a1dddf9fb34b90738f8b2e58760d6cd20ccceb1bb9c516"

definition P224_SHA384_TV02_Sig :: "nat × nat" where
  "P224_SHA384_TV02_Sig = (P224_SHA384_TV02_R, P224_SHA384_TV02_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA384_TV02_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV02_d" 
  by eval

lemma P224_SHA384_TV02_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA384_TV02_Q" 
  by eval

lemma P224_SHA384_TV02_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA384_TV02_d = P224_SHA384_TV02_Q" 
  by eval

lemma P224_SHA384_TV02_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA384_TV02_d P224_SHA384_TV02_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA384_TV02_dQvalidPair' P224_SHA384_TV02_d_valid) 

lemma P224_SHA384_TV02_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV02_k" 
  by eval

lemma P224_SHA384_TV02_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA384octets P224_SHA384_TV02_d P224_SHA384_TV02_Msg P224_SHA384_TV02_k = Some P224_SHA384_TV02_Sig" 
  by eval

lemma P224_SHA384_TV02_Verify: 
  "SEC1_P224_ECDSA_Verify SHA384octets P224_SHA384_TV02_Q P224_SHA384_TV02_Msg P224_SHA384_TV02_Sig" 
  by eval

subsection‹Test Vector: P224_SHA384_TV03›

subsubsection‹Given Test Vector Values›

definition P224_SHA384_TV03_Msg :: octets where
  "P224_SHA384_TV03_Msg = nat_to_octets_len 0x019df05929321ecea7ee1de4f412aba1c8d3c24437db04b194a68a0a59dd871be10bd3a4be6edf551350ea49fc7155a4d887e1221486291abe77a30633a4c4f7868fe2df24311cba0c73804883954460e122387ed414111ff96ff1aebac8b6a6491d8a0d16e48a63bf3d027c0f68ee4a4b234d73b412196706af8ea022b4dcef 128"

definition P224_SHA384_TV03_d :: nat where
  "P224_SHA384_TV03_d = 0xb0a203438e2586d7575bc417a4a798e47abc22aa3955b58fc2789f17"

definition P224_SHA384_TV03_Qx :: nat where
  "P224_SHA384_TV03_Qx = 0xdc5d217862a1e5b00c95affa9d8b925a72b9beaeb7a86dc397e788d8"

definition P224_SHA384_TV03_Qy :: nat where
  "P224_SHA384_TV03_Qy = 0x5f05f8e976ae1eb1036eca6d683a82850795bf9127dee5f8b2859445"

definition P224_SHA384_TV03_Q :: "int point" where
  "P224_SHA384_TV03_Q = Point (int P224_SHA384_TV03_Qx) (int P224_SHA384_TV03_Qy)"

definition P224_SHA384_TV03_k :: nat where
  "P224_SHA384_TV03_k = 0x408e9c8b1f33136d6ddb93ff3a498bc09d4eee99bf69cdd5af0aa5a2"

definition P224_SHA384_TV03_R :: nat where
  "P224_SHA384_TV03_R = 0x1b5a964c8b1fc634c6e2b82322499df1d7f0c12a4d2a77723c816ab8"

definition P224_SHA384_TV03_S :: nat where
  "P224_SHA384_TV03_S = 0xcf54599a36ca064fae0aa936de5266f87704409d22a15d28c01b7f2a"

definition P224_SHA384_TV03_Sig :: "nat × nat" where
  "P224_SHA384_TV03_Sig = (P224_SHA384_TV03_R, P224_SHA384_TV03_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA384_TV03_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV03_d" 
  by eval

lemma P224_SHA384_TV03_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA384_TV03_Q" 
  by eval

lemma P224_SHA384_TV03_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA384_TV03_d = P224_SHA384_TV03_Q" 
  by eval

lemma P224_SHA384_TV03_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA384_TV03_d P224_SHA384_TV03_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA384_TV03_dQvalidPair' P224_SHA384_TV03_d_valid) 

lemma P224_SHA384_TV03_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV03_k" 
  by eval

lemma P224_SHA384_TV03_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA384octets P224_SHA384_TV03_d P224_SHA384_TV03_Msg P224_SHA384_TV03_k = Some P224_SHA384_TV03_Sig" 
  by eval

lemma P224_SHA384_TV03_Verify: 
  "SEC1_P224_ECDSA_Verify SHA384octets P224_SHA384_TV03_Q P224_SHA384_TV03_Msg P224_SHA384_TV03_Sig" 
  by eval

subsection‹Test Vector: P224_SHA384_TV04›

subsubsection‹Given Test Vector Values›

definition P224_SHA384_TV04_Msg :: octets where
  "P224_SHA384_TV04_Msg = nat_to_octets_len 0x5d09d2b1d3fa6e12c10d8b26dc9aabc8dc02bd06e63ff33f8bb91ede4b8694592a69e4ed4cdf6820069e2b9c7803658949e877ffe23bf90bcf5ce1409c06c71d86885a94048b05ac0ec9db193e489a5a2bfa367caf6aa8ecdb032be366174343f6875d2fe1785e8d77334f5f469cec64998e08d3303e5c9a1923b34fdc105d65 128"

definition P224_SHA384_TV04_d :: nat where
  "P224_SHA384_TV04_d = 0xefcfa50fad6fb2065f9a55f28c0c42fa24c809ccb19b6fc6d8ffb085"

definition P224_SHA384_TV04_Qx :: nat where
  "P224_SHA384_TV04_Qx = 0x61521a0cfb72be77ba33cb3b8e022743cd9130ff49e97093b71aa178"

definition P224_SHA384_TV04_Qy :: nat where
  "P224_SHA384_TV04_Qy = 0xce0819aedaf6fce639d0e593f8ab0147eeb6058f5f2b448231584ea9"

definition P224_SHA384_TV04_Q :: "int point" where
  "P224_SHA384_TV04_Q = Point (int P224_SHA384_TV04_Qx) (int P224_SHA384_TV04_Qy)"

definition P224_SHA384_TV04_k :: nat where
  "P224_SHA384_TV04_k = 0xd1eea821f286eae6ebc1f61b08f9ad4323a3787e94af4c32cd31351b"

definition P224_SHA384_TV04_R :: nat where
  "P224_SHA384_TV04_R = 0xb37caaa71103752ac559f9eb4943324409ebfa8b585f684dcaa5c411"

definition P224_SHA384_TV04_S :: nat where
  "P224_SHA384_TV04_S = 0x7c28e7619e2944ab4b7be022878c8052ebdf2cae5dff4f976c49686a"

definition P224_SHA384_TV04_Sig :: "nat × nat" where
  "P224_SHA384_TV04_Sig = (P224_SHA384_TV04_R, P224_SHA384_TV04_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA384_TV04_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV04_d" 
  by eval

lemma P224_SHA384_TV04_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA384_TV04_Q" 
  by eval

lemma P224_SHA384_TV04_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA384_TV04_d = P224_SHA384_TV04_Q" 
  by eval

lemma P224_SHA384_TV04_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA384_TV04_d P224_SHA384_TV04_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA384_TV04_dQvalidPair' P224_SHA384_TV04_d_valid) 

lemma P224_SHA384_TV04_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV04_k" 
  by eval

lemma P224_SHA384_TV04_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA384octets P224_SHA384_TV04_d P224_SHA384_TV04_Msg P224_SHA384_TV04_k = Some P224_SHA384_TV04_Sig" 
  by eval

lemma P224_SHA384_TV04_Verify: 
  "SEC1_P224_ECDSA_Verify SHA384octets P224_SHA384_TV04_Q P224_SHA384_TV04_Msg P224_SHA384_TV04_Sig" 
  by eval

subsection‹Test Vector: P224_SHA384_TV05›

subsubsection‹Given Test Vector Values›

definition P224_SHA384_TV05_Msg :: octets where
  "P224_SHA384_TV05_Msg = nat_to_octets_len 0x50f6dfc81c6cf189e0a310f992907fe93356cee9dea9a41c7671a8daf3f4cfe0c459ce6122c1e731dbf7593419d7114cb73b46956158a982c5d52c72f43f0f822046093c69aeff1f7e4cd8af00ba655c5baa2e7b6a400b4be1f6fd51b3e4cfb35a69c80a28c5cafb771b6c2e52e0aeef0e3fd045e8d40745f3f8b74fd969f816 128"

definition P224_SHA384_TV05_d :: nat where
  "P224_SHA384_TV05_d = 0x61a17816937987764cdc064dc7b5b4f5b16db1023acdfe25902957dd"

definition P224_SHA384_TV05_Qx :: nat where
  "P224_SHA384_TV05_Qx = 0xa7e975c0a8f87c683bb8e31bc160843a7b69c945f4850bd60e1c08c0"

definition P224_SHA384_TV05_Qy :: nat where
  "P224_SHA384_TV05_Qy = 0x8930a454dcc2aa13bed7ea89368b2c9d689d816b2acf4e52585ee9c4"

definition P224_SHA384_TV05_Q :: "int point" where
  "P224_SHA384_TV05_Q = Point (int P224_SHA384_TV05_Qx) (int P224_SHA384_TV05_Qy)"

definition P224_SHA384_TV05_k :: nat where
  "P224_SHA384_TV05_k = 0x44b1fdec2629f9075f89c134ac28ff19bfddaa9db02a5d7f853582b4"

definition P224_SHA384_TV05_R :: nat where
  "P224_SHA384_TV05_R = 0xb0f5635d8bc9c53a1d54a3ec63de59ed66e6b2358d4ab79755414326"

definition P224_SHA384_TV05_S :: nat where
  "P224_SHA384_TV05_S = 0x67c68fe265c7e5aba4232deeafb88545a2aa266fb9f2c2bb3f3ae8d2"

definition P224_SHA384_TV05_Sig :: "nat × nat" where
  "P224_SHA384_TV05_Sig = (P224_SHA384_TV05_R, P224_SHA384_TV05_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA384_TV05_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV05_d" 
  by eval

lemma P224_SHA384_TV05_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA384_TV05_Q" 
  by eval

lemma P224_SHA384_TV05_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA384_TV05_d = P224_SHA384_TV05_Q" 
  by eval

lemma P224_SHA384_TV05_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA384_TV05_d P224_SHA384_TV05_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA384_TV05_dQvalidPair' P224_SHA384_TV05_d_valid) 

lemma P224_SHA384_TV05_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV05_k" 
  by eval

lemma P224_SHA384_TV05_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA384octets P224_SHA384_TV05_d P224_SHA384_TV05_Msg P224_SHA384_TV05_k = Some P224_SHA384_TV05_Sig" 
  by eval

lemma P224_SHA384_TV05_Verify: 
  "SEC1_P224_ECDSA_Verify SHA384octets P224_SHA384_TV05_Q P224_SHA384_TV05_Msg P224_SHA384_TV05_Sig" 
  by eval

subsection‹Test Vector: P224_SHA384_TV06›

subsubsection‹Given Test Vector Values›

definition P224_SHA384_TV06_Msg :: octets where
  "P224_SHA384_TV06_Msg = nat_to_octets_len 0xe90129ac6672c85bb7b6b18e9dc199c96c81fd65034b53c77818364d512366fb9cd1bc7c82404c451e561fc1ed916c0948f6ac561b33a1ccca093f07684b8c2bafa9e966377bd208556018a5bafb9edcecf70498c7140fe9c8cf3ad8b8c3b0aa489df797944465047465415bb0e24333235fcdd59a98829a3941eaaf62033e82 128"

definition P224_SHA384_TV06_d :: nat where
  "P224_SHA384_TV06_d = 0x79d5367314ec664aa0f6ca36f95549502a05bf8400bf532d669fab8d"

definition P224_SHA384_TV06_Qx :: nat where
  "P224_SHA384_TV06_Qx = 0x3191f0237102dac159032ab2dde53cf56c9ec827b5caddfe9e83c02a"

definition P224_SHA384_TV06_Qy :: nat where
  "P224_SHA384_TV06_Qy = 0xb496b1bdcca4434ac0d0d91ea38ff3bc33f9f54095bfe17796d5a9e2"

definition P224_SHA384_TV06_Q :: "int point" where
  "P224_SHA384_TV06_Q = Point (int P224_SHA384_TV06_Qx) (int P224_SHA384_TV06_Qy)"

definition P224_SHA384_TV06_k :: nat where
  "P224_SHA384_TV06_k = 0xda529c52f5cc1f435d873109cd991d6cd7e1631d9ff1dd9521dd5db6"

definition P224_SHA384_TV06_R :: nat where
  "P224_SHA384_TV06_R = 0x8e0ac63903f4921755430572c3f08bc272790639bdf1009fe2a9a714"

definition P224_SHA384_TV06_S :: nat where
  "P224_SHA384_TV06_S = 0x6278c841a2d0a270791fe54b36c49d426d67907aa4e4f59c8638ad97"

definition P224_SHA384_TV06_Sig :: "nat × nat" where
  "P224_SHA384_TV06_Sig = (P224_SHA384_TV06_R, P224_SHA384_TV06_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA384_TV06_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV06_d" 
  by eval

lemma P224_SHA384_TV06_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA384_TV06_Q" 
  by eval

lemma P224_SHA384_TV06_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA384_TV06_d = P224_SHA384_TV06_Q" 
  by eval

lemma P224_SHA384_TV06_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA384_TV06_d P224_SHA384_TV06_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA384_TV06_dQvalidPair' P224_SHA384_TV06_d_valid) 

lemma P224_SHA384_TV06_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV06_k" 
  by eval

lemma P224_SHA384_TV06_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA384octets P224_SHA384_TV06_d P224_SHA384_TV06_Msg P224_SHA384_TV06_k = Some P224_SHA384_TV06_Sig" 
  by eval

lemma P224_SHA384_TV06_Verify: 
  "SEC1_P224_ECDSA_Verify SHA384octets P224_SHA384_TV06_Q P224_SHA384_TV06_Msg P224_SHA384_TV06_Sig" 
  by eval

subsection‹Test Vector: P224_SHA384_TV07›

subsubsection‹Given Test Vector Values›

definition P224_SHA384_TV07_Msg :: octets where
  "P224_SHA384_TV07_Msg = nat_to_octets_len 0x3c9a483c9bee33b601549c592a82e95b4319b1e74b777877f0971bcb4273716b268e8f99f876e42f942f4cf08284896bbc1ffbf094ac0956c3cedfc3580cffa8c74fc6db29a371f2da2d05edb9185ece741fe0d3fabfe9d5b4d373755ebed13dc6840cfa3283b9ea46ec8b95c434f253ae86998182e9cc0e95ee64f323fc74b0 128"

definition P224_SHA384_TV07_d :: nat where
  "P224_SHA384_TV07_d = 0x1320eedad4745121793a7eaf732b0b4498f7cb456cac8cf45a1f66f0"

definition P224_SHA384_TV07_Qx :: nat where
  "P224_SHA384_TV07_Qx = 0x9fdd99906ab77fd29e9021bde947d05a7a9eb153612269bfb0899bc9"

definition P224_SHA384_TV07_Qy :: nat where
  "P224_SHA384_TV07_Qy = 0x681b65b9ac8e4c2899bb622dafb253b7bf5a6e38e5f6595f997c291a"

definition P224_SHA384_TV07_Q :: "int point" where
  "P224_SHA384_TV07_Q = Point (int P224_SHA384_TV07_Qx) (int P224_SHA384_TV07_Qy)"

definition P224_SHA384_TV07_k :: nat where
  "P224_SHA384_TV07_k = 0x66ed8d8934633f4125f593cf1b1d3745c4db1f15dde60cf46ca1c7f2"

definition P224_SHA384_TV07_R :: nat where
  "P224_SHA384_TV07_R = 0x80199485a3a96447b39f7679cd47412a78675ba17dcbd10465dc5b48"

definition P224_SHA384_TV07_S :: nat where
  "P224_SHA384_TV07_S = 0xa251fd9f136a3cb0dd0bc80659ae032e4a761ba7045da0034553fb8c"

definition P224_SHA384_TV07_Sig :: "nat × nat" where
  "P224_SHA384_TV07_Sig = (P224_SHA384_TV07_R, P224_SHA384_TV07_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA384_TV07_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV07_d" 
  by eval

lemma P224_SHA384_TV07_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA384_TV07_Q" 
  by eval

lemma P224_SHA384_TV07_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA384_TV07_d = P224_SHA384_TV07_Q" 
  by eval

lemma P224_SHA384_TV07_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA384_TV07_d P224_SHA384_TV07_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA384_TV07_dQvalidPair' P224_SHA384_TV07_d_valid) 

lemma P224_SHA384_TV07_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV07_k" 
  by eval

lemma P224_SHA384_TV07_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA384octets P224_SHA384_TV07_d P224_SHA384_TV07_Msg P224_SHA384_TV07_k = Some P224_SHA384_TV07_Sig" 
  by eval

lemma P224_SHA384_TV07_Verify: 
  "SEC1_P224_ECDSA_Verify SHA384octets P224_SHA384_TV07_Q P224_SHA384_TV07_Msg P224_SHA384_TV07_Sig" 
  by eval

subsection‹Test Vector: P224_SHA384_TV08›

subsubsection‹Given Test Vector Values›

definition P224_SHA384_TV08_Msg :: octets where
  "P224_SHA384_TV08_Msg = nat_to_octets_len 0xbfc073fdda63c5fccaa0ca8770c293e8154e7aec56128bbac4fdbd541d602216ebf7ca1e02b514d6e396f20683802ba3f334310a9226576926e3bb19ceee27738d13377cbafeb09d091043501702a07aa31d1f29d50ddc55adcf16ffd40578e734a4e6cb6535f26ad48e0c62ad90e79720000e87d419e92dca3e11f943655b03 128"

definition P224_SHA384_TV08_d :: nat where
  "P224_SHA384_TV08_d = 0xe18821329447d3f65ba7279e96bd4624ffa1b32b90f6e8331b1e876d"

definition P224_SHA384_TV08_Qx :: nat where
  "P224_SHA384_TV08_Qx = 0x46c9ed837232c47022df2f1a1578fbe65ac9f2e81c98a74cc22ea31a"

definition P224_SHA384_TV08_Qy :: nat where
  "P224_SHA384_TV08_Qy = 0x6fc5e9568ae62b31412a0b0b367242e9fd7e518c83aa06a069e1d90d"

definition P224_SHA384_TV08_Q :: "int point" where
  "P224_SHA384_TV08_Q = Point (int P224_SHA384_TV08_Qx) (int P224_SHA384_TV08_Qy)"

definition P224_SHA384_TV08_k :: nat where
  "P224_SHA384_TV08_k = 0xa4c1eb402a2fb3af26e0e14a3d2fc8ed3bc1a8b2475270356a79fdd3"

definition P224_SHA384_TV08_R :: nat where
  "P224_SHA384_TV08_R = 0xd478b68733d8ad44be46766e7b66af782fbdc7ff7ed0b191176da98a"

definition P224_SHA384_TV08_S :: nat where
  "P224_SHA384_TV08_S = 0x5eae9160ccf71fd1d359d89cecce72ef8afaeee2365f6ba828aa450a"

definition P224_SHA384_TV08_Sig :: "nat × nat" where
  "P224_SHA384_TV08_Sig = (P224_SHA384_TV08_R, P224_SHA384_TV08_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA384_TV08_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV08_d" 
  by eval

lemma P224_SHA384_TV08_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA384_TV08_Q" 
  by eval

lemma P224_SHA384_TV08_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA384_TV08_d = P224_SHA384_TV08_Q" 
  by eval

lemma P224_SHA384_TV08_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA384_TV08_d P224_SHA384_TV08_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA384_TV08_dQvalidPair' P224_SHA384_TV08_d_valid) 

lemma P224_SHA384_TV08_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV08_k" 
  by eval

lemma P224_SHA384_TV08_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA384octets P224_SHA384_TV08_d P224_SHA384_TV08_Msg P224_SHA384_TV08_k = Some P224_SHA384_TV08_Sig" 
  by eval

lemma P224_SHA384_TV08_Verify: 
  "SEC1_P224_ECDSA_Verify SHA384octets P224_SHA384_TV08_Q P224_SHA384_TV08_Msg P224_SHA384_TV08_Sig" 
  by eval

subsection‹Test Vector: P224_SHA384_TV09›

subsubsection‹Given Test Vector Values›

definition P224_SHA384_TV09_Msg :: octets where
  "P224_SHA384_TV09_Msg = nat_to_octets_len 0x08079955d1a1f33728128c73673ec9f21a6ce138dcab5adc4dc068e6ab57314b9fbd8b013123b2fdafa9524fbdd0288777a233de8055cccfad83046ada6a19f01c47817496667bba8fc8b9456fc0e044a562d931dab1adcb66af8b66325bdf28d83ded3e2937958ccd19da540d70ef2c189f55a506c9c0d63406394c5bd3823b 128"

definition P224_SHA384_TV09_d :: nat where
  "P224_SHA384_TV09_d = 0xf73e030d5a696b358986d3efaca121cf71f775f8835a21e6135145d7"

definition P224_SHA384_TV09_Qx :: nat where
  "P224_SHA384_TV09_Qx = 0x9ca2c6ea87ac8dd3a23a5b4010841a7c8af309038882ae44634bcf55"

definition P224_SHA384_TV09_Qy :: nat where
  "P224_SHA384_TV09_Qy = 0xb0a347dbd5ded3b8702ac5a457e8b32bd4de06fd315095fa1b7d5fe1"

definition P224_SHA384_TV09_Q :: "int point" where
  "P224_SHA384_TV09_Q = Point (int P224_SHA384_TV09_Qx) (int P224_SHA384_TV09_Qy)"

definition P224_SHA384_TV09_k :: nat where
  "P224_SHA384_TV09_k = 0xe3cc786c1288ea567836c51d6d69dd0cab5c015987d936ccc3a4beb3"

definition P224_SHA384_TV09_R :: nat where
  "P224_SHA384_TV09_R = 0xf1234da71761b7a0f49e661a419d2a739bdc4544bf87690e3d2f96db"

definition P224_SHA384_TV09_S :: nat where
  "P224_SHA384_TV09_S = 0x096d16bf8020c3d3c233894ad8eb81206010e62c6e692a215e088fd4"

definition P224_SHA384_TV09_Sig :: "nat × nat" where
  "P224_SHA384_TV09_Sig = (P224_SHA384_TV09_R, P224_SHA384_TV09_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA384_TV09_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV09_d" 
  by eval

lemma P224_SHA384_TV09_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA384_TV09_Q" 
  by eval

lemma P224_SHA384_TV09_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA384_TV09_d = P224_SHA384_TV09_Q" 
  by eval

lemma P224_SHA384_TV09_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA384_TV09_d P224_SHA384_TV09_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA384_TV09_dQvalidPair' P224_SHA384_TV09_d_valid) 

lemma P224_SHA384_TV09_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV09_k" 
  by eval

lemma P224_SHA384_TV09_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA384octets P224_SHA384_TV09_d P224_SHA384_TV09_Msg P224_SHA384_TV09_k = Some P224_SHA384_TV09_Sig" 
  by eval

lemma P224_SHA384_TV09_Verify: 
  "SEC1_P224_ECDSA_Verify SHA384octets P224_SHA384_TV09_Q P224_SHA384_TV09_Msg P224_SHA384_TV09_Sig" 
  by eval

subsection‹Test Vector: P224_SHA384_TV10›

subsubsection‹Given Test Vector Values›

definition P224_SHA384_TV10_Msg :: octets where
  "P224_SHA384_TV10_Msg = nat_to_octets_len 0x23900b768f6cd42b8a8df0dcbc9cb5daec8de36b9d5c619adcc1ba2b649103d5af123746cdf19c3fd0665a6fb9338156182aa06181e3c6e37ce56979612af2927440424f89cef43fc754854b8a5c43370808cf5f9929cf47712512ce2f8a2a20d2e9f568c2848b27dfbe09142843c83905ffa5da3b15501761b03dbc2c5398b6 128"

definition P224_SHA384_TV10_d :: nat where
  "P224_SHA384_TV10_d = 0x7a0789323f8741c157a1753ae165ecaf8e8b03a60561f8b80cee467c"

definition P224_SHA384_TV10_Qx :: nat where
  "P224_SHA384_TV10_Qx = 0x101271a9addd4bd1f19d00bf116c8524f52cefd598e85dc381597acb"

definition P224_SHA384_TV10_Qy :: nat where
  "P224_SHA384_TV10_Qy = 0x2f17d14f4d8ccb28b216553718152ba7c104646d8eca986dd9ddea39"

definition P224_SHA384_TV10_Q :: "int point" where
  "P224_SHA384_TV10_Q = Point (int P224_SHA384_TV10_Qx) (int P224_SHA384_TV10_Qy)"

definition P224_SHA384_TV10_k :: nat where
  "P224_SHA384_TV10_k = 0xd169f04f05b60c625cda864d187938863964dab7bb3b9dfc04b05519"

definition P224_SHA384_TV10_R :: nat where
  "P224_SHA384_TV10_R = 0xe4a51be686a764b709da23ab48b1985e153c6ee238d945e743907afc"

definition P224_SHA384_TV10_S :: nat where
  "P224_SHA384_TV10_S = 0x118a8f1ffe3cd556ce6345bd1a398dd9cc3729b7fd6d8af9bfd82f40"

definition P224_SHA384_TV10_Sig :: "nat × nat" where
  "P224_SHA384_TV10_Sig = (P224_SHA384_TV10_R, P224_SHA384_TV10_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA384_TV10_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV10_d" 
  by eval

lemma P224_SHA384_TV10_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA384_TV10_Q" 
  by eval

lemma P224_SHA384_TV10_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA384_TV10_d = P224_SHA384_TV10_Q" 
  by eval

lemma P224_SHA384_TV10_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA384_TV10_d P224_SHA384_TV10_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA384_TV10_dQvalidPair' P224_SHA384_TV10_d_valid) 

lemma P224_SHA384_TV10_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV10_k" 
  by eval

lemma P224_SHA384_TV10_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA384octets P224_SHA384_TV10_d P224_SHA384_TV10_Msg P224_SHA384_TV10_k = Some P224_SHA384_TV10_Sig" 
  by eval

lemma P224_SHA384_TV10_Verify: 
  "SEC1_P224_ECDSA_Verify SHA384octets P224_SHA384_TV10_Q P224_SHA384_TV10_Msg P224_SHA384_TV10_Sig" 
  by eval

subsection‹Test Vector: P224_SHA384_TV11›

subsubsection‹Given Test Vector Values›

definition P224_SHA384_TV11_Msg :: octets where
  "P224_SHA384_TV11_Msg = nat_to_octets_len 0x1eb28c0bcdd18f73e347f957ece15b4cc83a771b0877e1feaac38e24028fb38ccea8b54ee017dc7c3d5a1327bc6f40b294aa65d7dc487f278846cd101ee84202f14b38aa2c275046aa2577f65ebaea41cd383e8def2fd0b4444dcf426fa75c4082cd7fa035cdb1e0d34a3c79d42130f5b0273eae75bc701dda3aebe7358f41b5 128"

definition P224_SHA384_TV11_d :: nat where
  "P224_SHA384_TV11_d = 0x78e795d0edb11fd9e28dc26b21e751aa89bea0d87932ef11c95c0e18"

definition P224_SHA384_TV11_Qx :: nat where
  "P224_SHA384_TV11_Qx = 0x9edd544107977134bf6360d43ccabb3c94d627c03963c0a04b439627"

definition P224_SHA384_TV11_Qy :: nat where
  "P224_SHA384_TV11_Qy = 0xece4c61d319a0e41f3de7863e7c355bac94395aaa74cdb5f74a87a5b"

definition P224_SHA384_TV11_Q :: "int point" where
  "P224_SHA384_TV11_Q = Point (int P224_SHA384_TV11_Qx) (int P224_SHA384_TV11_Qy)"

definition P224_SHA384_TV11_k :: nat where
  "P224_SHA384_TV11_k = 0x36f7c0f76808b826a0a974a1fd6e155e00a73f1d34674a8f88be405a"

definition P224_SHA384_TV11_R :: nat where
  "P224_SHA384_TV11_R = 0x3e319444438bc2cc92f323ea842cb402b3c3c2448c89869ef7998edb"

definition P224_SHA384_TV11_S :: nat where
  "P224_SHA384_TV11_S = 0x3420cc38f058f41c31e71f4b1ad488f801111c73541de69fcee60695"

definition P224_SHA384_TV11_Sig :: "nat × nat" where
  "P224_SHA384_TV11_Sig = (P224_SHA384_TV11_R, P224_SHA384_TV11_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA384_TV11_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV11_d" 
  by eval

lemma P224_SHA384_TV11_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA384_TV11_Q" 
  by eval

lemma P224_SHA384_TV11_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA384_TV11_d = P224_SHA384_TV11_Q" 
  by eval

lemma P224_SHA384_TV11_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA384_TV11_d P224_SHA384_TV11_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA384_TV11_dQvalidPair' P224_SHA384_TV11_d_valid) 

lemma P224_SHA384_TV11_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV11_k" 
  by eval

lemma P224_SHA384_TV11_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA384octets P224_SHA384_TV11_d P224_SHA384_TV11_Msg P224_SHA384_TV11_k = Some P224_SHA384_TV11_Sig" 
  by eval

lemma P224_SHA384_TV11_Verify: 
  "SEC1_P224_ECDSA_Verify SHA384octets P224_SHA384_TV11_Q P224_SHA384_TV11_Msg P224_SHA384_TV11_Sig" 
  by eval

subsection‹Test Vector: P224_SHA384_TV12›

subsubsection‹Given Test Vector Values›

definition P224_SHA384_TV12_Msg :: octets where
  "P224_SHA384_TV12_Msg = nat_to_octets_len 0xefab51855407438fd5c250670366bca3c026ecec4a59394f00d8a4b51746d0c4564366656d507e3e13e62fe7abeb976b8859895848dbaecf6582f1898ea06f00d4247702ed9721bd375aa83ae4c67c2eaa6e080777ea5ecf2cf787d785389560ac91cf63a52f0373c3185e18a3b8a466e21b61a239f1b77624eb1acacc76c4e1 128"

definition P224_SHA384_TV12_d :: nat where
  "P224_SHA384_TV12_d = 0xbee02d8bc5bffb3fd3b4c9d6f686409f02662d10150d1e58d689966a"

definition P224_SHA384_TV12_Qx :: nat where
  "P224_SHA384_TV12_Qx = 0x8848f964c847fe9dddc774618d4588c9cd56bbe588d7b1fb369c8bfa"

definition P224_SHA384_TV12_Qy :: nat where
  "P224_SHA384_TV12_Qy = 0xebbb699fbd0dc08859fe9132285fe20dff3b9d561c0640b6e0717607"

definition P224_SHA384_TV12_Q :: "int point" where
  "P224_SHA384_TV12_Q = Point (int P224_SHA384_TV12_Qx) (int P224_SHA384_TV12_Qy)"

definition P224_SHA384_TV12_k :: nat where
  "P224_SHA384_TV12_k = 0x59f1450d857b40e5552a4b8cd4ab0df2f01716635d172c1106840f21"

definition P224_SHA384_TV12_R :: nat where
  "P224_SHA384_TV12_R = 0xa206d8398a16a991bc217f77f23c6f648384f254f255a8a876404444"

definition P224_SHA384_TV12_S :: nat where
  "P224_SHA384_TV12_S = 0xeb1169cb5b1423dc0bfaffe565ae57f986e00de06405e3e7b605862e"

definition P224_SHA384_TV12_Sig :: "nat × nat" where
  "P224_SHA384_TV12_Sig = (P224_SHA384_TV12_R, P224_SHA384_TV12_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA384_TV12_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV12_d" 
  by eval

lemma P224_SHA384_TV12_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA384_TV12_Q" 
  by eval

lemma P224_SHA384_TV12_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA384_TV12_d = P224_SHA384_TV12_Q" 
  by eval

lemma P224_SHA384_TV12_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA384_TV12_d P224_SHA384_TV12_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA384_TV12_dQvalidPair' P224_SHA384_TV12_d_valid) 

lemma P224_SHA384_TV12_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV12_k" 
  by eval

lemma P224_SHA384_TV12_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA384octets P224_SHA384_TV12_d P224_SHA384_TV12_Msg P224_SHA384_TV12_k = Some P224_SHA384_TV12_Sig" 
  by eval

lemma P224_SHA384_TV12_Verify: 
  "SEC1_P224_ECDSA_Verify SHA384octets P224_SHA384_TV12_Q P224_SHA384_TV12_Msg P224_SHA384_TV12_Sig" 
  by eval

subsection‹Test Vector: P224_SHA384_TV13›

subsubsection‹Given Test Vector Values›

definition P224_SHA384_TV13_Msg :: octets where
  "P224_SHA384_TV13_Msg = nat_to_octets_len 0x31c29ca10279a417f0cc9b1382cf54dbfdfc89f2e6ef08c403c11f580cbf8674b141ed1a417563282d99a55fc616d836421cde9424815c95e7fb7668bf3f137b29937f14882d74e034b732d78d91af7721aac4950734f5fa5d4b4d35534974f8cab6d2e6dca75ddb57e99148c8a59df9fc5bcd723e546e8356f671cf2f65640a 128"

definition P224_SHA384_TV13_d :: nat where
  "P224_SHA384_TV13_d = 0xdc0ddf6e501418bb8eafc5d7ccc143369e2aa441df8fc57d5f94a738"

definition P224_SHA384_TV13_Qx :: nat where
  "P224_SHA384_TV13_Qx = 0x063a5d632f4144376e14cfb03ad8ccf1489b613acd184d20dff66545"

definition P224_SHA384_TV13_Qy :: nat where
  "P224_SHA384_TV13_Qy = 0xe77727f057b043d8a0f7458196b72e92d11f85b0891c6aaa9d915f58"

definition P224_SHA384_TV13_Q :: "int point" where
  "P224_SHA384_TV13_Q = Point (int P224_SHA384_TV13_Qx) (int P224_SHA384_TV13_Qy)"

definition P224_SHA384_TV13_k :: nat where
  "P224_SHA384_TV13_k = 0xff0e5cae2671db7a1b90e22c63e7570bdd27352d45bac31e338debe0"

definition P224_SHA384_TV13_R :: nat where
  "P224_SHA384_TV13_R = 0x5bc0b4998481ecbd3b6609184a84ca41d69b08c37138097f559259f8"

definition P224_SHA384_TV13_S :: nat where
  "P224_SHA384_TV13_S = 0x0df8828eb1ca85e46405b94e1a2972c34c5e620a54e2f640f04aecc5"

definition P224_SHA384_TV13_Sig :: "nat × nat" where
  "P224_SHA384_TV13_Sig = (P224_SHA384_TV13_R, P224_SHA384_TV13_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA384_TV13_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV13_d" 
  by eval

lemma P224_SHA384_TV13_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA384_TV13_Q" 
  by eval

lemma P224_SHA384_TV13_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA384_TV13_d = P224_SHA384_TV13_Q" 
  by eval

lemma P224_SHA384_TV13_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA384_TV13_d P224_SHA384_TV13_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA384_TV13_dQvalidPair' P224_SHA384_TV13_d_valid) 

lemma P224_SHA384_TV13_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV13_k" 
  by eval

lemma P224_SHA384_TV13_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA384octets P224_SHA384_TV13_d P224_SHA384_TV13_Msg P224_SHA384_TV13_k = Some P224_SHA384_TV13_Sig" 
  by eval

lemma P224_SHA384_TV13_Verify: 
  "SEC1_P224_ECDSA_Verify SHA384octets P224_SHA384_TV13_Q P224_SHA384_TV13_Msg P224_SHA384_TV13_Sig" 
  by eval

subsection‹Test Vector: P224_SHA384_TV14›

subsubsection‹Given Test Vector Values›

definition P224_SHA384_TV14_Msg :: octets where
  "P224_SHA384_TV14_Msg = nat_to_octets_len 0x8db476f92e332519c1a0ece5d8deded6efbd2d8e8784eea0a6b4c3b4296c35f5f8de4317e5c1627b91fb1973fee86c06e4992aa5a20cb7475c8808ff1da354d07a488dffa7838c6ec1e3f99e3acba831f27bee8434eeda3eb36d0c6df3658883cd40068b1bed841310f6eb38d4a3d07d85848770ff7933c054cd8b34662660b1 128"

definition P224_SHA384_TV14_d :: nat where
  "P224_SHA384_TV14_d = 0x229d89b2fcf8441ffc95ebb2ac2ef156e25825782044b2b8bd6a3e01"

definition P224_SHA384_TV14_Qx :: nat where
  "P224_SHA384_TV14_Qx = 0xde616848d8044a44789ef1ba3a6dd66fe9257ddc57f7534e59a701be"

definition P224_SHA384_TV14_Qy :: nat where
  "P224_SHA384_TV14_Qy = 0x26cbf74a6d25e5b34b96d30f327abd574cff7f7dbe6686573a7d6c5c"

definition P224_SHA384_TV14_Q :: "int point" where
  "P224_SHA384_TV14_Q = Point (int P224_SHA384_TV14_Qx) (int P224_SHA384_TV14_Qy)"

definition P224_SHA384_TV14_k :: nat where
  "P224_SHA384_TV14_k = 0x3b18ca6ec8e8e255ac88f64302745ca0b73ff94b2b2d48be95b4aaee"

definition P224_SHA384_TV14_R :: nat where
  "P224_SHA384_TV14_R = 0xfa94fd8b827c06115c1eefd50afc02ce5926ee0e789667783c01c34b"

definition P224_SHA384_TV14_S :: nat where
  "P224_SHA384_TV14_S = 0xedf766a66973cfc33e4159966c07321a7f6549c3c60e8586ef41402b"

definition P224_SHA384_TV14_Sig :: "nat × nat" where
  "P224_SHA384_TV14_Sig = (P224_SHA384_TV14_R, P224_SHA384_TV14_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA384_TV14_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV14_d" 
  by eval

lemma P224_SHA384_TV14_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA384_TV14_Q" 
  by eval

lemma P224_SHA384_TV14_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA384_TV14_d = P224_SHA384_TV14_Q" 
  by eval

lemma P224_SHA384_TV14_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA384_TV14_d P224_SHA384_TV14_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA384_TV14_dQvalidPair' P224_SHA384_TV14_d_valid) 

lemma P224_SHA384_TV14_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV14_k" 
  by eval

lemma P224_SHA384_TV14_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA384octets P224_SHA384_TV14_d P224_SHA384_TV14_Msg P224_SHA384_TV14_k = Some P224_SHA384_TV14_Sig" 
  by eval

lemma P224_SHA384_TV14_Verify: 
  "SEC1_P224_ECDSA_Verify SHA384octets P224_SHA384_TV14_Q P224_SHA384_TV14_Msg P224_SHA384_TV14_Sig" 
  by eval

subsection‹Test Vector: P224_SHA384_TV15›

subsubsection‹Given Test Vector Values›

definition P224_SHA384_TV15_Msg :: octets where
  "P224_SHA384_TV15_Msg = nat_to_octets_len 0xfcb272c828fe8fd3c6f8de9410c7b6e2b36717c1b0e5e359e9109bd7fc378978aa98182a9d99961898ed88999b050d3b64d1457d7a899d6d273b9f4dde2aafa36d76329d62509043c338f265fc4c7d938459b7fa3b230a9f6cb632b61489546bb4181a5ad7f0d7369b8caced48eb374b075b2b325bc86add0f3b680cd9e80acd 128"

definition P224_SHA384_TV15_d :: nat where
  "P224_SHA384_TV15_d = 0x97d747068147c0393a0bb5c159e2c9f1bd538f6204823294883abe28"

definition P224_SHA384_TV15_Qx :: nat where
  "P224_SHA384_TV15_Qx = 0x3858a576eef2ce24d01766997fb81b3f3f78b6104cd188610be221d7"

definition P224_SHA384_TV15_Qy :: nat where
  "P224_SHA384_TV15_Qy = 0x95ffc677ac7bfe3e0bb4cffb17355a964c8356a807151b3cba5d1f4e"

definition P224_SHA384_TV15_Q :: "int point" where
  "P224_SHA384_TV15_Q = Point (int P224_SHA384_TV15_Qx) (int P224_SHA384_TV15_Qy)"

definition P224_SHA384_TV15_k :: nat where
  "P224_SHA384_TV15_k = 0xc1a2ec1ef16cfd5107c892790daefbed061be78bd8576696b60f64d5"

definition P224_SHA384_TV15_R :: nat where
  "P224_SHA384_TV15_R = 0x18c908541843fcdac99b9ff6bb397f3f8094d16b42670216e4eaa2d7"

definition P224_SHA384_TV15_S :: nat where
  "P224_SHA384_TV15_S = 0xc107a8a508ff57c5d4f78f86cc37e129c864d1c44ed5e73909613b74"

definition P224_SHA384_TV15_Sig :: "nat × nat" where
  "P224_SHA384_TV15_Sig = (P224_SHA384_TV15_R, P224_SHA384_TV15_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA384_TV15_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV15_d" 
  by eval

lemma P224_SHA384_TV15_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA384_TV15_Q" 
  by eval

lemma P224_SHA384_TV15_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA384_TV15_d = P224_SHA384_TV15_Q" 
  by eval

lemma P224_SHA384_TV15_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA384_TV15_d P224_SHA384_TV15_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA384_TV15_dQvalidPair' P224_SHA384_TV15_d_valid) 

lemma P224_SHA384_TV15_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA384_TV15_k" 
  by eval

lemma P224_SHA384_TV15_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA384octets P224_SHA384_TV15_d P224_SHA384_TV15_Msg P224_SHA384_TV15_k = Some P224_SHA384_TV15_Sig" 
  by eval

lemma P224_SHA384_TV15_Verify: 
  "SEC1_P224_ECDSA_Verify SHA384octets P224_SHA384_TV15_Q P224_SHA384_TV15_Msg P224_SHA384_TV15_Sig" 
  by eval

section‹NIST Test Vectors for ECDSA: Curve P-224, Hash SHA-512›
  
subsection‹Test Vector: P224_SHA512_TV01›

subsubsection‹Given Test Vector Values›

definition P224_SHA512_TV01_Msg :: octets where
  "P224_SHA512_TV01_Msg = nat_to_octets_len 0x7522492bdb916a597b8121f3e5c273b1d2800ef8c1db4f7dcbae633b60d7da5193ba53a63d7a377b351897c3b24903ae1cd1994211b259be3e6ae2cbc8970e4957fdf782c7d1bc7a91c80c8ef65468d4ef35428f26e2940ae8b0bd9b8074236bf6c00d0ebe83f9ddb2ade0f835138d39f33b59f244e0037c171f1ba7045a96f5 128"

definition P224_SHA512_TV01_d :: nat where
  "P224_SHA512_TV01_d = 0xba5374541c13597bded6880849184a593d69d3d4f0b1cb4d0919cbd6"

definition P224_SHA512_TV01_Qx :: nat where
  "P224_SHA512_TV01_Qx = 0xac635fe00e8b7a3c8ef5655bdfb7f83e8532e59c0cc0b6534d810ffa"

definition P224_SHA512_TV01_Qy :: nat where
  "P224_SHA512_TV01_Qy = 0x1d067aebeba66e79b28ecfe59ac6fdf5e1970dc3a84499c9d90cd8e2"

definition P224_SHA512_TV01_Q :: "int point" where
  "P224_SHA512_TV01_Q = Point (int P224_SHA512_TV01_Qx) (int P224_SHA512_TV01_Qy)"

definition P224_SHA512_TV01_k :: nat where
  "P224_SHA512_TV01_k = 0x187ed1f45c466cbafcd4b9577fb222408c011225dcccfd20f08b8d89"

definition P224_SHA512_TV01_R :: nat where
  "P224_SHA512_TV01_R = 0xf83d54945997584c923c09662c34cf9ad1e987da8bfd9be600e7a098"

definition P224_SHA512_TV01_S :: nat where
  "P224_SHA512_TV01_S = 0x4ff2dba9dba992c98a095b1144a539310e1a570e20c88b7d0aa1955c"

definition P224_SHA512_TV01_Sig :: "nat × nat" where
  "P224_SHA512_TV01_Sig = (P224_SHA512_TV01_R, P224_SHA512_TV01_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA512_TV01_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV01_d" 
  by eval

lemma P224_SHA512_TV01_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA512_TV01_Q" 
  by eval

lemma P224_SHA512_TV01_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA512_TV01_d = P224_SHA512_TV01_Q" 
  by eval

lemma P224_SHA512_TV01_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA512_TV01_d P224_SHA512_TV01_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA512_TV01_dQvalidPair' P224_SHA512_TV01_d_valid) 

lemma P224_SHA512_TV01_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV01_k" 
  by eval

lemma P224_SHA512_TV01_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA512octets P224_SHA512_TV01_d P224_SHA512_TV01_Msg P224_SHA512_TV01_k = Some P224_SHA512_TV01_Sig" 
  by eval

lemma P224_SHA512_TV01_Verify: 
  "SEC1_P224_ECDSA_Verify SHA512octets P224_SHA512_TV01_Q P224_SHA512_TV01_Msg P224_SHA512_TV01_Sig" 
  by eval

subsection‹Test Vector: P224_SHA512_TV02›

subsubsection‹Given Test Vector Values›

definition P224_SHA512_TV02_Msg :: octets where
  "P224_SHA512_TV02_Msg = nat_to_octets_len 0x61097114ff855c3e34a62d9b853f8982d35f29cfa4a89893badbca7849e5fb437a1a38d6451bf0ca5a0d528e352b8e4b57f2ea359a7fc8841d49dd3e570f9b016f14156b0bbc4be822e260bd147ec081454969e11cb0034b7450ef4deb7ed6edb977e2f4ed60121aa095fb0ab40240dc329ecc917f5c64b4410612af065ee9dd 128"

definition P224_SHA512_TV02_d :: nat where
  "P224_SHA512_TV02_d = 0x1e27187134d0a63542adf4665fba22f00cfc7b0a1e02effe913ceedc"

definition P224_SHA512_TV02_Qx :: nat where
  "P224_SHA512_TV02_Qx = 0xecaea8ceea55c3bd418fd34a4ff2499e25e66a104eed846bc00c31d2"

definition P224_SHA512_TV02_Qy :: nat where
  "P224_SHA512_TV02_Qy = 0x3933a356ab1f2dabc303ff0a5d076131e77032e6f502336883bf78a7"

definition P224_SHA512_TV02_Q :: "int point" where
  "P224_SHA512_TV02_Q = Point (int P224_SHA512_TV02_Qx) (int P224_SHA512_TV02_Qy)"

definition P224_SHA512_TV02_k :: nat where
  "P224_SHA512_TV02_k = 0x34cb597deae9a3b1cada937abcd247161b19b2b336b20e2e42ae01f1"

definition P224_SHA512_TV02_R :: nat where
  "P224_SHA512_TV02_R = 0x58177ba46fb291490b39368774accf72736412c1fb5ee0f27b9b1e02"

definition P224_SHA512_TV02_S :: nat where
  "P224_SHA512_TV02_S = 0x58337d78b95a080bfcabb5809bee012501b4da84b8ef310a4628f11c"

definition P224_SHA512_TV02_Sig :: "nat × nat" where
  "P224_SHA512_TV02_Sig = (P224_SHA512_TV02_R, P224_SHA512_TV02_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA512_TV02_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV02_d" 
  by eval

lemma P224_SHA512_TV02_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA512_TV02_Q" 
  by eval

lemma P224_SHA512_TV02_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA512_TV02_d = P224_SHA512_TV02_Q" 
  by eval

lemma P224_SHA512_TV02_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA512_TV02_d P224_SHA512_TV02_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA512_TV02_dQvalidPair' P224_SHA512_TV02_d_valid) 

lemma P224_SHA512_TV02_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV02_k" 
  by eval

lemma P224_SHA512_TV02_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA512octets P224_SHA512_TV02_d P224_SHA512_TV02_Msg P224_SHA512_TV02_k = Some P224_SHA512_TV02_Sig" 
  by eval

lemma P224_SHA512_TV02_Verify: 
  "SEC1_P224_ECDSA_Verify SHA512octets P224_SHA512_TV02_Q P224_SHA512_TV02_Msg P224_SHA512_TV02_Sig" 
  by eval

subsection‹Test Vector: P224_SHA512_TV03›

subsubsection‹Given Test Vector Values›

definition P224_SHA512_TV03_Msg :: octets where
  "P224_SHA512_TV03_Msg = nat_to_octets_len 0xdd09ae6c982bb1440ca175a87766fefeacc49393ff797c446200662744f37a6e30c5d33ba70cbd8f12277fd6cc0704c17478bbab2a3047469e9618e3c340a9c8caaff5ce7c8a4d90ecae6a9b84b813419dec14460298e7521c9b7fdb7a2089328005bd51d57f92a1bcbeecd34aa40482b549e006bbf6c4ce66d34a22dda4e0e0 128"

definition P224_SHA512_TV03_d :: nat where
  "P224_SHA512_TV03_d = 0x0905b40e6c29bfcbf55e04266f68f10ca8d3905001d68bb61a27749b"

definition P224_SHA512_TV03_Qx :: nat where
  "P224_SHA512_TV03_Qx = 0xd656b73b131aa4c6336a57849ce0d3682b6ab2113d013711e8c29762"

definition P224_SHA512_TV03_Qy :: nat where
  "P224_SHA512_TV03_Qy = 0x6328335ffc2029afbfe2a15cc5636978778c3f9dab84840b05f2e705"

definition P224_SHA512_TV03_Q :: "int point" where
  "P224_SHA512_TV03_Q = Point (int P224_SHA512_TV03_Qx) (int P224_SHA512_TV03_Qy)"

definition P224_SHA512_TV03_k :: nat where
  "P224_SHA512_TV03_k = 0xdc82840d147f893497a82f023d7d2cbf0a3a5b2ac6cc1b9b23e504be"

definition P224_SHA512_TV03_R :: nat where
  "P224_SHA512_TV03_R = 0x583af080e0ec7c1ba5a491a84889b7b7b11ccfe18927c7c219b11757"

definition P224_SHA512_TV03_S :: nat where
  "P224_SHA512_TV03_S = 0xb23700035349df25d839f0973bef78a7515287de6c83707907074fa6"

definition P224_SHA512_TV03_Sig :: "nat × nat" where
  "P224_SHA512_TV03_Sig = (P224_SHA512_TV03_R, P224_SHA512_TV03_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA512_TV03_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV03_d" 
  by eval

lemma P224_SHA512_TV03_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA512_TV03_Q" 
  by eval

lemma P224_SHA512_TV03_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA512_TV03_d = P224_SHA512_TV03_Q" 
  by eval

lemma P224_SHA512_TV03_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA512_TV03_d P224_SHA512_TV03_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA512_TV03_dQvalidPair' P224_SHA512_TV03_d_valid) 

lemma P224_SHA512_TV03_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV03_k" 
  by eval

lemma P224_SHA512_TV03_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA512octets P224_SHA512_TV03_d P224_SHA512_TV03_Msg P224_SHA512_TV03_k = Some P224_SHA512_TV03_Sig" 
  by eval

lemma P224_SHA512_TV03_Verify: 
  "SEC1_P224_ECDSA_Verify SHA512octets P224_SHA512_TV03_Q P224_SHA512_TV03_Msg P224_SHA512_TV03_Sig" 
  by eval

subsection‹Test Vector: P224_SHA512_TV04›

subsubsection‹Given Test Vector Values›

definition P224_SHA512_TV04_Msg :: octets where
  "P224_SHA512_TV04_Msg = nat_to_octets_len 0x37a73e2774d3b274db426c89b945696daa96035031f72cea01894b24508c7f81961ec254d36ed6a0f448e11cf7950af769dc6cd2c47e52c6caf0ea92c270974f0214b4db436c36a60fb722060a6bb544462a82e1714f5906ec32886f7d59ebf289541c3a00ec1e004892ef2b1286a0194f55d083c6ec92c64b8fd1452e1c68ba 128"

definition P224_SHA512_TV04_d :: nat where
  "P224_SHA512_TV04_d = 0xafbaede5d75e4f241dd5b53220f3f5b9c1aa1d5d298e2d43236452dc"

definition P224_SHA512_TV04_Qx :: nat where
  "P224_SHA512_TV04_Qx = 0xfe83e59fc8ea8b939355d3258fe53a64d45f63031a0716b7cc416173"

definition P224_SHA512_TV04_Qy :: nat where
  "P224_SHA512_TV04_Qy = 0xf151d23060f1c856eb7f1f58be72a7228c3af89e43b56e9695b558c7"

definition P224_SHA512_TV04_Q :: "int point" where
  "P224_SHA512_TV04_Q = Point (int P224_SHA512_TV04_Qx) (int P224_SHA512_TV04_Qy)"

definition P224_SHA512_TV04_k :: nat where
  "P224_SHA512_TV04_k = 0x0fbbe7b40136c81a8fb894498d5502157a1cf5a89d0643de92cd38f6"

definition P224_SHA512_TV04_R :: nat where
  "P224_SHA512_TV04_R = 0x24f3f457c7b72b7e759d5a8afbf330e31c5d8d2e36f92c0e79c5d87d"

definition P224_SHA512_TV04_S :: nat where
  "P224_SHA512_TV04_S = 0x36fd1193def34f12a960740fd79fb38bf2b480726ccad540eb42cdf8"

definition P224_SHA512_TV04_Sig :: "nat × nat" where
  "P224_SHA512_TV04_Sig = (P224_SHA512_TV04_R, P224_SHA512_TV04_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA512_TV04_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV04_d" 
  by eval

lemma P224_SHA512_TV04_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA512_TV04_Q" 
  by eval

lemma P224_SHA512_TV04_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA512_TV04_d = P224_SHA512_TV04_Q" 
  by eval

lemma P224_SHA512_TV04_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA512_TV04_d P224_SHA512_TV04_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA512_TV04_dQvalidPair' P224_SHA512_TV04_d_valid) 

lemma P224_SHA512_TV04_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV04_k" 
  by eval

lemma P224_SHA512_TV04_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA512octets P224_SHA512_TV04_d P224_SHA512_TV04_Msg P224_SHA512_TV04_k = Some P224_SHA512_TV04_Sig" 
  by eval

lemma P224_SHA512_TV04_Verify: 
  "SEC1_P224_ECDSA_Verify SHA512octets P224_SHA512_TV04_Q P224_SHA512_TV04_Msg P224_SHA512_TV04_Sig" 
  by eval

subsection‹Test Vector: P224_SHA512_TV05›

subsubsection‹Given Test Vector Values›

definition P224_SHA512_TV05_Msg :: octets where
  "P224_SHA512_TV05_Msg = nat_to_octets_len 0x9dc2046ffdc6804544db964481abe5d2d276a2a9eeec4c7ad40215b1de23561d402db69bd0f6eec2254711eea4487c64d9a6b62c3ebaf5ffa8db6e7e3a6e17154d126967a47a853a6f8339bdca9be306a13c7f992ded7619b0da59909a49b1e0930360e05b47f18628a36d69b2f87f2bfddd6a5d4a72f84dc76dbdd43f3a6a35 128"

definition P224_SHA512_TV05_d :: nat where
  "P224_SHA512_TV05_d = 0x950b07b0c2b7539a21b5135bfede214733f2e009647d38d8b21d760c"

definition P224_SHA512_TV05_Qx :: nat where
  "P224_SHA512_TV05_Qx = 0xf43d13bbfcee3b724063b3910fea49fd591b81e86fdb813b1a492d0c"

definition P224_SHA512_TV05_Qy :: nat where
  "P224_SHA512_TV05_Qy = 0x6b4c8d6fa5dc661889e3cf5ec64997a78222837885f85d2fe9b684fb"

definition P224_SHA512_TV05_Q :: "int point" where
  "P224_SHA512_TV05_Q = Point (int P224_SHA512_TV05_Qx) (int P224_SHA512_TV05_Qy)"

definition P224_SHA512_TV05_k :: nat where
  "P224_SHA512_TV05_k = 0x83e110d0d1e700d2f36543028737d2a2f1474aa3b4b28998a39e4793"

definition P224_SHA512_TV05_R :: nat where
  "P224_SHA512_TV05_R = 0x2685265bc878e85d10ab13293dec190881a57c4a467f8fc2170432ea"

definition P224_SHA512_TV05_S :: nat where
  "P224_SHA512_TV05_S = 0x80a347bb49036522369339bd6485a967cdda818915d8eb947302fcf9"

definition P224_SHA512_TV05_Sig :: "nat × nat" where
  "P224_SHA512_TV05_Sig = (P224_SHA512_TV05_R, P224_SHA512_TV05_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA512_TV05_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV05_d" 
  by eval

lemma P224_SHA512_TV05_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA512_TV05_Q" 
  by eval

lemma P224_SHA512_TV05_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA512_TV05_d = P224_SHA512_TV05_Q" 
  by eval

lemma P224_SHA512_TV05_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA512_TV05_d P224_SHA512_TV05_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA512_TV05_dQvalidPair' P224_SHA512_TV05_d_valid) 

lemma P224_SHA512_TV05_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV05_k" 
  by eval

lemma P224_SHA512_TV05_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA512octets P224_SHA512_TV05_d P224_SHA512_TV05_Msg P224_SHA512_TV05_k = Some P224_SHA512_TV05_Sig" 
  by eval

lemma P224_SHA512_TV05_Verify: 
  "SEC1_P224_ECDSA_Verify SHA512octets P224_SHA512_TV05_Q P224_SHA512_TV05_Msg P224_SHA512_TV05_Sig" 
  by eval

subsection‹Test Vector: P224_SHA512_TV06›

subsubsection‹Given Test Vector Values›

definition P224_SHA512_TV06_Msg :: octets where
  "P224_SHA512_TV06_Msg = nat_to_octets_len 0xd9c6847fce688c5e7525a1098b545cb6c15dcd21a02761fc82fc664372a667390680135f91c01a2fa5430c634b1a6d1cd6002d8aa021e7bf5956a7901c2f81bc25d502ba5f55a55f30c0323dc68205cbefec0538e68654e7b327ac1743641896c3e740d8f66f400902b304eafaa4e0d8cffae140536f0922444cc3216a675697 128"

definition P224_SHA512_TV06_d :: nat where
  "P224_SHA512_TV06_d = 0x015bd9f5dfef393b431c3c7fced24385d861ccb563542574a5d2a9bc"

definition P224_SHA512_TV06_Qx :: nat where
  "P224_SHA512_TV06_Qx = 0xe868690641e2cda13b289a6c5d2fb175940396044d9cf27b4f2240af"

definition P224_SHA512_TV06_Qy :: nat where
  "P224_SHA512_TV06_Qy = 0x4c78c9abdf2b7fc67ed4497001d7bcf1daca1739dc14a661f91d7c40"

definition P224_SHA512_TV06_Q :: "int point" where
  "P224_SHA512_TV06_Q = Point (int P224_SHA512_TV06_Qx) (int P224_SHA512_TV06_Qy)"

definition P224_SHA512_TV06_k :: nat where
  "P224_SHA512_TV06_k = 0xe2374350f47c08f3c1359d4edf87e61d1ba4e7dd1540d8d9062efa79"

definition P224_SHA512_TV06_R :: nat where
  "P224_SHA512_TV06_R = 0xe12dc088d2bc032bb214c77d0e0fb749fc8e61ebe1ed72996f1084b6"

definition P224_SHA512_TV06_S :: nat where
  "P224_SHA512_TV06_S = 0x0ab58aa31e0bba5fbc76855e6549f1036fba0a589aeab978ab01b8fb"

definition P224_SHA512_TV06_Sig :: "nat × nat" where
  "P224_SHA512_TV06_Sig = (P224_SHA512_TV06_R, P224_SHA512_TV06_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA512_TV06_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV06_d" 
  by eval

lemma P224_SHA512_TV06_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA512_TV06_Q" 
  by eval

lemma P224_SHA512_TV06_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA512_TV06_d = P224_SHA512_TV06_Q" 
  by eval

lemma P224_SHA512_TV06_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA512_TV06_d P224_SHA512_TV06_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA512_TV06_dQvalidPair' P224_SHA512_TV06_d_valid) 

lemma P224_SHA512_TV06_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV06_k" 
  by eval

lemma P224_SHA512_TV06_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA512octets P224_SHA512_TV06_d P224_SHA512_TV06_Msg P224_SHA512_TV06_k = Some P224_SHA512_TV06_Sig" 
  by eval

lemma P224_SHA512_TV06_Verify: 
  "SEC1_P224_ECDSA_Verify SHA512octets P224_SHA512_TV06_Q P224_SHA512_TV06_Msg P224_SHA512_TV06_Sig" 
  by eval

subsection‹Test Vector: P224_SHA512_TV07›

subsubsection‹Given Test Vector Values›

definition P224_SHA512_TV07_Msg :: octets where
  "P224_SHA512_TV07_Msg = nat_to_octets_len 0x69df8a01b66f04930efd2012ff2243874f256ca8758145d2a9e4ecc84d0dbdbd0dc494ae06db0ccbe819918137c90957114558580d6623efbafdd342b38dad9f08708084d32f874fba04782ce26aaab78de2102ad171f8a8f2b30b5bd3d55fdac5fa3acd6f7def7e61c2533938572b331ba6d1c02bd74bfdbf7337ade8f4a190 128"

definition P224_SHA512_TV07_d :: nat where
  "P224_SHA512_TV07_d = 0x0a3c259df933247445acffb6d8265b601d597fb9997dc2a1eb4deef4"

definition P224_SHA512_TV07_Qx :: nat where
  "P224_SHA512_TV07_Qx = 0xe67f4385a9da54253cc371ee9bc6739ae6385a4b87669c7baf0c460d"

definition P224_SHA512_TV07_Qy :: nat where
  "P224_SHA512_TV07_Qy = 0x2bb00b6ddd7b67d9ac5653ec04ca8529fbf16f815c04da3c2e58e82d"

definition P224_SHA512_TV07_Q :: "int point" where
  "P224_SHA512_TV07_Q = Point (int P224_SHA512_TV07_Qx) (int P224_SHA512_TV07_Qy)"

definition P224_SHA512_TV07_k :: nat where
  "P224_SHA512_TV07_k = 0x8bf5859665b6a23e6b05a311580f60187ba1c4ae89e44877fb48af66"

definition P224_SHA512_TV07_R :: nat where
  "P224_SHA512_TV07_R = 0x653675fb993c3fa9e57b32e33029ec230b966e8077c72c1ec90ddefc"

definition P224_SHA512_TV07_S :: nat where
  "P224_SHA512_TV07_S = 0x792723bf87e315147cd4303de7f1dfe95cd7658ebb95c38c1a196140"

definition P224_SHA512_TV07_Sig :: "nat × nat" where
  "P224_SHA512_TV07_Sig = (P224_SHA512_TV07_R, P224_SHA512_TV07_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA512_TV07_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV07_d" 
  by eval

lemma P224_SHA512_TV07_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA512_TV07_Q" 
  by eval

lemma P224_SHA512_TV07_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA512_TV07_d = P224_SHA512_TV07_Q" 
  by eval

lemma P224_SHA512_TV07_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA512_TV07_d P224_SHA512_TV07_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA512_TV07_dQvalidPair' P224_SHA512_TV07_d_valid) 

lemma P224_SHA512_TV07_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV07_k" 
  by eval

lemma P224_SHA512_TV07_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA512octets P224_SHA512_TV07_d P224_SHA512_TV07_Msg P224_SHA512_TV07_k = Some P224_SHA512_TV07_Sig" 
  by eval

lemma P224_SHA512_TV07_Verify: 
  "SEC1_P224_ECDSA_Verify SHA512octets P224_SHA512_TV07_Q P224_SHA512_TV07_Msg P224_SHA512_TV07_Sig" 
  by eval

subsection‹Test Vector: P224_SHA512_TV08›

subsubsection‹Given Test Vector Values›

definition P224_SHA512_TV08_Msg :: octets where
  "P224_SHA512_TV08_Msg = nat_to_octets_len 0x927524982b8d60777c1105c86fac05f634abf58c73f84fb95d81ba0b86e1e43592c4fcad2e395a40fbe7005697d86088e2fb3bb7287eb3f917d4f2dc281f5cbe65d05b4f9623bca849b10a03beca6aa2056a12ebb91cf257ac448c5e9a78f8349a6a29b17c8978bef43a443cbb8a149eb23f794844fc41693f2dbb97181444be 128"

definition P224_SHA512_TV08_d :: nat where
  "P224_SHA512_TV08_d = 0xa1c8ef463f9e7e3dd63e677412f87cf9ea4ac9a6a2dae629da5b9916"

definition P224_SHA512_TV08_Qx :: nat where
  "P224_SHA512_TV08_Qx = 0x400e5cd4b315ceb309545cd3277acb70bdae2073fda6ad896ea14b27"

definition P224_SHA512_TV08_Qy :: nat where
  "P224_SHA512_TV08_Qy = 0xfbe1d2466cd2e116f38248bd5cabaa6cbe6c4a2694d998abd7b0c991"

definition P224_SHA512_TV08_Q :: "int point" where
  "P224_SHA512_TV08_Q = Point (int P224_SHA512_TV08_Qx) (int P224_SHA512_TV08_Qy)"

definition P224_SHA512_TV08_k :: nat where
  "P224_SHA512_TV08_k = 0x82f55a25d3ed6e47c22a6eed0fa52ed0818b87d6ea7950281dfefc09"

definition P224_SHA512_TV08_R :: nat where
  "P224_SHA512_TV08_R = 0x16305a46a3f6f9e216ef8f6a6f5f0760d064a885657c864e1c1ea035"

definition P224_SHA512_TV08_S :: nat where
  "P224_SHA512_TV08_S = 0x58fd97050bfbca6f87e64e1458c4ad80bae26e280356da344ad3b25d"

definition P224_SHA512_TV08_Sig :: "nat × nat" where
  "P224_SHA512_TV08_Sig = (P224_SHA512_TV08_R, P224_SHA512_TV08_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA512_TV08_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV08_d" 
  by eval

lemma P224_SHA512_TV08_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA512_TV08_Q" 
  by eval

lemma P224_SHA512_TV08_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA512_TV08_d = P224_SHA512_TV08_Q" 
  by eval

lemma P224_SHA512_TV08_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA512_TV08_d P224_SHA512_TV08_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA512_TV08_dQvalidPair' P224_SHA512_TV08_d_valid) 

lemma P224_SHA512_TV08_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV08_k" 
  by eval

lemma P224_SHA512_TV08_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA512octets P224_SHA512_TV08_d P224_SHA512_TV08_Msg P224_SHA512_TV08_k = Some P224_SHA512_TV08_Sig" 
  by eval

lemma P224_SHA512_TV08_Verify: 
  "SEC1_P224_ECDSA_Verify SHA512octets P224_SHA512_TV08_Q P224_SHA512_TV08_Msg P224_SHA512_TV08_Sig" 
  by eval

subsection‹Test Vector: P224_SHA512_TV09›

subsubsection‹Given Test Vector Values›

definition P224_SHA512_TV09_Msg :: octets where
  "P224_SHA512_TV09_Msg = nat_to_octets_len 0x5f9042283561e7f19a436d01c7ef5a950a6d77ede5629cd7e43c0a5d58e8c5673c37945a453291d12938253c71dbe12c8b022ba7276eda6be034ef5ec1ec77dbd1e08f0d7b8e7725b7ec671c075e008a20f77f4ab266f97079b0aa6337df59a33b881954084057b21f294dd14bcb0869a4a6f1f597955ec7bf9d19bb3537a66a 128"

definition P224_SHA512_TV09_d :: nat where
  "P224_SHA512_TV09_d = 0xfa511dbf6fef7e5e9c73e4555eb75d435f7884322d9faf5d78cacc0b"

definition P224_SHA512_TV09_Qx :: nat where
  "P224_SHA512_TV09_Qx = 0xe8dccd706c31f895f2f261ab979cbab51b8ae28196bcc12a42046380"

definition P224_SHA512_TV09_Qy :: nat where
  "P224_SHA512_TV09_Qy = 0xec246be8e71ea3859cb717a59990fe22e4b76858ff49becd70739a01"

definition P224_SHA512_TV09_Q :: "int point" where
  "P224_SHA512_TV09_Q = Point (int P224_SHA512_TV09_Qx) (int P224_SHA512_TV09_Qy)"

definition P224_SHA512_TV09_k :: nat where
  "P224_SHA512_TV09_k = 0xa37d665fe4314aa4cd03eb8e6a1f366b43e11fdb419c96b48f787b62"

definition P224_SHA512_TV09_R :: nat where
  "P224_SHA512_TV09_R = 0x05e4909bcc172ab4140be291aad4660e375032bce2d762b6269ba764"

definition P224_SHA512_TV09_S :: nat where
  "P224_SHA512_TV09_S = 0xe347a1c9d3670690e1d8d1d4cd9579848f442199c10526488da5cebf"

definition P224_SHA512_TV09_Sig :: "nat × nat" where
  "P224_SHA512_TV09_Sig = (P224_SHA512_TV09_R, P224_SHA512_TV09_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA512_TV09_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV09_d" 
  by eval

lemma P224_SHA512_TV09_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA512_TV09_Q" 
  by eval

lemma P224_SHA512_TV09_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA512_TV09_d = P224_SHA512_TV09_Q" 
  by eval

lemma P224_SHA512_TV09_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA512_TV09_d P224_SHA512_TV09_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA512_TV09_dQvalidPair' P224_SHA512_TV09_d_valid) 

lemma P224_SHA512_TV09_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV09_k" 
  by eval

lemma P224_SHA512_TV09_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA512octets P224_SHA512_TV09_d P224_SHA512_TV09_Msg P224_SHA512_TV09_k = Some P224_SHA512_TV09_Sig" 
  by eval

lemma P224_SHA512_TV09_Verify: 
  "SEC1_P224_ECDSA_Verify SHA512octets P224_SHA512_TV09_Q P224_SHA512_TV09_Msg P224_SHA512_TV09_Sig" 
  by eval

subsection‹Test Vector: P224_SHA512_TV10›

subsubsection‹Given Test Vector Values›

definition P224_SHA512_TV10_Msg :: octets where
  "P224_SHA512_TV10_Msg = nat_to_octets_len 0xc2ae5573d3bf396523bfb703db8502fd0760cd1be528f6ddbfb95aad399e0b19f3bd9e0fabdb05d49e3f893dffec5b627c9c2f7ad5f32e92e4e27a38cb5c28657657377fdfa1b66cd7ac3d15c6d49df92d284db99f69744f37dc7cb4e7d52920fdb200a7942623a7057ba82e467dcccaa5da416b48510d8364446a6a5e2a5aa8 128"

definition P224_SHA512_TV10_d :: nat where
  "P224_SHA512_TV10_d = 0xa58bd53646400a646f0e4208320dc679a9664d1c6bfb27fdc8eac7ea"

definition P224_SHA512_TV10_Qx :: nat where
  "P224_SHA512_TV10_Qx = 0xe22e0dc4ecd96eb0071b72ba4b4988bf784f3fe73cb81bfb93d9ac4f"

definition P224_SHA512_TV10_Qy :: nat where
  "P224_SHA512_TV10_Qy = 0xb3e213e518bee1367a4fb3703b9008bac9d95a1fc4aa61225fff9f3c"

definition P224_SHA512_TV10_Q :: "int point" where
  "P224_SHA512_TV10_Q = Point (int P224_SHA512_TV10_Qx) (int P224_SHA512_TV10_Qy)"

definition P224_SHA512_TV10_k :: nat where
  "P224_SHA512_TV10_k = 0x42c5b6f87d3bb1ed74f5ee8398d8f8c61e9e50ffa7a1da12d39893f9"

definition P224_SHA512_TV10_R :: nat where
  "P224_SHA512_TV10_R = 0x5c0e5c6f057de1e99ef5d237a60d7a07fa9a42b120a82f573d9fb7b2"

definition P224_SHA512_TV10_S :: nat where
  "P224_SHA512_TV10_S = 0x2fffc0bf550bd2f650fed085a84501cacfa6a1bb984df1f9237eaa59"

definition P224_SHA512_TV10_Sig :: "nat × nat" where
  "P224_SHA512_TV10_Sig = (P224_SHA512_TV10_R, P224_SHA512_TV10_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA512_TV10_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV10_d" 
  by eval

lemma P224_SHA512_TV10_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA512_TV10_Q" 
  by eval

lemma P224_SHA512_TV10_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA512_TV10_d = P224_SHA512_TV10_Q" 
  by eval

lemma P224_SHA512_TV10_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA512_TV10_d P224_SHA512_TV10_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA512_TV10_dQvalidPair' P224_SHA512_TV10_d_valid) 

lemma P224_SHA512_TV10_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV10_k" 
  by eval

lemma P224_SHA512_TV10_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA512octets P224_SHA512_TV10_d P224_SHA512_TV10_Msg P224_SHA512_TV10_k = Some P224_SHA512_TV10_Sig" 
  by eval

lemma P224_SHA512_TV10_Verify: 
  "SEC1_P224_ECDSA_Verify SHA512octets P224_SHA512_TV10_Q P224_SHA512_TV10_Msg P224_SHA512_TV10_Sig" 
  by eval

subsection‹Test Vector: P224_SHA512_TV11›

subsubsection‹Given Test Vector Values›

definition P224_SHA512_TV11_Msg :: octets where
  "P224_SHA512_TV11_Msg = nat_to_octets_len 0x03c1a1cd30a039d0dcb22fee2450a7fa79495a0d0f4f43d2de4d75bce003c0334a8860f5c164dbd94888a9f751235a3e570d31070e3e1293a7be616af7176600585d36ac013600157d2569d491da4b8a3bf3630c26e0b9925412189f50b0ae6f04c86477932e2ecd8c3546106ae1ebc684cc3adb27ed665eddece886adea4ce3 128"

definition P224_SHA512_TV11_d :: nat where
  "P224_SHA512_TV11_d = 0x64bd4452b572cc95510ac2e572f41136299ff17f6e8448f4ffb571d0"

definition P224_SHA512_TV11_Qx :: nat where
  "P224_SHA512_TV11_Qx = 0x92521fa25c2e034d127e0921efdb167f0b2ff8b20504487ed87fa264"

definition P224_SHA512_TV11_Qy :: nat where
  "P224_SHA512_TV11_Qy = 0xe72c770e37375ad7dc2c4e63e5701826f6606f6ffb9461ee61b4e872"

definition P224_SHA512_TV11_Q :: "int point" where
  "P224_SHA512_TV11_Q = Point (int P224_SHA512_TV11_Qx) (int P224_SHA512_TV11_Qy)"

definition P224_SHA512_TV11_k :: nat where
  "P224_SHA512_TV11_k = 0xeaf76ee4d7e00d13d8a6d03dffd07ad9a8bb6dc8176c9f93059b1b7f"

definition P224_SHA512_TV11_R :: nat where
  "P224_SHA512_TV11_R = 0xcf5058e2a6cf5e61a138b013eb292f38a1b9f07239ae5941dbce8919"

definition P224_SHA512_TV11_S :: nat where
  "P224_SHA512_TV11_S = 0xd14198621650d985d270bc997da6e78588fd0ef843b874c66a3de3c3"

definition P224_SHA512_TV11_Sig :: "nat × nat" where
  "P224_SHA512_TV11_Sig = (P224_SHA512_TV11_R, P224_SHA512_TV11_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA512_TV11_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV11_d" 
  by eval

lemma P224_SHA512_TV11_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA512_TV11_Q" 
  by eval

lemma P224_SHA512_TV11_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA512_TV11_d = P224_SHA512_TV11_Q" 
  by eval

lemma P224_SHA512_TV11_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA512_TV11_d P224_SHA512_TV11_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA512_TV11_dQvalidPair' P224_SHA512_TV11_d_valid) 

lemma P224_SHA512_TV11_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV11_k" 
  by eval

lemma P224_SHA512_TV11_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA512octets P224_SHA512_TV11_d P224_SHA512_TV11_Msg P224_SHA512_TV11_k = Some P224_SHA512_TV11_Sig" 
  by eval

lemma P224_SHA512_TV11_Verify: 
  "SEC1_P224_ECDSA_Verify SHA512octets P224_SHA512_TV11_Q P224_SHA512_TV11_Msg P224_SHA512_TV11_Sig" 
  by eval

subsection‹Test Vector: P224_SHA512_TV12›

subsubsection‹Given Test Vector Values›

definition P224_SHA512_TV12_Msg :: octets where
  "P224_SHA512_TV12_Msg = nat_to_octets_len 0x888f6d9bc7c86c0079fbfd42d8c08d6958f40f6e570fb0b1f03d2f8f8a63df4fcc87b379a222cf835820a999d34996e08961f13b86b075e7fd1c303cd3baa44de42168561589012f7e5300da4f8bdf470c07119a5d9f7ba7293568cd7c6a1b7fc1e41cda40bed7d46e5a28af67ae2aabfefe67a86a1c601e6f5ee543e09bd7b6 128"

definition P224_SHA512_TV12_d :: nat where
  "P224_SHA512_TV12_d = 0x7f3edb710df9d982f486233d0c176aa88f5a0ee81efa9b8145020294"

definition P224_SHA512_TV12_Qx :: nat where
  "P224_SHA512_TV12_Qx = 0xe7611e013e7b43ff5b8b57ad83333bffcc9e469ad23070b5791dc594"

definition P224_SHA512_TV12_Qy :: nat where
  "P224_SHA512_TV12_Qy = 0x7784da0a11dbe16208c6e0b6d5029e71fbec4dffc9fa046d3eeb71c9"

definition P224_SHA512_TV12_Q :: "int point" where
  "P224_SHA512_TV12_Q = Point (int P224_SHA512_TV12_Qx) (int P224_SHA512_TV12_Qy)"

definition P224_SHA512_TV12_k :: nat where
  "P224_SHA512_TV12_k = 0x94db7ef9a232593091eb9a74f289529c7e0d7fef21f80b3c8556b75e"

definition P224_SHA512_TV12_R :: nat where
  "P224_SHA512_TV12_R = 0xa971f45bab10b1d16d7234ca8e4ec987da20d9e867f28aa063296e23"

definition P224_SHA512_TV12_S :: nat where
  "P224_SHA512_TV12_S = 0xe38c538d65a7e1a28fd3ec53f015a7e5beb60e9d309f1e3ba4b2c3d2"

definition P224_SHA512_TV12_Sig :: "nat × nat" where
  "P224_SHA512_TV12_Sig = (P224_SHA512_TV12_R, P224_SHA512_TV12_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA512_TV12_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV12_d" 
  by eval

lemma P224_SHA512_TV12_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA512_TV12_Q" 
  by eval

lemma P224_SHA512_TV12_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA512_TV12_d = P224_SHA512_TV12_Q" 
  by eval

lemma P224_SHA512_TV12_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA512_TV12_d P224_SHA512_TV12_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA512_TV12_dQvalidPair' P224_SHA512_TV12_d_valid) 

lemma P224_SHA512_TV12_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV12_k" 
  by eval

lemma P224_SHA512_TV12_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA512octets P224_SHA512_TV12_d P224_SHA512_TV12_Msg P224_SHA512_TV12_k = Some P224_SHA512_TV12_Sig" 
  by eval

lemma P224_SHA512_TV12_Verify: 
  "SEC1_P224_ECDSA_Verify SHA512octets P224_SHA512_TV12_Q P224_SHA512_TV12_Msg P224_SHA512_TV12_Sig" 
  by eval

subsection‹Test Vector: P224_SHA512_TV13›

subsubsection‹Given Test Vector Values›

definition P224_SHA512_TV13_Msg :: octets where
  "P224_SHA512_TV13_Msg = nat_to_octets_len 0x48453340f1317769e6ee6e103153714365731163dc18f84e9f2fa4b120f9c5a9645ee2f9b66c84c26d95912b422b009b64af96aa418b2427a4209f2e7513ba8e43ec8cf20b34e7529b22eb1199545afe9a9f7d9bcb320aec9ee0162f91c0d1dd9674c9c284f25199c5e109f6f84d7ed0d269cc6413edb81bc2c83e37d644d8b9 128"

definition P224_SHA512_TV13_d :: nat where
  "P224_SHA512_TV13_d = 0xb569f8296ff1d9cc01fffd9919016e5730c1858bdb7b99527153751a"

definition P224_SHA512_TV13_Qx :: nat where
  "P224_SHA512_TV13_Qx = 0x242f34959516a4706172f7dede23110efa314bff22eb320ab88feeff"

definition P224_SHA512_TV13_Qy :: nat where
  "P224_SHA512_TV13_Qy = 0x45e3227710900a8acfc9bcce728119d042f64ca40876c2b380ee46e0"

definition P224_SHA512_TV13_Q :: "int point" where
  "P224_SHA512_TV13_Q = Point (int P224_SHA512_TV13_Qx) (int P224_SHA512_TV13_Qy)"

definition P224_SHA512_TV13_k :: nat where
  "P224_SHA512_TV13_k = 0xae61523866a8f43e6cdd42ba27a34ed06527e8a5842901a64c393f76"

definition P224_SHA512_TV13_R :: nat where
  "P224_SHA512_TV13_R = 0xc2732a4e0815f9f785500e80147e9486994446beccf8a6a352b97585"

definition P224_SHA512_TV13_S :: nat where
  "P224_SHA512_TV13_S = 0x6ecaece6487d7920e398f7f951ab7c7aba5832dabf03704106ad1244"

definition P224_SHA512_TV13_Sig :: "nat × nat" where
  "P224_SHA512_TV13_Sig = (P224_SHA512_TV13_R, P224_SHA512_TV13_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA512_TV13_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV13_d" 
  by eval

lemma P224_SHA512_TV13_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA512_TV13_Q" 
  by eval

lemma P224_SHA512_TV13_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA512_TV13_d = P224_SHA512_TV13_Q" 
  by eval

lemma P224_SHA512_TV13_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA512_TV13_d P224_SHA512_TV13_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA512_TV13_dQvalidPair' P224_SHA512_TV13_d_valid) 

lemma P224_SHA512_TV13_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV13_k" 
  by eval

lemma P224_SHA512_TV13_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA512octets P224_SHA512_TV13_d P224_SHA512_TV13_Msg P224_SHA512_TV13_k = Some P224_SHA512_TV13_Sig" 
  by eval

lemma P224_SHA512_TV13_Verify: 
  "SEC1_P224_ECDSA_Verify SHA512octets P224_SHA512_TV13_Q P224_SHA512_TV13_Msg P224_SHA512_TV13_Sig" 
  by eval

subsection‹Test Vector: P224_SHA512_TV14›

subsubsection‹Given Test Vector Values›

definition P224_SHA512_TV14_Msg :: octets where
  "P224_SHA512_TV14_Msg = nat_to_octets_len 0x4bdfd3b91d83108409ad765b256e0c9b9937ecf647f8e6f9fc807e2e72af8246178b3fe046b4ea10170450d71a4eec790ecb05f03d7077341de26c4db7eeae24d55c9a9093e837dfdb38168fe8230cb9605825a1282fecd741989bfcdb34678fe077477927f66bd26d003e5dda22043341a14dd31841ba483ad5ce2701e0f68e 128"

definition P224_SHA512_TV14_d :: nat where
  "P224_SHA512_TV14_d = 0x41a4dd8eee39232b728516e2f21e66011e7426a6b25986c3ffa237e4"

definition P224_SHA512_TV14_Qx :: nat where
  "P224_SHA512_TV14_Qx = 0xc32988171caab178bf50dc7310bc7f604df5a9d19a8e602519c72d8a"

definition P224_SHA512_TV14_Qy :: nat where
  "P224_SHA512_TV14_Qy = 0xf8985d112ad9de05969e5364d943c1cc5cd198359f4c62b19da0e117"

definition P224_SHA512_TV14_Q :: "int point" where
  "P224_SHA512_TV14_Q = Point (int P224_SHA512_TV14_Qx) (int P224_SHA512_TV14_Qy)"

definition P224_SHA512_TV14_k :: nat where
  "P224_SHA512_TV14_k = 0x827d4999da81fa920c8492ccc1e2d5cdafed9754cf7382a859952071"

definition P224_SHA512_TV14_R :: nat where
  "P224_SHA512_TV14_R = 0x89c61da7422ccd676baec07e2185c12e947a2374eede87847304be6c"

definition P224_SHA512_TV14_S :: nat where
  "P224_SHA512_TV14_S = 0x2685379624717ea28422e8d001c090405a130b4ef9f1ac726c3ca502"

definition P224_SHA512_TV14_Sig :: "nat × nat" where
  "P224_SHA512_TV14_Sig = (P224_SHA512_TV14_R, P224_SHA512_TV14_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA512_TV14_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV14_d" 
  by eval

lemma P224_SHA512_TV14_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA512_TV14_Q" 
  by eval

lemma P224_SHA512_TV14_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA512_TV14_d = P224_SHA512_TV14_Q" 
  by eval

lemma P224_SHA512_TV14_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA512_TV14_d P224_SHA512_TV14_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA512_TV14_dQvalidPair' P224_SHA512_TV14_d_valid) 

lemma P224_SHA512_TV14_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV14_k" 
  by eval

lemma P224_SHA512_TV14_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA512octets P224_SHA512_TV14_d P224_SHA512_TV14_Msg P224_SHA512_TV14_k = Some P224_SHA512_TV14_Sig" 
  by eval

lemma P224_SHA512_TV14_Verify: 
  "SEC1_P224_ECDSA_Verify SHA512octets P224_SHA512_TV14_Q P224_SHA512_TV14_Msg P224_SHA512_TV14_Sig" 
  by eval

subsection‹Test Vector: P224_SHA512_TV15›

subsubsection‹Given Test Vector Values›

definition P224_SHA512_TV15_Msg :: octets where
  "P224_SHA512_TV15_Msg = nat_to_octets_len 0xe6cdee8558bc1eacc24e82f0624ce8d02cc8d925b4dd3dec3a72f4a4e0fb76076bfa3ef2e2c33bdd7c27b322bdc09bbfee8fe46f75dbd7bbd2af09690b7137943efe21706e0a1b6d3089540fc58d85ddb55ea836616db573e36c521be008893f40a0a7c349602cc178ea43be59d31ec6449e7ff2c5379379f7d7645134df1bc3 128"

definition P224_SHA512_TV15_d :: nat where
  "P224_SHA512_TV15_d = 0x67fa50569257c8cc89ac0325db4902003a62f30b917f53e4035a7e04"

definition P224_SHA512_TV15_Qx :: nat where
  "P224_SHA512_TV15_Qx = 0x6773a0436a9c42635730413b19aa4166f08c69c0e5002953da42253b"

definition P224_SHA512_TV15_Qy :: nat where
  "P224_SHA512_TV15_Qy = 0x555138290b093bf2fe79acda9131d920cd1e7ac43fb8775776cd713c"

definition P224_SHA512_TV15_Q :: "int point" where
  "P224_SHA512_TV15_Q = Point (int P224_SHA512_TV15_Qx) (int P224_SHA512_TV15_Qy)"

definition P224_SHA512_TV15_k :: nat where
  "P224_SHA512_TV15_k = 0x557cb45fd3a30b3bdbf08c56eabbd4478736024aaa52bf8448096453"

definition P224_SHA512_TV15_R :: nat where
  "P224_SHA512_TV15_R = 0x8e92cf7a674aa5f7542dd95c695589a05747431692edd04804299b8f"

definition P224_SHA512_TV15_S :: nat where
  "P224_SHA512_TV15_S = 0xaf4908b41f8180b71a6ff10fd51f3d143147af6ddddf7534d3284ed9"

definition P224_SHA512_TV15_Sig :: "nat × nat" where
  "P224_SHA512_TV15_Sig = (P224_SHA512_TV15_R, P224_SHA512_TV15_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P224_SHA512_TV15_d_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV15_d" 
  by eval

lemma P224_SHA512_TV15_QonCurve: 
  "SEC1_P224_on_curve' P224_SHA512_TV15_Q" 
  by eval

lemma P224_SHA512_TV15_dQvalidPair': 
  "SEC1_P224_ECkeyGen P224_SHA512_TV15_d = P224_SHA512_TV15_Q" 
  by eval

lemma P224_SHA512_TV15_dQvalidPair: 
  "SEC1_P224_ECkeyPairValid P224_SHA512_TV15_d P224_SHA512_TV15_Q" 
  by (simp add: SEC1_P224.ECkeyPairEqKeyGen P224_SHA512_TV15_dQvalidPair' P224_SHA512_TV15_d_valid) 

lemma P224_SHA512_TV15_k_valid: 
  "SEC1_P224_ECprivateKeyValid P224_SHA512_TV15_k" 
  by eval

lemma P224_SHA512_TV15_Sign: 
  "SEC1_P224_ECDSA_Sign' SHA512octets P224_SHA512_TV15_d P224_SHA512_TV15_Msg P224_SHA512_TV15_k = Some P224_SHA512_TV15_Sig" 
  by eval

lemma P224_SHA512_TV15_Verify: 
  "SEC1_P224_ECDSA_Verify SHA512octets P224_SHA512_TV15_Q P224_SHA512_TV15_Msg P224_SHA512_TV15_Sig" 
  by eval

section‹NIST Test Vectors for ECDSA: Curve P-256, Hash SHA-224›
  
subsection‹Test Vector: P256_SHA224_TV01›

subsubsection‹Given Test Vector Values›

definition P256_SHA224_TV01_Msg :: octets where
  "P256_SHA224_TV01_Msg = nat_to_octets_len 0xff624d0ba02c7b6370c1622eec3fa2186ea681d1659e0a845448e777b75a8e77a77bb26e5733179d58ef9bc8a4e8b6971aef2539f77ab0963a3415bbd6258339bd1bf55de65db520c63f5b8eab3d55debd05e9494212170f5d65b3286b8b668705b1e2b2b5568610617abb51d2dd0cb450ef59df4b907da90cfa7b268de8c4c2 128"

definition P256_SHA224_TV01_d :: nat where
  "P256_SHA224_TV01_d = 0x708309a7449e156b0db70e5b52e606c7e094ed676ce8953bf6c14757c826f590"

definition P256_SHA224_TV01_Qx :: nat where
  "P256_SHA224_TV01_Qx = 0x29578c7ab6ce0d11493c95d5ea05d299d536801ca9cbd50e9924e43b733b83ab"

definition P256_SHA224_TV01_Qy :: nat where
  "P256_SHA224_TV01_Qy = 0x08c8049879c6278b2273348474158515accaa38344106ef96803c5a05adc4800"

definition P256_SHA224_TV01_Q :: "int point" where
  "P256_SHA224_TV01_Q = Point (int P256_SHA224_TV01_Qx) (int P256_SHA224_TV01_Qy)"

definition P256_SHA224_TV01_k :: nat where
  "P256_SHA224_TV01_k = 0x58f741771620bdc428e91a32d86d230873e9140336fcfb1e122892ee1d501bdc"

definition P256_SHA224_TV01_R :: nat where
  "P256_SHA224_TV01_R = 0x4a19274429e40522234b8785dc25fc524f179dcc95ff09b3c9770fc71f54ca0d"

definition P256_SHA224_TV01_S :: nat where
  "P256_SHA224_TV01_S = 0x58982b79a65b7320f5b92d13bdaecdd1259e760f0f718ba933fd098f6f75d4b7"

definition P256_SHA224_TV01_Sig :: "nat × nat" where
  "P256_SHA224_TV01_Sig = (P256_SHA224_TV01_R, P256_SHA224_TV01_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA224_TV01_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV01_d" 
  by eval

lemma P256_SHA224_TV01_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA224_TV01_Q" 
  by eval

lemma P256_SHA224_TV01_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA224_TV01_d = P256_SHA224_TV01_Q" 
  by eval

lemma P256_SHA224_TV01_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA224_TV01_d P256_SHA224_TV01_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA224_TV01_dQvalidPair' P256_SHA224_TV01_d_valid) 

lemma P256_SHA224_TV01_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV01_k" 
  by eval

lemma P256_SHA224_TV01_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA224octets P256_SHA224_TV01_d P256_SHA224_TV01_Msg P256_SHA224_TV01_k = Some P256_SHA224_TV01_Sig" 
  by eval

lemma P256_SHA224_TV01_Verify: 
  "SEC1_P256_ECDSA_Verify SHA224octets P256_SHA224_TV01_Q P256_SHA224_TV01_Msg P256_SHA224_TV01_Sig" 
  by eval

subsection‹Test Vector: P256_SHA224_TV02›

subsubsection‹Given Test Vector Values›

definition P256_SHA224_TV02_Msg :: octets where
  "P256_SHA224_TV02_Msg = nat_to_octets_len 0x9155e91fd9155eeed15afd83487ea1a3af04c5998b77c0fe8c43dcc479440a8a9a89efe883d9385cb9edfde10b43bce61fb63669935ad39419cf29ef3a936931733bfc2378e253e73b7ae9a3ec7a6a7932ab10f1e5b94d05160c053988f3bdc9167155d069337d42c9a7056619efc031fa5ec7310d29bd28980b1e3559757578 128"

definition P256_SHA224_TV02_d :: nat where
  "P256_SHA224_TV02_d = 0x90c5386100b137a75b0bb495002b28697a451add2f1f22cb65f735e8aaeace98"

definition P256_SHA224_TV02_Qx :: nat where
  "P256_SHA224_TV02_Qx = 0x4a92396ff7930b1da9a873a479a28a9896af6cc3d39345b949b726dc3cd978b5"

definition P256_SHA224_TV02_Qy :: nat where
  "P256_SHA224_TV02_Qy = 0x475abb18eaed948879b9c1453e3ef2755dd90f77519ec7b6a30297aad08e4931"

definition P256_SHA224_TV02_Q :: "int point" where
  "P256_SHA224_TV02_Q = Point (int P256_SHA224_TV02_Qx) (int P256_SHA224_TV02_Qy)"

definition P256_SHA224_TV02_k :: nat where
  "P256_SHA224_TV02_k = 0x36f853b5c54b1ec61588c9c6137eb56e7a708f09c57513093e4ecf6d739900e5"

definition P256_SHA224_TV02_R :: nat where
  "P256_SHA224_TV02_R = 0x38b29558511061cfabdc8e5bb65ac2976d1aa2ba9a5deab8074097b2172bb9ad"

definition P256_SHA224_TV02_S :: nat where
  "P256_SHA224_TV02_S = 0x0de2cde610502b6e03c0b23602eafbcd3faf886c81d111d156b7aa550f5bcd51"

definition P256_SHA224_TV02_Sig :: "nat × nat" where
  "P256_SHA224_TV02_Sig = (P256_SHA224_TV02_R, P256_SHA224_TV02_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA224_TV02_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV02_d" 
  by eval

lemma P256_SHA224_TV02_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA224_TV02_Q" 
  by eval

lemma P256_SHA224_TV02_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA224_TV02_d = P256_SHA224_TV02_Q" 
  by eval

lemma P256_SHA224_TV02_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA224_TV02_d P256_SHA224_TV02_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA224_TV02_dQvalidPair' P256_SHA224_TV02_d_valid) 

lemma P256_SHA224_TV02_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV02_k" 
  by eval

lemma P256_SHA224_TV02_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA224octets P256_SHA224_TV02_d P256_SHA224_TV02_Msg P256_SHA224_TV02_k = Some P256_SHA224_TV02_Sig" 
  by eval

lemma P256_SHA224_TV02_Verify: 
  "SEC1_P256_ECDSA_Verify SHA224octets P256_SHA224_TV02_Q P256_SHA224_TV02_Msg P256_SHA224_TV02_Sig" 
  by eval

subsection‹Test Vector: P256_SHA224_TV03›

subsubsection‹Given Test Vector Values›

definition P256_SHA224_TV03_Msg :: octets where
  "P256_SHA224_TV03_Msg = nat_to_octets_len 0xb242a7586a1383368a33c88264889adfa3be45422fbef4a2df4e3c5325a9c7757017e0d5cf4bbf4de7f99d189f81f1fd2f0dd645574d1eb0d547eead9375677819297c1abe62526ae29fc54cdd11bfe17714f2fbd2d0d0e8d297ff98535980482dd5c1ebdc5a7274aabf1382c9f2315ca61391e3943856e4c5e616c2f1f7be0d 128"

definition P256_SHA224_TV03_d :: nat where
  "P256_SHA224_TV03_d = 0xa3a43cece9c1abeff81099fb344d01f7d8df66447b95a667ee368f924bccf870"

definition P256_SHA224_TV03_Qx :: nat where
  "P256_SHA224_TV03_Qx = 0x5775174deb0248112e069cb86f1546ac7a78bc2127d0cb953bad46384dd6be5b"

definition P256_SHA224_TV03_Qy :: nat where
  "P256_SHA224_TV03_Qy = 0xa27020952971cc0b0c3abd06e9ca3e141a4943f560564eba31e5288928bc7ce7"

definition P256_SHA224_TV03_Q :: "int point" where
  "P256_SHA224_TV03_Q = Point (int P256_SHA224_TV03_Qx) (int P256_SHA224_TV03_Qy)"

definition P256_SHA224_TV03_k :: nat where
  "P256_SHA224_TV03_k = 0xa0d9a7a245bd9b9aa86cecb89341c9de2e4f9b5d095a8150826c7ba7fb3e7df7"

definition P256_SHA224_TV03_R :: nat where
  "P256_SHA224_TV03_R = 0xb02a440add66a9ff9c3c0e9acf1be678f6bd48a10cbdec2ad6d186ffe05f3f2a"

definition P256_SHA224_TV03_S :: nat where
  "P256_SHA224_TV03_S = 0xa98bea42aec56a1fcecec00a1cc69b01fcbcf5de7ac1b2f2dcc09b6db064f92b"

definition P256_SHA224_TV03_Sig :: "nat × nat" where
  "P256_SHA224_TV03_Sig = (P256_SHA224_TV03_R, P256_SHA224_TV03_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA224_TV03_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV03_d" 
  by eval

lemma P256_SHA224_TV03_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA224_TV03_Q" 
  by eval

lemma P256_SHA224_TV03_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA224_TV03_d = P256_SHA224_TV03_Q" 
  by eval

lemma P256_SHA224_TV03_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA224_TV03_d P256_SHA224_TV03_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA224_TV03_dQvalidPair' P256_SHA224_TV03_d_valid) 

lemma P256_SHA224_TV03_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV03_k" 
  by eval

lemma P256_SHA224_TV03_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA224octets P256_SHA224_TV03_d P256_SHA224_TV03_Msg P256_SHA224_TV03_k = Some P256_SHA224_TV03_Sig" 
  by eval

lemma P256_SHA224_TV03_Verify: 
  "SEC1_P256_ECDSA_Verify SHA224octets P256_SHA224_TV03_Q P256_SHA224_TV03_Msg P256_SHA224_TV03_Sig" 
  by eval

subsection‹Test Vector: P256_SHA224_TV04›

subsubsection‹Given Test Vector Values›

definition P256_SHA224_TV04_Msg :: octets where
  "P256_SHA224_TV04_Msg = nat_to_octets_len 0xb64005da76b24715880af94dba379acc25a047b06066c9bedc8f17b8c74e74f4fc720d9f4ef0e2a659e0756931c080587ebdcd0f85e819aea6dacb327a9d96496da53ea21aef3b2e793a9c0def5196acec99891f46ead78a85bc7ab644765781d3543da9fbf9fec916dca975ef3b4271e50ecc68bf79b2d8935e2b25fc063358 128"

definition P256_SHA224_TV04_d :: nat where
  "P256_SHA224_TV04_d = 0x7bbc8ff13f6f921f21e949b224c16b7176c5984d312b671cf6c2e4841135fc7f"

definition P256_SHA224_TV04_Qx :: nat where
  "P256_SHA224_TV04_Qx = 0xf888e913ec6f3cd8b31eb89e4f8aaa8887d30ae5348ed7118696949d5b8cc7c1"

definition P256_SHA224_TV04_Qy :: nat where
  "P256_SHA224_TV04_Qy = 0x08895d09620500d244e5035e262dea3f2867cd8967b226324d5c05220d8b410c"

definition P256_SHA224_TV04_Q :: "int point" where
  "P256_SHA224_TV04_Q = Point (int P256_SHA224_TV04_Qx) (int P256_SHA224_TV04_Qy)"

definition P256_SHA224_TV04_k :: nat where
  "P256_SHA224_TV04_k = 0x21c942f3b487accbf7fadc1c4b7a6c7567ce876c195022459fa1ebf6d04ffbaa"

definition P256_SHA224_TV04_R :: nat where
  "P256_SHA224_TV04_R = 0x2e6cc883b8acc904ee9691ef4a9f1f5a9e5fbfde847cda3be833f949fb9c7182"

definition P256_SHA224_TV04_S :: nat where
  "P256_SHA224_TV04_S = 0x2ac48f7a930912131a8b4e3ab495307817c465d638c2a9ea5ae9e2808806e20a"

definition P256_SHA224_TV04_Sig :: "nat × nat" where
  "P256_SHA224_TV04_Sig = (P256_SHA224_TV04_R, P256_SHA224_TV04_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA224_TV04_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV04_d" 
  by eval

lemma P256_SHA224_TV04_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA224_TV04_Q" 
  by eval

lemma P256_SHA224_TV04_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA224_TV04_d = P256_SHA224_TV04_Q" 
  by eval

lemma P256_SHA224_TV04_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA224_TV04_d P256_SHA224_TV04_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA224_TV04_dQvalidPair' P256_SHA224_TV04_d_valid) 

lemma P256_SHA224_TV04_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV04_k" 
  by eval

lemma P256_SHA224_TV04_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA224octets P256_SHA224_TV04_d P256_SHA224_TV04_Msg P256_SHA224_TV04_k = Some P256_SHA224_TV04_Sig" 
  by eval

lemma P256_SHA224_TV04_Verify: 
  "SEC1_P256_ECDSA_Verify SHA224octets P256_SHA224_TV04_Q P256_SHA224_TV04_Msg P256_SHA224_TV04_Sig" 
  by eval

subsection‹Test Vector: P256_SHA224_TV05›

subsubsection‹Given Test Vector Values›

definition P256_SHA224_TV05_Msg :: octets where
  "P256_SHA224_TV05_Msg = nat_to_octets_len 0xfe6e1ea477640655eaa1f6e3352d4bce53eb3d95424df7f238e93d8531da8f36bc35fa6be4bf5a6a382e06e855139eb617a9cc9376b4dafacbd80876343b12628619d7cbe1bff6757e3706111ed53898c0219823adbc044eaf8c6ad449df8f6aab9d444dadb5c3380eec0d91694df5fc4b30280d4b87d27e67ae58a1df828963 128"

definition P256_SHA224_TV05_d :: nat where
  "P256_SHA224_TV05_d = 0xdaf5ec7a4eebc20d9485796c355b4a65ad254fe19b998d0507e91ea24135f45d"

definition P256_SHA224_TV05_Qx :: nat where
  "P256_SHA224_TV05_Qx = 0x137c465085c1b1b8cccbe9fccbe9d0295a331aaf332f3ed2e285d16e574b943b"

definition P256_SHA224_TV05_Qy :: nat where
  "P256_SHA224_TV05_Qy = 0xd3e8d5a24cd218c19760b0e85b35a8569945aa857cbf0fd6a3ce127581b217b6"

definition P256_SHA224_TV05_Q :: "int point" where
  "P256_SHA224_TV05_Q = Point (int P256_SHA224_TV05_Qx) (int P256_SHA224_TV05_Qy)"

definition P256_SHA224_TV05_k :: nat where
  "P256_SHA224_TV05_k = 0x343251dffa56e6a612fec7b078f9c3819eab402a72686b894a47a08fd97e6c23"

definition P256_SHA224_TV05_R :: nat where
  "P256_SHA224_TV05_R = 0x775e25a296bd259510ae9375f548997bec8a744900022945281dc8c4d94f2b5b"

definition P256_SHA224_TV05_S :: nat where
  "P256_SHA224_TV05_S = 0xd87592ceab773ae103daebbb56a04144aaccb1e14efc1024dc36c0e382df1f70"

definition P256_SHA224_TV05_Sig :: "nat × nat" where
  "P256_SHA224_TV05_Sig = (P256_SHA224_TV05_R, P256_SHA224_TV05_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA224_TV05_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV05_d" 
  by eval

lemma P256_SHA224_TV05_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA224_TV05_Q" 
  by eval

lemma P256_SHA224_TV05_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA224_TV05_d = P256_SHA224_TV05_Q" 
  by eval

lemma P256_SHA224_TV05_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA224_TV05_d P256_SHA224_TV05_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA224_TV05_dQvalidPair' P256_SHA224_TV05_d_valid) 

lemma P256_SHA224_TV05_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV05_k" 
  by eval

lemma P256_SHA224_TV05_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA224octets P256_SHA224_TV05_d P256_SHA224_TV05_Msg P256_SHA224_TV05_k = Some P256_SHA224_TV05_Sig" 
  by eval

lemma P256_SHA224_TV05_Verify: 
  "SEC1_P256_ECDSA_Verify SHA224octets P256_SHA224_TV05_Q P256_SHA224_TV05_Msg P256_SHA224_TV05_Sig" 
  by eval

subsection‹Test Vector: P256_SHA224_TV06›

subsubsection‹Given Test Vector Values›

definition P256_SHA224_TV06_Msg :: octets where
  "P256_SHA224_TV06_Msg = nat_to_octets_len 0x907c0c00dc080a688548957b5b8b1f33ba378de1368023dcad43242411f554eb7d392d3e5c1668fad3944ff9634105343d83b8c85d2a988da5f5dc60ee0518327caed6dd5cf4e9bc6222deb46d00abde745f9b71d6e7aee6c7fdfc9ed053f2c0b611d4c6863088bd012ea9810ee94f8e58905970ebd07353f1f409a371ed03e3 128"

definition P256_SHA224_TV06_d :: nat where
  "P256_SHA224_TV06_d = 0x8729a8396f262dabd991aa404cc1753581cea405f0d19222a0b3f210de8ee3c5"

definition P256_SHA224_TV06_Qx :: nat where
  "P256_SHA224_TV06_Qx = 0x82b1f1a7af9b48ca8452613d7032beb0e4f28fe710306aeccc959e4d03662a35"

definition P256_SHA224_TV06_Qy :: nat where
  "P256_SHA224_TV06_Qy = 0x5e39f33574097b8d32b471a591972496f5d44db344c037d13f06fafc75f016fd"

definition P256_SHA224_TV06_Q :: "int point" where
  "P256_SHA224_TV06_Q = Point (int P256_SHA224_TV06_Qx) (int P256_SHA224_TV06_Qy)"

definition P256_SHA224_TV06_k :: nat where
  "P256_SHA224_TV06_k = 0x6de9e21f0b2cacc1762b3558fd44d3cf156b85dbef430dd28d59713bfb9cfa0b"

definition P256_SHA224_TV06_R :: nat where
  "P256_SHA224_TV06_R = 0xa754b42720e71925d51fcef76151405a3696cc8f9fc9ca7b46d0b16edd7fb699"

definition P256_SHA224_TV06_S :: nat where
  "P256_SHA224_TV06_S = 0x603924780439cc16ac4cf97c2c3065bc95353aa9179d0ab5f0322ca82f851cf2"

definition P256_SHA224_TV06_Sig :: "nat × nat" where
  "P256_SHA224_TV06_Sig = (P256_SHA224_TV06_R, P256_SHA224_TV06_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA224_TV06_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV06_d" 
  by eval

lemma P256_SHA224_TV06_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA224_TV06_Q" 
  by eval

lemma P256_SHA224_TV06_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA224_TV06_d = P256_SHA224_TV06_Q" 
  by eval

lemma P256_SHA224_TV06_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA224_TV06_d P256_SHA224_TV06_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA224_TV06_dQvalidPair' P256_SHA224_TV06_d_valid) 

lemma P256_SHA224_TV06_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV06_k" 
  by eval

lemma P256_SHA224_TV06_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA224octets P256_SHA224_TV06_d P256_SHA224_TV06_Msg P256_SHA224_TV06_k = Some P256_SHA224_TV06_Sig" 
  by eval

lemma P256_SHA224_TV06_Verify: 
  "SEC1_P256_ECDSA_Verify SHA224octets P256_SHA224_TV06_Q P256_SHA224_TV06_Msg P256_SHA224_TV06_Sig" 
  by eval

subsection‹Test Vector: P256_SHA224_TV07›

subsubsection‹Given Test Vector Values›

definition P256_SHA224_TV07_Msg :: octets where
  "P256_SHA224_TV07_Msg = nat_to_octets_len 0x771c4d7bce05610a3e71b272096b57f0d1efcce33a1cb4f714d6ebc0865b2773ec5eedc25fae81dee1d256474dbd9676623614c150916e6ed92ce4430b26037d28fa5252ef6b10c09dc2f7ee5a36a1ea7897b69f389d9f5075e271d92f4eb97b148f3abcb1e5be0b4feb8278613d18abf6da60bfe448238aa04d7f11b71f44c5 128"

definition P256_SHA224_TV07_d :: nat where
  "P256_SHA224_TV07_d = 0xf1b62413935fc589ad2280f6892599ad994dae8ca3655ed4f7318cc89b61aa96"

definition P256_SHA224_TV07_Qx :: nat where
  "P256_SHA224_TV07_Qx = 0xe0bbfe4016eea93e6f509518cbffc25d492de6ebbf80465a461caa5bdc018159"

definition P256_SHA224_TV07_Qy :: nat where
  "P256_SHA224_TV07_Qy = 0x3231ee7a119d84fa56e3034d50fea85929aec2eb437abc7646821e1bf805fb50"

definition P256_SHA224_TV07_Q :: "int point" where
  "P256_SHA224_TV07_Q = Point (int P256_SHA224_TV07_Qx) (int P256_SHA224_TV07_Qy)"

definition P256_SHA224_TV07_k :: nat where
  "P256_SHA224_TV07_k = 0x7a33eeb9f469afd55de2fb786847a1d3e7797929305c0f90d953b6f143bb8fc6"

definition P256_SHA224_TV07_R :: nat where
  "P256_SHA224_TV07_R = 0x96d1c9399948254ea381631fc0f43ea808110506db8aacf081df5535ac5eb8ad"

definition P256_SHA224_TV07_S :: nat where
  "P256_SHA224_TV07_S = 0x73bf3691260dddd9997c97313f2a70783eacf8d15bdfb34bb13025cdfae72f70"

definition P256_SHA224_TV07_Sig :: "nat × nat" where
  "P256_SHA224_TV07_Sig = (P256_SHA224_TV07_R, P256_SHA224_TV07_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA224_TV07_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV07_d" 
  by eval

lemma P256_SHA224_TV07_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA224_TV07_Q" 
  by eval

lemma P256_SHA224_TV07_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA224_TV07_d = P256_SHA224_TV07_Q" 
  by eval

lemma P256_SHA224_TV07_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA224_TV07_d P256_SHA224_TV07_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA224_TV07_dQvalidPair' P256_SHA224_TV07_d_valid) 

lemma P256_SHA224_TV07_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV07_k" 
  by eval

lemma P256_SHA224_TV07_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA224octets P256_SHA224_TV07_d P256_SHA224_TV07_Msg P256_SHA224_TV07_k = Some P256_SHA224_TV07_Sig" 
  by eval

lemma P256_SHA224_TV07_Verify: 
  "SEC1_P256_ECDSA_Verify SHA224octets P256_SHA224_TV07_Q P256_SHA224_TV07_Msg P256_SHA224_TV07_Sig" 
  by eval

subsection‹Test Vector: P256_SHA224_TV08›

subsubsection‹Given Test Vector Values›

definition P256_SHA224_TV08_Msg :: octets where
  "P256_SHA224_TV08_Msg = nat_to_octets_len 0xa3b2825235718fc679b942e8ac38fb4f54415a213c65875b5453d18ca012320ddfbbc58b991eaebadfc2d1a28d4f0cd82652b12e4d5bfda89eda3be12ac52188e38e8cce32a264a300c0e463631f525ae501348594f980392c76b4a12ddc88e5ca086cb8685d03895919a8627725a3e00c4728e2b7c6f6a14fc342b2937fc3dd 128"

definition P256_SHA224_TV08_d :: nat where
  "P256_SHA224_TV08_d = 0x4caaa26f93f009682bbba6db6b265aec17b7ec1542bda458e8550b9e68eed18d"

definition P256_SHA224_TV08_Qx :: nat where
  "P256_SHA224_TV08_Qx = 0xe3c58c1c254d11c7e781ad133e4c36dd1b5de362120d336a58e7b68813f3fbee"

definition P256_SHA224_TV08_Qy :: nat where
  "P256_SHA224_TV08_Qy = 0x59760db66120afe0d962c81a8e5586588fd19de2f40556371611c73af22c8a68"

definition P256_SHA224_TV08_Q :: "int point" where
  "P256_SHA224_TV08_Q = Point (int P256_SHA224_TV08_Qx) (int P256_SHA224_TV08_Qy)"

definition P256_SHA224_TV08_k :: nat where
  "P256_SHA224_TV08_k = 0xc0d37142dc8b0d614fad20c4d35af6eb819e259e513ddeac1e1c273e7e1dc1bb"

definition P256_SHA224_TV08_R :: nat where
  "P256_SHA224_TV08_R = 0x25dd8e4086c62a40d2a310e2f90f6af5cb7e677b4dfdb4dc4e99e23ea2f0e6dc"

definition P256_SHA224_TV08_S :: nat where
  "P256_SHA224_TV08_S = 0x90ad62c179b0c9d61f521dde1cd762bfd224b5525c39c3706f2549313ddb4f39"

definition P256_SHA224_TV08_Sig :: "nat × nat" where
  "P256_SHA224_TV08_Sig = (P256_SHA224_TV08_R, P256_SHA224_TV08_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA224_TV08_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV08_d" 
  by eval

lemma P256_SHA224_TV08_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA224_TV08_Q" 
  by eval

lemma P256_SHA224_TV08_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA224_TV08_d = P256_SHA224_TV08_Q" 
  by eval

lemma P256_SHA224_TV08_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA224_TV08_d P256_SHA224_TV08_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA224_TV08_dQvalidPair' P256_SHA224_TV08_d_valid) 

lemma P256_SHA224_TV08_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV08_k" 
  by eval

lemma P256_SHA224_TV08_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA224octets P256_SHA224_TV08_d P256_SHA224_TV08_Msg P256_SHA224_TV08_k = Some P256_SHA224_TV08_Sig" 
  by eval

lemma P256_SHA224_TV08_Verify: 
  "SEC1_P256_ECDSA_Verify SHA224octets P256_SHA224_TV08_Q P256_SHA224_TV08_Msg P256_SHA224_TV08_Sig" 
  by eval

subsection‹Test Vector: P256_SHA224_TV09›

subsubsection‹Given Test Vector Values›

definition P256_SHA224_TV09_Msg :: octets where
  "P256_SHA224_TV09_Msg = nat_to_octets_len 0x3e6e2a9bffd729ee5d4807849cd4250021d8184cda723df6ab0e5c939d39237c8e58af9d869fe62d3c97b3298a99e891e5e11aa68b11a087573a40a3e83c7965e7910d72f81cad0f42accc5c25a4fd3cdd8cee63757bbbfbdae98be2bc867d3bcb1333c4632cb0a55dffeb77d8b119c466cd889ec468454fabe6fbee7102deaf 128"

definition P256_SHA224_TV09_d :: nat where
  "P256_SHA224_TV09_d = 0x7af4b150bb7167cb68037f280d0823ce5320c01a92b1b56ee1b88547481b1de9"

definition P256_SHA224_TV09_Qx :: nat where
  "P256_SHA224_TV09_Qx = 0xcb3634ec4f0cbb99986be788f889e586026d5a851e80d15382f1bdb1bda2bc75"

definition P256_SHA224_TV09_Qy :: nat where
  "P256_SHA224_TV09_Qy = 0x51e4e43bc16fb114896b18198a1aebe6054ba20ed0c0317c1b8776158c0e6bfb"

definition P256_SHA224_TV09_Q :: "int point" where
  "P256_SHA224_TV09_Q = Point (int P256_SHA224_TV09_Qx) (int P256_SHA224_TV09_Qy)"

definition P256_SHA224_TV09_k :: nat where
  "P256_SHA224_TV09_k = 0x98edd59fafbcaee5f64e84eb5ed59fff45d14aabada47cee2fa674377173627a"

definition P256_SHA224_TV09_R :: nat where
  "P256_SHA224_TV09_R = 0x261a1cdb0fd93c0fb06ea6068b6b03c330a12f621a7eba76682a1d152c0e8d08"

definition P256_SHA224_TV09_S :: nat where
  "P256_SHA224_TV09_S = 0x7ca049bad54feee101d6db807635ffb8bdb05a38e445c8c3d65d60df143514c5"

definition P256_SHA224_TV09_Sig :: "nat × nat" where
  "P256_SHA224_TV09_Sig = (P256_SHA224_TV09_R, P256_SHA224_TV09_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA224_TV09_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV09_d" 
  by eval

lemma P256_SHA224_TV09_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA224_TV09_Q" 
  by eval

lemma P256_SHA224_TV09_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA224_TV09_d = P256_SHA224_TV09_Q" 
  by eval

lemma P256_SHA224_TV09_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA224_TV09_d P256_SHA224_TV09_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA224_TV09_dQvalidPair' P256_SHA224_TV09_d_valid) 

lemma P256_SHA224_TV09_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV09_k" 
  by eval

lemma P256_SHA224_TV09_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA224octets P256_SHA224_TV09_d P256_SHA224_TV09_Msg P256_SHA224_TV09_k = Some P256_SHA224_TV09_Sig" 
  by eval

lemma P256_SHA224_TV09_Verify: 
  "SEC1_P256_ECDSA_Verify SHA224octets P256_SHA224_TV09_Q P256_SHA224_TV09_Msg P256_SHA224_TV09_Sig" 
  by eval

subsection‹Test Vector: P256_SHA224_TV10›

subsubsection‹Given Test Vector Values›

definition P256_SHA224_TV10_Msg :: octets where
  "P256_SHA224_TV10_Msg = nat_to_octets_len 0x52e5c308e70329a17c71eaedb66bbee303c8ec48a6f1a2efb235d308563cd58553d434e12f353227a9ea28608ec9c820ed83c95124e7a886f7e832a2de1032e78dc059208f9ec354170b2b1cab992b52ac01e6c0e4e1b0112686962edc53ab226dafcc9fc7baed2cd9307160e8572edb125935db49289b178f35a8ad23f4f801 128"

definition P256_SHA224_TV10_d :: nat where
  "P256_SHA224_TV10_d = 0x52ad53e849e30bec0e6345c3e9d98ebc808b19496c1ef16d72ab4a00bbb8c634"

definition P256_SHA224_TV10_Qx :: nat where
  "P256_SHA224_TV10_Qx = 0x7cca1334bfc2a78728c50b370399be3f9690d445aa03c701da643eeb0b0f7fa8"

definition P256_SHA224_TV10_Qy :: nat where
  "P256_SHA224_TV10_Qy = 0x3f7522238668e615405e49b2f63faee58286000a30cdb4b564ac0df99bc8950f"

definition P256_SHA224_TV10_Q :: "int point" where
  "P256_SHA224_TV10_Q = Point (int P256_SHA224_TV10_Qx) (int P256_SHA224_TV10_Qy)"

definition P256_SHA224_TV10_k :: nat where
  "P256_SHA224_TV10_k = 0x8650c30712fc253610884fbba4a332a4574d4b7822f7776cab1df8f5fa05442a"

definition P256_SHA224_TV10_R :: nat where
  "P256_SHA224_TV10_R = 0xa18194c7ac5829afc408d78dde19542837e7be82706c3941b2d9c5e036bb51e0"

definition P256_SHA224_TV10_S :: nat where
  "P256_SHA224_TV10_S = 0x188ead1cdf7c1d21114ff56d0421ffd501ab978ef58337462c0fa736d86299af"

definition P256_SHA224_TV10_Sig :: "nat × nat" where
  "P256_SHA224_TV10_Sig = (P256_SHA224_TV10_R, P256_SHA224_TV10_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA224_TV10_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV10_d" 
  by eval

lemma P256_SHA224_TV10_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA224_TV10_Q" 
  by eval

lemma P256_SHA224_TV10_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA224_TV10_d = P256_SHA224_TV10_Q" 
  by eval

lemma P256_SHA224_TV10_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA224_TV10_d P256_SHA224_TV10_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA224_TV10_dQvalidPair' P256_SHA224_TV10_d_valid) 

lemma P256_SHA224_TV10_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV10_k" 
  by eval

lemma P256_SHA224_TV10_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA224octets P256_SHA224_TV10_d P256_SHA224_TV10_Msg P256_SHA224_TV10_k = Some P256_SHA224_TV10_Sig" 
  by eval

lemma P256_SHA224_TV10_Verify: 
  "SEC1_P256_ECDSA_Verify SHA224octets P256_SHA224_TV10_Q P256_SHA224_TV10_Msg P256_SHA224_TV10_Sig" 
  by eval

subsection‹Test Vector: P256_SHA224_TV11›

subsubsection‹Given Test Vector Values›

definition P256_SHA224_TV11_Msg :: octets where
  "P256_SHA224_TV11_Msg = nat_to_octets_len 0xd3e9e82051d4c84d699453c9ff44c7c09f6523bb92232bcf30bf3c380224249de2964e871d56a364d6955c81ef91d06482a6c7c61bc70f66ef22fad128d15416e7174312619134f968f1009f92cbf99248932efb533ff113fb6d949e21d6b80dfbbe69010c8d1ccb0f3808ea309bb0bac1a222168c95b088847e613749b19d04 128"

definition P256_SHA224_TV11_d :: nat where
  "P256_SHA224_TV11_d = 0x80754962a864be1803bc441fa331e126005bfc6d8b09ed38b7e69d9a030a5d27"

definition P256_SHA224_TV11_Qx :: nat where
  "P256_SHA224_TV11_Qx = 0x0aaeed6dd1ae020d6eefc98ec4241ac93cbd3c8afed05bb28007e7da5727571b"

definition P256_SHA224_TV11_Qy :: nat where
  "P256_SHA224_TV11_Qy = 0x2dda1d5b7872eb94dfffb456115037ff8d3e72f8ebdd8fcfc42391f96809be69"

definition P256_SHA224_TV11_Q :: "int point" where
  "P256_SHA224_TV11_Q = Point (int P256_SHA224_TV11_Qx) (int P256_SHA224_TV11_Qy)"

definition P256_SHA224_TV11_k :: nat where
  "P256_SHA224_TV11_k = 0x738e050aeefe54ecba5be5f93a97bbcb7557d701f9da2d7e88483454b97b55a8"

definition P256_SHA224_TV11_R :: nat where
  "P256_SHA224_TV11_R = 0x8cb9f41dfdcb9604e0725ac9b78fc0db916dc071186ee982f6dba3da36f02efa"

definition P256_SHA224_TV11_S :: nat where
  "P256_SHA224_TV11_S = 0x5c87fe868fd4282fb114f5d70e9590a10a5d35cedf3ff6402ba5c4344738a32e"

definition P256_SHA224_TV11_Sig :: "nat × nat" where
  "P256_SHA224_TV11_Sig = (P256_SHA224_TV11_R, P256_SHA224_TV11_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA224_TV11_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV11_d" 
  by eval

lemma P256_SHA224_TV11_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA224_TV11_Q" 
  by eval

lemma P256_SHA224_TV11_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA224_TV11_d = P256_SHA224_TV11_Q" 
  by eval

lemma P256_SHA224_TV11_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA224_TV11_d P256_SHA224_TV11_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA224_TV11_dQvalidPair' P256_SHA224_TV11_d_valid) 

lemma P256_SHA224_TV11_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV11_k" 
  by eval

lemma P256_SHA224_TV11_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA224octets P256_SHA224_TV11_d P256_SHA224_TV11_Msg P256_SHA224_TV11_k = Some P256_SHA224_TV11_Sig" 
  by eval

lemma P256_SHA224_TV11_Verify: 
  "SEC1_P256_ECDSA_Verify SHA224octets P256_SHA224_TV11_Q P256_SHA224_TV11_Msg P256_SHA224_TV11_Sig" 
  by eval

subsection‹Test Vector: P256_SHA224_TV12›

subsubsection‹Given Test Vector Values›

definition P256_SHA224_TV12_Msg :: octets where
  "P256_SHA224_TV12_Msg = nat_to_octets_len 0x968951c2c1918436fe19fa2fe2152656a08f9a6b8aa6201920f1b424da98cee71928897ff087620cc5c551320b1e75a1e98d7d98a5bd5361c9393759614a6087cc0f7fb01fcb173783eb4c4c23961a8231ac4a07d72e683b0c1bd4c51ef1b031df875e7b8d5a6e0628949f5b8f157f43dccaea3b2a4fc11181e6b451e06ceb37 128"

definition P256_SHA224_TV12_d :: nat where
  "P256_SHA224_TV12_d = 0xcfa8c8bd810eb0d73585f36280ecdd296ee098511be8ad5eac68984eca8eb19d"

definition P256_SHA224_TV12_Qx :: nat where
  "P256_SHA224_TV12_Qx = 0xc227a2af15dfa8734e11c0c50f77e24e77ed58dd8cccf1b0e9fa06bee1c64766"

definition P256_SHA224_TV12_Qy :: nat where
  "P256_SHA224_TV12_Qy = 0xb686592ce3745eb300d2704083db55e1fa8274e4cb7e256889ccc0bb34a60570"

definition P256_SHA224_TV12_Q :: "int point" where
  "P256_SHA224_TV12_Q = Point (int P256_SHA224_TV12_Qx) (int P256_SHA224_TV12_Qy)"

definition P256_SHA224_TV12_k :: nat where
  "P256_SHA224_TV12_k = 0x2d6b449bb38b543d6b6d34ff8cb053f5e5b337f949b069b21f421995ebb28823"

definition P256_SHA224_TV12_R :: nat where
  "P256_SHA224_TV12_R = 0x5e89d3c9b103c2fa3cb8cebeec23640acda0257d63ffbe2d509bfc49fab1dca6"

definition P256_SHA224_TV12_S :: nat where
  "P256_SHA224_TV12_S = 0xd70c5b1eeb29e016af9925798d24e166c23d58fedd2f1a3bbdb1ef78cdbfb63a"

definition P256_SHA224_TV12_Sig :: "nat × nat" where
  "P256_SHA224_TV12_Sig = (P256_SHA224_TV12_R, P256_SHA224_TV12_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA224_TV12_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV12_d" 
  by eval

lemma P256_SHA224_TV12_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA224_TV12_Q" 
  by eval

lemma P256_SHA224_TV12_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA224_TV12_d = P256_SHA224_TV12_Q" 
  by eval

lemma P256_SHA224_TV12_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA224_TV12_d P256_SHA224_TV12_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA224_TV12_dQvalidPair' P256_SHA224_TV12_d_valid) 

lemma P256_SHA224_TV12_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV12_k" 
  by eval

lemma P256_SHA224_TV12_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA224octets P256_SHA224_TV12_d P256_SHA224_TV12_Msg P256_SHA224_TV12_k = Some P256_SHA224_TV12_Sig" 
  by eval

lemma P256_SHA224_TV12_Verify: 
  "SEC1_P256_ECDSA_Verify SHA224octets P256_SHA224_TV12_Q P256_SHA224_TV12_Msg P256_SHA224_TV12_Sig" 
  by eval

subsection‹Test Vector: P256_SHA224_TV13›

subsubsection‹Given Test Vector Values›

definition P256_SHA224_TV13_Msg :: octets where
  "P256_SHA224_TV13_Msg = nat_to_octets_len 0x78048628932e1c1cdd1e70932bd7b76f704ba08d7e7d825d3de763bf1a062315f4af16eccefe0b6ebadccaf403d013f50833ce2c54e24eea8345e25f93b69bb048988d102240225ceacf5003e2abdcc90299f4bf2c101585d36ecdd7a155953c674789d070480d1ef47cc7858e97a6d87c41c6922a00ea12539f251826e141b4 128"

definition P256_SHA224_TV13_d :: nat where
  "P256_SHA224_TV13_d = 0xb2021e2665ce543b7feadd0cd5a4bd57ffcc5b32deb860b4d736d9880855da3c"

definition P256_SHA224_TV13_Qx :: nat where
  "P256_SHA224_TV13_Qx = 0x722e0abad4504b7832a148746153777694714eca220eced2b2156ca64cfed3dd"

definition P256_SHA224_TV13_Qy :: nat where
  "P256_SHA224_TV13_Qy = 0xf0351b357b3081e859c46cad5328c5afa10546e92bc6c3fd541796ac30397a75"

definition P256_SHA224_TV13_Q :: "int point" where
  "P256_SHA224_TV13_Q = Point (int P256_SHA224_TV13_Qx) (int P256_SHA224_TV13_Qy)"

definition P256_SHA224_TV13_k :: nat where
  "P256_SHA224_TV13_k = 0xb15bbce4b382145de7ecd670d947e77555ef7cd1693bd53c694e2b52b04d10e1"

definition P256_SHA224_TV13_R :: nat where
  "P256_SHA224_TV13_R = 0x9d086dcd22da165a43091991bede9c1c14515e656633cb759ec2c17f51c35253"

definition P256_SHA224_TV13_S :: nat where
  "P256_SHA224_TV13_S = 0x23595ad1cb714559faaecaf946beb9a71e584616030ceaed8a8470f4bf62768f"

definition P256_SHA224_TV13_Sig :: "nat × nat" where
  "P256_SHA224_TV13_Sig = (P256_SHA224_TV13_R, P256_SHA224_TV13_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA224_TV13_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV13_d" 
  by eval

lemma P256_SHA224_TV13_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA224_TV13_Q" 
  by eval

lemma P256_SHA224_TV13_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA224_TV13_d = P256_SHA224_TV13_Q" 
  by eval

lemma P256_SHA224_TV13_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA224_TV13_d P256_SHA224_TV13_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA224_TV13_dQvalidPair' P256_SHA224_TV13_d_valid) 

lemma P256_SHA224_TV13_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV13_k" 
  by eval

lemma P256_SHA224_TV13_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA224octets P256_SHA224_TV13_d P256_SHA224_TV13_Msg P256_SHA224_TV13_k = Some P256_SHA224_TV13_Sig" 
  by eval

lemma P256_SHA224_TV13_Verify: 
  "SEC1_P256_ECDSA_Verify SHA224octets P256_SHA224_TV13_Q P256_SHA224_TV13_Msg P256_SHA224_TV13_Sig" 
  by eval

subsection‹Test Vector: P256_SHA224_TV14›

subsubsection‹Given Test Vector Values›

definition P256_SHA224_TV14_Msg :: octets where
  "P256_SHA224_TV14_Msg = nat_to_octets_len 0x9b0800c443e693067591737fdbcf0966fdfa50872d41d0c189d87cbc34c2771ee5e1255fd604f09fcf167fda16437c245d299147299c69046895d22482db29aba37ff57f756716cd3d6223077f747c4caffbecc0a7c9dfaaafd9a9817470ded8777e6355838ac54d11b2f0fc3f43668ff949cc31de0c2d15af5ef17884e4d66a 128"

definition P256_SHA224_TV14_d :: nat where
  "P256_SHA224_TV14_d = 0x0c9bce6a568ca239395fc3552755575cbcdddb1d89f6f5ab354517a057b17b48"

definition P256_SHA224_TV14_Qx :: nat where
  "P256_SHA224_TV14_Qx = 0x4814d454495df7103e2da383aba55f7842fd84f1750ee5801ad32c10d0be6c7d"

definition P256_SHA224_TV14_Qy :: nat where
  "P256_SHA224_TV14_Qy = 0xa0bd039d5097c8f0770477f6b18d247876e88e528bf0453eab515ffab8a9eda3"

definition P256_SHA224_TV14_Q :: "int point" where
  "P256_SHA224_TV14_Q = Point (int P256_SHA224_TV14_Qx) (int P256_SHA224_TV14_Qy)"

definition P256_SHA224_TV14_k :: nat where
  "P256_SHA224_TV14_k = 0xd414f1525cdcc41eba1652de017c034ebcc7946cb2efe4713d09f67c85b83153"

definition P256_SHA224_TV14_R :: nat where
  "P256_SHA224_TV14_R = 0x84db02c678f9a21208cec8564d145a35ba8c6f26b4eb7e19522e439720dae44c"

definition P256_SHA224_TV14_S :: nat where
  "P256_SHA224_TV14_S = 0x537c564da0d2dc5ac4376c5f0ca3b628d01d48df47a83d842c927e4d6db1e16d"

definition P256_SHA224_TV14_Sig :: "nat × nat" where
  "P256_SHA224_TV14_Sig = (P256_SHA224_TV14_R, P256_SHA224_TV14_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA224_TV14_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV14_d" 
  by eval

lemma P256_SHA224_TV14_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA224_TV14_Q" 
  by eval

lemma P256_SHA224_TV14_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA224_TV14_d = P256_SHA224_TV14_Q" 
  by eval

lemma P256_SHA224_TV14_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA224_TV14_d P256_SHA224_TV14_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA224_TV14_dQvalidPair' P256_SHA224_TV14_d_valid) 

lemma P256_SHA224_TV14_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV14_k" 
  by eval

lemma P256_SHA224_TV14_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA224octets P256_SHA224_TV14_d P256_SHA224_TV14_Msg P256_SHA224_TV14_k = Some P256_SHA224_TV14_Sig" 
  by eval

lemma P256_SHA224_TV14_Verify: 
  "SEC1_P256_ECDSA_Verify SHA224octets P256_SHA224_TV14_Q P256_SHA224_TV14_Msg P256_SHA224_TV14_Sig" 
  by eval

subsection‹Test Vector: P256_SHA224_TV15›

subsubsection‹Given Test Vector Values›

definition P256_SHA224_TV15_Msg :: octets where
  "P256_SHA224_TV15_Msg = nat_to_octets_len 0xfc3b8291c172dae635a6859f525beaf01cf683765d7c86f1a4d768df7cae055f639eccc08d7a0272394d949f82d5e12d69c08e2483e11a1d28a4c61f18193106e12e5de4a9d0b4bf341e2acd6b715dc83ae5ff63328f8346f35521ca378b311299947f63ec593a5e32e6bd11ec4edb0e75302a9f54d21226d23314729e061016 128"

definition P256_SHA224_TV15_d :: nat where
  "P256_SHA224_TV15_d = 0x1daa385ec7c7f8a09adfcaea42801a4de4c889fb5c6eb4e92bc611d596d68e3f"

definition P256_SHA224_TV15_Qx :: nat where
  "P256_SHA224_TV15_Qx = 0xf04e9f2831d9697ae146c7d4552e5f91085cc46778400b75b76f00205252941d"

definition P256_SHA224_TV15_Qy :: nat where
  "P256_SHA224_TV15_Qy = 0xbd267148174cd0c2b019cd0a5256e2f3f889d1e597160372b5a1339c8d787f10"

definition P256_SHA224_TV15_Q :: "int point" where
  "P256_SHA224_TV15_Q = Point (int P256_SHA224_TV15_Qx) (int P256_SHA224_TV15_Qy)"

definition P256_SHA224_TV15_k :: nat where
  "P256_SHA224_TV15_k = 0x7707db348ee6f60365b43a2a994e9b40ed56fe03c2c31c7e781bc4ffadcba760"

definition P256_SHA224_TV15_R :: nat where
  "P256_SHA224_TV15_R = 0x5d95c385eeba0f15db0b80ae151912409128c9c80e554246067b8f6a36d85ea5"

definition P256_SHA224_TV15_S :: nat where
  "P256_SHA224_TV15_S = 0xdb5d8a1e345f883e4fcb3871276f170b783c1a1e9da6b6615913368a8526f1c3"

definition P256_SHA224_TV15_Sig :: "nat × nat" where
  "P256_SHA224_TV15_Sig = (P256_SHA224_TV15_R, P256_SHA224_TV15_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA224_TV15_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV15_d" 
  by eval

lemma P256_SHA224_TV15_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA224_TV15_Q" 
  by eval

lemma P256_SHA224_TV15_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA224_TV15_d = P256_SHA224_TV15_Q" 
  by eval

lemma P256_SHA224_TV15_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA224_TV15_d P256_SHA224_TV15_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA224_TV15_dQvalidPair' P256_SHA224_TV15_d_valid) 

lemma P256_SHA224_TV15_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA224_TV15_k" 
  by eval

lemma P256_SHA224_TV15_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA224octets P256_SHA224_TV15_d P256_SHA224_TV15_Msg P256_SHA224_TV15_k = Some P256_SHA224_TV15_Sig" 
  by eval

lemma P256_SHA224_TV15_Verify: 
  "SEC1_P256_ECDSA_Verify SHA224octets P256_SHA224_TV15_Q P256_SHA224_TV15_Msg P256_SHA224_TV15_Sig" 
  by eval

section‹NIST Test Vectors for ECDSA: Curve P-256, Hash SHA-256›
  
subsection‹Test Vector: P256_SHA256_TV01›

subsubsection‹Given Test Vector Values›

definition P256_SHA256_TV01_Msg :: octets where
  "P256_SHA256_TV01_Msg = nat_to_octets_len 0x5905238877c77421f73e43ee3da6f2d9e2ccad5fc942dcec0cbd25482935faaf416983fe165b1a045ee2bcd2e6dca3bdf46c4310a7461f9a37960ca672d3feb5473e253605fb1ddfd28065b53cb5858a8ad28175bf9bd386a5e471ea7a65c17cc934a9d791e91491eb3754d03799790fe2d308d16146d5c9b0d0debd97d79ce8 128"

definition P256_SHA256_TV01_d :: nat where
  "P256_SHA256_TV01_d = 0x519b423d715f8b581f4fa8ee59f4771a5b44c8130b4e3eacca54a56dda72b464"

definition P256_SHA256_TV01_Qx :: nat where
  "P256_SHA256_TV01_Qx = 0x1ccbe91c075fc7f4f033bfa248db8fccd3565de94bbfb12f3c59ff46c271bf83"

definition P256_SHA256_TV01_Qy :: nat where
  "P256_SHA256_TV01_Qy = 0xce4014c68811f9a21a1fdb2c0e6113e06db7ca93b7404e78dc7ccd5ca89a4ca9"

definition P256_SHA256_TV01_Q :: "int point" where
  "P256_SHA256_TV01_Q = Point (int P256_SHA256_TV01_Qx) (int P256_SHA256_TV01_Qy)"

definition P256_SHA256_TV01_k :: nat where
  "P256_SHA256_TV01_k = 0x94a1bbb14b906a61a280f245f9e93c7f3b4a6247824f5d33b9670787642a68de"

definition P256_SHA256_TV01_R :: nat where
  "P256_SHA256_TV01_R = 0xf3ac8061b514795b8843e3d6629527ed2afd6b1f6a555a7acabb5e6f79c8c2ac"

definition P256_SHA256_TV01_S :: nat where
  "P256_SHA256_TV01_S = 0x8bf77819ca05a6b2786c76262bf7371cef97b218e96f175a3ccdda2acc058903"

definition P256_SHA256_TV01_Sig :: "nat × nat" where
  "P256_SHA256_TV01_Sig = (P256_SHA256_TV01_R, P256_SHA256_TV01_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA256_TV01_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV01_d" 
  by eval

lemma P256_SHA256_TV01_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA256_TV01_Q" 
  by eval

lemma P256_SHA256_TV01_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA256_TV01_d = P256_SHA256_TV01_Q" 
  by eval

lemma P256_SHA256_TV01_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA256_TV01_d P256_SHA256_TV01_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA256_TV01_dQvalidPair' P256_SHA256_TV01_d_valid) 

lemma P256_SHA256_TV01_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV01_k" 
  by eval

lemma P256_SHA256_TV01_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA256octets P256_SHA256_TV01_d P256_SHA256_TV01_Msg P256_SHA256_TV01_k = Some P256_SHA256_TV01_Sig" 
  by eval

lemma P256_SHA256_TV01_Verify: 
  "SEC1_P256_ECDSA_Verify SHA256octets P256_SHA256_TV01_Q P256_SHA256_TV01_Msg P256_SHA256_TV01_Sig" 
  by eval

subsection‹Test Vector: P256_SHA256_TV02›

subsubsection‹Given Test Vector Values›

definition P256_SHA256_TV02_Msg :: octets where
  "P256_SHA256_TV02_Msg = nat_to_octets_len 0xc35e2f092553c55772926bdbe87c9796827d17024dbb9233a545366e2e5987dd344deb72df987144b8c6c43bc41b654b94cc856e16b96d7a821c8ec039b503e3d86728c494a967d83011a0e090b5d54cd47f4e366c0912bc808fbb2ea96efac88fb3ebec9342738e225f7c7c2b011ce375b56621a20642b4d36e060db4524af1 128"

definition P256_SHA256_TV02_d :: nat where
  "P256_SHA256_TV02_d = 0x0f56db78ca460b055c500064824bed999a25aaf48ebb519ac201537b85479813"

definition P256_SHA256_TV02_Qx :: nat where
  "P256_SHA256_TV02_Qx = 0xe266ddfdc12668db30d4ca3e8f7749432c416044f2d2b8c10bf3d4012aeffa8a"

definition P256_SHA256_TV02_Qy :: nat where
  "P256_SHA256_TV02_Qy = 0xbfa86404a2e9ffe67d47c587ef7a97a7f456b863b4d02cfc6928973ab5b1cb39"

definition P256_SHA256_TV02_Q :: "int point" where
  "P256_SHA256_TV02_Q = Point (int P256_SHA256_TV02_Qx) (int P256_SHA256_TV02_Qy)"

definition P256_SHA256_TV02_k :: nat where
  "P256_SHA256_TV02_k = 0x6d3e71882c3b83b156bb14e0ab184aa9fb728068d3ae9fac421187ae0b2f34c6"

definition P256_SHA256_TV02_R :: nat where
  "P256_SHA256_TV02_R = 0x976d3a4e9d23326dc0baa9fa560b7c4e53f42864f508483a6473b6a11079b2db"

definition P256_SHA256_TV02_S :: nat where
  "P256_SHA256_TV02_S = 0x1b766e9ceb71ba6c01dcd46e0af462cd4cfa652ae5017d4555b8eeefe36e1932"

definition P256_SHA256_TV02_Sig :: "nat × nat" where
  "P256_SHA256_TV02_Sig = (P256_SHA256_TV02_R, P256_SHA256_TV02_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA256_TV02_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV02_d" 
  by eval

lemma P256_SHA256_TV02_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA256_TV02_Q" 
  by eval

lemma P256_SHA256_TV02_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA256_TV02_d = P256_SHA256_TV02_Q" 
  by eval

lemma P256_SHA256_TV02_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA256_TV02_d P256_SHA256_TV02_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA256_TV02_dQvalidPair' P256_SHA256_TV02_d_valid) 

lemma P256_SHA256_TV02_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV02_k" 
  by eval

lemma P256_SHA256_TV02_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA256octets P256_SHA256_TV02_d P256_SHA256_TV02_Msg P256_SHA256_TV02_k = Some P256_SHA256_TV02_Sig" 
  by eval

lemma P256_SHA256_TV02_Verify: 
  "SEC1_P256_ECDSA_Verify SHA256octets P256_SHA256_TV02_Q P256_SHA256_TV02_Msg P256_SHA256_TV02_Sig" 
  by eval

subsection‹Test Vector: P256_SHA256_TV03›

subsubsection‹Given Test Vector Values›

definition P256_SHA256_TV03_Msg :: octets where
  "P256_SHA256_TV03_Msg = nat_to_octets_len 0x3c054e333a94259c36af09ab5b4ff9beb3492f8d5b4282d16801daccb29f70fe61a0b37ffef5c04cd1b70e85b1f549a1c4dc672985e50f43ea037efa9964f096b5f62f7ffdf8d6bfb2cc859558f5a393cb949dbd48f269343b5263dcdb9c556eca074f2e98e6d94c2c29a677afaf806edf79b15a3fcd46e7067b7669f83188ee 128"

definition P256_SHA256_TV03_d :: nat where
  "P256_SHA256_TV03_d = 0xe283871239837e13b95f789e6e1af63bf61c918c992e62bca040d64cad1fc2ef"

definition P256_SHA256_TV03_Qx :: nat where
  "P256_SHA256_TV03_Qx = 0x74ccd8a62fba0e667c50929a53f78c21b8ff0c3c737b0b40b1750b2302b0bde8"

definition P256_SHA256_TV03_Qy :: nat where
  "P256_SHA256_TV03_Qy = 0x29074e21f3a0ef88b9efdf10d06aa4c295cc1671f758ca0e4cd108803d0f2614"

definition P256_SHA256_TV03_Q :: "int point" where
  "P256_SHA256_TV03_Q = Point (int P256_SHA256_TV03_Qx) (int P256_SHA256_TV03_Qy)"

definition P256_SHA256_TV03_k :: nat where
  "P256_SHA256_TV03_k = 0xad5e887eb2b380b8d8280ad6e5ff8a60f4d26243e0124c2f31a297b5d0835de2"

definition P256_SHA256_TV03_R :: nat where
  "P256_SHA256_TV03_R = 0x35fb60f5ca0f3ca08542fb3cc641c8263a2cab7a90ee6a5e1583fac2bb6f6bd1"

definition P256_SHA256_TV03_S :: nat where
  "P256_SHA256_TV03_S = 0xee59d81bc9db1055cc0ed97b159d8784af04e98511d0a9a407b99bb292572e96"

definition P256_SHA256_TV03_Sig :: "nat × nat" where
  "P256_SHA256_TV03_Sig = (P256_SHA256_TV03_R, P256_SHA256_TV03_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA256_TV03_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV03_d" 
  by eval

lemma P256_SHA256_TV03_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA256_TV03_Q" 
  by eval

lemma P256_SHA256_TV03_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA256_TV03_d = P256_SHA256_TV03_Q" 
  by eval

lemma P256_SHA256_TV03_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA256_TV03_d P256_SHA256_TV03_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA256_TV03_dQvalidPair' P256_SHA256_TV03_d_valid) 

lemma P256_SHA256_TV03_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV03_k" 
  by eval

lemma P256_SHA256_TV03_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA256octets P256_SHA256_TV03_d P256_SHA256_TV03_Msg P256_SHA256_TV03_k = Some P256_SHA256_TV03_Sig" 
  by eval

lemma P256_SHA256_TV03_Verify: 
  "SEC1_P256_ECDSA_Verify SHA256octets P256_SHA256_TV03_Q P256_SHA256_TV03_Msg P256_SHA256_TV03_Sig" 
  by eval

subsection‹Test Vector: P256_SHA256_TV04›

subsubsection‹Given Test Vector Values›

definition P256_SHA256_TV04_Msg :: octets where
  "P256_SHA256_TV04_Msg = nat_to_octets_len 0x0989122410d522af64ceb07da2c865219046b4c3d9d99b01278c07ff63eaf1039cb787ae9e2dd46436cc0415f280c562bebb83a23e639e476a02ec8cff7ea06cd12c86dcc3adefbf1a9e9a9b6646c7599ec631b0da9a60debeb9b3e19324977f3b4f36892c8a38671c8e1cc8e50fcd50f9e51deaf98272f9266fc702e4e57c30 128"

definition P256_SHA256_TV04_d :: nat where
  "P256_SHA256_TV04_d = 0xa3d2d3b7596f6592ce98b4bfe10d41837f10027a90d7bb75349490018cf72d07"

definition P256_SHA256_TV04_Qx :: nat where
  "P256_SHA256_TV04_Qx = 0x322f80371bf6e044bc49391d97c1714ab87f990b949bc178cb7c43b7c22d89e1"

definition P256_SHA256_TV04_Qy :: nat where
  "P256_SHA256_TV04_Qy = 0x3c15d54a5cc6b9f09de8457e873eb3deb1fceb54b0b295da6050294fae7fd999"

definition P256_SHA256_TV04_Q :: "int point" where
  "P256_SHA256_TV04_Q = Point (int P256_SHA256_TV04_Qx) (int P256_SHA256_TV04_Qy)"

definition P256_SHA256_TV04_k :: nat where
  "P256_SHA256_TV04_k = 0x24fc90e1da13f17ef9fe84cc96b9471ed1aaac17e3a4bae33a115df4e5834f18"

definition P256_SHA256_TV04_R :: nat where
  "P256_SHA256_TV04_R = 0xd7c562370af617b581c84a2468cc8bd50bb1cbf322de41b7887ce07c0e5884ca"

definition P256_SHA256_TV04_S :: nat where
  "P256_SHA256_TV04_S = 0xb46d9f2d8c4bf83546ff178f1d78937c008d64e8ecc5cbb825cb21d94d670d89"

definition P256_SHA256_TV04_Sig :: "nat × nat" where
  "P256_SHA256_TV04_Sig = (P256_SHA256_TV04_R, P256_SHA256_TV04_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA256_TV04_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV04_d" 
  by eval

lemma P256_SHA256_TV04_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA256_TV04_Q" 
  by eval

lemma P256_SHA256_TV04_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA256_TV04_d = P256_SHA256_TV04_Q" 
  by eval

lemma P256_SHA256_TV04_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA256_TV04_d P256_SHA256_TV04_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA256_TV04_dQvalidPair' P256_SHA256_TV04_d_valid) 

lemma P256_SHA256_TV04_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV04_k" 
  by eval

lemma P256_SHA256_TV04_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA256octets P256_SHA256_TV04_d P256_SHA256_TV04_Msg P256_SHA256_TV04_k = Some P256_SHA256_TV04_Sig" 
  by eval

lemma P256_SHA256_TV04_Verify: 
  "SEC1_P256_ECDSA_Verify SHA256octets P256_SHA256_TV04_Q P256_SHA256_TV04_Msg P256_SHA256_TV04_Sig" 
  by eval

subsection‹Test Vector: P256_SHA256_TV05›

subsubsection‹Given Test Vector Values›

definition P256_SHA256_TV05_Msg :: octets where
  "P256_SHA256_TV05_Msg = nat_to_octets_len 0xdc66e39f9bbfd9865318531ffe9207f934fa615a5b285708a5e9c46b7775150e818d7f24d2a123df3672fff2094e3fd3df6fbe259e3989dd5edfcccbe7d45e26a775a5c4329a084f057c42c13f3248e3fd6f0c76678f890f513c32292dd306eaa84a59abe34b16cb5e38d0e885525d10336ca443e1682aa04a7af832b0eee4e7 128"

definition P256_SHA256_TV05_d :: nat where
  "P256_SHA256_TV05_d = 0x53a0e8a8fe93db01e7ae94e1a9882a102ebd079b3a535827d583626c272d280d"

definition P256_SHA256_TV05_Qx :: nat where
  "P256_SHA256_TV05_Qx = 0x1bcec4570e1ec2436596b8ded58f60c3b1ebc6a403bc5543040ba82963057244"

definition P256_SHA256_TV05_Qy :: nat where
  "P256_SHA256_TV05_Qy = 0x8af62a4c683f096b28558320737bf83b9959a46ad2521004ef74cf85e67494e1"

definition P256_SHA256_TV05_Q :: "int point" where
  "P256_SHA256_TV05_Q = Point (int P256_SHA256_TV05_Qx) (int P256_SHA256_TV05_Qy)"

definition P256_SHA256_TV05_k :: nat where
  "P256_SHA256_TV05_k = 0x5d833e8d24cc7a402d7ee7ec852a3587cddeb48358cea71b0bedb8fabe84e0c4"

definition P256_SHA256_TV05_R :: nat where
  "P256_SHA256_TV05_R = 0x18caaf7b663507a8bcd992b836dec9dc5703c080af5e51dfa3a9a7c387182604"

definition P256_SHA256_TV05_S :: nat where
  "P256_SHA256_TV05_S = 0x77c68928ac3b88d985fb43fb615fb7ff45c18ba5c81af796c613dfa98352d29c"

definition P256_SHA256_TV05_Sig :: "nat × nat" where
  "P256_SHA256_TV05_Sig = (P256_SHA256_TV05_R, P256_SHA256_TV05_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA256_TV05_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV05_d" 
  by eval

lemma P256_SHA256_TV05_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA256_TV05_Q" 
  by eval

lemma P256_SHA256_TV05_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA256_TV05_d = P256_SHA256_TV05_Q" 
  by eval

lemma P256_SHA256_TV05_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA256_TV05_d P256_SHA256_TV05_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA256_TV05_dQvalidPair' P256_SHA256_TV05_d_valid) 

lemma P256_SHA256_TV05_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV05_k" 
  by eval

lemma P256_SHA256_TV05_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA256octets P256_SHA256_TV05_d P256_SHA256_TV05_Msg P256_SHA256_TV05_k = Some P256_SHA256_TV05_Sig" 
  by eval

lemma P256_SHA256_TV05_Verify: 
  "SEC1_P256_ECDSA_Verify SHA256octets P256_SHA256_TV05_Q P256_SHA256_TV05_Msg P256_SHA256_TV05_Sig" 
  by eval

subsection‹Test Vector: P256_SHA256_TV06›

subsubsection‹Given Test Vector Values›

definition P256_SHA256_TV06_Msg :: octets where
  "P256_SHA256_TV06_Msg = nat_to_octets_len 0x600974e7d8c5508e2c1aab0783ad0d7c4494ab2b4da265c2fe496421c4df238b0be25f25659157c8a225fb03953607f7df996acfd402f147e37aee2f1693e3bf1c35eab3ae360a2bd91d04622ea47f83d863d2dfecb618e8b8bdc39e17d15d672eee03bb4ce2cc5cf6b217e5faf3f336fdd87d972d3a8b8a593ba85955cc9d71 128"

definition P256_SHA256_TV06_d :: nat where
  "P256_SHA256_TV06_d = 0x4af107e8e2194c830ffb712a65511bc9186a133007855b49ab4b3833aefc4a1d"

definition P256_SHA256_TV06_Qx :: nat where
  "P256_SHA256_TV06_Qx = 0xa32e50be3dae2c8ba3f5e4bdae14cf7645420d425ead94036c22dd6c4fc59e00"

definition P256_SHA256_TV06_Qy :: nat where
  "P256_SHA256_TV06_Qy = 0xd623bf641160c289d6742c6257ae6ba574446dd1d0e74db3aaa80900b78d4ae9"

definition P256_SHA256_TV06_Q :: "int point" where
  "P256_SHA256_TV06_Q = Point (int P256_SHA256_TV06_Qx) (int P256_SHA256_TV06_Qy)"

definition P256_SHA256_TV06_k :: nat where
  "P256_SHA256_TV06_k = 0xe18f96f84dfa2fd3cdfaec9159d4c338cd54ad314134f0b31e20591fc238d0ab"

definition P256_SHA256_TV06_R :: nat where
  "P256_SHA256_TV06_R = 0x8524c5024e2d9a73bde8c72d9129f57873bbad0ed05215a372a84fdbc78f2e68"

definition P256_SHA256_TV06_S :: nat where
  "P256_SHA256_TV06_S = 0xd18c2caf3b1072f87064ec5e8953f51301cada03469c640244760328eb5a05cb"

definition P256_SHA256_TV06_Sig :: "nat × nat" where
  "P256_SHA256_TV06_Sig = (P256_SHA256_TV06_R, P256_SHA256_TV06_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA256_TV06_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV06_d" 
  by eval

lemma P256_SHA256_TV06_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA256_TV06_Q" 
  by eval

lemma P256_SHA256_TV06_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA256_TV06_d = P256_SHA256_TV06_Q" 
  by eval

lemma P256_SHA256_TV06_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA256_TV06_d P256_SHA256_TV06_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA256_TV06_dQvalidPair' P256_SHA256_TV06_d_valid) 

lemma P256_SHA256_TV06_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV06_k" 
  by eval

lemma P256_SHA256_TV06_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA256octets P256_SHA256_TV06_d P256_SHA256_TV06_Msg P256_SHA256_TV06_k = Some P256_SHA256_TV06_Sig" 
  by eval

lemma P256_SHA256_TV06_Verify: 
  "SEC1_P256_ECDSA_Verify SHA256octets P256_SHA256_TV06_Q P256_SHA256_TV06_Msg P256_SHA256_TV06_Sig" 
  by eval

subsection‹Test Vector: P256_SHA256_TV07›

subsubsection‹Given Test Vector Values›

definition P256_SHA256_TV07_Msg :: octets where
  "P256_SHA256_TV07_Msg = nat_to_octets_len 0xdfa6cb9b39adda6c74cc8b2a8b53a12c499ab9dee01b4123642b4f11af336a91a5c9ce0520eb2395a6190ecbf6169c4cba81941de8e76c9c908eb843b98ce95e0da29c5d4388040264e05e07030a577cc5d176387154eabae2af52a83e85c61c7c61da930c9b19e45d7e34c8516dc3c238fddd6e450a77455d534c48a152010b 128"

definition P256_SHA256_TV07_d :: nat where
  "P256_SHA256_TV07_d = 0x78dfaa09f1076850b3e206e477494cddcfb822aaa0128475053592c48ebaf4ab"

definition P256_SHA256_TV07_Qx :: nat where
  "P256_SHA256_TV07_Qx = 0x8bcfe2a721ca6d753968f564ec4315be4857e28bef1908f61a366b1f03c97479"

definition P256_SHA256_TV07_Qy :: nat where
  "P256_SHA256_TV07_Qy = 0x0f67576a30b8e20d4232d8530b52fb4c89cbc589ede291e499ddd15fe870ab96"

definition P256_SHA256_TV07_Q :: "int point" where
  "P256_SHA256_TV07_Q = Point (int P256_SHA256_TV07_Qx) (int P256_SHA256_TV07_Qy)"

definition P256_SHA256_TV07_k :: nat where
  "P256_SHA256_TV07_k = 0x295544dbb2da3da170741c9b2c6551d40af7ed4e891445f11a02b66a5c258a77"

definition P256_SHA256_TV07_R :: nat where
  "P256_SHA256_TV07_R = 0xc5a186d72df452015480f7f338970bfe825087f05c0088d95305f87aacc9b254"

definition P256_SHA256_TV07_S :: nat where
  "P256_SHA256_TV07_S = 0x84a58f9e9d9e735344b316b1aa1ab5185665b85147dc82d92e969d7bee31ca30"

definition P256_SHA256_TV07_Sig :: "nat × nat" where
  "P256_SHA256_TV07_Sig = (P256_SHA256_TV07_R, P256_SHA256_TV07_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA256_TV07_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV07_d" 
  by eval

lemma P256_SHA256_TV07_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA256_TV07_Q" 
  by eval

lemma P256_SHA256_TV07_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA256_TV07_d = P256_SHA256_TV07_Q" 
  by eval

lemma P256_SHA256_TV07_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA256_TV07_d P256_SHA256_TV07_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA256_TV07_dQvalidPair' P256_SHA256_TV07_d_valid) 

lemma P256_SHA256_TV07_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV07_k" 
  by eval

lemma P256_SHA256_TV07_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA256octets P256_SHA256_TV07_d P256_SHA256_TV07_Msg P256_SHA256_TV07_k = Some P256_SHA256_TV07_Sig" 
  by eval

lemma P256_SHA256_TV07_Verify: 
  "SEC1_P256_ECDSA_Verify SHA256octets P256_SHA256_TV07_Q P256_SHA256_TV07_Msg P256_SHA256_TV07_Sig" 
  by eval

subsection‹Test Vector: P256_SHA256_TV08›

subsubsection‹Given Test Vector Values›

definition P256_SHA256_TV08_Msg :: octets where
  "P256_SHA256_TV08_Msg = nat_to_octets_len 0x51d2547cbff92431174aa7fc7302139519d98071c755ff1c92e4694b58587ea560f72f32fc6dd4dee7d22bb7387381d0256e2862d0644cdf2c277c5d740fa089830eb52bf79d1e75b8596ecf0ea58a0b9df61e0c9754bfcd62efab6ea1bd216bf181c5593da79f10135a9bc6e164f1854bc8859734341aad237ba29a81a3fc8b 128"

definition P256_SHA256_TV08_d :: nat where
  "P256_SHA256_TV08_d = 0x80e692e3eb9fcd8c7d44e7de9f7a5952686407f90025a1d87e52c7096a62618a"

definition P256_SHA256_TV08_Qx :: nat where
  "P256_SHA256_TV08_Qx = 0xa88bc8430279c8c0400a77d751f26c0abc93e5de4ad9a4166357952fe041e767"

definition P256_SHA256_TV08_Qy :: nat where
  "P256_SHA256_TV08_Qy = 0x2d365a1eef25ead579cc9a069b6abc1b16b81c35f18785ce26a10ba6d1381185"

definition P256_SHA256_TV08_Q :: "int point" where
  "P256_SHA256_TV08_Q = Point (int P256_SHA256_TV08_Qx) (int P256_SHA256_TV08_Qy)"

definition P256_SHA256_TV08_k :: nat where
  "P256_SHA256_TV08_k = 0x7c80fd66d62cc076cef2d030c17c0a69c99611549cb32c4ff662475adbe84b22"

definition P256_SHA256_TV08_R :: nat where
  "P256_SHA256_TV08_R = 0x9d0c6afb6df3bced455b459cc21387e14929392664bb8741a3693a1795ca6902"

definition P256_SHA256_TV08_S :: nat where
  "P256_SHA256_TV08_S = 0xd7f9ddd191f1f412869429209ee3814c75c72fa46a9cccf804a2f5cc0b7e739f"

definition P256_SHA256_TV08_Sig :: "nat × nat" where
  "P256_SHA256_TV08_Sig = (P256_SHA256_TV08_R, P256_SHA256_TV08_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA256_TV08_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV08_d" 
  by eval

lemma P256_SHA256_TV08_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA256_TV08_Q" 
  by eval

lemma P256_SHA256_TV08_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA256_TV08_d = P256_SHA256_TV08_Q" 
  by eval

lemma P256_SHA256_TV08_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA256_TV08_d P256_SHA256_TV08_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA256_TV08_dQvalidPair' P256_SHA256_TV08_d_valid) 

lemma P256_SHA256_TV08_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV08_k" 
  by eval

lemma P256_SHA256_TV08_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA256octets P256_SHA256_TV08_d P256_SHA256_TV08_Msg P256_SHA256_TV08_k = Some P256_SHA256_TV08_Sig" 
  by eval

lemma P256_SHA256_TV08_Verify: 
  "SEC1_P256_ECDSA_Verify SHA256octets P256_SHA256_TV08_Q P256_SHA256_TV08_Msg P256_SHA256_TV08_Sig" 
  by eval

subsection‹Test Vector: P256_SHA256_TV09›

subsubsection‹Given Test Vector Values›

definition P256_SHA256_TV09_Msg :: octets where
  "P256_SHA256_TV09_Msg = nat_to_octets_len 0x558c2ac13026402bad4a0a83ebc9468e50f7ffab06d6f981e5db1d082098065bcff6f21a7a74558b1e8612914b8b5a0aa28ed5b574c36ac4ea5868432a62bb8ef0695d27c1e3ceaf75c7b251c65ddb268696f07c16d2767973d85beb443f211e6445e7fe5d46f0dce70d58a4cd9fe70688c035688ea8c6baec65a5fc7e2c93e8 128"

definition P256_SHA256_TV09_d :: nat where
  "P256_SHA256_TV09_d = 0x5e666c0db0214c3b627a8e48541cc84a8b6fd15f300da4dff5d18aec6c55b881"

definition P256_SHA256_TV09_Qx :: nat where
  "P256_SHA256_TV09_Qx = 0x1bc487570f040dc94196c9befe8ab2b6de77208b1f38bdaae28f9645c4d2bc3a"

definition P256_SHA256_TV09_Qy :: nat where
  "P256_SHA256_TV09_Qy = 0xec81602abd8345e71867c8210313737865b8aa186851e1b48eaca140320f5d8f"

definition P256_SHA256_TV09_Q :: "int point" where
  "P256_SHA256_TV09_Q = Point (int P256_SHA256_TV09_Qx) (int P256_SHA256_TV09_Qy)"

definition P256_SHA256_TV09_k :: nat where
  "P256_SHA256_TV09_k = 0x2e7625a48874d86c9e467f890aaa7cd6ebdf71c0102bfdcfa24565d6af3fdce9"

definition P256_SHA256_TV09_R :: nat where
  "P256_SHA256_TV09_R = 0x2f9e2b4e9f747c657f705bffd124ee178bbc5391c86d056717b140c153570fd9"

definition P256_SHA256_TV09_S :: nat where
  "P256_SHA256_TV09_S = 0xf5413bfd85949da8d83de83ab0d19b2986613e224d1901d76919de23ccd03199"

definition P256_SHA256_TV09_Sig :: "nat × nat" where
  "P256_SHA256_TV09_Sig = (P256_SHA256_TV09_R, P256_SHA256_TV09_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA256_TV09_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV09_d" 
  by eval

lemma P256_SHA256_TV09_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA256_TV09_Q" 
  by eval

lemma P256_SHA256_TV09_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA256_TV09_d = P256_SHA256_TV09_Q" 
  by eval

lemma P256_SHA256_TV09_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA256_TV09_d P256_SHA256_TV09_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA256_TV09_dQvalidPair' P256_SHA256_TV09_d_valid) 

lemma P256_SHA256_TV09_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV09_k" 
  by eval

lemma P256_SHA256_TV09_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA256octets P256_SHA256_TV09_d P256_SHA256_TV09_Msg P256_SHA256_TV09_k = Some P256_SHA256_TV09_Sig" 
  by eval

lemma P256_SHA256_TV09_Verify: 
  "SEC1_P256_ECDSA_Verify SHA256octets P256_SHA256_TV09_Q P256_SHA256_TV09_Msg P256_SHA256_TV09_Sig" 
  by eval

subsection‹Test Vector: P256_SHA256_TV10›

subsubsection‹Given Test Vector Values›

definition P256_SHA256_TV10_Msg :: octets where
  "P256_SHA256_TV10_Msg = nat_to_octets_len 0x4d55c99ef6bd54621662c3d110c3cb627c03d6311393b264ab97b90a4b15214a5593ba2510a53d63fb34be251facb697c973e11b665cb7920f1684b0031b4dd370cb927ca7168b0bf8ad285e05e9e31e34bc24024739fdc10b78586f29eff94412034e3b606ed850ec2c1900e8e68151fc4aee5adebb066eb6da4eaa5681378e 128"

definition P256_SHA256_TV10_d :: nat where
  "P256_SHA256_TV10_d = 0xf73f455271c877c4d5334627e37c278f68d143014b0a05aa62f308b2101c5308"

definition P256_SHA256_TV10_Qx :: nat where
  "P256_SHA256_TV10_Qx = 0xb8188bd68701fc396dab53125d4d28ea33a91daf6d21485f4770f6ea8c565dde"

definition P256_SHA256_TV10_Qy :: nat where
  "P256_SHA256_TV10_Qy = 0x423f058810f277f8fe076f6db56e9285a1bf2c2a1dae145095edd9c04970bc4a"

definition P256_SHA256_TV10_Q :: "int point" where
  "P256_SHA256_TV10_Q = Point (int P256_SHA256_TV10_Qx) (int P256_SHA256_TV10_Qy)"

definition P256_SHA256_TV10_k :: nat where
  "P256_SHA256_TV10_k = 0x62f8665fd6e26b3fa069e85281777a9b1f0dfd2c0b9f54a086d0c109ff9fd615"

definition P256_SHA256_TV10_R :: nat where
  "P256_SHA256_TV10_R = 0x1cc628533d0004b2b20e7f4baad0b8bb5e0673db159bbccf92491aef61fc9620"

definition P256_SHA256_TV10_S :: nat where
  "P256_SHA256_TV10_S = 0x880e0bbf82a8cf818ed46ba03cf0fc6c898e36fca36cc7fdb1d2db7503634430"

definition P256_SHA256_TV10_Sig :: "nat × nat" where
  "P256_SHA256_TV10_Sig = (P256_SHA256_TV10_R, P256_SHA256_TV10_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA256_TV10_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV10_d" 
  by eval

lemma P256_SHA256_TV10_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA256_TV10_Q" 
  by eval

lemma P256_SHA256_TV10_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA256_TV10_d = P256_SHA256_TV10_Q" 
  by eval

lemma P256_SHA256_TV10_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA256_TV10_d P256_SHA256_TV10_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA256_TV10_dQvalidPair' P256_SHA256_TV10_d_valid) 

lemma P256_SHA256_TV10_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV10_k" 
  by eval

lemma P256_SHA256_TV10_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA256octets P256_SHA256_TV10_d P256_SHA256_TV10_Msg P256_SHA256_TV10_k = Some P256_SHA256_TV10_Sig" 
  by eval

lemma P256_SHA256_TV10_Verify: 
  "SEC1_P256_ECDSA_Verify SHA256octets P256_SHA256_TV10_Q P256_SHA256_TV10_Msg P256_SHA256_TV10_Sig" 
  by eval

subsection‹Test Vector: P256_SHA256_TV11›

subsubsection‹Given Test Vector Values›

definition P256_SHA256_TV11_Msg :: octets where
  "P256_SHA256_TV11_Msg = nat_to_octets_len 0xf8248ad47d97c18c984f1f5c10950dc1404713c56b6ea397e01e6dd925e903b4fadfe2c9e877169e71ce3c7fe5ce70ee4255d9cdc26f6943bf48687874de64f6cf30a012512e787b88059bbf561162bdcc23a3742c835ac144cc14167b1bd6727e940540a9c99f3cbb41fb1dcb00d76dda04995847c657f4c19d303eb09eb48a 128"

definition P256_SHA256_TV11_d :: nat where
  "P256_SHA256_TV11_d = 0xb20d705d9bd7c2b8dc60393a5357f632990e599a0975573ac67fd89b49187906"

definition P256_SHA256_TV11_Qx :: nat where
  "P256_SHA256_TV11_Qx = 0x51f99d2d52d4a6e734484a018b7ca2f895c2929b6754a3a03224d07ae61166ce"

definition P256_SHA256_TV11_Qy :: nat where
  "P256_SHA256_TV11_Qy = 0x4737da963c6ef7247fb88d19f9b0c667cac7fe12837fdab88c66f10d3c14cad1"

definition P256_SHA256_TV11_Q :: "int point" where
  "P256_SHA256_TV11_Q = Point (int P256_SHA256_TV11_Qx) (int P256_SHA256_TV11_Qy)"

definition P256_SHA256_TV11_k :: nat where
  "P256_SHA256_TV11_k = 0x72b656f6b35b9ccbc712c9f1f3b1a14cbbebaec41c4bca8da18f492a062d6f6f"

definition P256_SHA256_TV11_R :: nat where
  "P256_SHA256_TV11_R = 0x9886ae46c1415c3bc959e82b760ad760aab66885a84e620aa339fdf102465c42"

definition P256_SHA256_TV11_S :: nat where
  "P256_SHA256_TV11_S = 0x2bf3a80bc04faa35ebecc0f4864ac02d349f6f126e0f988501b8d3075409a26c"

definition P256_SHA256_TV11_Sig :: "nat × nat" where
  "P256_SHA256_TV11_Sig = (P256_SHA256_TV11_R, P256_SHA256_TV11_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA256_TV11_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV11_d" 
  by eval

lemma P256_SHA256_TV11_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA256_TV11_Q" 
  by eval

lemma P256_SHA256_TV11_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA256_TV11_d = P256_SHA256_TV11_Q" 
  by eval

lemma P256_SHA256_TV11_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA256_TV11_d P256_SHA256_TV11_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA256_TV11_dQvalidPair' P256_SHA256_TV11_d_valid) 

lemma P256_SHA256_TV11_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV11_k" 
  by eval

lemma P256_SHA256_TV11_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA256octets P256_SHA256_TV11_d P256_SHA256_TV11_Msg P256_SHA256_TV11_k = Some P256_SHA256_TV11_Sig" 
  by eval

lemma P256_SHA256_TV11_Verify: 
  "SEC1_P256_ECDSA_Verify SHA256octets P256_SHA256_TV11_Q P256_SHA256_TV11_Msg P256_SHA256_TV11_Sig" 
  by eval

subsection‹Test Vector: P256_SHA256_TV12›

subsubsection‹Given Test Vector Values›

definition P256_SHA256_TV12_Msg :: octets where
  "P256_SHA256_TV12_Msg = nat_to_octets_len 0x3b6ee2425940b3d240d35b97b6dcd61ed3423d8e71a0ada35d47b322d17b35ea0472f35edd1d252f87b8b65ef4b716669fc9ac28b00d34a9d66ad118c9d94e7f46d0b4f6c2b2d339fd6bcd351241a387cc82609057048c12c4ec3d85c661975c45b300cb96930d89370a327c98b67defaa89497aa8ef994c77f1130f752f94a4 128"

definition P256_SHA256_TV12_d :: nat where
  "P256_SHA256_TV12_d = 0xd4234bebfbc821050341a37e1240efe5e33763cbbb2ef76a1c79e24724e5a5e7"

definition P256_SHA256_TV12_Qx :: nat where
  "P256_SHA256_TV12_Qx = 0x8fb287f0202ad57ae841aea35f29b2e1d53e196d0ddd9aec24813d64c0922fb7"

definition P256_SHA256_TV12_Qy :: nat where
  "P256_SHA256_TV12_Qy = 0x1f6daff1aa2dd2d6d3741623eecb5e7b612997a1039aab2e5cf2de969cfea573"

definition P256_SHA256_TV12_Q :: "int point" where
  "P256_SHA256_TV12_Q = Point (int P256_SHA256_TV12_Qx) (int P256_SHA256_TV12_Qy)"

definition P256_SHA256_TV12_k :: nat where
  "P256_SHA256_TV12_k = 0xd926fe10f1bfd9855610f4f5a3d666b1a149344057e35537373372ead8b1a778"

definition P256_SHA256_TV12_R :: nat where
  "P256_SHA256_TV12_R = 0x490efd106be11fc365c7467eb89b8d39e15d65175356775deab211163c2504cb"

definition P256_SHA256_TV12_S :: nat where
  "P256_SHA256_TV12_S = 0x644300fc0da4d40fb8c6ead510d14f0bd4e1321a469e9c0a581464c7186b7aa7"

definition P256_SHA256_TV12_Sig :: "nat × nat" where
  "P256_SHA256_TV12_Sig = (P256_SHA256_TV12_R, P256_SHA256_TV12_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA256_TV12_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV12_d" 
  by eval

lemma P256_SHA256_TV12_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA256_TV12_Q" 
  by eval

lemma P256_SHA256_TV12_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA256_TV12_d = P256_SHA256_TV12_Q" 
  by eval

lemma P256_SHA256_TV12_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA256_TV12_d P256_SHA256_TV12_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA256_TV12_dQvalidPair' P256_SHA256_TV12_d_valid) 

lemma P256_SHA256_TV12_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV12_k" 
  by eval

lemma P256_SHA256_TV12_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA256octets P256_SHA256_TV12_d P256_SHA256_TV12_Msg P256_SHA256_TV12_k = Some P256_SHA256_TV12_Sig" 
  by eval

lemma P256_SHA256_TV12_Verify: 
  "SEC1_P256_ECDSA_Verify SHA256octets P256_SHA256_TV12_Q P256_SHA256_TV12_Msg P256_SHA256_TV12_Sig" 
  by eval

subsection‹Test Vector: P256_SHA256_TV13›

subsubsection‹Given Test Vector Values›

definition P256_SHA256_TV13_Msg :: octets where
  "P256_SHA256_TV13_Msg = nat_to_octets_len 0xc5204b81ec0a4df5b7e9fda3dc245f98082ae7f4efe81998dcaa286bd4507ca840a53d21b01e904f55e38f78c3757d5a5a4a44b1d5d4e480be3afb5b394a5d2840af42b1b4083d40afbfe22d702f370d32dbfd392e128ea4724d66a3701da41ae2f03bb4d91bb946c7969404cb544f71eb7a49eb4c4ec55799bda1eb545143a7 128"

definition P256_SHA256_TV13_d :: nat where
  "P256_SHA256_TV13_d = 0xb58f5211dff440626bb56d0ad483193d606cf21f36d9830543327292f4d25d8c"

definition P256_SHA256_TV13_Qx :: nat where
  "P256_SHA256_TV13_Qx = 0x68229b48c2fe19d3db034e4c15077eb7471a66031f28a980821873915298ba76"

definition P256_SHA256_TV13_Qy :: nat where
  "P256_SHA256_TV13_Qy = 0x303e8ee3742a893f78b810991da697083dd8f11128c47651c27a56740a80c24c"

definition P256_SHA256_TV13_Q :: "int point" where
  "P256_SHA256_TV13_Q = Point (int P256_SHA256_TV13_Qx) (int P256_SHA256_TV13_Qy)"

definition P256_SHA256_TV13_k :: nat where
  "P256_SHA256_TV13_k = 0xe158bf4a2d19a99149d9cdb879294ccb7aaeae03d75ddd616ef8ae51a6dc1071"

definition P256_SHA256_TV13_R :: nat where
  "P256_SHA256_TV13_R = 0xe67a9717ccf96841489d6541f4f6adb12d17b59a6bef847b6183b8fcf16a32eb"

definition P256_SHA256_TV13_S :: nat where
  "P256_SHA256_TV13_S = 0x9ae6ba6d637706849a6a9fc388cf0232d85c26ea0d1fe7437adb48de58364333"

definition P256_SHA256_TV13_Sig :: "nat × nat" where
  "P256_SHA256_TV13_Sig = (P256_SHA256_TV13_R, P256_SHA256_TV13_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA256_TV13_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV13_d" 
  by eval

lemma P256_SHA256_TV13_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA256_TV13_Q" 
  by eval

lemma P256_SHA256_TV13_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA256_TV13_d = P256_SHA256_TV13_Q" 
  by eval

lemma P256_SHA256_TV13_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA256_TV13_d P256_SHA256_TV13_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA256_TV13_dQvalidPair' P256_SHA256_TV13_d_valid) 

lemma P256_SHA256_TV13_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV13_k" 
  by eval

lemma P256_SHA256_TV13_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA256octets P256_SHA256_TV13_d P256_SHA256_TV13_Msg P256_SHA256_TV13_k = Some P256_SHA256_TV13_Sig" 
  by eval

lemma P256_SHA256_TV13_Verify: 
  "SEC1_P256_ECDSA_Verify SHA256octets P256_SHA256_TV13_Q P256_SHA256_TV13_Msg P256_SHA256_TV13_Sig" 
  by eval

subsection‹Test Vector: P256_SHA256_TV14›

subsubsection‹Given Test Vector Values›

definition P256_SHA256_TV14_Msg :: octets where
  "P256_SHA256_TV14_Msg = nat_to_octets_len 0x72e81fe221fb402148d8b7ab03549f1180bcc03d41ca59d7653801f0ba853add1f6d29edd7f9abc621b2d548f8dbf8979bd16608d2d8fc3260b4ebc0dd42482481d548c7075711b5759649c41f439fad69954956c9326841ea6492956829f9e0dc789f73633b40f6ac77bcae6dfc7930cfe89e526d1684365c5b0be2437fdb01 128"

definition P256_SHA256_TV14_d :: nat where
  "P256_SHA256_TV14_d = 0x54c066711cdb061eda07e5275f7e95a9962c6764b84f6f1f3ab5a588e0a2afb1"

definition P256_SHA256_TV14_Qx :: nat where
  "P256_SHA256_TV14_Qx = 0x0a7dbb8bf50cb605eb2268b081f26d6b08e012f952c4b70a5a1e6e7d46af98bb"

definition P256_SHA256_TV14_Qy :: nat where
  "P256_SHA256_TV14_Qy = 0xf26dd7d799930062480849962ccf5004edcfd307c044f4e8f667c9baa834eeae"

definition P256_SHA256_TV14_Q :: "int point" where
  "P256_SHA256_TV14_Q = Point (int P256_SHA256_TV14_Qx) (int P256_SHA256_TV14_Qy)"

definition P256_SHA256_TV14_k :: nat where
  "P256_SHA256_TV14_k = 0x646fe933e96c3b8f9f507498e907fdd201f08478d0202c752a7c2cfebf4d061a"

definition P256_SHA256_TV14_R :: nat where
  "P256_SHA256_TV14_R = 0xb53ce4da1aa7c0dc77a1896ab716b921499aed78df725b1504aba1597ba0c64b"

definition P256_SHA256_TV14_S :: nat where
  "P256_SHA256_TV14_S = 0xd7c246dc7ad0e67700c373edcfdd1c0a0495fc954549ad579df6ed1438840851"

definition P256_SHA256_TV14_Sig :: "nat × nat" where
  "P256_SHA256_TV14_Sig = (P256_SHA256_TV14_R, P256_SHA256_TV14_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA256_TV14_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV14_d" 
  by eval

lemma P256_SHA256_TV14_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA256_TV14_Q" 
  by eval

lemma P256_SHA256_TV14_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA256_TV14_d = P256_SHA256_TV14_Q" 
  by eval

lemma P256_SHA256_TV14_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA256_TV14_d P256_SHA256_TV14_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA256_TV14_dQvalidPair' P256_SHA256_TV14_d_valid) 

lemma P256_SHA256_TV14_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV14_k" 
  by eval

lemma P256_SHA256_TV14_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA256octets P256_SHA256_TV14_d P256_SHA256_TV14_Msg P256_SHA256_TV14_k = Some P256_SHA256_TV14_Sig" 
  by eval

lemma P256_SHA256_TV14_Verify: 
  "SEC1_P256_ECDSA_Verify SHA256octets P256_SHA256_TV14_Q P256_SHA256_TV14_Msg P256_SHA256_TV14_Sig" 
  by eval

subsection‹Test Vector: P256_SHA256_TV15›

subsubsection‹Given Test Vector Values›

definition P256_SHA256_TV15_Msg :: octets where
  "P256_SHA256_TV15_Msg = nat_to_octets_len 0x21188c3edd5de088dacc1076b9e1bcecd79de1003c2414c3866173054dc82dde85169baa77993adb20c269f60a5226111828578bcc7c29e6e8d2dae81806152c8ba0c6ada1986a1983ebeec1473a73a04795b6319d48662d40881c1723a706f516fe75300f92408aa1dc6ae4288d2046f23c1aa2e54b7fb6448a0da922bd7f34 128"

definition P256_SHA256_TV15_d :: nat where
  "P256_SHA256_TV15_d = 0x34fa4682bf6cb5b16783adcd18f0e6879b92185f76d7c920409f904f522db4b1"

definition P256_SHA256_TV15_Qx :: nat where
  "P256_SHA256_TV15_Qx = 0x105d22d9c626520faca13e7ced382dcbe93498315f00cc0ac39c4821d0d73737"

definition P256_SHA256_TV15_Qy :: nat where
  "P256_SHA256_TV15_Qy = 0x6c47f3cbbfa97dfcebe16270b8c7d5d3a5900b888c42520d751e8faf3b401ef4"

definition P256_SHA256_TV15_Q :: "int point" where
  "P256_SHA256_TV15_Q = Point (int P256_SHA256_TV15_Qx) (int P256_SHA256_TV15_Qy)"

definition P256_SHA256_TV15_k :: nat where
  "P256_SHA256_TV15_k = 0xa6f463ee72c9492bc792fe98163112837aebd07bab7a84aaed05be64db3086f4"

definition P256_SHA256_TV15_R :: nat where
  "P256_SHA256_TV15_R = 0x542c40a18140a6266d6f0286e24e9a7bad7650e72ef0e2131e629c076d962663"

definition P256_SHA256_TV15_S :: nat where
  "P256_SHA256_TV15_S = 0x4f7f65305e24a6bbb5cff714ba8f5a2cee5bdc89ba8d75dcbf21966ce38eb66f"

definition P256_SHA256_TV15_Sig :: "nat × nat" where
  "P256_SHA256_TV15_Sig = (P256_SHA256_TV15_R, P256_SHA256_TV15_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA256_TV15_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV15_d" 
  by eval

lemma P256_SHA256_TV15_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA256_TV15_Q" 
  by eval

lemma P256_SHA256_TV15_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA256_TV15_d = P256_SHA256_TV15_Q" 
  by eval

lemma P256_SHA256_TV15_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA256_TV15_d P256_SHA256_TV15_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA256_TV15_dQvalidPair' P256_SHA256_TV15_d_valid) 

lemma P256_SHA256_TV15_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA256_TV15_k" 
  by eval

lemma P256_SHA256_TV15_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA256octets P256_SHA256_TV15_d P256_SHA256_TV15_Msg P256_SHA256_TV15_k = Some P256_SHA256_TV15_Sig" 
  by eval

lemma P256_SHA256_TV15_Verify: 
  "SEC1_P256_ECDSA_Verify SHA256octets P256_SHA256_TV15_Q P256_SHA256_TV15_Msg P256_SHA256_TV15_Sig" 
  by eval

section‹NIST Test Vectors for ECDSA: Curve P-256, Hash SHA-384›
  
subsection‹Test Vector: P256_SHA384_TV01›

subsubsection‹Given Test Vector Values›

definition P256_SHA384_TV01_Msg :: octets where
  "P256_SHA384_TV01_Msg = nat_to_octets_len 0xe0b8596b375f3306bbc6e77a0b42f7469d7e83635990e74aa6d713594a3a24498feff5006790742d9c2e9b47d714bee932435db747c6e733e3d8de41f2f91311f2e9fd8e025651631ffd84f66732d3473fbd1627e63dc7194048ebec93c95c159b5039ab5e79e42c80b484a943f125de3da1e04e5bf9c16671ad55a1117d3306 128"

definition P256_SHA384_TV01_d :: nat where
  "P256_SHA384_TV01_d = 0xb6faf2c8922235c589c27368a3b3e6e2f42eb6073bf9507f19eed0746c79dced"

definition P256_SHA384_TV01_Qx :: nat where
  "P256_SHA384_TV01_Qx = 0xe0e7b99bc62d8dd67883e39ed9fa0657789c5ff556cc1fd8dd1e2a55e9e3f243"

definition P256_SHA384_TV01_Qy :: nat where
  "P256_SHA384_TV01_Qy = 0x63fbfd0232b95578075c903a4dbf85ad58f8350516e1ec89b0ee1f5e1362da69"

definition P256_SHA384_TV01_Q :: "int point" where
  "P256_SHA384_TV01_Q = Point (int P256_SHA384_TV01_Qx) (int P256_SHA384_TV01_Qy)"

definition P256_SHA384_TV01_k :: nat where
  "P256_SHA384_TV01_k = 0x9980b9cdfcef3ab8e219b9827ed6afdd4dbf20bd927e9cd01f15762703487007"

definition P256_SHA384_TV01_R :: nat where
  "P256_SHA384_TV01_R = 0xf5087878e212b703578f5c66f434883f3ef414dc23e2e8d8ab6a8d159ed5ad83"

definition P256_SHA384_TV01_S :: nat where
  "P256_SHA384_TV01_S = 0x306b4c6c20213707982dffbb30fba99b96e792163dd59dbe606e734328dd7c8a"

definition P256_SHA384_TV01_Sig :: "nat × nat" where
  "P256_SHA384_TV01_Sig = (P256_SHA384_TV01_R, P256_SHA384_TV01_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA384_TV01_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV01_d" 
  by eval

lemma P256_SHA384_TV01_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA384_TV01_Q" 
  by eval

lemma P256_SHA384_TV01_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA384_TV01_d = P256_SHA384_TV01_Q" 
  by eval

lemma P256_SHA384_TV01_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA384_TV01_d P256_SHA384_TV01_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA384_TV01_dQvalidPair' P256_SHA384_TV01_d_valid) 

lemma P256_SHA384_TV01_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV01_k" 
  by eval

lemma P256_SHA384_TV01_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA384octets P256_SHA384_TV01_d P256_SHA384_TV01_Msg P256_SHA384_TV01_k = Some P256_SHA384_TV01_Sig" 
  by eval

lemma P256_SHA384_TV01_Verify: 
  "SEC1_P256_ECDSA_Verify SHA384octets P256_SHA384_TV01_Q P256_SHA384_TV01_Msg P256_SHA384_TV01_Sig" 
  by eval

subsection‹Test Vector: P256_SHA384_TV02›

subsubsection‹Given Test Vector Values›

definition P256_SHA384_TV02_Msg :: octets where
  "P256_SHA384_TV02_Msg = nat_to_octets_len 0x099a0131179fff4c6928e49886d2fdb3a9f239b7dd5fa828a52cbbe3fcfabecfbba3e192159b887b5d13aa1e14e6a07ccbb21f6ad8b7e88fee6bea9b86dea40ffb962f38554056fb7c5bb486418915f7e7e9b9033fe3baaf9a069db98bc02fa8af3d3d1859a11375d6f98aa2ce632606d0800dff7f55b40f971a8586ed6b39e9 128"

definition P256_SHA384_TV02_d :: nat where
  "P256_SHA384_TV02_d = 0x118958fd0ff0f0b0ed11d3cf8fa664bc17cdb5fed1f4a8fc52d0b1ae30412181"

definition P256_SHA384_TV02_Qx :: nat where
  "P256_SHA384_TV02_Qx = 0xafda82260c9f42122a3f11c6058839488f6d7977f6f2a263c67d06e27ea2c355"

definition P256_SHA384_TV02_Qy :: nat where
  "P256_SHA384_TV02_Qy = 0x0ae2bbdd2207c590332c5bfeb4c8b5b16622134bd4dc55382ae806435468058b"

definition P256_SHA384_TV02_Q :: "int point" where
  "P256_SHA384_TV02_Q = Point (int P256_SHA384_TV02_Qx) (int P256_SHA384_TV02_Qy)"

definition P256_SHA384_TV02_k :: nat where
  "P256_SHA384_TV02_k = 0x23129a99eeda3d99a44a5778a46e8e7568b91c31fb7a8628c5d9820d4bed4a6b"

definition P256_SHA384_TV02_R :: nat where
  "P256_SHA384_TV02_R = 0xe446600cab1286ebc3bb332012a2f5cc33b0a5ef7291d5a62a84de5969d77946"

definition P256_SHA384_TV02_S :: nat where
  "P256_SHA384_TV02_S = 0xcf89b12793ee1792eb26283b48fa0bdcb45ae6f6ad4b02564bf786bb97057d5a"

definition P256_SHA384_TV02_Sig :: "nat × nat" where
  "P256_SHA384_TV02_Sig = (P256_SHA384_TV02_R, P256_SHA384_TV02_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA384_TV02_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV02_d" 
  by eval

lemma P256_SHA384_TV02_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA384_TV02_Q" 
  by eval

lemma P256_SHA384_TV02_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA384_TV02_d = P256_SHA384_TV02_Q" 
  by eval

lemma P256_SHA384_TV02_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA384_TV02_d P256_SHA384_TV02_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA384_TV02_dQvalidPair' P256_SHA384_TV02_d_valid) 

lemma P256_SHA384_TV02_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV02_k" 
  by eval

lemma P256_SHA384_TV02_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA384octets P256_SHA384_TV02_d P256_SHA384_TV02_Msg P256_SHA384_TV02_k = Some P256_SHA384_TV02_Sig" 
  by eval

lemma P256_SHA384_TV02_Verify: 
  "SEC1_P256_ECDSA_Verify SHA384octets P256_SHA384_TV02_Q P256_SHA384_TV02_Msg P256_SHA384_TV02_Sig" 
  by eval

subsection‹Test Vector: P256_SHA384_TV03›

subsubsection‹Given Test Vector Values›

definition P256_SHA384_TV03_Msg :: octets where
  "P256_SHA384_TV03_Msg = nat_to_octets_len 0x0fbc07ea947c946bea26afa10c51511039b94ddbc4e2e4184ca3559260da24a14522d1497ca5e77a5d1a8e86583aeea1f5d4ff9b04a6aa0de79cd88fdb85e01f171143535f2f7c23b050289d7e05cebccdd131888572534bae0061bdcc3015206b9270b0d5af9f1da2f9de91772d178a632c3261a1e7b3fb255608b3801962f9 128"

definition P256_SHA384_TV03_d :: nat where
  "P256_SHA384_TV03_d = 0x3e647357cd5b754fad0fdb876eaf9b1abd7b60536f383c81ce5745ec80826431"

definition P256_SHA384_TV03_Qx :: nat where
  "P256_SHA384_TV03_Qx = 0x702b2c94d039e590dd5c8f9736e753cf5824aacf33ee3de74fe1f5f7c858d5ed"

definition P256_SHA384_TV03_Qy :: nat where
  "P256_SHA384_TV03_Qy = 0x0c28894e907af99fb0d18c9e98f19ac80dd77abfa4bebe45055c0857b82a0f4d"

definition P256_SHA384_TV03_Q :: "int point" where
  "P256_SHA384_TV03_Q = Point (int P256_SHA384_TV03_Qx) (int P256_SHA384_TV03_Qy)"

definition P256_SHA384_TV03_k :: nat where
  "P256_SHA384_TV03_k = 0x9beab7722f0bcb468e5f234e074170a60225255de494108459abdf603c6e8b35"

definition P256_SHA384_TV03_R :: nat where
  "P256_SHA384_TV03_R = 0xc4021fb7185a07096547af1fb06932e37cf8bd90cf593dea48d48614fa237e5e"

definition P256_SHA384_TV03_S :: nat where
  "P256_SHA384_TV03_S = 0x7fb45d09e2172bec8d3e330aa06c43fbb5f625525485234e7714b7f6e92ba8f1"

definition P256_SHA384_TV03_Sig :: "nat × nat" where
  "P256_SHA384_TV03_Sig = (P256_SHA384_TV03_R, P256_SHA384_TV03_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA384_TV03_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV03_d" 
  by eval

lemma P256_SHA384_TV03_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA384_TV03_Q" 
  by eval

lemma P256_SHA384_TV03_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA384_TV03_d = P256_SHA384_TV03_Q" 
  by eval

lemma P256_SHA384_TV03_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA384_TV03_d P256_SHA384_TV03_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA384_TV03_dQvalidPair' P256_SHA384_TV03_d_valid) 

lemma P256_SHA384_TV03_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV03_k" 
  by eval

lemma P256_SHA384_TV03_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA384octets P256_SHA384_TV03_d P256_SHA384_TV03_Msg P256_SHA384_TV03_k = Some P256_SHA384_TV03_Sig" 
  by eval

lemma P256_SHA384_TV03_Verify: 
  "SEC1_P256_ECDSA_Verify SHA384octets P256_SHA384_TV03_Q P256_SHA384_TV03_Msg P256_SHA384_TV03_Sig" 
  by eval

subsection‹Test Vector: P256_SHA384_TV04›

subsubsection‹Given Test Vector Values›

definition P256_SHA384_TV04_Msg :: octets where
  "P256_SHA384_TV04_Msg = nat_to_octets_len 0x1e38d750d936d8522e9db1873fb4996bef97f8da3c6674a1223d29263f1234a90b751785316444e9ba698bc8ab6cd010638d182c9adad4e334b2bd7529f0ae8e9a52ad60f59804b2d780ed52bdd33b0bf5400147c28b4304e5e3434505ae7ce30d4b239e7e6f0ecf058badd5b388eddbad64d24d2430dd04b4ddee98f972988f 128"

definition P256_SHA384_TV04_d :: nat where
  "P256_SHA384_TV04_d = 0x76c17c2efc99891f3697ba4d71850e5816a1b65562cc39a13da4b6da9051b0fd"

definition P256_SHA384_TV04_Qx :: nat where
  "P256_SHA384_TV04_Qx = 0xd12512e934c367e4c4384dbd010e93416840288a0ba00b299b4e7c0d91578b57"

definition P256_SHA384_TV04_Qy :: nat where
  "P256_SHA384_TV04_Qy = 0xebf8835661d9b578f18d14ae4acf9c357c0dc8b7112fc32824a685ed72754e23"

definition P256_SHA384_TV04_Q :: "int point" where
  "P256_SHA384_TV04_Q = Point (int P256_SHA384_TV04_Qx) (int P256_SHA384_TV04_Qy)"

definition P256_SHA384_TV04_k :: nat where
  "P256_SHA384_TV04_k = 0x77cffa6f9a73904306f9fcd3f6bbb37f52d71e39931bb4aec28f9b076e436ccf"

definition P256_SHA384_TV04_R :: nat where
  "P256_SHA384_TV04_R = 0x4d5a9d95b0f09ce8704b0f457b39059ee606092310df65d3f8ae7a2a424cf232"

definition P256_SHA384_TV04_S :: nat where
  "P256_SHA384_TV04_S = 0x7d3c014ca470a73cef1d1da86f2a541148ad542fbccaf9149d1b0b030441a7eb"

definition P256_SHA384_TV04_Sig :: "nat × nat" where
  "P256_SHA384_TV04_Sig = (P256_SHA384_TV04_R, P256_SHA384_TV04_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA384_TV04_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV04_d" 
  by eval

lemma P256_SHA384_TV04_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA384_TV04_Q" 
  by eval

lemma P256_SHA384_TV04_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA384_TV04_d = P256_SHA384_TV04_Q" 
  by eval

lemma P256_SHA384_TV04_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA384_TV04_d P256_SHA384_TV04_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA384_TV04_dQvalidPair' P256_SHA384_TV04_d_valid) 

lemma P256_SHA384_TV04_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV04_k" 
  by eval

lemma P256_SHA384_TV04_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA384octets P256_SHA384_TV04_d P256_SHA384_TV04_Msg P256_SHA384_TV04_k = Some P256_SHA384_TV04_Sig" 
  by eval

lemma P256_SHA384_TV04_Verify: 
  "SEC1_P256_ECDSA_Verify SHA384octets P256_SHA384_TV04_Q P256_SHA384_TV04_Msg P256_SHA384_TV04_Sig" 
  by eval

subsection‹Test Vector: P256_SHA384_TV05›

subsubsection‹Given Test Vector Values›

definition P256_SHA384_TV05_Msg :: octets where
  "P256_SHA384_TV05_Msg = nat_to_octets_len 0xabcf0e0f046b2e0672d1cc6c0a114905627cbbdefdf9752f0c31660aa95f2d0ede72d17919a9e9b1add3213164e0c9b5ae3c76f1a2f79d3eeb444e6741521019d8bd5ca391b28c1063347f07afcfbb705be4b52261c19ebaf1d6f054a74d86fb5d091fa7f229450996b76f0ada5f977b09b58488eebfb5f5e9539a8fd89662ab 128"

definition P256_SHA384_TV05_d :: nat where
  "P256_SHA384_TV05_d = 0x67b9dea6a575b5103999efffce29cca688c781782a41129fdecbce76608174de"

definition P256_SHA384_TV05_Qx :: nat where
  "P256_SHA384_TV05_Qx = 0xb4238b029fc0b7d9a5286d8c29b6f3d5a569e9108d44d889cd795c4a385905be"

definition P256_SHA384_TV05_Qy :: nat where
  "P256_SHA384_TV05_Qy = 0x8cb3fff8f6cca7187c6a9ad0a2b1d9f40ae01b32a7e8f8c4ca75d71a1fffb309"

definition P256_SHA384_TV05_Q :: "int point" where
  "P256_SHA384_TV05_Q = Point (int P256_SHA384_TV05_Qx) (int P256_SHA384_TV05_Qy)"

definition P256_SHA384_TV05_k :: nat where
  "P256_SHA384_TV05_k = 0xd02617f26ede3584f0afcfc89554cdfb2ae188c192092fdde3436335fafe43f1"

definition P256_SHA384_TV05_R :: nat where
  "P256_SHA384_TV05_R = 0x26fd9147d0c86440689ff2d75569795650140506970791c90ace0924b44f1586"

definition P256_SHA384_TV05_S :: nat where
  "P256_SHA384_TV05_S = 0x00a34b00c20a8099df4b0a757cbef8fea1cb3ea7ced5fbf7e987f70b25ee6d4f"

definition P256_SHA384_TV05_Sig :: "nat × nat" where
  "P256_SHA384_TV05_Sig = (P256_SHA384_TV05_R, P256_SHA384_TV05_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA384_TV05_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV05_d" 
  by eval

lemma P256_SHA384_TV05_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA384_TV05_Q" 
  by eval

lemma P256_SHA384_TV05_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA384_TV05_d = P256_SHA384_TV05_Q" 
  by eval

lemma P256_SHA384_TV05_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA384_TV05_d P256_SHA384_TV05_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA384_TV05_dQvalidPair' P256_SHA384_TV05_d_valid) 

lemma P256_SHA384_TV05_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV05_k" 
  by eval

lemma P256_SHA384_TV05_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA384octets P256_SHA384_TV05_d P256_SHA384_TV05_Msg P256_SHA384_TV05_k = Some P256_SHA384_TV05_Sig" 
  by eval

lemma P256_SHA384_TV05_Verify: 
  "SEC1_P256_ECDSA_Verify SHA384octets P256_SHA384_TV05_Q P256_SHA384_TV05_Msg P256_SHA384_TV05_Sig" 
  by eval

subsection‹Test Vector: P256_SHA384_TV06›

subsubsection‹Given Test Vector Values›

definition P256_SHA384_TV06_Msg :: octets where
  "P256_SHA384_TV06_Msg = nat_to_octets_len 0xdc3d4884c741a4a687593c79fb4e35c5c13c781dca16db561d7e393577f7b62ca41a6e259fc1fb8d0c4e1e062517a0fdf95558b7799f20c211796167953e6372c11829beec64869d67bf3ee1f1455dd87acfbdbcc597056e7fb347a17688ad32fda7ccc3572da7677d7255c261738f07763cd45973c728c6e9adbeecadc3d961 128"

definition P256_SHA384_TV06_d :: nat where
  "P256_SHA384_TV06_d = 0xecf644ea9b6c3a04fdfe2de4fdcb55fdcdfcf738c0b3176575fa91515194b566"

definition P256_SHA384_TV06_Qx :: nat where
  "P256_SHA384_TV06_Qx = 0xc3bdc7c795ec94620a2cfff614c13a3390a5e86c892e53a24d3ed22228bc85bf"

definition P256_SHA384_TV06_Qy :: nat where
  "P256_SHA384_TV06_Qy = 0x70480fc5cf4aacd73e24618b61b5c56c1ced8c4f1b869580ea538e68c7a61ca3"

definition P256_SHA384_TV06_Q :: "int point" where
  "P256_SHA384_TV06_Q = Point (int P256_SHA384_TV06_Qx) (int P256_SHA384_TV06_Qy)"

definition P256_SHA384_TV06_k :: nat where
  "P256_SHA384_TV06_k = 0x53291d51f68d9a12d1dcdc58892b2f786cc15f631f16997d2a49bace513557d4"

definition P256_SHA384_TV06_R :: nat where
  "P256_SHA384_TV06_R = 0xa860c8b286edf973ce4ce4cf6e70dc9bbf3818c36c023a845677a9963705df8b"

definition P256_SHA384_TV06_S :: nat where
  "P256_SHA384_TV06_S = 0x5630f986b1c45e36e127dd7932221c4272a8cc6e255e89f0f0ca4ec3a9f76494"

definition P256_SHA384_TV06_Sig :: "nat × nat" where
  "P256_SHA384_TV06_Sig = (P256_SHA384_TV06_R, P256_SHA384_TV06_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA384_TV06_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV06_d" 
  by eval

lemma P256_SHA384_TV06_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA384_TV06_Q" 
  by eval

lemma P256_SHA384_TV06_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA384_TV06_d = P256_SHA384_TV06_Q" 
  by eval

lemma P256_SHA384_TV06_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA384_TV06_d P256_SHA384_TV06_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA384_TV06_dQvalidPair' P256_SHA384_TV06_d_valid) 

lemma P256_SHA384_TV06_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV06_k" 
  by eval

lemma P256_SHA384_TV06_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA384octets P256_SHA384_TV06_d P256_SHA384_TV06_Msg P256_SHA384_TV06_k = Some P256_SHA384_TV06_Sig" 
  by eval

lemma P256_SHA384_TV06_Verify: 
  "SEC1_P256_ECDSA_Verify SHA384octets P256_SHA384_TV06_Q P256_SHA384_TV06_Msg P256_SHA384_TV06_Sig" 
  by eval

subsection‹Test Vector: P256_SHA384_TV07›

subsubsection‹Given Test Vector Values›

definition P256_SHA384_TV07_Msg :: octets where
  "P256_SHA384_TV07_Msg = nat_to_octets_len 0x719bf1911ae5b5e08f1d97b92a5089c0ab9d6f1c175ac7199086aeeaa416a17e6d6f8486c711d386f284f096296689a54d330c8efb0f5fa1c5ba128d3234a3da856c2a94667ef7103616a64c913135f4e1dc50e38daa60610f732ad1bedfcc396f87169392520314a6b6b9af6793dbabad4599525228cc7c9c32c4d8e097ddf6 128"

definition P256_SHA384_TV07_d :: nat where
  "P256_SHA384_TV07_d = 0x4961485cbc978f8456ec5ac7cfc9f7d9298f99415ecae69c8491b258c029bfee"

definition P256_SHA384_TV07_Qx :: nat where
  "P256_SHA384_TV07_Qx = 0x8d40bf2299e05d758d421972e81cfb0cce68b949240dc30f315836acc70bef03"

definition P256_SHA384_TV07_Qy :: nat where
  "P256_SHA384_TV07_Qy = 0x5674e6f77f8b46f46cca937d83b128dffbe9bd7e0d3d08aa2cbbfdfb16f72c9a"

definition P256_SHA384_TV07_Q :: "int point" where
  "P256_SHA384_TV07_Q = Point (int P256_SHA384_TV07_Qx) (int P256_SHA384_TV07_Qy)"

definition P256_SHA384_TV07_k :: nat where
  "P256_SHA384_TV07_k = 0x373a825b5a74b7b9e02f8d4d876b577b4c3984168d704ba9f95b19c05ed590af"

definition P256_SHA384_TV07_R :: nat where
  "P256_SHA384_TV07_R = 0xef6fb386ad044b63feb7445fa16b10319018e9cea9ef42bca83bdad01992234a"

definition P256_SHA384_TV07_S :: nat where
  "P256_SHA384_TV07_S = 0xac1f42f652eb1786e57be01d847c81f7efa072ba566d4583af4f1551a3f76c65"

definition P256_SHA384_TV07_Sig :: "nat × nat" where
  "P256_SHA384_TV07_Sig = (P256_SHA384_TV07_R, P256_SHA384_TV07_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA384_TV07_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV07_d" 
  by eval

lemma P256_SHA384_TV07_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA384_TV07_Q" 
  by eval

lemma P256_SHA384_TV07_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA384_TV07_d = P256_SHA384_TV07_Q" 
  by eval

lemma P256_SHA384_TV07_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA384_TV07_d P256_SHA384_TV07_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA384_TV07_dQvalidPair' P256_SHA384_TV07_d_valid) 

lemma P256_SHA384_TV07_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV07_k" 
  by eval

lemma P256_SHA384_TV07_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA384octets P256_SHA384_TV07_d P256_SHA384_TV07_Msg P256_SHA384_TV07_k = Some P256_SHA384_TV07_Sig" 
  by eval

lemma P256_SHA384_TV07_Verify: 
  "SEC1_P256_ECDSA_Verify SHA384octets P256_SHA384_TV07_Q P256_SHA384_TV07_Msg P256_SHA384_TV07_Sig" 
  by eval

subsection‹Test Vector: P256_SHA384_TV08›

subsubsection‹Given Test Vector Values›

definition P256_SHA384_TV08_Msg :: octets where
  "P256_SHA384_TV08_Msg = nat_to_octets_len 0x7cf19f4c851e97c5bca11a39f0074c3b7bd3274e7dd75d0447b7b84995dfc9f716bf08c25347f56fcc5e5149cb3f9cfb39d408ace5a5c47e75f7a827fa0bb9921bb5b23a6053dbe1fa2bba341ac874d9b1333fc4dc224854949f5c8d8a5fedd02fb26fdfcd3be351aec0fcbef18972956c6ec0effaf057eb4420b6d28e0c008c 128"

definition P256_SHA384_TV08_d :: nat where
  "P256_SHA384_TV08_d = 0x587907e7f215cf0d2cb2c9e6963d45b6e535ed426c828a6ea2fb637cca4c5cbd"

definition P256_SHA384_TV08_Qx :: nat where
  "P256_SHA384_TV08_Qx = 0x660da45c413cc9c9526202c16b402af602d30daaa7c342f1e722f15199407f31"

definition P256_SHA384_TV08_Qy :: nat where
  "P256_SHA384_TV08_Qy = 0xe6f8cbb06913cc718f2d69ba2fb3137f04a41c27c676d1a80fbf30ea3ca46439"

definition P256_SHA384_TV08_Q :: "int point" where
  "P256_SHA384_TV08_Q = Point (int P256_SHA384_TV08_Qx) (int P256_SHA384_TV08_Qy)"

definition P256_SHA384_TV08_k :: nat where
  "P256_SHA384_TV08_k = 0x6b8eb7c0d8af9456b95dd70561a0e902863e6dfa1c28d0fd4a0509f1c2a647b2"

definition P256_SHA384_TV08_R :: nat where
  "P256_SHA384_TV08_R = 0x08fabf9b57de81875bfa7a4118e3e44cfb38ec6a9b2014940207ba3b1c583038"

definition P256_SHA384_TV08_S :: nat where
  "P256_SHA384_TV08_S = 0xa58d199b1deba7350616230d867b2747a3459421811c291836abee715b8f67b4"

definition P256_SHA384_TV08_Sig :: "nat × nat" where
  "P256_SHA384_TV08_Sig = (P256_SHA384_TV08_R, P256_SHA384_TV08_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA384_TV08_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV08_d" 
  by eval

lemma P256_SHA384_TV08_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA384_TV08_Q" 
  by eval

lemma P256_SHA384_TV08_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA384_TV08_d = P256_SHA384_TV08_Q" 
  by eval

lemma P256_SHA384_TV08_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA384_TV08_d P256_SHA384_TV08_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA384_TV08_dQvalidPair' P256_SHA384_TV08_d_valid) 

lemma P256_SHA384_TV08_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV08_k" 
  by eval

lemma P256_SHA384_TV08_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA384octets P256_SHA384_TV08_d P256_SHA384_TV08_Msg P256_SHA384_TV08_k = Some P256_SHA384_TV08_Sig" 
  by eval

lemma P256_SHA384_TV08_Verify: 
  "SEC1_P256_ECDSA_Verify SHA384octets P256_SHA384_TV08_Q P256_SHA384_TV08_Msg P256_SHA384_TV08_Sig" 
  by eval

subsection‹Test Vector: P256_SHA384_TV09›

subsubsection‹Given Test Vector Values›

definition P256_SHA384_TV09_Msg :: octets where
  "P256_SHA384_TV09_Msg = nat_to_octets_len 0xb892ffabb809e98a99b0a79895445fc734fa1b6159f9cddb6d21e510708bdab6076633ac30aaef43db566c0d21f4381db46711fe3812c5ce0fb4a40e3d5d8ab24e4e82d3560c6dc7c37794ee17d4a144065ef99c8d1c88bc22ad8c4c27d85ad518fa5747ae35276fc104829d3f5c72fc2a9ea55a1c3a87007cd133263f79e405 128"

definition P256_SHA384_TV09_d :: nat where
  "P256_SHA384_TV09_d = 0x24b1e5676d1a9d6b645a984141a157c124531feeb92d915110aef474b1e27666"

definition P256_SHA384_TV09_Qx :: nat where
  "P256_SHA384_TV09_Qx = 0xb4909a5bdf25f7659f4ef35e4b811429fb2c59126e3dad09100b46aea6ebe7a6"

definition P256_SHA384_TV09_Qy :: nat where
  "P256_SHA384_TV09_Qy = 0x760ae015fa6af5c9749c4030fdb5de6e58c6b5b1944829105cf7edf7d3a22cfb"

definition P256_SHA384_TV09_Q :: "int point" where
  "P256_SHA384_TV09_Q = Point (int P256_SHA384_TV09_Qx) (int P256_SHA384_TV09_Qy)"

definition P256_SHA384_TV09_k :: nat where
  "P256_SHA384_TV09_k = 0x88794923d8943b5dbcc7a7a76503880ff7da632b0883aaa60a9fcc71bf880fd6"

definition P256_SHA384_TV09_R :: nat where
  "P256_SHA384_TV09_R = 0x6ec9a340b77fae3c7827fa96d997e92722ff2a928217b6dd3c628f3d49ae4ce6"

definition P256_SHA384_TV09_S :: nat where
  "P256_SHA384_TV09_S = 0x637b54bbcfb7e7d8a41ea317fcfca8ad74eb3bb6b778bc7ef9dec009281976f7"

definition P256_SHA384_TV09_Sig :: "nat × nat" where
  "P256_SHA384_TV09_Sig = (P256_SHA384_TV09_R, P256_SHA384_TV09_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA384_TV09_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV09_d" 
  by eval

lemma P256_SHA384_TV09_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA384_TV09_Q" 
  by eval

lemma P256_SHA384_TV09_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA384_TV09_d = P256_SHA384_TV09_Q" 
  by eval

lemma P256_SHA384_TV09_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA384_TV09_d P256_SHA384_TV09_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA384_TV09_dQvalidPair' P256_SHA384_TV09_d_valid) 

lemma P256_SHA384_TV09_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV09_k" 
  by eval

lemma P256_SHA384_TV09_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA384octets P256_SHA384_TV09_d P256_SHA384_TV09_Msg P256_SHA384_TV09_k = Some P256_SHA384_TV09_Sig" 
  by eval

lemma P256_SHA384_TV09_Verify: 
  "SEC1_P256_ECDSA_Verify SHA384octets P256_SHA384_TV09_Q P256_SHA384_TV09_Msg P256_SHA384_TV09_Sig" 
  by eval

subsection‹Test Vector: P256_SHA384_TV10›

subsubsection‹Given Test Vector Values›

definition P256_SHA384_TV10_Msg :: octets where
  "P256_SHA384_TV10_Msg = nat_to_octets_len 0x8144e37014c95e13231cbd6fa64772771f93b44e37f7b02f592099cc146343edd4f4ec9fa1bc68d7f2e9ee78fc370443aa2803ff4ca52ee49a2f4daf2c8181ea7b8475b3a0f608fc3279d09e2d057fbe3f2ffbe5133796124781299c6da60cfe7ecea3abc30706ded2cdf18f9d788e59f2c31662df3abe01a9b12304fb8d5c8c 128"

definition P256_SHA384_TV10_d :: nat where
  "P256_SHA384_TV10_d = 0xbce49c7b03dcdc72393b0a67cf5aa5df870f5aaa6137ada1edc7862e0981ec67"

definition P256_SHA384_TV10_Qx :: nat where
  "P256_SHA384_TV10_Qx = 0xc786d9421d67b72b922cf3def2a25eeb5e73f34543eb50b152e738a98afb0ca5"

definition P256_SHA384_TV10_Qy :: nat where
  "P256_SHA384_TV10_Qy = 0x6796271e79e2496f9e74b126b1123a3d067de56b5605d6f51c8f6e1d5bb93aba"

definition P256_SHA384_TV10_Q :: "int point" where
  "P256_SHA384_TV10_Q = Point (int P256_SHA384_TV10_Qx) (int P256_SHA384_TV10_Qy)"

definition P256_SHA384_TV10_k :: nat where
  "P256_SHA384_TV10_k = 0x89e690d78a5e0d2b8ce9f7fcbf34e2605fd9584760fa7729043397612dd21f94"

definition P256_SHA384_TV10_R :: nat where
  "P256_SHA384_TV10_R = 0x07e5054c384839584624e8d730454dc27e673c4a90cbf129d88b91250341854d"

definition P256_SHA384_TV10_S :: nat where
  "P256_SHA384_TV10_S = 0xf7e665b88614d0c5cbb3007cafe713763d81831525971f1747d92e4d1ca263a7"

definition P256_SHA384_TV10_Sig :: "nat × nat" where
  "P256_SHA384_TV10_Sig = (P256_SHA384_TV10_R, P256_SHA384_TV10_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA384_TV10_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV10_d" 
  by eval

lemma P256_SHA384_TV10_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA384_TV10_Q" 
  by eval

lemma P256_SHA384_TV10_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA384_TV10_d = P256_SHA384_TV10_Q" 
  by eval

lemma P256_SHA384_TV10_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA384_TV10_d P256_SHA384_TV10_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA384_TV10_dQvalidPair' P256_SHA384_TV10_d_valid) 

lemma P256_SHA384_TV10_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV10_k" 
  by eval

lemma P256_SHA384_TV10_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA384octets P256_SHA384_TV10_d P256_SHA384_TV10_Msg P256_SHA384_TV10_k = Some P256_SHA384_TV10_Sig" 
  by eval

lemma P256_SHA384_TV10_Verify: 
  "SEC1_P256_ECDSA_Verify SHA384octets P256_SHA384_TV10_Q P256_SHA384_TV10_Msg P256_SHA384_TV10_Sig" 
  by eval

subsection‹Test Vector: P256_SHA384_TV11›

subsubsection‹Given Test Vector Values›

definition P256_SHA384_TV11_Msg :: octets where
  "P256_SHA384_TV11_Msg = nat_to_octets_len 0xa3683d120807f0a030feed679785326698c3702f1983eaba1b70ddfa7f0b3188060b845e2b67ed57ee68087746710450f7427cb34655d719c0acbc09ac696adb4b22aba1b9322b7111076e67053a55f62b501a4bca0ad9d50a868f51aeeb4ef27823236f5267e8da83e143047422ce140d66e05e44dc84fb3a4506b2a5d7caa8 128"

definition P256_SHA384_TV11_d :: nat where
  "P256_SHA384_TV11_d = 0x73188a923bc0b289e81c3db48d826917910f1b957700f8925425c1fb27cabab9"

definition P256_SHA384_TV11_Qx :: nat where
  "P256_SHA384_TV11_Qx = 0x86662c014ab666ee770723be8da38c5cd299efc6480fc6f8c3603438fa8397b9"

definition P256_SHA384_TV11_Qy :: nat where
  "P256_SHA384_TV11_Qy = 0xf26b3307a650c3863faaa5f642f3ba1384c3d3a02edd3d48c657c269609cc3fc"

definition P256_SHA384_TV11_Q :: "int point" where
  "P256_SHA384_TV11_Q = Point (int P256_SHA384_TV11_Qx) (int P256_SHA384_TV11_Qy)"

definition P256_SHA384_TV11_k :: nat where
  "P256_SHA384_TV11_k = 0xec90584ab3b383b590626f36ed4f5110e49888aec7ae7a9c5ea62dd2dc378666"

definition P256_SHA384_TV11_R :: nat where
  "P256_SHA384_TV11_R = 0x13e9ad59112fde3af4163eb5c2400b5e9a602576d5869ac1c569075f08c90ff6"

definition P256_SHA384_TV11_S :: nat where
  "P256_SHA384_TV11_S = 0x708ac65ff2b0baaccc6dd954e2a93df46016bd04457636de06798fcc17f02be5"

definition P256_SHA384_TV11_Sig :: "nat × nat" where
  "P256_SHA384_TV11_Sig = (P256_SHA384_TV11_R, P256_SHA384_TV11_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA384_TV11_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV11_d" 
  by eval

lemma P256_SHA384_TV11_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA384_TV11_Q" 
  by eval

lemma P256_SHA384_TV11_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA384_TV11_d = P256_SHA384_TV11_Q" 
  by eval

lemma P256_SHA384_TV11_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA384_TV11_d P256_SHA384_TV11_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA384_TV11_dQvalidPair' P256_SHA384_TV11_d_valid) 

lemma P256_SHA384_TV11_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV11_k" 
  by eval

lemma P256_SHA384_TV11_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA384octets P256_SHA384_TV11_d P256_SHA384_TV11_Msg P256_SHA384_TV11_k = Some P256_SHA384_TV11_Sig" 
  by eval

lemma P256_SHA384_TV11_Verify: 
  "SEC1_P256_ECDSA_Verify SHA384octets P256_SHA384_TV11_Q P256_SHA384_TV11_Msg P256_SHA384_TV11_Sig" 
  by eval

subsection‹Test Vector: P256_SHA384_TV12›

subsubsection‹Given Test Vector Values›

definition P256_SHA384_TV12_Msg :: octets where
  "P256_SHA384_TV12_Msg = nat_to_octets_len 0xb1df8051b213fc5f636537e37e212eb20b2423e6467a9c7081336a870e6373fc835899d59e546c0ac668cc81ce4921e88f42e6da2a109a03b4f4e819a17c955b8d099ec6b282fb495258dca13ec779c459da909475519a3477223c06b99afbd77f9922e7cbef844b93f3ce5f50db816b2e0d8b1575d2e17a6b8db9111d6da578 128"

definition P256_SHA384_TV12_d :: nat where
  "P256_SHA384_TV12_d = 0xf637d55763fe819541588e0c603f288a693cc66823c6bb7b8e003bd38580ebce"

definition P256_SHA384_TV12_Qx :: nat where
  "P256_SHA384_TV12_Qx = 0x74a4620c578601475fc169a9b84be613b4a16cb6acab8fd98848a6ec9fbd133d"

definition P256_SHA384_TV12_Qy :: nat where
  "P256_SHA384_TV12_Qy = 0x42b9e35d347c107e63bd55f525f915bcf1e3d2b81d002d3c39acf10fc30645a1"

definition P256_SHA384_TV12_Q :: "int point" where
  "P256_SHA384_TV12_Q = Point (int P256_SHA384_TV12_Qx) (int P256_SHA384_TV12_Qy)"

definition P256_SHA384_TV12_k :: nat where
  "P256_SHA384_TV12_k = 0x4d578f5099636234d9c1d566f1215d5d887ae5d47022be17dbf32a11a03f053b"

definition P256_SHA384_TV12_R :: nat where
  "P256_SHA384_TV12_R = 0x113a933ebc4d94ce1cef781e4829df0c493b0685d39fb2048ce01b21c398dbba"

definition P256_SHA384_TV12_S :: nat where
  "P256_SHA384_TV12_S = 0x3005bd4ec63dbd04ce9ff0c6246ad65d27fcf62edb2b7e461589f9f0e7446ffd"

definition P256_SHA384_TV12_Sig :: "nat × nat" where
  "P256_SHA384_TV12_Sig = (P256_SHA384_TV12_R, P256_SHA384_TV12_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA384_TV12_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV12_d" 
  by eval

lemma P256_SHA384_TV12_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA384_TV12_Q" 
  by eval

lemma P256_SHA384_TV12_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA384_TV12_d = P256_SHA384_TV12_Q" 
  by eval

lemma P256_SHA384_TV12_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA384_TV12_d P256_SHA384_TV12_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA384_TV12_dQvalidPair' P256_SHA384_TV12_d_valid) 

lemma P256_SHA384_TV12_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV12_k" 
  by eval

lemma P256_SHA384_TV12_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA384octets P256_SHA384_TV12_d P256_SHA384_TV12_Msg P256_SHA384_TV12_k = Some P256_SHA384_TV12_Sig" 
  by eval

lemma P256_SHA384_TV12_Verify: 
  "SEC1_P256_ECDSA_Verify SHA384octets P256_SHA384_TV12_Q P256_SHA384_TV12_Msg P256_SHA384_TV12_Sig" 
  by eval

subsection‹Test Vector: P256_SHA384_TV13›

subsubsection‹Given Test Vector Values›

definition P256_SHA384_TV13_Msg :: octets where
  "P256_SHA384_TV13_Msg = nat_to_octets_len 0x0b918ede985b5c491797d0a81446b2933be312f419b212e3aae9ba5914c00af431747a9d287a7c7761e9bcbc8a12aaf9d4a76d13dad59fc742f8f218ef66eb67035220a07acc1a357c5b562ecb6b895cf725c4230412fefac72097f2c2b829ed58742d7c327cad0f1058df1bddd4ae9c6d2aba25480424308684cecd6517cdd8 128"

definition P256_SHA384_TV13_d :: nat where
  "P256_SHA384_TV13_d = 0x2e357d51517ff93b821f895932fddded8347f32596b812308e6f1baf7dd8a47f"

definition P256_SHA384_TV13_Qx :: nat where
  "P256_SHA384_TV13_Qx = 0x7e4078a1d50c669fb2996dd9bacb0c3ac7ede4f58fa0fa1222e78dbf5d1f4186"

definition P256_SHA384_TV13_Qy :: nat where
  "P256_SHA384_TV13_Qy = 0x0014e46e90cc171fbb83ea34c6b78202ea8137a7d926f0169147ed5ae3d6596f"

definition P256_SHA384_TV13_Q :: "int point" where
  "P256_SHA384_TV13_Q = Point (int P256_SHA384_TV13_Qx) (int P256_SHA384_TV13_Qy)"

definition P256_SHA384_TV13_k :: nat where
  "P256_SHA384_TV13_k = 0xbe522b0940b9a40d84bf790fe6abdc252877e671f2efa63a33a65a512fc2aa5c"

definition P256_SHA384_TV13_R :: nat where
  "P256_SHA384_TV13_R = 0xa26b9ad775ac37ff4c7f042cdc4872c5e4e5e800485f488ddfaaed379f468090"

definition P256_SHA384_TV13_S :: nat where
  "P256_SHA384_TV13_S = 0xf88eae2019bebbba62b453b8ee3472ca5c67c267964cffe0cf2d2933c1723dff"

definition P256_SHA384_TV13_Sig :: "nat × nat" where
  "P256_SHA384_TV13_Sig = (P256_SHA384_TV13_R, P256_SHA384_TV13_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA384_TV13_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV13_d" 
  by eval

lemma P256_SHA384_TV13_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA384_TV13_Q" 
  by eval

lemma P256_SHA384_TV13_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA384_TV13_d = P256_SHA384_TV13_Q" 
  by eval

lemma P256_SHA384_TV13_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA384_TV13_d P256_SHA384_TV13_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA384_TV13_dQvalidPair' P256_SHA384_TV13_d_valid) 

lemma P256_SHA384_TV13_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV13_k" 
  by eval

lemma P256_SHA384_TV13_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA384octets P256_SHA384_TV13_d P256_SHA384_TV13_Msg P256_SHA384_TV13_k = Some P256_SHA384_TV13_Sig" 
  by eval

lemma P256_SHA384_TV13_Verify: 
  "SEC1_P256_ECDSA_Verify SHA384octets P256_SHA384_TV13_Q P256_SHA384_TV13_Msg P256_SHA384_TV13_Sig" 
  by eval

subsection‹Test Vector: P256_SHA384_TV14›

subsubsection‹Given Test Vector Values›

definition P256_SHA384_TV14_Msg :: octets where
  "P256_SHA384_TV14_Msg = nat_to_octets_len 0x0fab26fde1a4467ca930dbe513ccc3452b70313cccde2994eead2fde85c8da1db84d7d06a024c9e88629d5344224a4eae01b21a2665d5f7f36d5524bf5367d7f8b6a71ea05d413d4afde33777f0a3be49c9e6aa29ea447746a9e77ce27232a550b31dd4e7c9bc8913485f2dc83a56298051c92461fd46b14cc895c300a4fb874 128"

definition P256_SHA384_TV14_d :: nat where
  "P256_SHA384_TV14_d = 0x77d60cacbbac86ab89009403c97289b5900466856887d3e6112af427f7f0f50b"

definition P256_SHA384_TV14_Qx :: nat where
  "P256_SHA384_TV14_Qx = 0xa62032dfdb87e25ed0c70cad20d927c7effeb2638e6c88ddd670f74df16090e5"

definition P256_SHA384_TV14_Qy :: nat where
  "P256_SHA384_TV14_Qy = 0x44c5ee2cf740ded468f5d2efe13daa7c5234645a37c073af35330d03a4fed976"

definition P256_SHA384_TV14_Q :: "int point" where
  "P256_SHA384_TV14_Q = Point (int P256_SHA384_TV14_Qx) (int P256_SHA384_TV14_Qy)"

definition P256_SHA384_TV14_k :: nat where
  "P256_SHA384_TV14_k = 0x06c1e692b045f425a21347ecf72833d0242906c7c1094f805566cdcb1256e394"

definition P256_SHA384_TV14_R :: nat where
  "P256_SHA384_TV14_R = 0xeb173b51fb0aec318950d097e7fda5c34e529519631c3e2c9b4550b903da417d"

definition P256_SHA384_TV14_S :: nat where
  "P256_SHA384_TV14_S = 0xca2c13574bf1b7d56e9dc18315036a31b8bceddf3e2c2902dcb40f0cc9e31b45"

definition P256_SHA384_TV14_Sig :: "nat × nat" where
  "P256_SHA384_TV14_Sig = (P256_SHA384_TV14_R, P256_SHA384_TV14_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA384_TV14_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV14_d" 
  by eval

lemma P256_SHA384_TV14_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA384_TV14_Q" 
  by eval

lemma P256_SHA384_TV14_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA384_TV14_d = P256_SHA384_TV14_Q" 
  by eval

lemma P256_SHA384_TV14_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA384_TV14_d P256_SHA384_TV14_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA384_TV14_dQvalidPair' P256_SHA384_TV14_d_valid) 

lemma P256_SHA384_TV14_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV14_k" 
  by eval

lemma P256_SHA384_TV14_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA384octets P256_SHA384_TV14_d P256_SHA384_TV14_Msg P256_SHA384_TV14_k = Some P256_SHA384_TV14_Sig" 
  by eval

lemma P256_SHA384_TV14_Verify: 
  "SEC1_P256_ECDSA_Verify SHA384octets P256_SHA384_TV14_Q P256_SHA384_TV14_Msg P256_SHA384_TV14_Sig" 
  by eval

subsection‹Test Vector: P256_SHA384_TV15›

subsubsection‹Given Test Vector Values›

definition P256_SHA384_TV15_Msg :: octets where
  "P256_SHA384_TV15_Msg = nat_to_octets_len 0x7843f157ef8566722a7d69da67de7599ee65cb3975508f70c612b3289190e364141781e0b832f2d9627122742f4b5871ceeafcd09ba5ec90cae6bcc01ae32b50f13f63918dfb5177df9797c6273b92d103c3f7a3fc2050d2b196cc872c57b77f9bdb1782d4195445fcc6236dd8bd14c8bcbc8223a6739f6a17c9a861e8c821a6 128"

definition P256_SHA384_TV15_d :: nat where
  "P256_SHA384_TV15_d = 0x486854e77962117f49e09378de6c9e3b3522fa752b10b2c810bf48db584d7388"

definition P256_SHA384_TV15_Qx :: nat where
  "P256_SHA384_TV15_Qx = 0x760b5624bd64d19c866e54ccd74ad7f98851afdbc3ddeae3ec2c52a135be9cfa"

definition P256_SHA384_TV15_Qy :: nat where
  "P256_SHA384_TV15_Qy = 0xfeca15ce9350877102eee0f5af18b2fed89dc86b7df0bf7bc2963c1638e36fe8"

definition P256_SHA384_TV15_Q :: "int point" where
  "P256_SHA384_TV15_Q = Point (int P256_SHA384_TV15_Qx) (int P256_SHA384_TV15_Qy)"

definition P256_SHA384_TV15_k :: nat where
  "P256_SHA384_TV15_k = 0xe4f77c6442eca239b01b0254e11a4182782d96f48ab521cc3d1d68df12b5a41a"

definition P256_SHA384_TV15_R :: nat where
  "P256_SHA384_TV15_R = 0xbdff14e4600309c2c77f79a25963a955b5b500a7b2d34cb172cd6acd52905c7b"

definition P256_SHA384_TV15_S :: nat where
  "P256_SHA384_TV15_S = 0xb0479cdb3df79923ec36a104a129534c5d59f622be7d613aa04530ad2507d3a2"

definition P256_SHA384_TV15_Sig :: "nat × nat" where
  "P256_SHA384_TV15_Sig = (P256_SHA384_TV15_R, P256_SHA384_TV15_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA384_TV15_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV15_d" 
  by eval

lemma P256_SHA384_TV15_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA384_TV15_Q" 
  by eval

lemma P256_SHA384_TV15_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA384_TV15_d = P256_SHA384_TV15_Q" 
  by eval

lemma P256_SHA384_TV15_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA384_TV15_d P256_SHA384_TV15_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA384_TV15_dQvalidPair' P256_SHA384_TV15_d_valid) 

lemma P256_SHA384_TV15_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA384_TV15_k" 
  by eval

lemma P256_SHA384_TV15_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA384octets P256_SHA384_TV15_d P256_SHA384_TV15_Msg P256_SHA384_TV15_k = Some P256_SHA384_TV15_Sig" 
  by eval

lemma P256_SHA384_TV15_Verify: 
  "SEC1_P256_ECDSA_Verify SHA384octets P256_SHA384_TV15_Q P256_SHA384_TV15_Msg P256_SHA384_TV15_Sig" 
  by eval

section‹NIST Test Vectors for ECDSA: Curve P-256, Hash SHA-512›
  
subsection‹Test Vector: P256_SHA512_TV01›

subsubsection‹Given Test Vector Values›

definition P256_SHA512_TV01_Msg :: octets where
  "P256_SHA512_TV01_Msg = nat_to_octets_len 0x6c8572b6a3a4a9e8e03dbeed99334d41661b8a8417074f335ab1845f6cc852adb8c01d9820fcf8e10699cc827a8fbdca2cbd46cc66e4e6b7ba41ec3efa733587e4a30ec552cd8ddab8163e148e50f4d090782897f3ddac84a41e1fcfe8c56b6152c0097b0d634b41011471ffd004f43eb4aafc038197ec6bae2b4470e869bded 128"

definition P256_SHA512_TV01_d :: nat where
  "P256_SHA512_TV01_d = 0x9dd0d3a3d514c2a8adb162b81e3adfba3299309f7d2018f607bdb15b1a25f499"

definition P256_SHA512_TV01_Qx :: nat where
  "P256_SHA512_TV01_Qx = 0x6b738de3398b6ac57b9591f9d7985dd4f32137ad3460dcf8970c1390cb9eaf8d"

definition P256_SHA512_TV01_Qy :: nat where
  "P256_SHA512_TV01_Qy = 0x83bc61e26d2bbbd3cf2d2ab445a2bc4ab5dde41f4a13078fd1d3cc36ab596d57"

definition P256_SHA512_TV01_Q :: "int point" where
  "P256_SHA512_TV01_Q = Point (int P256_SHA512_TV01_Qx) (int P256_SHA512_TV01_Qy)"

definition P256_SHA512_TV01_k :: nat where
  "P256_SHA512_TV01_k = 0x9106192170ccb3c64684d48287bb81bbed51b40d503462c900e5c7aae43e380a"

definition P256_SHA512_TV01_R :: nat where
  "P256_SHA512_TV01_R = 0x275fa760878b4dc05e9d157fedfd8e9b1c9c861222a712748cb4b7754c043fb1"

definition P256_SHA512_TV01_S :: nat where
  "P256_SHA512_TV01_S = 0x699d906bb8435a05345af3b37e3b357786939e94caae257852f0503adb1e0f7e"

definition P256_SHA512_TV01_Sig :: "nat × nat" where
  "P256_SHA512_TV01_Sig = (P256_SHA512_TV01_R, P256_SHA512_TV01_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA512_TV01_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV01_d" 
  by eval

lemma P256_SHA512_TV01_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA512_TV01_Q" 
  by eval

lemma P256_SHA512_TV01_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA512_TV01_d = P256_SHA512_TV01_Q" 
  by eval

lemma P256_SHA512_TV01_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA512_TV01_d P256_SHA512_TV01_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA512_TV01_dQvalidPair' P256_SHA512_TV01_d_valid) 

lemma P256_SHA512_TV01_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV01_k" 
  by eval

lemma P256_SHA512_TV01_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA512octets P256_SHA512_TV01_d P256_SHA512_TV01_Msg P256_SHA512_TV01_k = Some P256_SHA512_TV01_Sig" 
  by eval

lemma P256_SHA512_TV01_Verify: 
  "SEC1_P256_ECDSA_Verify SHA512octets P256_SHA512_TV01_Q P256_SHA512_TV01_Msg P256_SHA512_TV01_Sig" 
  by eval

subsection‹Test Vector: P256_SHA512_TV02›

subsubsection‹Given Test Vector Values›

definition P256_SHA512_TV02_Msg :: octets where
  "P256_SHA512_TV02_Msg = nat_to_octets_len 0x7e3c8fe162d48cc8c5b11b5e5ebc05ebc45c439bdbc0b0902145921b8383037cb0812222031598cd1a56fa71694fbd304cc62938233465ec39c6e49f57dfe823983b6923c4e865633949183e6b90e9e06d8275f3907d97967d47b6239fe2847b7d49cf16ba69d2862083cf1bccf7afe34fdc90e21998964107b64abe6b89d126 128"

definition P256_SHA512_TV02_d :: nat where
  "P256_SHA512_TV02_d = 0xf9bf909b7973bf0e3dad0e43dcb2d7fa8bda49dbe6e5357f8f0e2bd119be30e6"

definition P256_SHA512_TV02_Qx :: nat where
  "P256_SHA512_TV02_Qx = 0xf2a6674d4e86152a527199bed293fa63acde1b4d8a92b62e552210ba45c38792"

definition P256_SHA512_TV02_Qy :: nat where
  "P256_SHA512_TV02_Qy = 0xc72565c24f0eee6a094af341ddd8579747b865f91c8ed5b44cda8a19cc93776f"

definition P256_SHA512_TV02_Q :: "int point" where
  "P256_SHA512_TV02_Q = Point (int P256_SHA512_TV02_Qx) (int P256_SHA512_TV02_Qy)"

definition P256_SHA512_TV02_k :: nat where
  "P256_SHA512_TV02_k = 0xe547791f7185850f03d0c58419648f65b9d29cdc22ed1de2a64280220cfcafba"

definition P256_SHA512_TV02_R :: nat where
  "P256_SHA512_TV02_R = 0x4782903d2aaf8b190dab5cae2223388d2d8bd845b3875d37485c54e1ded1d3d8"

definition P256_SHA512_TV02_S :: nat where
  "P256_SHA512_TV02_S = 0xdfb40e406bfa074f0bf832771b2b9f186e2211f0bca279644a0ca8559acf39da"

definition P256_SHA512_TV02_Sig :: "nat × nat" where
  "P256_SHA512_TV02_Sig = (P256_SHA512_TV02_R, P256_SHA512_TV02_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA512_TV02_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV02_d" 
  by eval

lemma P256_SHA512_TV02_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA512_TV02_Q" 
  by eval

lemma P256_SHA512_TV02_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA512_TV02_d = P256_SHA512_TV02_Q" 
  by eval

lemma P256_SHA512_TV02_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA512_TV02_d P256_SHA512_TV02_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA512_TV02_dQvalidPair' P256_SHA512_TV02_d_valid) 

lemma P256_SHA512_TV02_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV02_k" 
  by eval

lemma P256_SHA512_TV02_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA512octets P256_SHA512_TV02_d P256_SHA512_TV02_Msg P256_SHA512_TV02_k = Some P256_SHA512_TV02_Sig" 
  by eval

lemma P256_SHA512_TV02_Verify: 
  "SEC1_P256_ECDSA_Verify SHA512octets P256_SHA512_TV02_Q P256_SHA512_TV02_Msg P256_SHA512_TV02_Sig" 
  by eval

subsection‹Test Vector: P256_SHA512_TV03›

subsubsection‹Given Test Vector Values›

definition P256_SHA512_TV03_Msg :: octets where
  "P256_SHA512_TV03_Msg = nat_to_octets_len 0xd5aa8ac9218ca661cd177756af6fbb5a40a3fecfd4eea6d5872fbb9a2884784aa9b5f0c023a6e0da5cf6364754ee6465b4ee2d0ddc745b02994c98427a213c849537da5a4477b3abfe02648be67f26e80b56a33150490d062aaac137aa47f11cfeddba855bab9e4e028532a563326d927f9e6e3292b1fb248ee90b6f429798db 128"

definition P256_SHA512_TV03_d :: nat where
  "P256_SHA512_TV03_d = 0x724567d21ef682dfc6dc4d46853880cfa86fe6fea0efd51fac456f03c3d36ead"

definition P256_SHA512_TV03_Qx :: nat where
  "P256_SHA512_TV03_Qx = 0x70b877b5e365fcf08140b1eca119baba662879f38e059d074a2cb60b03ea5d39"

definition P256_SHA512_TV03_Qy :: nat where
  "P256_SHA512_TV03_Qy = 0x5f56f94d591df40b9f3b8763ac4b3dbe622c956d5bd0c55658b6f46fa3deb201"

definition P256_SHA512_TV03_Q :: "int point" where
  "P256_SHA512_TV03_Q = Point (int P256_SHA512_TV03_Qx) (int P256_SHA512_TV03_Qy)"

definition P256_SHA512_TV03_k :: nat where
  "P256_SHA512_TV03_k = 0x79d6c967ed23c763ece9ca4b026218004c84dc2d4ccc86cf05c5d0f791f6279b"

definition P256_SHA512_TV03_R :: nat where
  "P256_SHA512_TV03_R = 0x2ba2ea2d316f8937f184ad3028e364574d20a202e4e7513d7af57ac2456804d1"

definition P256_SHA512_TV03_S :: nat where
  "P256_SHA512_TV03_S = 0x64fe94968d18c5967c799e0349041b9e40e6c6c92ebb475e80dd82f51cf07320"

definition P256_SHA512_TV03_Sig :: "nat × nat" where
  "P256_SHA512_TV03_Sig = (P256_SHA512_TV03_R, P256_SHA512_TV03_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA512_TV03_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV03_d" 
  by eval

lemma P256_SHA512_TV03_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA512_TV03_Q" 
  by eval

lemma P256_SHA512_TV03_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA512_TV03_d = P256_SHA512_TV03_Q" 
  by eval

lemma P256_SHA512_TV03_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA512_TV03_d P256_SHA512_TV03_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA512_TV03_dQvalidPair' P256_SHA512_TV03_d_valid) 

lemma P256_SHA512_TV03_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV03_k" 
  by eval

lemma P256_SHA512_TV03_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA512octets P256_SHA512_TV03_d P256_SHA512_TV03_Msg P256_SHA512_TV03_k = Some P256_SHA512_TV03_Sig" 
  by eval

lemma P256_SHA512_TV03_Verify: 
  "SEC1_P256_ECDSA_Verify SHA512octets P256_SHA512_TV03_Q P256_SHA512_TV03_Msg P256_SHA512_TV03_Sig" 
  by eval

subsection‹Test Vector: P256_SHA512_TV04›

subsubsection‹Given Test Vector Values›

definition P256_SHA512_TV04_Msg :: octets where
  "P256_SHA512_TV04_Msg = nat_to_octets_len 0x790b06054afc9c3fc4dfe72df19dd5d68d108cfcfca6212804f6d534fd2fbe489bd8f64bf205ce04bcb50124a12ce5238fc3fe7dd76e6fa640206af52549f133d593a1bfd423ab737f3326fa79433cde293236f90d4238f0dd38ed69492ddbd9c3eae583b6325a95dec3166fe52b21658293d8c137830ef45297d67813b7a508 128"

definition P256_SHA512_TV04_d :: nat where
  "P256_SHA512_TV04_d = 0x29c5d54d7d1f099d50f949bfce8d6073dae059c5a19cc70834722f18a7199edd"

definition P256_SHA512_TV04_Qx :: nat where
  "P256_SHA512_TV04_Qx = 0x3088d4f45d274cc5f418c8ecc4cbcf96be87491f420250f8cbc01cdf2503ec47"

definition P256_SHA512_TV04_Qy :: nat where
  "P256_SHA512_TV04_Qy = 0x634db48198129237ed068c88ff5809f6211921a6258f548f4b64dd125921b78b"

definition P256_SHA512_TV04_Q :: "int point" where
  "P256_SHA512_TV04_Q = Point (int P256_SHA512_TV04_Qx) (int P256_SHA512_TV04_Qy)"

definition P256_SHA512_TV04_k :: nat where
  "P256_SHA512_TV04_k = 0x0508ad7774908b5705895fda5c3b7a3032bf85dab7232bf981177019f3d76460"

definition P256_SHA512_TV04_R :: nat where
  "P256_SHA512_TV04_R = 0xacd9f3b63626c5f32103e90e1dd1695907b1904aa9b14f2132caef331321971b"

definition P256_SHA512_TV04_S :: nat where
  "P256_SHA512_TV04_S = 0x15c04a8bd6c13ed5e9961814b2f406f064670153e4d5465dcef63c1d9dd52a87"

definition P256_SHA512_TV04_Sig :: "nat × nat" where
  "P256_SHA512_TV04_Sig = (P256_SHA512_TV04_R, P256_SHA512_TV04_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA512_TV04_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV04_d" 
  by eval

lemma P256_SHA512_TV04_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA512_TV04_Q" 
  by eval

lemma P256_SHA512_TV04_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA512_TV04_d = P256_SHA512_TV04_Q" 
  by eval

lemma P256_SHA512_TV04_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA512_TV04_d P256_SHA512_TV04_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA512_TV04_dQvalidPair' P256_SHA512_TV04_d_valid) 

lemma P256_SHA512_TV04_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV04_k" 
  by eval

lemma P256_SHA512_TV04_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA512octets P256_SHA512_TV04_d P256_SHA512_TV04_Msg P256_SHA512_TV04_k = Some P256_SHA512_TV04_Sig" 
  by eval

lemma P256_SHA512_TV04_Verify: 
  "SEC1_P256_ECDSA_Verify SHA512octets P256_SHA512_TV04_Q P256_SHA512_TV04_Msg P256_SHA512_TV04_Sig" 
  by eval

subsection‹Test Vector: P256_SHA512_TV05›

subsubsection‹Given Test Vector Values›

definition P256_SHA512_TV05_Msg :: octets where
  "P256_SHA512_TV05_Msg = nat_to_octets_len 0x6d549aa87afdb8bfa60d22a68e2783b27e8db46041e4df04be0c261c4734b608a96f198d1cdb8d082ae48579ec9defcf21fbc72803764a58c31e5323d5452b9fb57c8991d31749140da7ef067b18bf0d7dfbae6eefd0d8064f334bf7e9ec1e028daed4e86e17635ec2e409a3ed1238048a45882c5c57501b314e636b9bc81cbe 128"

definition P256_SHA512_TV05_d :: nat where
  "P256_SHA512_TV05_d = 0x0d8095da1abba06b0d349c226511f642dabbf1043ad41baa4e14297afe8a3117"

definition P256_SHA512_TV05_Qx :: nat where
  "P256_SHA512_TV05_Qx = 0x75a45758ced45ecf55f755cb56ca2601d794ebeaeb2e6107fe2fc443f580e23c"

definition P256_SHA512_TV05_Qy :: nat where
  "P256_SHA512_TV05_Qy = 0x5303d47d5a75ec821d51a2ee7548448208c699eca0cd89810ffc1aa4faf81ead"

definition P256_SHA512_TV05_Q :: "int point" where
  "P256_SHA512_TV05_Q = Point (int P256_SHA512_TV05_Qx) (int P256_SHA512_TV05_Qy)"

definition P256_SHA512_TV05_k :: nat where
  "P256_SHA512_TV05_k = 0x5165c54def4026ab648f7768c4f1488bcb183f6db7ffe02c7022a529a116482a"

definition P256_SHA512_TV05_R :: nat where
  "P256_SHA512_TV05_R = 0xebc85fc4176b446b3384ccc62fc2526b45665561a0e7e9404ac376c90e450b59"

definition P256_SHA512_TV05_S :: nat where
  "P256_SHA512_TV05_S = 0x8b2c09428e62c5109d17ed0cf8f9fd7c370d018a2a73f701effc9b17d04852c6"

definition P256_SHA512_TV05_Sig :: "nat × nat" where
  "P256_SHA512_TV05_Sig = (P256_SHA512_TV05_R, P256_SHA512_TV05_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA512_TV05_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV05_d" 
  by eval

lemma P256_SHA512_TV05_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA512_TV05_Q" 
  by eval

lemma P256_SHA512_TV05_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA512_TV05_d = P256_SHA512_TV05_Q" 
  by eval

lemma P256_SHA512_TV05_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA512_TV05_d P256_SHA512_TV05_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA512_TV05_dQvalidPair' P256_SHA512_TV05_d_valid) 

lemma P256_SHA512_TV05_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV05_k" 
  by eval

lemma P256_SHA512_TV05_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA512octets P256_SHA512_TV05_d P256_SHA512_TV05_Msg P256_SHA512_TV05_k = Some P256_SHA512_TV05_Sig" 
  by eval

lemma P256_SHA512_TV05_Verify: 
  "SEC1_P256_ECDSA_Verify SHA512octets P256_SHA512_TV05_Q P256_SHA512_TV05_Msg P256_SHA512_TV05_Sig" 
  by eval

subsection‹Test Vector: P256_SHA512_TV06›

subsubsection‹Given Test Vector Values›

definition P256_SHA512_TV06_Msg :: octets where
  "P256_SHA512_TV06_Msg = nat_to_octets_len 0x1906e48b7f889ee3ff7ab0807a7aa88f53f4018808870bfed6372a77330c737647961324c2b4d46f6ee8b01190474951a701b048ae86579ff8e3fc889fecf926b17f98958ac7534e6e781ca2db2baa380dec766cfb2a3eca2a9d5818967d64dfab84f768d24ec122eebacaab0a4dc3a75f37331bb1c43dd8966cc09ec4945bbd 128"

definition P256_SHA512_TV06_d :: nat where
  "P256_SHA512_TV06_d = 0x52fe57da3427b1a75cb816f61c4e8e0e0551b94c01382b1a80837940ed579e61"

definition P256_SHA512_TV06_Qx :: nat where
  "P256_SHA512_TV06_Qx = 0x2177e20a2092a46667debdcc21e7e45d6da72f124adecbc5ada6a7bcc7b401d5"

definition P256_SHA512_TV06_Qy :: nat where
  "P256_SHA512_TV06_Qy = 0x550e468f2626070a080afeeb98edd75a721eb773c8e62149f3e903cf9c4d7b61"

definition P256_SHA512_TV06_Q :: "int point" where
  "P256_SHA512_TV06_Q = Point (int P256_SHA512_TV06_Qx) (int P256_SHA512_TV06_Qy)"

definition P256_SHA512_TV06_k :: nat where
  "P256_SHA512_TV06_k = 0x0464fe9674b01ff5bd8be21af3399fad66f90ad30f4e8ee6e2eb9bcccfd5185c"

definition P256_SHA512_TV06_R :: nat where
  "P256_SHA512_TV06_R = 0xf8250f073f34034c1cde58f69a85e2f5a030703ebdd4dbfb98d3b3690db7d114"

definition P256_SHA512_TV06_S :: nat where
  "P256_SHA512_TV06_S = 0xa9e83e05f1d6e0fef782f186bedf43684c825ac480174d48b0e4d31505e27498"

definition P256_SHA512_TV06_Sig :: "nat × nat" where
  "P256_SHA512_TV06_Sig = (P256_SHA512_TV06_R, P256_SHA512_TV06_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA512_TV06_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV06_d" 
  by eval

lemma P256_SHA512_TV06_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA512_TV06_Q" 
  by eval

lemma P256_SHA512_TV06_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA512_TV06_d = P256_SHA512_TV06_Q" 
  by eval

lemma P256_SHA512_TV06_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA512_TV06_d P256_SHA512_TV06_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA512_TV06_dQvalidPair' P256_SHA512_TV06_d_valid) 

lemma P256_SHA512_TV06_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV06_k" 
  by eval

lemma P256_SHA512_TV06_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA512octets P256_SHA512_TV06_d P256_SHA512_TV06_Msg P256_SHA512_TV06_k = Some P256_SHA512_TV06_Sig" 
  by eval

lemma P256_SHA512_TV06_Verify: 
  "SEC1_P256_ECDSA_Verify SHA512octets P256_SHA512_TV06_Q P256_SHA512_TV06_Msg P256_SHA512_TV06_Sig" 
  by eval

subsection‹Test Vector: P256_SHA512_TV07›

subsubsection‹Given Test Vector Values›

definition P256_SHA512_TV07_Msg :: octets where
  "P256_SHA512_TV07_Msg = nat_to_octets_len 0x7b59fef13daf01afec35dea3276541be681c4916767f34d4e874464d20979863ee77ad0fd1635bcdf93e9f62ed69ae52ec90aab5bbf87f8951213747ccec9f38c775c1df1e9d7f735c2ce39b42edb3b0c5086247556cfea539995c5d9689765288ec600848ecf085c01ca738bbef11f5d12d4457db988b4add90be00781024ad 128"

definition P256_SHA512_TV07_d :: nat where
  "P256_SHA512_TV07_d = 0x003d91611445919f59bfe3ca71fe0bfdeb0e39a7195e83ac03a37c7eceef0df2"

definition P256_SHA512_TV07_Qx :: nat where
  "P256_SHA512_TV07_Qx = 0x7b9c592f61aae0555855d0b9ebb6fd00fb6746e8842e2523565c858630b9ba00"

definition P256_SHA512_TV07_Qy :: nat where
  "P256_SHA512_TV07_Qy = 0xd35b2e168b1875bbc563bea5e8d63c4e38957c774a65e762959a349eaf263ba0"

definition P256_SHA512_TV07_Q :: "int point" where
  "P256_SHA512_TV07_Q = Point (int P256_SHA512_TV07_Qx) (int P256_SHA512_TV07_Qy)"

definition P256_SHA512_TV07_k :: nat where
  "P256_SHA512_TV07_k = 0xef9df291ea27a4b45708f7608723c27d7d56b7df0599a54bc2c2fabbff373b40"

definition P256_SHA512_TV07_R :: nat where
  "P256_SHA512_TV07_R = 0x66d057fd39958b0e4932bacd70a1769bbadcb62e4470937b45497a3d4500fabb"

definition P256_SHA512_TV07_S :: nat where
  "P256_SHA512_TV07_S = 0x6c853b889e18b5a49ee54b54dd1aaedfdd642e30eba171c5cab677f0df9e7318"

definition P256_SHA512_TV07_Sig :: "nat × nat" where
  "P256_SHA512_TV07_Sig = (P256_SHA512_TV07_R, P256_SHA512_TV07_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA512_TV07_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV07_d" 
  by eval

lemma P256_SHA512_TV07_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA512_TV07_Q" 
  by eval

lemma P256_SHA512_TV07_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA512_TV07_d = P256_SHA512_TV07_Q" 
  by eval

lemma P256_SHA512_TV07_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA512_TV07_d P256_SHA512_TV07_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA512_TV07_dQvalidPair' P256_SHA512_TV07_d_valid) 

lemma P256_SHA512_TV07_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV07_k" 
  by eval

lemma P256_SHA512_TV07_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA512octets P256_SHA512_TV07_d P256_SHA512_TV07_Msg P256_SHA512_TV07_k = Some P256_SHA512_TV07_Sig" 
  by eval

lemma P256_SHA512_TV07_Verify: 
  "SEC1_P256_ECDSA_Verify SHA512octets P256_SHA512_TV07_Q P256_SHA512_TV07_Msg P256_SHA512_TV07_Sig" 
  by eval

subsection‹Test Vector: P256_SHA512_TV08›

subsubsection‹Given Test Vector Values›

definition P256_SHA512_TV08_Msg :: octets where
  "P256_SHA512_TV08_Msg = nat_to_octets_len 0x041a6767a935dc3d8985eb4e608b0cbfebe7f93789d4200bcfe595277ac2b0f402889b580b72def5da778a680fd380c955421f626d52dd9a83ea180187b850e1b72a4ec6dd63235e598fd15a9b19f8ce9aec1d23f0bd6ea4d92360d50f951152bc9a01354732ba0cf90aaed33c307c1de8fa3d14f9489151b8377b57c7215f0b 128"

definition P256_SHA512_TV08_d :: nat where
  "P256_SHA512_TV08_d = 0x48f13d393899cd835c4193670ec62f28e4c4903e0bbe5817bf0996831a720bb7"

definition P256_SHA512_TV08_Qx :: nat where
  "P256_SHA512_TV08_Qx = 0x82a1a96f4648393c5e42633ecdeb1d8245c78c5ea236b5bab460dedcc8924bc0"

definition P256_SHA512_TV08_Qy :: nat where
  "P256_SHA512_TV08_Qy = 0xe8cbf03c34b5154f876de19f3bb6fd43cd2eabf6e7c95467bcfa8c8fc42d76fd"

definition P256_SHA512_TV08_Q :: "int point" where
  "P256_SHA512_TV08_Q = Point (int P256_SHA512_TV08_Qx) (int P256_SHA512_TV08_Qy)"

definition P256_SHA512_TV08_k :: nat where
  "P256_SHA512_TV08_k = 0xefed736e627899fea944007eea39a4a63c0c2e26491cd12adb546be3e5c68f7d"

definition P256_SHA512_TV08_R :: nat where
  "P256_SHA512_TV08_R = 0xcf7fc24bdaa09ac0cca8497e13298b961380668613c7493954048c06385a7044"

definition P256_SHA512_TV08_S :: nat where
  "P256_SHA512_TV08_S = 0xf38b1c8306cf82ab76ee3a772b14416b49993fe11f986e9b0f0593c52ec91525"

definition P256_SHA512_TV08_Sig :: "nat × nat" where
  "P256_SHA512_TV08_Sig = (P256_SHA512_TV08_R, P256_SHA512_TV08_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA512_TV08_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV08_d" 
  by eval

lemma P256_SHA512_TV08_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA512_TV08_Q" 
  by eval

lemma P256_SHA512_TV08_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA512_TV08_d = P256_SHA512_TV08_Q" 
  by eval

lemma P256_SHA512_TV08_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA512_TV08_d P256_SHA512_TV08_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA512_TV08_dQvalidPair' P256_SHA512_TV08_d_valid) 

lemma P256_SHA512_TV08_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV08_k" 
  by eval

lemma P256_SHA512_TV08_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA512octets P256_SHA512_TV08_d P256_SHA512_TV08_Msg P256_SHA512_TV08_k = Some P256_SHA512_TV08_Sig" 
  by eval

lemma P256_SHA512_TV08_Verify: 
  "SEC1_P256_ECDSA_Verify SHA512octets P256_SHA512_TV08_Q P256_SHA512_TV08_Msg P256_SHA512_TV08_Sig" 
  by eval

subsection‹Test Vector: P256_SHA512_TV09›

subsubsection‹Given Test Vector Values›

definition P256_SHA512_TV09_Msg :: octets where
  "P256_SHA512_TV09_Msg = nat_to_octets_len 0x7905a9036e022c78b2c9efd40b77b0a194fbc1d45462779b0b76ad30dc52c564e48a493d8249a061e62f26f453ba566538a4d43c64fb9fdbd1f36409316433c6f074e1b47b544a847de25fc67d81ac801ed9f7371a43da39001c90766f943e629d74d0436ba1240c3d7fab990d586a6d6ef1771786722df56448815f2feda48f 128"

definition P256_SHA512_TV09_d :: nat where
  "P256_SHA512_TV09_d = 0x95c99cf9ec26480275f23de419e41bb779590f0eab5cf9095d37dd70cb75e870"

definition P256_SHA512_TV09_Qx :: nat where
  "P256_SHA512_TV09_Qx = 0x42c292b0fbcc9f457ae361d940a9d45ad9427431a105a6e5cd90a345fe3507f7"

definition P256_SHA512_TV09_Qy :: nat where
  "P256_SHA512_TV09_Qy = 0x313b08fd2fa351908b3178051ee782cc62b9954ad95d4119aa564900f8ade70c"

definition P256_SHA512_TV09_Q :: "int point" where
  "P256_SHA512_TV09_Q = Point (int P256_SHA512_TV09_Qx) (int P256_SHA512_TV09_Qy)"

definition P256_SHA512_TV09_k :: nat where
  "P256_SHA512_TV09_k = 0x4c08dd0f8b72ae9c674e1e448d4e2afe3a1ee69927fa23bbff3716f0b99553b7"

definition P256_SHA512_TV09_R :: nat where
  "P256_SHA512_TV09_R = 0xf2bc35eb1b8488b9e8d4a1dbb200e1abcb855458e1557dc1bf988278a174eb3b"

definition P256_SHA512_TV09_S :: nat where
  "P256_SHA512_TV09_S = 0xed9a2ec043a1d578e8eba6f57217976310e8674385ad2da08d6146c629de1cd9"

definition P256_SHA512_TV09_Sig :: "nat × nat" where
  "P256_SHA512_TV09_Sig = (P256_SHA512_TV09_R, P256_SHA512_TV09_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA512_TV09_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV09_d" 
  by eval

lemma P256_SHA512_TV09_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA512_TV09_Q" 
  by eval

lemma P256_SHA512_TV09_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA512_TV09_d = P256_SHA512_TV09_Q" 
  by eval

lemma P256_SHA512_TV09_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA512_TV09_d P256_SHA512_TV09_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA512_TV09_dQvalidPair' P256_SHA512_TV09_d_valid) 

lemma P256_SHA512_TV09_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV09_k" 
  by eval

lemma P256_SHA512_TV09_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA512octets P256_SHA512_TV09_d P256_SHA512_TV09_Msg P256_SHA512_TV09_k = Some P256_SHA512_TV09_Sig" 
  by eval

lemma P256_SHA512_TV09_Verify: 
  "SEC1_P256_ECDSA_Verify SHA512octets P256_SHA512_TV09_Q P256_SHA512_TV09_Msg P256_SHA512_TV09_Sig" 
  by eval

subsection‹Test Vector: P256_SHA512_TV10›

subsubsection‹Given Test Vector Values›

definition P256_SHA512_TV10_Msg :: octets where
  "P256_SHA512_TV10_Msg = nat_to_octets_len 0xcf25e4642d4f39d15afb7aec79469d82fc9aedb8f89964e79b749a852d931d37436502804e39555f5a3c75dd958fd5291ada647c1a5e38fe7b1048f16f2b711fdd5d39acc0812ca65bd50d7f8119f2fd195ab16633503a78ee9102c1f9c4c22568e0b54bd4fa3f5ff7b49160bf23e7e2231b1ebebbdaf0e4a7d4484158a87e07 128"

definition P256_SHA512_TV10_d :: nat where
  "P256_SHA512_TV10_d = 0xe15e835d0e2217bc7c6f05a498f20af1cd56f2f165c23d225eb3360aa2c5cbcf"

definition P256_SHA512_TV10_Qx :: nat where
  "P256_SHA512_TV10_Qx = 0x89dd22052ec3ab4840206a62f2270c21e7836d1a9109a3407dd0974c7802b9ae"

definition P256_SHA512_TV10_Qy :: nat where
  "P256_SHA512_TV10_Qy = 0xe91609ba35c7008b080c77a9068d97a14ca77b97299e74945217672b2fd5faf0"

definition P256_SHA512_TV10_Q :: "int point" where
  "P256_SHA512_TV10_Q = Point (int P256_SHA512_TV10_Qx) (int P256_SHA512_TV10_Qy)"

definition P256_SHA512_TV10_k :: nat where
  "P256_SHA512_TV10_k = 0xc9f621441c235fc47ec34eef4c08625df1ec74918e1f86075b753f2589f4c60b"

definition P256_SHA512_TV10_R :: nat where
  "P256_SHA512_TV10_R = 0xa70d1a2d555d599bfb8c9b1f0d43725341151d17a8d0845fa56f3563703528a7"

definition P256_SHA512_TV10_S :: nat where
  "P256_SHA512_TV10_S = 0x4e05c45adf41783e394a5312f86e66871c4be4896948c85966879d5c66d54b37"

definition P256_SHA512_TV10_Sig :: "nat × nat" where
  "P256_SHA512_TV10_Sig = (P256_SHA512_TV10_R, P256_SHA512_TV10_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA512_TV10_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV10_d" 
  by eval

lemma P256_SHA512_TV10_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA512_TV10_Q" 
  by eval

lemma P256_SHA512_TV10_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA512_TV10_d = P256_SHA512_TV10_Q" 
  by eval

lemma P256_SHA512_TV10_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA512_TV10_d P256_SHA512_TV10_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA512_TV10_dQvalidPair' P256_SHA512_TV10_d_valid) 

lemma P256_SHA512_TV10_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV10_k" 
  by eval

lemma P256_SHA512_TV10_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA512octets P256_SHA512_TV10_d P256_SHA512_TV10_Msg P256_SHA512_TV10_k = Some P256_SHA512_TV10_Sig" 
  by eval

lemma P256_SHA512_TV10_Verify: 
  "SEC1_P256_ECDSA_Verify SHA512octets P256_SHA512_TV10_Q P256_SHA512_TV10_Msg P256_SHA512_TV10_Sig" 
  by eval

subsection‹Test Vector: P256_SHA512_TV11›

subsubsection‹Given Test Vector Values›

definition P256_SHA512_TV11_Msg :: octets where
  "P256_SHA512_TV11_Msg = nat_to_octets_len 0x7562c445b35883cc937be6349b4cefc3556a80255d70f09e28c3f393daac19442a7eecedcdfbe8f7628e30cd8939537ec56d5c9645d43340eb4e78fc5dd4322de8a07966b262770d7ff13a071ff3dce560718e60ed3086b7e0003a6abafe91af90af86733ce8689440bf73d2aa0acfe9776036e877599acbabfcb03bb3b50faa 128"

definition P256_SHA512_TV11_d :: nat where
  "P256_SHA512_TV11_d = 0x808c08c0d77423a6feaaffc8f98a2948f17726e67c15eeae4e672edbe388f98c"

definition P256_SHA512_TV11_Qx :: nat where
  "P256_SHA512_TV11_Qx = 0xb0c0ad5e1f6001d8e9018ec611b2e3b91923e69fa6c98690ab644d650f640c42"

definition P256_SHA512_TV11_Qy :: nat where
  "P256_SHA512_TV11_Qy = 0x610539c0b9ed21ac0a2f27527c1a61d9b47cbf033187b1a6ada006eb5b2662ed"

definition P256_SHA512_TV11_Q :: "int point" where
  "P256_SHA512_TV11_Q = Point (int P256_SHA512_TV11_Qx) (int P256_SHA512_TV11_Qy)"

definition P256_SHA512_TV11_k :: nat where
  "P256_SHA512_TV11_k = 0x1f6d4a905c761a53d54c362976717d0d7fc94d222bb5489e4830080a1a67535d"

definition P256_SHA512_TV11_R :: nat where
  "P256_SHA512_TV11_R = 0x83404dcf8320baf206381800071e6a75160342d19743b4f176960d669dd03d07"

definition P256_SHA512_TV11_S :: nat where
  "P256_SHA512_TV11_S = 0x3f75dcf102008b2989f81683ae45e9f1d4b67a6ef6fd5c8af44828af80e1cfb5"

definition P256_SHA512_TV11_Sig :: "nat × nat" where
  "P256_SHA512_TV11_Sig = (P256_SHA512_TV11_R, P256_SHA512_TV11_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA512_TV11_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV11_d" 
  by eval

lemma P256_SHA512_TV11_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA512_TV11_Q" 
  by eval

lemma P256_SHA512_TV11_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA512_TV11_d = P256_SHA512_TV11_Q" 
  by eval

lemma P256_SHA512_TV11_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA512_TV11_d P256_SHA512_TV11_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA512_TV11_dQvalidPair' P256_SHA512_TV11_d_valid) 

lemma P256_SHA512_TV11_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV11_k" 
  by eval

lemma P256_SHA512_TV11_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA512octets P256_SHA512_TV11_d P256_SHA512_TV11_Msg P256_SHA512_TV11_k = Some P256_SHA512_TV11_Sig" 
  by eval

lemma P256_SHA512_TV11_Verify: 
  "SEC1_P256_ECDSA_Verify SHA512octets P256_SHA512_TV11_Q P256_SHA512_TV11_Msg P256_SHA512_TV11_Sig" 
  by eval

subsection‹Test Vector: P256_SHA512_TV12›

subsubsection‹Given Test Vector Values›

definition P256_SHA512_TV12_Msg :: octets where
  "P256_SHA512_TV12_Msg = nat_to_octets_len 0x051c2db8e71e44653ea1cb0afc9e0abdf12658e9e761bfb767c20c7ab4adfcb18ed9b5c372a3ac11d8a43c55f7f99b33355437891686d42362abd71db8b6d84dd694d6982f0612178a937aa934b9ac3c0794c39027bdd767841c4370666c80dbc0f8132ca27474f553d266deefd7c9dbad6d734f9006bb557567701bb7e6a7c9 128"

definition P256_SHA512_TV12_d :: nat where
  "P256_SHA512_TV12_d = 0xf7c6315f0081acd8f09c7a2c3ec1b7ece20180b0a6365a27dcd8f71b729558f9"

definition P256_SHA512_TV12_Qx :: nat where
  "P256_SHA512_TV12_Qx = 0x250f7112d381c1751860045d9bcaf20dbeb25a001431f96ac6f19109362ffebb"

definition P256_SHA512_TV12_Qy :: nat where
  "P256_SHA512_TV12_Qy = 0x49fba9efe73546135a5a31ab3753e247034741ce839d3d94bd73936c4a17e4aa"

definition P256_SHA512_TV12_Q :: "int point" where
  "P256_SHA512_TV12_Q = Point (int P256_SHA512_TV12_Qx) (int P256_SHA512_TV12_Qy)"

definition P256_SHA512_TV12_k :: nat where
  "P256_SHA512_TV12_k = 0x68c299be2c0c6d52d208d5d1a9e0ffa2af19b4833271404e5876e0aa93987866"

definition P256_SHA512_TV12_R :: nat where
  "P256_SHA512_TV12_R = 0x7b195e92d2ba95911cda7570607e112d02a1c847ddaa33924734b51f5d81adab"

definition P256_SHA512_TV12_S :: nat where
  "P256_SHA512_TV12_S = 0x10d9f206755cef70ab5143ac43f3f8d38aea2644f31d52eaf3b472ee816e11e5"

definition P256_SHA512_TV12_Sig :: "nat × nat" where
  "P256_SHA512_TV12_Sig = (P256_SHA512_TV12_R, P256_SHA512_TV12_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA512_TV12_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV12_d" 
  by eval

lemma P256_SHA512_TV12_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA512_TV12_Q" 
  by eval

lemma P256_SHA512_TV12_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA512_TV12_d = P256_SHA512_TV12_Q" 
  by eval

lemma P256_SHA512_TV12_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA512_TV12_d P256_SHA512_TV12_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA512_TV12_dQvalidPair' P256_SHA512_TV12_d_valid) 

lemma P256_SHA512_TV12_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV12_k" 
  by eval

lemma P256_SHA512_TV12_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA512octets P256_SHA512_TV12_d P256_SHA512_TV12_Msg P256_SHA512_TV12_k = Some P256_SHA512_TV12_Sig" 
  by eval

lemma P256_SHA512_TV12_Verify: 
  "SEC1_P256_ECDSA_Verify SHA512octets P256_SHA512_TV12_Q P256_SHA512_TV12_Msg P256_SHA512_TV12_Sig" 
  by eval

subsection‹Test Vector: P256_SHA512_TV13›

subsubsection‹Given Test Vector Values›

definition P256_SHA512_TV13_Msg :: octets where
  "P256_SHA512_TV13_Msg = nat_to_octets_len 0x4dcb7b62ba31b866fce7c1feedf0be1f67bf611dbc2e2e86f004422f67b3bc1839c6958eb1dc3ead137c3d7f88aa97244577a775c8021b1642a8647bba82871e3c15d0749ed343ea6cad38f123835d8ef66b0719273105e924e8685b65fd5dc430efbc35b05a6097f17ebc5943cdcd9abcba752b7f8f37027409bd6e11cd158f 128"

definition P256_SHA512_TV13_d :: nat where
  "P256_SHA512_TV13_d = 0xf547735a9409386dbff719ce2dae03c50cb437d6b30cc7fa3ea20d9aec17e5a5"

definition P256_SHA512_TV13_Qx :: nat where
  "P256_SHA512_TV13_Qx = 0x4ca87c5845fb04c2f76ae3273073b0523e356a445e4e95737260eba9e2d021db"

definition P256_SHA512_TV13_Qy :: nat where
  "P256_SHA512_TV13_Qy = 0x0f86475d07f82655320fdf2cd8db23b21905b1b1f2f9c48e2df87e24119c4880"

definition P256_SHA512_TV13_Q :: "int point" where
  "P256_SHA512_TV13_Q = Point (int P256_SHA512_TV13_Qx) (int P256_SHA512_TV13_Qy)"

definition P256_SHA512_TV13_k :: nat where
  "P256_SHA512_TV13_k = 0x91bd7d97f7ed3253cedefc144771bb8acbbda6eb24f9d752bbe1dd018e1384c7"

definition P256_SHA512_TV13_R :: nat where
  "P256_SHA512_TV13_R = 0x008c1755d3df81e64e25270dbaa9396641556df7ffc7ac9add6739c382705397"

definition P256_SHA512_TV13_S :: nat where
  "P256_SHA512_TV13_S = 0x77df443c729b039aded5b516b1077fecdd9986402d2c4b01734ba91e055e87fc"

definition P256_SHA512_TV13_Sig :: "nat × nat" where
  "P256_SHA512_TV13_Sig = (P256_SHA512_TV13_R, P256_SHA512_TV13_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA512_TV13_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV13_d" 
  by eval

lemma P256_SHA512_TV13_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA512_TV13_Q" 
  by eval

lemma P256_SHA512_TV13_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA512_TV13_d = P256_SHA512_TV13_Q" 
  by eval

lemma P256_SHA512_TV13_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA512_TV13_d P256_SHA512_TV13_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA512_TV13_dQvalidPair' P256_SHA512_TV13_d_valid) 

lemma P256_SHA512_TV13_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV13_k" 
  by eval

lemma P256_SHA512_TV13_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA512octets P256_SHA512_TV13_d P256_SHA512_TV13_Msg P256_SHA512_TV13_k = Some P256_SHA512_TV13_Sig" 
  by eval

lemma P256_SHA512_TV13_Verify: 
  "SEC1_P256_ECDSA_Verify SHA512octets P256_SHA512_TV13_Q P256_SHA512_TV13_Msg P256_SHA512_TV13_Sig" 
  by eval

subsection‹Test Vector: P256_SHA512_TV14›

subsubsection‹Given Test Vector Values›

definition P256_SHA512_TV14_Msg :: octets where
  "P256_SHA512_TV14_Msg = nat_to_octets_len 0xefe55737771070d5ac79236b04e3fbaf4f2e9bed187d1930680fcf1aba769674bf426310f21245006f528779347d28b8aeacd2b1d5e3456dcbf188b2be8c07f19219e4067c1e7c9714784285d8bac79a76b56f2e2676ea93994f11eb573af1d03fc8ed1118eafc7f07a82f3263c33eb85e497e18f435d4076a774f42d276c323 128"

definition P256_SHA512_TV14_d :: nat where
  "P256_SHA512_TV14_d = 0x26a1aa4b927a516b661986895aff58f40b78cc5d0c767eda7eaa3dbb835b5628"

definition P256_SHA512_TV14_Qx :: nat where
  "P256_SHA512_TV14_Qx = 0x28afa3b0f81a0e95ad302f487a9b679fcdef8d3f40236ec4d4dbf4bb0cbba8b2"

definition P256_SHA512_TV14_Qy :: nat where
  "P256_SHA512_TV14_Qy = 0xbb4ac1be8405cbae8a553fbc28e29e2e689fabe7def26d653a1dafc023f3cecf"

definition P256_SHA512_TV14_Q :: "int point" where
  "P256_SHA512_TV14_Q = Point (int P256_SHA512_TV14_Qx) (int P256_SHA512_TV14_Qy)"

definition P256_SHA512_TV14_k :: nat where
  "P256_SHA512_TV14_k = 0xf98e1933c7fad4acbe94d95c1b013e1d6931fa8f67e6dbb677b564ef7c3e56ce"

definition P256_SHA512_TV14_R :: nat where
  "P256_SHA512_TV14_R = 0x15a9a5412d6a03edd71b84c121ce9a94cdd166e40da9ce4d79f1afff6a395a53"

definition P256_SHA512_TV14_S :: nat where
  "P256_SHA512_TV14_S = 0x86bbc2b6c63bad706ec0b093578e3f064736ec69c0dba59b9e3e7f73762a4dc3"

definition P256_SHA512_TV14_Sig :: "nat × nat" where
  "P256_SHA512_TV14_Sig = (P256_SHA512_TV14_R, P256_SHA512_TV14_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA512_TV14_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV14_d" 
  by eval

lemma P256_SHA512_TV14_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA512_TV14_Q" 
  by eval

lemma P256_SHA512_TV14_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA512_TV14_d = P256_SHA512_TV14_Q" 
  by eval

lemma P256_SHA512_TV14_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA512_TV14_d P256_SHA512_TV14_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA512_TV14_dQvalidPair' P256_SHA512_TV14_d_valid) 

lemma P256_SHA512_TV14_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV14_k" 
  by eval

lemma P256_SHA512_TV14_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA512octets P256_SHA512_TV14_d P256_SHA512_TV14_Msg P256_SHA512_TV14_k = Some P256_SHA512_TV14_Sig" 
  by eval

lemma P256_SHA512_TV14_Verify: 
  "SEC1_P256_ECDSA_Verify SHA512octets P256_SHA512_TV14_Q P256_SHA512_TV14_Msg P256_SHA512_TV14_Sig" 
  by eval

subsection‹Test Vector: P256_SHA512_TV15›

subsubsection‹Given Test Vector Values›

definition P256_SHA512_TV15_Msg :: octets where
  "P256_SHA512_TV15_Msg = nat_to_octets_len 0xea95859cc13cccb37198d919803be89c2ee10befdcaf5d5afa09dcc529d333ae1e4ffd3bd8ba8642203badd7a80a3f77eeee9402eed365d53f05c1a995c536f8236ba6b6ff8897393506660cc8ea82b2163aa6a1855251c87d935e23857fe35b889427b449de7274d7754bdeace960b4303c5dd5f745a5cfd580293d6548c832 128"

definition P256_SHA512_TV15_d :: nat where
  "P256_SHA512_TV15_d = 0x6a5ca39aae2d45aa331f18a8598a3f2db32781f7c92efd4f64ee3bbe0c4c4e49"

definition P256_SHA512_TV15_Qx :: nat where
  "P256_SHA512_TV15_Qx = 0xc62cc4a39ace01006ad48cf49a3e71466955bbeeca5d318d672695df926b3aa4"

definition P256_SHA512_TV15_Qy :: nat where
  "P256_SHA512_TV15_Qy = 0xc85ccf517bf2ebd9ad6a9e99254def0d74d1d2fd611e328b4a3988d4f045fe6f"

definition P256_SHA512_TV15_Q :: "int point" where
  "P256_SHA512_TV15_Q = Point (int P256_SHA512_TV15_Qx) (int P256_SHA512_TV15_Qy)"

definition P256_SHA512_TV15_k :: nat where
  "P256_SHA512_TV15_k = 0xdac00c462bc85bf39c31b5e01df33e2ec1569e6efcb334bf18f0951992ac6160"

definition P256_SHA512_TV15_R :: nat where
  "P256_SHA512_TV15_R = 0x6e7ff8ec7a5c48e0877224a9fa8481283de45fcbee23b4c252b0c622442c26ad"

definition P256_SHA512_TV15_S :: nat where
  "P256_SHA512_TV15_S = 0x3dfac320b9c873318117da6bd856000a392b815659e5aa2a6a1852ccb2501df3"

definition P256_SHA512_TV15_Sig :: "nat × nat" where
  "P256_SHA512_TV15_Sig = (P256_SHA512_TV15_R, P256_SHA512_TV15_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P256_SHA512_TV15_d_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV15_d" 
  by eval

lemma P256_SHA512_TV15_QonCurve: 
  "SEC1_P256_on_curve' P256_SHA512_TV15_Q" 
  by eval

lemma P256_SHA512_TV15_dQvalidPair': 
  "SEC1_P256_ECkeyGen P256_SHA512_TV15_d = P256_SHA512_TV15_Q" 
  by eval

lemma P256_SHA512_TV15_dQvalidPair: 
  "SEC1_P256_ECkeyPairValid P256_SHA512_TV15_d P256_SHA512_TV15_Q" 
  by (simp add: SEC1_P256.ECkeyPairEqKeyGen P256_SHA512_TV15_dQvalidPair' P256_SHA512_TV15_d_valid) 

lemma P256_SHA512_TV15_k_valid: 
  "SEC1_P256_ECprivateKeyValid P256_SHA512_TV15_k" 
  by eval

lemma P256_SHA512_TV15_Sign: 
  "SEC1_P256_ECDSA_Sign' SHA512octets P256_SHA512_TV15_d P256_SHA512_TV15_Msg P256_SHA512_TV15_k = Some P256_SHA512_TV15_Sig" 
  by eval

lemma P256_SHA512_TV15_Verify: 
  "SEC1_P256_ECDSA_Verify SHA512octets P256_SHA512_TV15_Q P256_SHA512_TV15_Msg P256_SHA512_TV15_Sig" 
  by eval

section‹NIST Test Vectors for ECDSA: Curve P-384, Hash SHA-224›
  
subsection‹Test Vector: P384_SHA224_TV01›

subsubsection‹Given Test Vector Values›

definition P384_SHA224_TV01_Msg :: octets where
  "P384_SHA224_TV01_Msg = nat_to_octets_len 0x39f0b25d4c15b09a0692b22fbacbb5f8aee184cb75887e2ebe0cd3be5d3815d29f9b587e10b3168c939054a89df11068e5c3fac21af742bf4c3e9512f5569674e7ad8b39042bcd73e4b7ce3e64fbea1c434ed01ad4ad8b5b569f6a0b9a1144f94097925672e59ba97bc4d33be2fa21b46c3dadbfb3a1f89afa199d4b44189938 128"

definition P384_SHA224_TV01_d :: nat where
  "P384_SHA224_TV01_d = 0x0af857beff08046f23b03c4299eda86490393bde88e4f74348886b200555276b93b37d4f6fdec17c0ea581a30c59c727"

definition P384_SHA224_TV01_Qx :: nat where
  "P384_SHA224_TV01_Qx = 0x00ea9d109dbaa3900461a9236453952b1f1c2a5aa12f6d500ac774acdff84ab7cb71a0f91bcd55aaa57cb8b4fbb3087d"

definition P384_SHA224_TV01_Qy :: nat where
  "P384_SHA224_TV01_Qy = 0x0fc0e3116c9e94be583b02b21b1eb168d8facf3955279360cbcd86e04ee50751054cfaebcf542538ac113d56ccc38b3e"

definition P384_SHA224_TV01_Q :: "int point" where
  "P384_SHA224_TV01_Q = Point (int P384_SHA224_TV01_Qx) (int P384_SHA224_TV01_Qy)"

definition P384_SHA224_TV01_k :: nat where
  "P384_SHA224_TV01_k = 0xe2f0ce83c5bbef3a6eccd1744f893bb52952475d2531a2854a88ff0aa9b12c65961e2e517fb334ef40e0c0d7a31ed5f5"

definition P384_SHA224_TV01_R :: nat where
  "P384_SHA224_TV01_R = 0xc36e5f0d3de71411e6e519f63e0f56cff432330a04fefef2993fdb56343e49f2f7db5fcab7728acc1e33d4692553c02e"

definition P384_SHA224_TV01_S :: nat where
  "P384_SHA224_TV01_S = 0x0d4064399d58cd771ab9420d438757f5936c3808e97081e457bc862a0c905295dca60ee94f4537591c6c7d217453909b"

definition P384_SHA224_TV01_Sig :: "nat × nat" where
  "P384_SHA224_TV01_Sig = (P384_SHA224_TV01_R, P384_SHA224_TV01_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA224_TV01_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV01_d" 
  by eval

lemma P384_SHA224_TV01_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA224_TV01_Q" 
  by eval

lemma P384_SHA224_TV01_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA224_TV01_d = P384_SHA224_TV01_Q" 
  by eval

lemma P384_SHA224_TV01_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA224_TV01_d P384_SHA224_TV01_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA224_TV01_dQvalidPair' P384_SHA224_TV01_d_valid) 

lemma P384_SHA224_TV01_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV01_k" 
  by eval

lemma P384_SHA224_TV01_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA224octets P384_SHA224_TV01_d P384_SHA224_TV01_Msg P384_SHA224_TV01_k = Some P384_SHA224_TV01_Sig" 
  by eval

lemma P384_SHA224_TV01_Verify: 
  "SEC1_P384_ECDSA_Verify SHA224octets P384_SHA224_TV01_Q P384_SHA224_TV01_Msg P384_SHA224_TV01_Sig" 
  by eval

subsection‹Test Vector: P384_SHA224_TV02›

subsubsection‹Given Test Vector Values›

definition P384_SHA224_TV02_Msg :: octets where
  "P384_SHA224_TV02_Msg = nat_to_octets_len 0x5a3c80e608ed3ac75a6e45f6e94d374271a6d42b67a481860d5d309cc8b37c79cb61f1716dc8aa84cb309ef9d68eb7fc6cf4b42333f316a5c30e74198c8b340926e340c5de47674a707293c4aa2a1a2274a602f01c26b156e895499c60b38ef53fc2032e7485c168d73700d6fa14232596a0e4997854a0b05d02e351b9d3de96 128"

definition P384_SHA224_TV02_d :: nat where
  "P384_SHA224_TV02_d = 0x047dd5baab23f439ec23b58b7e6ff4cc37813cccb4ea73bb2308e6b82b3170edfe0e131eca50841bf1b686e651c57246"

definition P384_SHA224_TV02_Qx :: nat where
  "P384_SHA224_TV02_Qx = 0xde92ff09af2950854a70f2178d2ed50cc7042a7188301a1ea81d9629ad3c29795cb7f0d56630a401e4d6e5bed0068d1e"

definition P384_SHA224_TV02_Qy :: nat where
  "P384_SHA224_TV02_Qy = 0x6135adbd8624130735e64e65ecbd43770dcc12b28e737b5ed033666f34c918eb5589508e4a13b9243374a118a628dd0b"

definition P384_SHA224_TV02_Q :: "int point" where
  "P384_SHA224_TV02_Q = Point (int P384_SHA224_TV02_Qx) (int P384_SHA224_TV02_Qy)"

definition P384_SHA224_TV02_k :: nat where
  "P384_SHA224_TV02_k = 0xf3922351d14f1e5af84faab12fe57ded30f185afe5547aeb3061104740ecc42a8df0c27f3877b4d855642b78938c4e05"

definition P384_SHA224_TV02_R :: nat where
  "P384_SHA224_TV02_R = 0x38e181870cb797c1f4e6598cfd032add1cb60447d33473038d06df73919f844eddd16f40f911075f8a4bacc0d924e684"

definition P384_SHA224_TV02_S :: nat where
  "P384_SHA224_TV02_S = 0xa58dd1ca18aa31277de66c30c3bb7a14b53705ce6c547ed2cb0e336f63c42809422efffcc722d1155f2254330a02b278"

definition P384_SHA224_TV02_Sig :: "nat × nat" where
  "P384_SHA224_TV02_Sig = (P384_SHA224_TV02_R, P384_SHA224_TV02_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA224_TV02_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV02_d" 
  by eval

lemma P384_SHA224_TV02_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA224_TV02_Q" 
  by eval

lemma P384_SHA224_TV02_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA224_TV02_d = P384_SHA224_TV02_Q" 
  by eval

lemma P384_SHA224_TV02_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA224_TV02_d P384_SHA224_TV02_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA224_TV02_dQvalidPair' P384_SHA224_TV02_d_valid) 

lemma P384_SHA224_TV02_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV02_k" 
  by eval

lemma P384_SHA224_TV02_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA224octets P384_SHA224_TV02_d P384_SHA224_TV02_Msg P384_SHA224_TV02_k = Some P384_SHA224_TV02_Sig" 
  by eval

lemma P384_SHA224_TV02_Verify: 
  "SEC1_P384_ECDSA_Verify SHA224octets P384_SHA224_TV02_Q P384_SHA224_TV02_Msg P384_SHA224_TV02_Sig" 
  by eval

subsection‹Test Vector: P384_SHA224_TV03›

subsubsection‹Given Test Vector Values›

definition P384_SHA224_TV03_Msg :: octets where
  "P384_SHA224_TV03_Msg = nat_to_octets_len 0xe7d974c5dbd3bfb8a2fb92fdd782f997d04be79e9713944ce13c5eb6f75dfdec811b7ee4b3859114b07f263846ae13f795eec8f3cb5b7565baff68e0fdd5e09ba8b176d5a71cb03fbc5546e6937fba560acb4db24bd42de1851432b96e8ca4078313cb849bce29c9d805258601d67cd0259e255f3048682e8fdbdda3398c3e31 128"

definition P384_SHA224_TV03_d :: nat where
  "P384_SHA224_TV03_d = 0x54ba9c740535574cebc41ca5dc950629674ee94730353ac521aafd1c342d3f8ac52046ed804264e1440d7fe409c45c83"

definition P384_SHA224_TV03_Qx :: nat where
  "P384_SHA224_TV03_Qx = 0x3db95ded500b2506b627270bac75688dd7d44f47029adeff99397ab4b6329a38dbb278a0fc58fe4914e6ae31721a6875"

definition P384_SHA224_TV03_Qy :: nat where
  "P384_SHA224_TV03_Qy = 0x049288341553a9ac3dc2d9e18e7a92c43dd3c25ca866f0cb4c68127bef6b0e4ba85713d27d45c7d0dc57e5782a6bf733"

definition P384_SHA224_TV03_Q :: "int point" where
  "P384_SHA224_TV03_Q = Point (int P384_SHA224_TV03_Qx) (int P384_SHA224_TV03_Qy)"

definition P384_SHA224_TV03_k :: nat where
  "P384_SHA224_TV03_k = 0x04324bd078807f6b18507a93ee60da02031717217ee5ce569750737be912be72da087ac00f50e13fdf7249a6ae33f73e"

definition P384_SHA224_TV03_R :: nat where
  "P384_SHA224_TV03_R = 0xb2752aa7abc1e5a29421c9c76620bcc3049ecc97e6bc39fcca126f505a9a1bfae3bde89fb751a1aa7b66fa8db3891ef0"

definition P384_SHA224_TV03_S :: nat where
  "P384_SHA224_TV03_S = 0xf1c69e6d818ca7ae3a477049b46420cebd910c0a9a477fd1a67a38d628d6edaac123aebfca67c53a5c80fe454dba7a9d"

definition P384_SHA224_TV03_Sig :: "nat × nat" where
  "P384_SHA224_TV03_Sig = (P384_SHA224_TV03_R, P384_SHA224_TV03_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA224_TV03_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV03_d" 
  by eval

lemma P384_SHA224_TV03_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA224_TV03_Q" 
  by eval

lemma P384_SHA224_TV03_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA224_TV03_d = P384_SHA224_TV03_Q" 
  by eval

lemma P384_SHA224_TV03_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA224_TV03_d P384_SHA224_TV03_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA224_TV03_dQvalidPair' P384_SHA224_TV03_d_valid) 

lemma P384_SHA224_TV03_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV03_k" 
  by eval

lemma P384_SHA224_TV03_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA224octets P384_SHA224_TV03_d P384_SHA224_TV03_Msg P384_SHA224_TV03_k = Some P384_SHA224_TV03_Sig" 
  by eval

lemma P384_SHA224_TV03_Verify: 
  "SEC1_P384_ECDSA_Verify SHA224octets P384_SHA224_TV03_Q P384_SHA224_TV03_Msg P384_SHA224_TV03_Sig" 
  by eval

subsection‹Test Vector: P384_SHA224_TV04›

subsubsection‹Given Test Vector Values›

definition P384_SHA224_TV04_Msg :: octets where
  "P384_SHA224_TV04_Msg = nat_to_octets_len 0xa670fda4d1d56c70de1d8680328043b2b7029633caf0ee59ffe1421c914bb937133d5a0f9214846b2e0b350455a74c4ab434c56de65a17139bb8212bf1c76071a37536fa29348f871dbb26baa92eb93d97e923a6d2ffd9be25cbc33075e494e6db657bd8dc053fe4e17148d8cf6e2058164f2b5766750eb01bbe7b361cdb848c 128"

definition P384_SHA224_TV04_d :: nat where
  "P384_SHA224_TV04_d = 0xdabe87bbe95499bac23bc83c8b7307fe04be198f00059e2bf67c9611feaffb2c8f274f6aa50eb99c3074186d8067d659"

definition P384_SHA224_TV04_Qx :: nat where
  "P384_SHA224_TV04_Qx = 0xc2aa0a695125279705917e02a4f258cade4c3ff9140a071414babf87764f426f7f36ffda9d5f3394375d24864235476f"

definition P384_SHA224_TV04_Qy :: nat where
  "P384_SHA224_TV04_Qy = 0x8f9808da0ce0227cf453f9e456f557db9752e23b45cce4baad5fee3844ddd7e1112bcec01ea9d67c7a76f3535bd0cb58"

definition P384_SHA224_TV04_Q :: "int point" where
  "P384_SHA224_TV04_Q = Point (int P384_SHA224_TV04_Qx) (int P384_SHA224_TV04_Qy)"

definition P384_SHA224_TV04_k :: nat where
  "P384_SHA224_TV04_k = 0x65a0305854033cbc6fe3ca139c40ca354d45801ecb59f4a923c251dc6b25d12d452d99b5d6711fdb5efac812aa464cc4"

definition P384_SHA224_TV04_R :: nat where
  "P384_SHA224_TV04_R = 0xc7fc32997d17ac79baf5789e4503f5f1a8863872bc350a91f12dd3ef8cf78c254e829217809e8e00b6b8d4d85be3f1fd"

definition P384_SHA224_TV04_S :: nat where
  "P384_SHA224_TV04_S = 0x1422e1838a22496df93486bce1142961dbd8478ae844b8dda54e210afdae0d9e930d587c91bb600b0bde7237186d94e6"

definition P384_SHA224_TV04_Sig :: "nat × nat" where
  "P384_SHA224_TV04_Sig = (P384_SHA224_TV04_R, P384_SHA224_TV04_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA224_TV04_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV04_d" 
  by eval

lemma P384_SHA224_TV04_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA224_TV04_Q" 
  by eval

lemma P384_SHA224_TV04_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA224_TV04_d = P384_SHA224_TV04_Q" 
  by eval

lemma P384_SHA224_TV04_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA224_TV04_d P384_SHA224_TV04_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA224_TV04_dQvalidPair' P384_SHA224_TV04_d_valid) 

lemma P384_SHA224_TV04_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV04_k" 
  by eval

lemma P384_SHA224_TV04_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA224octets P384_SHA224_TV04_d P384_SHA224_TV04_Msg P384_SHA224_TV04_k = Some P384_SHA224_TV04_Sig" 
  by eval

lemma P384_SHA224_TV04_Verify: 
  "SEC1_P384_ECDSA_Verify SHA224octets P384_SHA224_TV04_Q P384_SHA224_TV04_Msg P384_SHA224_TV04_Sig" 
  by eval

subsection‹Test Vector: P384_SHA224_TV05›

subsubsection‹Given Test Vector Values›

definition P384_SHA224_TV05_Msg :: octets where
  "P384_SHA224_TV05_Msg = nat_to_octets_len 0x7843f918fe2588bcfe756e1f05b491d913523255aa006818be20b676c957f4edb8df863c6f5f8c15b3b80c7a2aa277b70d53f210bdfb856337980c406ea140e439dd321471407f374f69877b2d82367eed51e3c82c13948616dcb301d0c31f8f0352f2846abd9e72071f446a2f1bd3339a09ae41b84e150fd18f4ba5d3c6bfa0 128"

definition P384_SHA224_TV05_d :: nat where
  "P384_SHA224_TV05_d = 0xdf43107a1deb24d02e31d479087bd669e2bc3e50f1f44b7db9484a7143cdca6a3391bddfea72dc940dbce8ec5efbd718"

definition P384_SHA224_TV05_Qx :: nat where
  "P384_SHA224_TV05_Qx = 0x76bd4be5d520471162cb5c36f80038301b325f845d9642204a84d78b3e721098932827bf872bde0a9f86383953667d29"

definition P384_SHA224_TV05_Qy :: nat where
  "P384_SHA224_TV05_Qy = 0x415116b8b878f896a5aa4dbbdc21076f27135d8bbcaaca02489ef639d742bd63f377da0c8e8ab36ff19b4a7cc5d4ceb4"

definition P384_SHA224_TV05_Q :: "int point" where
  "P384_SHA224_TV05_Q = Point (int P384_SHA224_TV05_Qx) (int P384_SHA224_TV05_Qy)"

definition P384_SHA224_TV05_k :: nat where
  "P384_SHA224_TV05_k = 0x798abad5a30d1805794540057388ee05e2422901c6335f985b9d4447b3ef75524751abfeab6409ad6bf77d4ae3014558"

definition P384_SHA224_TV05_R :: nat where
  "P384_SHA224_TV05_R = 0x98744e5c6742fa5118a74a70db4957647a3cc12add4e876b45974a6a8707809f871daadbfc0b865e01624f706b65f10c"

definition P384_SHA224_TV05_S :: nat where
  "P384_SHA224_TV05_S = 0x9e256e8da8eff5a0c83baaa1ef4f7be798eba9543bf97adb0fff8719f5406ea1207a0cf703d99aa8f02169724b492273"

definition P384_SHA224_TV05_Sig :: "nat × nat" where
  "P384_SHA224_TV05_Sig = (P384_SHA224_TV05_R, P384_SHA224_TV05_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA224_TV05_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV05_d" 
  by eval

lemma P384_SHA224_TV05_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA224_TV05_Q" 
  by eval

lemma P384_SHA224_TV05_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA224_TV05_d = P384_SHA224_TV05_Q" 
  by eval

lemma P384_SHA224_TV05_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA224_TV05_d P384_SHA224_TV05_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA224_TV05_dQvalidPair' P384_SHA224_TV05_d_valid) 

lemma P384_SHA224_TV05_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV05_k" 
  by eval

lemma P384_SHA224_TV05_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA224octets P384_SHA224_TV05_d P384_SHA224_TV05_Msg P384_SHA224_TV05_k = Some P384_SHA224_TV05_Sig" 
  by eval

lemma P384_SHA224_TV05_Verify: 
  "SEC1_P384_ECDSA_Verify SHA224octets P384_SHA224_TV05_Q P384_SHA224_TV05_Msg P384_SHA224_TV05_Sig" 
  by eval

subsection‹Test Vector: P384_SHA224_TV06›

subsubsection‹Given Test Vector Values›

definition P384_SHA224_TV06_Msg :: octets where
  "P384_SHA224_TV06_Msg = nat_to_octets_len 0xcaa83d5ab07febbd2e0fe2d63738b9b7b8752594bea7aaf50345b3d2f316653a8c9222f2b7877b64679e9573e81461a426029e45b8873a575094a1d572e0d32a9f0a9c6bcb9a2868543b7d8bbe4a69a09e7321f05f8366cced1b72df526f895b60aed2c39c249653c7839538770d4e5f47d3926ec0d168ab6a1af15bf1dca1f7 128"

definition P384_SHA224_TV06_d :: nat where
  "P384_SHA224_TV06_d = 0xea7a563ba2a7f5ab69973dca1f1a0d1572f0c59817cd3b62ad356c2099e2cdca1c553323563f9dfbb333b126d84abc7f"

definition P384_SHA224_TV06_Qx :: nat where
  "P384_SHA224_TV06_Qx = 0xcf4717c5f5de668b785f06bdc9845df5a09e4edd83f4669756407cbb60807305c632bc49f818f4a84b194369aa07736f"

definition P384_SHA224_TV06_Qy :: nat where
  "P384_SHA224_TV06_Qy = 0x7391e4982af8a2218f704f627d01f0508bfc8304992a2d598a420bf2eb519f33bd7caf79380793733b3dba0cc5e2b9d8"

definition P384_SHA224_TV06_Q :: "int point" where
  "P384_SHA224_TV06_Q = Point (int P384_SHA224_TV06_Qx) (int P384_SHA224_TV06_Qy)"

definition P384_SHA224_TV06_k :: nat where
  "P384_SHA224_TV06_k = 0x7b9606b3df7b2a340dbc68d9754de0734e1faeb5a0135578a97628d948702235c60b20c8002c8fcf906783e1b389e754"

definition P384_SHA224_TV06_R :: nat where
  "P384_SHA224_TV06_R = 0x0d680010bed373287f9767955b5d2850e150b6713b49e453eb280148e45230c853d99ea2d2f8fcbd3ddcba19aeec0af1"

definition P384_SHA224_TV06_S :: nat where
  "P384_SHA224_TV06_S = 0x64329763a930ab5452afdb0557fef16ff71810d6343dfc9c6ae18905c3d274db6554cdc69d6078a1ca03284474a94f30"

definition P384_SHA224_TV06_Sig :: "nat × nat" where
  "P384_SHA224_TV06_Sig = (P384_SHA224_TV06_R, P384_SHA224_TV06_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA224_TV06_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV06_d" 
  by eval

lemma P384_SHA224_TV06_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA224_TV06_Q" 
  by eval

lemma P384_SHA224_TV06_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA224_TV06_d = P384_SHA224_TV06_Q" 
  by eval

lemma P384_SHA224_TV06_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA224_TV06_d P384_SHA224_TV06_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA224_TV06_dQvalidPair' P384_SHA224_TV06_d_valid) 

lemma P384_SHA224_TV06_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV06_k" 
  by eval

lemma P384_SHA224_TV06_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA224octets P384_SHA224_TV06_d P384_SHA224_TV06_Msg P384_SHA224_TV06_k = Some P384_SHA224_TV06_Sig" 
  by eval

lemma P384_SHA224_TV06_Verify: 
  "SEC1_P384_ECDSA_Verify SHA224octets P384_SHA224_TV06_Q P384_SHA224_TV06_Msg P384_SHA224_TV06_Sig" 
  by eval

subsection‹Test Vector: P384_SHA224_TV07›

subsubsection‹Given Test Vector Values›

definition P384_SHA224_TV07_Msg :: octets where
  "P384_SHA224_TV07_Msg = nat_to_octets_len 0x594603458d6534974aeeafba919c4d0f4cb6843a3af41204bbb88aeb2fca2772d305163dba863da050aabedbaf89db521955d1715de95bbcef979ecdc0c976181ece00355385f8a8f8cce127c9eac15ce3e958a3ed686184674ec9a50eb63271606ee7fdcb1323da3c3db8e89cad1fb42139a32d08abcfbf0d4ccfca18c89a86 128"

definition P384_SHA224_TV07_d :: nat where
  "P384_SHA224_TV07_d = 0x4cc70cb35b3ddeb0df53a6bd7bd05f8ff4392a2db7344f2d443761484b3a468a4ee3d1a8b27113d57283fd18b05f7829"

definition P384_SHA224_TV07_Qx :: nat where
  "P384_SHA224_TV07_Qx = 0x40e1fe21df34bb85a642a0abe819ebd128f7e39b84d8dcc4a9a599b372fb9588da1484600ec28b1297bb685f9ae77831"

definition P384_SHA224_TV07_Qy :: nat where
  "P384_SHA224_TV07_Qy = 0xf3aa69ada57879fdcbe8df19cefabc308add7d03b17b1fac2f7783fece6a8dfe20bc36f518692677d96e3f730a67a671"

definition P384_SHA224_TV07_Q :: "int point" where
  "P384_SHA224_TV07_Q = Point (int P384_SHA224_TV07_Qx) (int P384_SHA224_TV07_Qy)"

definition P384_SHA224_TV07_k :: nat where
  "P384_SHA224_TV07_k = 0x8eda401d98f5688c34d8dbebcd3991c87c0442b0379154eaa2e5287dabe9a9e34cfc1305d11ff68781df25d5611b331d"

definition P384_SHA224_TV07_R :: nat where
  "P384_SHA224_TV07_R = 0xff2d772786e159448bba26afd8c3281941a4cb0c56fec6f5cccb4c292c4ee0f7af9bd39bbe2d88148732585e104fdb30"

definition P384_SHA224_TV07_S :: nat where
  "P384_SHA224_TV07_S = 0x07a1d890770daa949a17797dca7af3e8163da981ec330c03d63d1a8312c152be6a718163205ffa08da7dcc163ba261f4"

definition P384_SHA224_TV07_Sig :: "nat × nat" where
  "P384_SHA224_TV07_Sig = (P384_SHA224_TV07_R, P384_SHA224_TV07_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA224_TV07_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV07_d" 
  by eval

lemma P384_SHA224_TV07_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA224_TV07_Q" 
  by eval

lemma P384_SHA224_TV07_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA224_TV07_d = P384_SHA224_TV07_Q" 
  by eval

lemma P384_SHA224_TV07_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA224_TV07_d P384_SHA224_TV07_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA224_TV07_dQvalidPair' P384_SHA224_TV07_d_valid) 

lemma P384_SHA224_TV07_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV07_k" 
  by eval

lemma P384_SHA224_TV07_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA224octets P384_SHA224_TV07_d P384_SHA224_TV07_Msg P384_SHA224_TV07_k = Some P384_SHA224_TV07_Sig" 
  by eval

lemma P384_SHA224_TV07_Verify: 
  "SEC1_P384_ECDSA_Verify SHA224octets P384_SHA224_TV07_Q P384_SHA224_TV07_Msg P384_SHA224_TV07_Sig" 
  by eval

subsection‹Test Vector: P384_SHA224_TV08›

subsubsection‹Given Test Vector Values›

definition P384_SHA224_TV08_Msg :: octets where
  "P384_SHA224_TV08_Msg = nat_to_octets_len 0x733252d2bd35547838be22656cc7aa67eff0af0b13b428f77267a513c6824c3dbae533068b6817e82665f009560affcfe4b2ddb5b667a644fc1a42d24f24e0947e0dc50fb62c919bc1fe4e7ded5e28f2e6d80fcf66a081fb2763526f8def5a81a4ddd38be0b59ee839da1643eeeaee7b1927cec12cf3da67c02bc5465151e346 128"

definition P384_SHA224_TV08_d :: nat where
  "P384_SHA224_TV08_d = 0x366d15e4cd7605c71560a418bd0f382fd7cd7ad3090ff1b2dfbed74336166a905e1b760cf0bccee7a0e66c5ebfb831f1"

definition P384_SHA224_TV08_Qx :: nat where
  "P384_SHA224_TV08_Qx = 0xa143f277ab36a10b645ff6c58241ea67ffdc8acf12d60973068390f06b4d8f4d773b10c1ebf6889b1cfa73ebb90f6ca1"

definition P384_SHA224_TV08_Qy :: nat where
  "P384_SHA224_TV08_Qy = 0x7a17cad29bb507b309021f6f92cb5c10ba535f4a3e317fcc68cfd02d3ccd269f465169c73d30ff308f5350d881b08aec"

definition P384_SHA224_TV08_Q :: "int point" where
  "P384_SHA224_TV08_Q = Point (int P384_SHA224_TV08_Qx) (int P384_SHA224_TV08_Qy)"

definition P384_SHA224_TV08_k :: nat where
  "P384_SHA224_TV08_k = 0xdbe545f920bc3d704c43d834bab21e40df12ec9e16a619a3e6b3f08760c26aae6e4fd91fad00f745194794b74bb1baee"

definition P384_SHA224_TV08_R :: nat where
  "P384_SHA224_TV08_R = 0xcdc39b12bba30da66fe9554713c05880ddc27afa4d2d151440f124c351fb9496dc95046516b0921083347d64369846ac"

definition P384_SHA224_TV08_S :: nat where
  "P384_SHA224_TV08_S = 0x797d0344e49f9ba87a187c50f664e5015d449e346b1a7bd9427c5be559fc58173651880d5aadf053f81899d3368d6181"

definition P384_SHA224_TV08_Sig :: "nat × nat" where
  "P384_SHA224_TV08_Sig = (P384_SHA224_TV08_R, P384_SHA224_TV08_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA224_TV08_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV08_d" 
  by eval

lemma P384_SHA224_TV08_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA224_TV08_Q" 
  by eval

lemma P384_SHA224_TV08_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA224_TV08_d = P384_SHA224_TV08_Q" 
  by eval

lemma P384_SHA224_TV08_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA224_TV08_d P384_SHA224_TV08_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA224_TV08_dQvalidPair' P384_SHA224_TV08_d_valid) 

lemma P384_SHA224_TV08_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV08_k" 
  by eval

lemma P384_SHA224_TV08_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA224octets P384_SHA224_TV08_d P384_SHA224_TV08_Msg P384_SHA224_TV08_k = Some P384_SHA224_TV08_Sig" 
  by eval

lemma P384_SHA224_TV08_Verify: 
  "SEC1_P384_ECDSA_Verify SHA224octets P384_SHA224_TV08_Q P384_SHA224_TV08_Msg P384_SHA224_TV08_Sig" 
  by eval

subsection‹Test Vector: P384_SHA224_TV09›

subsubsection‹Given Test Vector Values›

definition P384_SHA224_TV09_Msg :: octets where
  "P384_SHA224_TV09_Msg = nat_to_octets_len 0x5a182bd174feb038dfae3346267156bf663167f713dea1ce936b0edb815cd9b8c8e4d411c786ba2494a81442617255db7158b142e720d86c9b56680fb9efd4298cdd69079a28153494c42a24251c7ad42ecf7e97eabc1b3997529b2a297cbad2474269b87a0b1e385f2d7f8b6eb8d1cd75eaf7e91d1acbecd45d7b2bfbbe3216 128"

definition P384_SHA224_TV09_d :: nat where
  "P384_SHA224_TV09_d = 0xe357d869857a52a06e1ece5593d16407022354780eb9a7cb8575cef327f877d22322c006b3c8c11e3d7d296a708bdb6d"

definition P384_SHA224_TV09_Qx :: nat where
  "P384_SHA224_TV09_Qx = 0xce9a2185a68d6094aa5849a6efe78b349946f7380f0c79aa9664246cfcc71a879e90ad78a0474f58644c6a208168150e"

definition P384_SHA224_TV09_Qy :: nat where
  "P384_SHA224_TV09_Qy = 0x8354fa47673cb3e07d446521345706c5515584b2602f921c3b9c44dded9e2c3f90ce47adb36d7e5f9f95a8c5ad8af397"

definition P384_SHA224_TV09_Q :: "int point" where
  "P384_SHA224_TV09_Q = Point (int P384_SHA224_TV09_Qx) (int P384_SHA224_TV09_Qy)"

definition P384_SHA224_TV09_k :: nat where
  "P384_SHA224_TV09_k = 0x1e77367ac4e10924854d135ad2f2507f39e2bafdbce33ff256bcbe9a7329b8d27185218bcc3550aafbe3390e84c77292"

definition P384_SHA224_TV09_R :: nat where
  "P384_SHA224_TV09_R = 0xdf3182d49ad70959fb0c95bc7312750ce70fc87f1a328d39d9b29ac05d31305ce7209d6c24d13225d9567b489f7a187b"

definition P384_SHA224_TV09_S :: nat where
  "P384_SHA224_TV09_S = 0xd812b05abab0e96de13291e1f0da6479444ed5cd9d959b76f6cb43d394769035364f7c831a104dc7b5bd9b4a8e64df64"

definition P384_SHA224_TV09_Sig :: "nat × nat" where
  "P384_SHA224_TV09_Sig = (P384_SHA224_TV09_R, P384_SHA224_TV09_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA224_TV09_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV09_d" 
  by eval

lemma P384_SHA224_TV09_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA224_TV09_Q" 
  by eval

lemma P384_SHA224_TV09_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA224_TV09_d = P384_SHA224_TV09_Q" 
  by eval

lemma P384_SHA224_TV09_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA224_TV09_d P384_SHA224_TV09_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA224_TV09_dQvalidPair' P384_SHA224_TV09_d_valid) 

lemma P384_SHA224_TV09_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV09_k" 
  by eval

lemma P384_SHA224_TV09_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA224octets P384_SHA224_TV09_d P384_SHA224_TV09_Msg P384_SHA224_TV09_k = Some P384_SHA224_TV09_Sig" 
  by eval

lemma P384_SHA224_TV09_Verify: 
  "SEC1_P384_ECDSA_Verify SHA224octets P384_SHA224_TV09_Q P384_SHA224_TV09_Msg P384_SHA224_TV09_Sig" 
  by eval

subsection‹Test Vector: P384_SHA224_TV10›

subsubsection‹Given Test Vector Values›

definition P384_SHA224_TV10_Msg :: octets where
  "P384_SHA224_TV10_Msg = nat_to_octets_len 0xaaa99fb1c71340d785a18f6f668e898c25cf7a0ac31d13c5b388b7233408493a5a109af6d07065376b96f4903df7aba2b2af671a18772bb0472490d1240cde28967680727dd4acd47e0308920a75da857a6eeedee5b6586d45dff3d8a680599665aa895c89dd7770b824b7dee477ac5e7602d409d3cc553090c970b50811dbab 128"

definition P384_SHA224_TV10_d :: nat where
  "P384_SHA224_TV10_d = 0x745a18db47324a3710b993d115b2834339315e84e7006eafd889fb49bd3cc5a8b50c90526e65e6c53bddd2916d14bead"

definition P384_SHA224_TV10_Qx :: nat where
  "P384_SHA224_TV10_Qx = 0xf692578c6f77531210aef55c9e004ce3b66cf268c6900dde31a8bbb76e7562e3fb76242de34ca330d2501030aa119466"

definition P384_SHA224_TV10_Qy :: nat where
  "P384_SHA224_TV10_Qy = 0x40965833b28de926c46de060aa25beaeda98f8415a6b1e3564aa77870cf4c89bd4fde92c8f5d9bf0eb41721586859d8e"

definition P384_SHA224_TV10_Q :: "int point" where
  "P384_SHA224_TV10_Q = Point (int P384_SHA224_TV10_Qx) (int P384_SHA224_TV10_Qy)"

definition P384_SHA224_TV10_k :: nat where
  "P384_SHA224_TV10_k = 0x11b9b36720abcac084efdb44c9f5b7d039e3250cb1e9c47850189ba3cfc1489d858b2a44df357772b61d919c7e729c0f"

definition P384_SHA224_TV10_R :: nat where
  "P384_SHA224_TV10_R = 0x02b252c99820cf50e6ce060ab55bd4f682276e29b4ae4197417432e6a7bfb8cf0bac89dfe105456af805d822cee77696"

definition P384_SHA224_TV10_S :: nat where
  "P384_SHA224_TV10_S = 0x8e248bbf7d7028d63177e565c9d1666ee5be4d1ffbfffc9c7814b0cd38f74b98f3f2cd59be42b9f132bfe5ee789cd96c"

definition P384_SHA224_TV10_Sig :: "nat × nat" where
  "P384_SHA224_TV10_Sig = (P384_SHA224_TV10_R, P384_SHA224_TV10_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA224_TV10_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV10_d" 
  by eval

lemma P384_SHA224_TV10_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA224_TV10_Q" 
  by eval

lemma P384_SHA224_TV10_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA224_TV10_d = P384_SHA224_TV10_Q" 
  by eval

lemma P384_SHA224_TV10_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA224_TV10_d P384_SHA224_TV10_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA224_TV10_dQvalidPair' P384_SHA224_TV10_d_valid) 

lemma P384_SHA224_TV10_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV10_k" 
  by eval

lemma P384_SHA224_TV10_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA224octets P384_SHA224_TV10_d P384_SHA224_TV10_Msg P384_SHA224_TV10_k = Some P384_SHA224_TV10_Sig" 
  by eval

lemma P384_SHA224_TV10_Verify: 
  "SEC1_P384_ECDSA_Verify SHA224octets P384_SHA224_TV10_Q P384_SHA224_TV10_Msg P384_SHA224_TV10_Sig" 
  by eval

subsection‹Test Vector: P384_SHA224_TV11›

subsubsection‹Given Test Vector Values›

definition P384_SHA224_TV11_Msg :: octets where
  "P384_SHA224_TV11_Msg = nat_to_octets_len 0x1fadfa8254d3a0b82d137cfdd82043d5dc1fef195d5297b09cc5cfb061f59c933451c0dc2a11b4037f34f88dacb803251f8880c4b72585c3c196e6fb23484ca43a191f8e41b9b9a37e2e6fcaab6738c3c62d1c98e1c620bb788b7b51a04f998a510efdba0d3418622fe8ce203b3fcd553b9b4206365a39031797ad11e49745ec 128"

definition P384_SHA224_TV11_d :: nat where
  "P384_SHA224_TV11_d = 0x93f20963ea5011ff4f26481e359309e634195f6289134087bd2e83eee008c962780a679784ee7ac6acda03d663ed27e0"

definition P384_SHA224_TV11_Qx :: nat where
  "P384_SHA224_TV11_Qx = 0x0edcde3533ea019e18f1a3cd97b7962e8823dda36c389f8f9287549f796d11376392b8a01c7a80f127a8f75795e04f54"

definition P384_SHA224_TV11_Qy :: nat where
  "P384_SHA224_TV11_Qy = 0x63d7c458dccfc02f5148d755d59f9bbc8e3c3ea34908777928440747795955741296abcdd5386676419ed8049fedb489"

definition P384_SHA224_TV11_Q :: "int point" where
  "P384_SHA224_TV11_Q = Point (int P384_SHA224_TV11_Qx) (int P384_SHA224_TV11_Qy)"

definition P384_SHA224_TV11_k :: nat where
  "P384_SHA224_TV11_k = 0x3ad308faf04c42ee5ac69d36bc0aa9a96aacf55ea0f27dac4f52e088f023d206340a6324874ffad169ff80624de24c96"

definition P384_SHA224_TV11_R :: nat where
  "P384_SHA224_TV11_R = 0x209b72f9aae72c4339813573c3a8408a9e0be641ca863d81d9d14c48d0bf4cd44a1a7985cff07b5d68f3f9478475645b"

definition P384_SHA224_TV11_S :: nat where
  "P384_SHA224_TV11_S = 0xf6292e599b22a76eda95393cf59f4745fa6c472effd1f781879ad9a4437a98080b0b07dadad0c249631c682d2836a977"

definition P384_SHA224_TV11_Sig :: "nat × nat" where
  "P384_SHA224_TV11_Sig = (P384_SHA224_TV11_R, P384_SHA224_TV11_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA224_TV11_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV11_d" 
  by eval

lemma P384_SHA224_TV11_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA224_TV11_Q" 
  by eval

lemma P384_SHA224_TV11_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA224_TV11_d = P384_SHA224_TV11_Q" 
  by eval

lemma P384_SHA224_TV11_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA224_TV11_d P384_SHA224_TV11_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA224_TV11_dQvalidPair' P384_SHA224_TV11_d_valid) 

lemma P384_SHA224_TV11_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV11_k" 
  by eval

lemma P384_SHA224_TV11_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA224octets P384_SHA224_TV11_d P384_SHA224_TV11_Msg P384_SHA224_TV11_k = Some P384_SHA224_TV11_Sig" 
  by eval

lemma P384_SHA224_TV11_Verify: 
  "SEC1_P384_ECDSA_Verify SHA224octets P384_SHA224_TV11_Q P384_SHA224_TV11_Msg P384_SHA224_TV11_Sig" 
  by eval

subsection‹Test Vector: P384_SHA224_TV12›

subsubsection‹Given Test Vector Values›

definition P384_SHA224_TV12_Msg :: octets where
  "P384_SHA224_TV12_Msg = nat_to_octets_len 0x9ecb6f5ed3ba666a8536a81ef65012c2cb8b433508798d84708abb06dfb75503886f78384fb8c7a4d2d49ef539d9b8a0b60938c7f07471dda91f258b0d99691b38a8403a2bb3f956bdfd09baba16d9b6877097a9b6213481b47a06e139d23ec7abad5668d21f912fdb70d31bb9adf9b3ce80e308252fa81a51674f88d02db72b 128"

definition P384_SHA224_TV12_d :: nat where
  "P384_SHA224_TV12_d = 0xf175e6ac42fd48ec9d652c10707c039c67c4cc61d8c45a373dcda6e4ca6c53e947e49c24e01b48e7cdf92edfe6d316a1"

definition P384_SHA224_TV12_Qx :: nat where
  "P384_SHA224_TV12_Qx = 0xa40c64f595491ce15790a5a87fbe64c1800247b42acd08fe5257700719f46afc8acce0e4ede0517a312092d5e3d089cd"

definition P384_SHA224_TV12_Qy :: nat where
  "P384_SHA224_TV12_Qy = 0xd565df9dc2f381cc0c5d84f382a43a98018524c0b4708a44b3e2817f9719f29fbf9c15803591ed9b4790c5adaba9f433"

definition P384_SHA224_TV12_Q :: "int point" where
  "P384_SHA224_TV12_Q = Point (int P384_SHA224_TV12_Qx) (int P384_SHA224_TV12_Qy)"

definition P384_SHA224_TV12_k :: nat where
  "P384_SHA224_TV12_k = 0x812dcaa6d4f9a43ccc553288065d13761581485aa903a500a690ccafbd330ba4818c977b98c4bb57f8a182a1afacfae9"

definition P384_SHA224_TV12_R :: nat where
  "P384_SHA224_TV12_R = 0xd000f18d3e4c162ff0d16f662e6703e7a6f5bff7a333ed266fa4f44c752415946c34945c342c20f739677186b1d80ab3"

definition P384_SHA224_TV12_S :: nat where
  "P384_SHA224_TV12_S = 0xae7f1271c89e0aaa238710d039ea73a69110cc28fcf426f2fe6754b63a59e417fa84f903cf7dccb5468b43ff083bbfd5"

definition P384_SHA224_TV12_Sig :: "nat × nat" where
  "P384_SHA224_TV12_Sig = (P384_SHA224_TV12_R, P384_SHA224_TV12_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA224_TV12_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV12_d" 
  by eval

lemma P384_SHA224_TV12_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA224_TV12_Q" 
  by eval

lemma P384_SHA224_TV12_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA224_TV12_d = P384_SHA224_TV12_Q" 
  by eval

lemma P384_SHA224_TV12_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA224_TV12_d P384_SHA224_TV12_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA224_TV12_dQvalidPair' P384_SHA224_TV12_d_valid) 

lemma P384_SHA224_TV12_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV12_k" 
  by eval

lemma P384_SHA224_TV12_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA224octets P384_SHA224_TV12_d P384_SHA224_TV12_Msg P384_SHA224_TV12_k = Some P384_SHA224_TV12_Sig" 
  by eval

lemma P384_SHA224_TV12_Verify: 
  "SEC1_P384_ECDSA_Verify SHA224octets P384_SHA224_TV12_Q P384_SHA224_TV12_Msg P384_SHA224_TV12_Sig" 
  by eval

subsection‹Test Vector: P384_SHA224_TV13›

subsubsection‹Given Test Vector Values›

definition P384_SHA224_TV13_Msg :: octets where
  "P384_SHA224_TV13_Msg = nat_to_octets_len 0xe55bfca78d98e68d1b63688db12485578f36c489766f4d0bfaa0088433ff12133aaca455805095f2e655940860958b3ead111d9070778ee3bbf3e47e43d9eba8b8d9b1fdf72f793fcde2bcaa334f3e35fa2cca531ea7cf27fe9ccba741e38ac26129b2d612bf54a34e0ae6c166c0fef07fcd2b9ac253d7e041a500f7be7b8369 128"

definition P384_SHA224_TV13_d :: nat where
  "P384_SHA224_TV13_d = 0x46c4f0b228b28aaa0ec8cfdf1d0ed3408b7ae049312fb9eaf5f3892720e68684cc8ad29844a3dc9d110edf6916dfb8bb"

definition P384_SHA224_TV13_Qx :: nat where
  "P384_SHA224_TV13_Qx = 0x13ddec844731b7e30c467451df08ca11d6c581cb64abd8a257671cffd26f5ccad4df7b9ee8924047a88a5d2d7567609c"

definition P384_SHA224_TV13_Qy :: nat where
  "P384_SHA224_TV13_Qy = 0xd74ca94f590fd1d13e190cc1e03c3da6c3faab15c7dda034af3deefee8aeec3628fa8b1978c54cfcd071baa319a46ec0"

definition P384_SHA224_TV13_Q :: "int point" where
  "P384_SHA224_TV13_Q = Point (int P384_SHA224_TV13_Qx) (int P384_SHA224_TV13_Qy)"

definition P384_SHA224_TV13_k :: nat where
  "P384_SHA224_TV13_k = 0x2a9dd520207c40a379cd4036adef9ee60fa8bc8c0d39b3ad91850ac93fd543f218b1688581f23481a090b0e4c73792ac"

definition P384_SHA224_TV13_R :: nat where
  "P384_SHA224_TV13_R = 0x94e08cca20fe3866f643f53ec65faf3f2b4d80cd9bcc8ff8f88bb28da9eada324fc2d048908dd3d08a9e0ebb547731bc"

definition P384_SHA224_TV13_S :: nat where
  "P384_SHA224_TV13_S = 0x8e6f82c4d3069b14f4c844b4ca133a9503493265c9f77a7d4775eda67de76798a23dd7ea48e0ac3c337dd62bf058319d"

definition P384_SHA224_TV13_Sig :: "nat × nat" where
  "P384_SHA224_TV13_Sig = (P384_SHA224_TV13_R, P384_SHA224_TV13_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA224_TV13_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV13_d" 
  by eval

lemma P384_SHA224_TV13_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA224_TV13_Q" 
  by eval

lemma P384_SHA224_TV13_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA224_TV13_d = P384_SHA224_TV13_Q" 
  by eval

lemma P384_SHA224_TV13_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA224_TV13_d P384_SHA224_TV13_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA224_TV13_dQvalidPair' P384_SHA224_TV13_d_valid) 

lemma P384_SHA224_TV13_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV13_k" 
  by eval

lemma P384_SHA224_TV13_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA224octets P384_SHA224_TV13_d P384_SHA224_TV13_Msg P384_SHA224_TV13_k = Some P384_SHA224_TV13_Sig" 
  by eval

lemma P384_SHA224_TV13_Verify: 
  "SEC1_P384_ECDSA_Verify SHA224octets P384_SHA224_TV13_Q P384_SHA224_TV13_Msg P384_SHA224_TV13_Sig" 
  by eval

subsection‹Test Vector: P384_SHA224_TV14›

subsubsection‹Given Test Vector Values›

definition P384_SHA224_TV14_Msg :: octets where
  "P384_SHA224_TV14_Msg = nat_to_octets_len 0x02c6b3c83bd34b288d96409162aa4ff114e9d134bf948046eb5ebcc0c7fe9dfceadda83ed69da2fac00c8840f6c702a3fc5e6959d70f7e8af923e99e4937232ae3b841ffefd2e62fab3671a7c94a0281b8ea5bc176add57c5c9b6893fe7f5d48ce7256b96510810c4e046168a3c5be9843b84d5268a50349b3444341aa5490dd 128"

definition P384_SHA224_TV14_d :: nat where
  "P384_SHA224_TV14_d = 0x1d7b71ef01d0d33a8513a3aed3cabb83829589c8021087a740ca65b570777089be721a61172b874a22a1f81aef3f8bb6"

definition P384_SHA224_TV14_Qx :: nat where
  "P384_SHA224_TV14_Qx = 0x8d2721370df8f097d5a69396249a315f6037dc7045b3da11eacae6d43036f779d5de7053d101768b42cc2b1283a3aaea"

definition P384_SHA224_TV14_Qy :: nat where
  "P384_SHA224_TV14_Qy = 0xa046039ae662141f9954d278183eaa2e03917fe58583e32d344074d59d60caa5b0949c53066525d5cca923e2f201502e"

definition P384_SHA224_TV14_Q :: "int point" where
  "P384_SHA224_TV14_Q = Point (int P384_SHA224_TV14_Qx) (int P384_SHA224_TV14_Qy)"

definition P384_SHA224_TV14_k :: nat where
  "P384_SHA224_TV14_k = 0xd1b25ad25581cad17e96f1d302251681fee5b2efbb71c3c15ff035b2145d015d18e0e52dc3187ab5a560277b3a3929b0"

definition P384_SHA224_TV14_R :: nat where
  "P384_SHA224_TV14_R = 0xd836f52b14c7391744868daa2d5cf27eb9380b9b6176195573d5b04842e9f2fc3794d6cf877feafee63d11b05f6a6bee"

definition P384_SHA224_TV14_S :: nat where
  "P384_SHA224_TV14_S = 0x8b89042fef2c04d4bd6c9d66a06a010514321d623a5f8d57ba5ac3686872eaabca9e0ba2d058ae7028e870acf03ca32d"

definition P384_SHA224_TV14_Sig :: "nat × nat" where
  "P384_SHA224_TV14_Sig = (P384_SHA224_TV14_R, P384_SHA224_TV14_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA224_TV14_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV14_d" 
  by eval

lemma P384_SHA224_TV14_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA224_TV14_Q" 
  by eval

lemma P384_SHA224_TV14_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA224_TV14_d = P384_SHA224_TV14_Q" 
  by eval

lemma P384_SHA224_TV14_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA224_TV14_d P384_SHA224_TV14_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA224_TV14_dQvalidPair' P384_SHA224_TV14_d_valid) 

lemma P384_SHA224_TV14_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV14_k" 
  by eval

lemma P384_SHA224_TV14_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA224octets P384_SHA224_TV14_d P384_SHA224_TV14_Msg P384_SHA224_TV14_k = Some P384_SHA224_TV14_Sig" 
  by eval

lemma P384_SHA224_TV14_Verify: 
  "SEC1_P384_ECDSA_Verify SHA224octets P384_SHA224_TV14_Q P384_SHA224_TV14_Msg P384_SHA224_TV14_Sig" 
  by eval

subsection‹Test Vector: P384_SHA224_TV15›

subsubsection‹Given Test Vector Values›

definition P384_SHA224_TV15_Msg :: octets where
  "P384_SHA224_TV15_Msg = nat_to_octets_len 0x94f8bfbb9dd6c9b6193e84c2023a27dea00fd48356909faec2161972439686c146184f80686bc09e1a698af7df9dea3d24d9e9fd6d7348a146339c839282cf8984345dc6a51096d74ad238c35233012ad729f262481ec7cd6488f13a6ebac3f3d23438c7ccb5a66e2bf820e92b71c730bb12fd64ea1770d1f892e5b1e14a9e5c 128"

definition P384_SHA224_TV15_d :: nat where
  "P384_SHA224_TV15_d = 0xcf53bdd4c91fe5aa4d82f116bd68153c907963fa3c9d478c9462bb03c79039493a8eaeb855773f2df37e4e551d509dcd"

definition P384_SHA224_TV15_Qx :: nat where
  "P384_SHA224_TV15_Qx = 0x3a65b26c08102b44838f8c2327ea080daf1e4fc45bb279ce03af13a2f9575f0fff9e2e4423a58594ce95d1e710b590ce"

definition P384_SHA224_TV15_Qy :: nat where
  "P384_SHA224_TV15_Qy = 0xfe9dcbcb2ec6e8bd8ed3af3ff0aa619e900cc8bab3f50f6e5f79fac09164fb6a2077cc4f1fed3e9ec6899e91db329bf3"

definition P384_SHA224_TV15_Q :: "int point" where
  "P384_SHA224_TV15_Q = Point (int P384_SHA224_TV15_Qx) (int P384_SHA224_TV15_Qy)"

definition P384_SHA224_TV15_k :: nat where
  "P384_SHA224_TV15_k = 0xdf31908c9289d1fe25e055df199591b23e266433ab8657cc82cb3bca96b88720e229f8dfd42d8b78af7db69342430bca"

definition P384_SHA224_TV15_R :: nat where
  "P384_SHA224_TV15_R = 0x6770eea9369d6718e60dd0b91aee845ff7ed7e0fcc91675f56d32e5227fd3a4612bbcb1556fe94a989b9e3bcc25bb20e"

definition P384_SHA224_TV15_S :: nat where
  "P384_SHA224_TV15_S = 0xc43072f706c98126d06a82b04251e3ecb0ba66c4bb6cd7c025919b9cc6019cdc635256d2a7fa017b806b1e88649d2c0d"

definition P384_SHA224_TV15_Sig :: "nat × nat" where
  "P384_SHA224_TV15_Sig = (P384_SHA224_TV15_R, P384_SHA224_TV15_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA224_TV15_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV15_d" 
  by eval

lemma P384_SHA224_TV15_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA224_TV15_Q" 
  by eval

lemma P384_SHA224_TV15_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA224_TV15_d = P384_SHA224_TV15_Q" 
  by eval

lemma P384_SHA224_TV15_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA224_TV15_d P384_SHA224_TV15_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA224_TV15_dQvalidPair' P384_SHA224_TV15_d_valid) 

lemma P384_SHA224_TV15_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA224_TV15_k" 
  by eval

lemma P384_SHA224_TV15_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA224octets P384_SHA224_TV15_d P384_SHA224_TV15_Msg P384_SHA224_TV15_k = Some P384_SHA224_TV15_Sig" 
  by eval

lemma P384_SHA224_TV15_Verify: 
  "SEC1_P384_ECDSA_Verify SHA224octets P384_SHA224_TV15_Q P384_SHA224_TV15_Msg P384_SHA224_TV15_Sig" 
  by eval

section‹NIST Test Vectors for ECDSA: Curve P-384, Hash SHA-256›
  
subsection‹Test Vector: P384_SHA256_TV01›

subsubsection‹Given Test Vector Values›

definition P384_SHA256_TV01_Msg :: octets where
  "P384_SHA256_TV01_Msg = nat_to_octets_len 0x663b12ebf44b7ed3872b385477381f4b11adeb0aec9e0e2478776313d536376dc8fd5f3c715bb6ddf32c01ee1d6f8b731785732c0d8441df636d8145577e7b3138e43c32a61bc1242e0e73d62d624cdc924856076bdbbf1ec04ad4420732ef0c53d42479a08235fcfc4db4d869c4eb2828c73928cdc3e3758362d1b770809997 128"

definition P384_SHA256_TV01_d :: nat where
  "P384_SHA256_TV01_d = 0xc602bc74a34592c311a6569661e0832c84f7207274676cc42a89f058162630184b52f0d99b855a7783c987476d7f9e6b"

definition P384_SHA256_TV01_Qx :: nat where
  "P384_SHA256_TV01_Qx = 0x0400193b21f07cd059826e9453d3e96dd145041c97d49ff6b7047f86bb0b0439e909274cb9c282bfab88674c0765bc75"

definition P384_SHA256_TV01_Qy :: nat where
  "P384_SHA256_TV01_Qy = 0xf70d89c52acbc70468d2c5ae75c76d7f69b76af62dcf95e99eba5dd11adf8f42ec9a425b0c5ec98e2f234a926b82a147"

definition P384_SHA256_TV01_Q :: "int point" where
  "P384_SHA256_TV01_Q = Point (int P384_SHA256_TV01_Qx) (int P384_SHA256_TV01_Qy)"

definition P384_SHA256_TV01_k :: nat where
  "P384_SHA256_TV01_k = 0xc10b5c25c4683d0b7827d0d88697cdc0932496b5299b798c0dd1e7af6cc757ccb30fcd3d36ead4a804877e24f3a32443"

definition P384_SHA256_TV01_R :: nat where
  "P384_SHA256_TV01_R = 0xb11db00cdaf53286d4483f38cd02785948477ed7ebc2ad609054551da0ab0359978c61851788aa2ec3267946d440e878"

definition P384_SHA256_TV01_S :: nat where
  "P384_SHA256_TV01_S = 0x16007873c5b0604ce68112a8fee973e8e2b6e3319c683a762ff5065a076512d7c98b27e74b7887671048ac027df8cbf2"

definition P384_SHA256_TV01_Sig :: "nat × nat" where
  "P384_SHA256_TV01_Sig = (P384_SHA256_TV01_R, P384_SHA256_TV01_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA256_TV01_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV01_d" 
  by eval

lemma P384_SHA256_TV01_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA256_TV01_Q" 
  by eval

lemma P384_SHA256_TV01_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA256_TV01_d = P384_SHA256_TV01_Q" 
  by eval

lemma P384_SHA256_TV01_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA256_TV01_d P384_SHA256_TV01_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA256_TV01_dQvalidPair' P384_SHA256_TV01_d_valid) 

lemma P384_SHA256_TV01_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV01_k" 
  by eval

lemma P384_SHA256_TV01_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA256octets P384_SHA256_TV01_d P384_SHA256_TV01_Msg P384_SHA256_TV01_k = Some P384_SHA256_TV01_Sig" 
  by eval

lemma P384_SHA256_TV01_Verify: 
  "SEC1_P384_ECDSA_Verify SHA256octets P384_SHA256_TV01_Q P384_SHA256_TV01_Msg P384_SHA256_TV01_Sig" 
  by eval

subsection‹Test Vector: P384_SHA256_TV02›

subsubsection‹Given Test Vector Values›

definition P384_SHA256_TV02_Msg :: octets where
  "P384_SHA256_TV02_Msg = nat_to_octets_len 0x784d7f4686c01bea32cb6cab8c089fb25c341080d9832e04feac6ea63a341079cbd562a75365c63cf7e63e7e1dddc9e99db75ccee59c5295340c2bba36f457690a8f05c62ab001e3d6b333780117d1456a9c8b27d6c2504db9c1428dad8ba797a4419914fcc636f0f14ede3fba49b023b12a77a2176b0b8ff55a895dcaf8dbce 128"

definition P384_SHA256_TV02_d :: nat where
  "P384_SHA256_TV02_d = 0x0287f62a5aa8432ff5e95618ec8f9ccaa870dde99c30b51b7673378efe4ccac598f4bbebbfd8993f9abb747b6ad638b9"

definition P384_SHA256_TV02_Qx :: nat where
  "P384_SHA256_TV02_Qx = 0xb36418a3014074ec9bbcc6a4b2367a4fb464cca7ec0a324cb68670d5c5e03e7a7eb07da117c5ea50b665ab62bd02a491"

definition P384_SHA256_TV02_Qy :: nat where
  "P384_SHA256_TV02_Qy = 0x4ea299c30e7d76e2c5905babada2d3bb4ee5eb35a5a23605cdb0d5133471a53eb9e6758e49105a4eaf29d2267ba84ef2"

definition P384_SHA256_TV02_Q :: "int point" where
  "P384_SHA256_TV02_Q = Point (int P384_SHA256_TV02_Qx) (int P384_SHA256_TV02_Qy)"

definition P384_SHA256_TV02_k :: nat where
  "P384_SHA256_TV02_k = 0x935eeab3edeb281fbd4eead0d9c0babd4b10ff18a31663ee9de3bfa9ae8f9d266441158ea31c889ded9b3c592da77fd7"

definition P384_SHA256_TV02_R :: nat where
  "P384_SHA256_TV02_R = 0x738f9cb28f3b991335ef17b62559255faf75cad370a222464a492e27bb173c7f16b22100ada6b695875c7e4b1a28f158"

definition P384_SHA256_TV02_S :: nat where
  "P384_SHA256_TV02_S = 0xbc998c30e1491cd5d60dc7d1c38333165efe036b2a78db9b8f0e85ee68619cfba654e11ae5ca5ee5a87099c27cf22442"

definition P384_SHA256_TV02_Sig :: "nat × nat" where
  "P384_SHA256_TV02_Sig = (P384_SHA256_TV02_R, P384_SHA256_TV02_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA256_TV02_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV02_d" 
  by eval

lemma P384_SHA256_TV02_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA256_TV02_Q" 
  by eval

lemma P384_SHA256_TV02_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA256_TV02_d = P384_SHA256_TV02_Q" 
  by eval

lemma P384_SHA256_TV02_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA256_TV02_d P384_SHA256_TV02_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA256_TV02_dQvalidPair' P384_SHA256_TV02_d_valid) 

lemma P384_SHA256_TV02_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV02_k" 
  by eval

lemma P384_SHA256_TV02_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA256octets P384_SHA256_TV02_d P384_SHA256_TV02_Msg P384_SHA256_TV02_k = Some P384_SHA256_TV02_Sig" 
  by eval

lemma P384_SHA256_TV02_Verify: 
  "SEC1_P384_ECDSA_Verify SHA256octets P384_SHA256_TV02_Q P384_SHA256_TV02_Msg P384_SHA256_TV02_Sig" 
  by eval

subsection‹Test Vector: P384_SHA256_TV03›

subsubsection‹Given Test Vector Values›

definition P384_SHA256_TV03_Msg :: octets where
  "P384_SHA256_TV03_Msg = nat_to_octets_len 0x45e47fccc5bd6801f237cdbeac8f66ebc75f8b71a6da556d2e002352bd85bf269b6bc7c928d7bb1b0422601e4dd80b29d5906f8fcac212fe0eaaf52eda552303259cbcbe532e60abd3d38d786a45e39a2875bce675800a3eaeb9e42983d9fd9031180abd9adccc9ba30c6c198b4202c4dd70f241e969a3c412724b9b595bc28a 128"

definition P384_SHA256_TV03_d :: nat where
  "P384_SHA256_TV03_d = 0xd44d3108873977036c9b97e03f914cba2f5775b68c425d550995574081191da764acc50196f6d2508082a150af5cd41f"

definition P384_SHA256_TV03_Qx :: nat where
  "P384_SHA256_TV03_Qx = 0xc703835d723c85c643260379d8445b0c816fe9534351921e14a8e147fe140ec7b0c4d704f8dc66a232b2333b28f03dee"

definition P384_SHA256_TV03_Qy :: nat where
  "P384_SHA256_TV03_Qy = 0xc5d0bb054053fd86c26f147c4966757aa04b00513a02d427b8d06c16055c607955efdc518d338abfe7927c195dc28588"

definition P384_SHA256_TV03_Q :: "int point" where
  "P384_SHA256_TV03_Q = Point (int P384_SHA256_TV03_Qx) (int P384_SHA256_TV03_Qy)"

definition P384_SHA256_TV03_k :: nat where
  "P384_SHA256_TV03_k = 0xc80f63e080650c8a21e4f63a62ec909adfb7d877f365d11ee1cb260baf112eb4730c161c1d99dba98fc0d5bbd00dc97d"

definition P384_SHA256_TV03_R :: nat where
  "P384_SHA256_TV03_R = 0x81de2810cde421997013513951a3d537c51a013110d6dbb29251410bcb5ba001a9686b8490f1e581e282fd2ed0974b22"

definition P384_SHA256_TV03_S :: nat where
  "P384_SHA256_TV03_S = 0x9cab0bbaffe91c7677ec3dd1f17060211a3cc0be574cbca064aa8c4b66ba6e64f3d80e83da895042ca32d311c388d950"

definition P384_SHA256_TV03_Sig :: "nat × nat" where
  "P384_SHA256_TV03_Sig = (P384_SHA256_TV03_R, P384_SHA256_TV03_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA256_TV03_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV03_d" 
  by eval

lemma P384_SHA256_TV03_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA256_TV03_Q" 
  by eval

lemma P384_SHA256_TV03_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA256_TV03_d = P384_SHA256_TV03_Q" 
  by eval

lemma P384_SHA256_TV03_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA256_TV03_d P384_SHA256_TV03_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA256_TV03_dQvalidPair' P384_SHA256_TV03_d_valid) 

lemma P384_SHA256_TV03_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV03_k" 
  by eval

lemma P384_SHA256_TV03_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA256octets P384_SHA256_TV03_d P384_SHA256_TV03_Msg P384_SHA256_TV03_k = Some P384_SHA256_TV03_Sig" 
  by eval

lemma P384_SHA256_TV03_Verify: 
  "SEC1_P384_ECDSA_Verify SHA256octets P384_SHA256_TV03_Q P384_SHA256_TV03_Msg P384_SHA256_TV03_Sig" 
  by eval

subsection‹Test Vector: P384_SHA256_TV04›

subsubsection‹Given Test Vector Values›

definition P384_SHA256_TV04_Msg :: octets where
  "P384_SHA256_TV04_Msg = nat_to_octets_len 0xc33ff63b4e6891e00b2349b3f2907c417ca355560544a91e24a7a0ee260d6850aeded29fc0176b6039ca6187e8333391047cceaf14b1077df8f147dad84d36b2dac5666dc2f69dc9b58b88cc73956efdb3b47f91831d5875051c76b0c4e9fc087012a1f03eeee85d6745b46aa50bd9cb0110c2c94508765cec162ee1aa841d73 128"

definition P384_SHA256_TV04_d :: nat where
  "P384_SHA256_TV04_d = 0xd5b72cbb6ec68aca46b9c27ad992afd8ffa02cb3067b234fcfa6e272e3b31be760695ff7df988b57663057ab19dd65e3"

definition P384_SHA256_TV04_Qx :: nat where
  "P384_SHA256_TV04_Qx = 0x135a6542612f1468d8a4d01ff1914e532b1dd64d3627db9d403dc325651d3f82b0f6f0fd1dbdeca2be967c4fb3793b5f"

definition P384_SHA256_TV04_Qy :: nat where
  "P384_SHA256_TV04_Qy = 0xcbbd40f6d3a38d0dfb64582ff4789d7b268241bc0c36de2884bccfaeeff3b7b2b46a30bb35719804e0d11124b4e7f480"

definition P384_SHA256_TV04_Q :: "int point" where
  "P384_SHA256_TV04_Q = Point (int P384_SHA256_TV04_Qx) (int P384_SHA256_TV04_Qy)"

definition P384_SHA256_TV04_k :: nat where
  "P384_SHA256_TV04_k = 0x9da6de7c87c101b68db64fea40d97f8ad974ceb88224c6796c690cbf61b8bd8eede8470b3caf6e6106b66cf3f0eebd55"

definition P384_SHA256_TV04_R :: nat where
  "P384_SHA256_TV04_R = 0x17840911ecdf6ae0428b2634f442163c2c11b8dbf0cc7a5596fbe4d33e3e52f9d99e99ad169867b1f39e89c9180cedc2"

definition P384_SHA256_TV04_S :: nat where
  "P384_SHA256_TV04_S = 0xdd7ed67e480866d0474379ea4afff72870746f4feef2153be42f13bf472b1613d7faa5c0abb7f7464070f94d7cf3f234"

definition P384_SHA256_TV04_Sig :: "nat × nat" where
  "P384_SHA256_TV04_Sig = (P384_SHA256_TV04_R, P384_SHA256_TV04_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA256_TV04_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV04_d" 
  by eval

lemma P384_SHA256_TV04_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA256_TV04_Q" 
  by eval

lemma P384_SHA256_TV04_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA256_TV04_d = P384_SHA256_TV04_Q" 
  by eval

lemma P384_SHA256_TV04_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA256_TV04_d P384_SHA256_TV04_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA256_TV04_dQvalidPair' P384_SHA256_TV04_d_valid) 

lemma P384_SHA256_TV04_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV04_k" 
  by eval

lemma P384_SHA256_TV04_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA256octets P384_SHA256_TV04_d P384_SHA256_TV04_Msg P384_SHA256_TV04_k = Some P384_SHA256_TV04_Sig" 
  by eval

lemma P384_SHA256_TV04_Verify: 
  "SEC1_P384_ECDSA_Verify SHA256octets P384_SHA256_TV04_Q P384_SHA256_TV04_Msg P384_SHA256_TV04_Sig" 
  by eval

subsection‹Test Vector: P384_SHA256_TV05›

subsubsection‹Given Test Vector Values›

definition P384_SHA256_TV05_Msg :: octets where
  "P384_SHA256_TV05_Msg = nat_to_octets_len 0xf562f2b9d84b0e96a52532c3b43c39c8018c738bd8dc3797a7de7353971b2729d522d6961b1f2e4df3f6a4bd3653e6d72b74fc0dba92ab939c4b542e994e5db6dd8ed4f56f651e699052e791237ae1f552f990ad156226ae8f7bf17fcbfa564f749604f97e9df0879d50985747d981422a23040fe52f5ec74caf1d4aaad8a710 128"

definition P384_SHA256_TV05_d :: nat where
  "P384_SHA256_TV05_d = 0x218ee54a71ef2ccf012aca231fee28a2c665fc395ff5cd20bde9b8df598c282664abf9159c5b3923132983f945056d93"

definition P384_SHA256_TV05_Qx :: nat where
  "P384_SHA256_TV05_Qx = 0x01989ff07a7a452d8084937448be946bfedac4049cea34b3db6f7c91d07d69e926cce0af3d6e88855a28120cf3dba8df"

definition P384_SHA256_TV05_Qy :: nat where
  "P384_SHA256_TV05_Qy = 0xeb064e029d7539d4b301aabafe8de8870162deffe6383bc63cc005add6ee1d5ced4a5761219c60cd58ad5b2a7c74aaa9"

definition P384_SHA256_TV05_Q :: "int point" where
  "P384_SHA256_TV05_Q = Point (int P384_SHA256_TV05_Qx) (int P384_SHA256_TV05_Qy)"

definition P384_SHA256_TV05_k :: nat where
  "P384_SHA256_TV05_k = 0xc5d39b436d851d94691f5f4aa9ef447f7989d984f279ae8b091aef5449ac062bcc0567740f914624ad5b99fc32f9af0b"

definition P384_SHA256_TV05_R :: nat where
  "P384_SHA256_TV05_R = 0x07d5b1b12877e8cb5e0aa5e71eeeb17bf0aa203064c7e98b3a1798a74dc9717252dc47c7f06aaf1d5fe15b868323bbb9"

definition P384_SHA256_TV05_S :: nat where
  "P384_SHA256_TV05_S = 0x69428cf101a7af5d08161a9fd7af212e02e33b6062aebdce4c96bf3a0684b5394cb902ca7c2dec6e2f01f40c4576009d"

definition P384_SHA256_TV05_Sig :: "nat × nat" where
  "P384_SHA256_TV05_Sig = (P384_SHA256_TV05_R, P384_SHA256_TV05_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA256_TV05_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV05_d" 
  by eval

lemma P384_SHA256_TV05_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA256_TV05_Q" 
  by eval

lemma P384_SHA256_TV05_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA256_TV05_d = P384_SHA256_TV05_Q" 
  by eval

lemma P384_SHA256_TV05_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA256_TV05_d P384_SHA256_TV05_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA256_TV05_dQvalidPair' P384_SHA256_TV05_d_valid) 

lemma P384_SHA256_TV05_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV05_k" 
  by eval

lemma P384_SHA256_TV05_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA256octets P384_SHA256_TV05_d P384_SHA256_TV05_Msg P384_SHA256_TV05_k = Some P384_SHA256_TV05_Sig" 
  by eval

lemma P384_SHA256_TV05_Verify: 
  "SEC1_P384_ECDSA_Verify SHA256octets P384_SHA256_TV05_Q P384_SHA256_TV05_Msg P384_SHA256_TV05_Sig" 
  by eval

subsection‹Test Vector: P384_SHA256_TV06›

subsubsection‹Given Test Vector Values›

definition P384_SHA256_TV06_Msg :: octets where
  "P384_SHA256_TV06_Msg = nat_to_octets_len 0xace953ae851f571d71779aa120915f27450b236da23e9106f8d0756abdd25861937941228d225d5fb1aa1b1ebf759b1e326aeb3b6cd0cd87edd2ab9f6a7ad67b63d2c501d6a550edb2e7c9d216cc8af78dd33546af64d00abed4d0d2cfc5c9a7b5a055dbe8f7547902d185cf46937314832bc5c602419a82ab83dbd9d3bd5aff 128"

definition P384_SHA256_TV06_d :: nat where
  "P384_SHA256_TV06_d = 0xe6ab171f6937c000e144950801ad91023ae8e8476856c2592d9f7d5bb7180fd729211803d39a412ead6c0be761cfa5d1"

definition P384_SHA256_TV06_Qx :: nat where
  "P384_SHA256_TV06_Qx = 0x38bc42b8c9d8866d09b214398d584b1b24a488dfacc3420d1e9506aa825b19fdf1ba74e7b8f547f47b571467fe8c4d1f"

definition P384_SHA256_TV06_Qy :: nat where
  "P384_SHA256_TV06_Qy = 0x5179d62668d3f6a7ab5c8e3761a685e12008fb87d0529a97645f65cfb5364376c1b6682e0ffcddd0bcd995c41d013ad3"

definition P384_SHA256_TV06_Q :: "int point" where
  "P384_SHA256_TV06_Q = Point (int P384_SHA256_TV06_Qx) (int P384_SHA256_TV06_Qy)"

definition P384_SHA256_TV06_k :: nat where
  "P384_SHA256_TV06_k = 0x05e9718aea9669c9e434f73866da5f252dec6d24c47a1c4ee3233450b6ec626de9746ebe095b285558dfc89fc1b622fe"

definition P384_SHA256_TV06_R :: nat where
  "P384_SHA256_TV06_R = 0xdf9bab9dd1f22ec6f27116f38831cb2089aa78aa8c073024a0faddd9a48e810a5e8e2cadd80fbf8dbd6088c71fe30b5b"

definition P384_SHA256_TV06_S :: nat where
  "P384_SHA256_TV06_S = 0x1e0e8718567d12d18558c57f9e87a755c309e4ffb497335a3adfc8d7475ce8fd882d5dc33a8f5a16274b7ad74bb7862a"

definition P384_SHA256_TV06_Sig :: "nat × nat" where
  "P384_SHA256_TV06_Sig = (P384_SHA256_TV06_R, P384_SHA256_TV06_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA256_TV06_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV06_d" 
  by eval

lemma P384_SHA256_TV06_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA256_TV06_Q" 
  by eval

lemma P384_SHA256_TV06_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA256_TV06_d = P384_SHA256_TV06_Q" 
  by eval

lemma P384_SHA256_TV06_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA256_TV06_d P384_SHA256_TV06_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA256_TV06_dQvalidPair' P384_SHA256_TV06_d_valid) 

lemma P384_SHA256_TV06_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV06_k" 
  by eval

lemma P384_SHA256_TV06_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA256octets P384_SHA256_TV06_d P384_SHA256_TV06_Msg P384_SHA256_TV06_k = Some P384_SHA256_TV06_Sig" 
  by eval

lemma P384_SHA256_TV06_Verify: 
  "SEC1_P384_ECDSA_Verify SHA256octets P384_SHA256_TV06_Q P384_SHA256_TV06_Msg P384_SHA256_TV06_Sig" 
  by eval

subsection‹Test Vector: P384_SHA256_TV07›

subsubsection‹Given Test Vector Values›

definition P384_SHA256_TV07_Msg :: octets where
  "P384_SHA256_TV07_Msg = nat_to_octets_len 0x9635ab832240be95301bedb94c5aec169eedc198cbbdfedcf41e9b586143d829b4597a6b2a81902828332825fd84a785f187a3894e21bd99d22c4f94dcf34453fc052f15ec64d1447c932cb38fcdd30b7be851963409c11881438cbaad7e96f9efbde317f2235d66af804477a5dfe9f0c51448383830050ecf228889f83631e1 128"

definition P384_SHA256_TV07_d :: nat where
  "P384_SHA256_TV07_d = 0x14acd516c7198798fd42ab0684d18df1cd1c99e304312752b3035bed6535a8975dff8acfc2ba1675787c817b5bff6960"

definition P384_SHA256_TV07_Qx :: nat where
  "P384_SHA256_TV07_Qx = 0x29909d143cf7ee9c74b11d52f1a8f3ebd4a720c135612ca5618d3f432f03a95602ee75a2057e1d7aab51d0648ac0b334"

definition P384_SHA256_TV07_Qy :: nat where
  "P384_SHA256_TV07_Qy = 0x404b6c5adffbadfa1b0380ae89fed96ec1ca16cc28661e623d0f1c8b130fbaa96dd7257eae2bf03c2d3dcbc3dbc82c58"

definition P384_SHA256_TV07_Q :: "int point" where
  "P384_SHA256_TV07_Q = Point (int P384_SHA256_TV07_Qx) (int P384_SHA256_TV07_Qy)"

definition P384_SHA256_TV07_k :: nat where
  "P384_SHA256_TV07_k = 0x7f623c103eaa9099a0462e55f80519c565adaeffcb57a29993f3a8a92e63a560be8f0fb9d23dc80bff1064bb41abad79"

definition P384_SHA256_TV07_R :: nat where
  "P384_SHA256_TV07_R = 0x932ab291950c16b2b19a8036cd2e905714c6229cb190a73b3ea49c48dd8e76063a453c7c3267a57597d2973678216296"

definition P384_SHA256_TV07_S :: nat where
  "P384_SHA256_TV07_S = 0xd17d4c5ddbb9c27beebf526f113b416c8abfad53d11c4224813c7f351ba41a77dd4e77d6e4a65bef2c9f62cc37a469a5"

definition P384_SHA256_TV07_Sig :: "nat × nat" where
  "P384_SHA256_TV07_Sig = (P384_SHA256_TV07_R, P384_SHA256_TV07_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA256_TV07_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV07_d" 
  by eval

lemma P384_SHA256_TV07_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA256_TV07_Q" 
  by eval

lemma P384_SHA256_TV07_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA256_TV07_d = P384_SHA256_TV07_Q" 
  by eval

lemma P384_SHA256_TV07_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA256_TV07_d P384_SHA256_TV07_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA256_TV07_dQvalidPair' P384_SHA256_TV07_d_valid) 

lemma P384_SHA256_TV07_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV07_k" 
  by eval

lemma P384_SHA256_TV07_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA256octets P384_SHA256_TV07_d P384_SHA256_TV07_Msg P384_SHA256_TV07_k = Some P384_SHA256_TV07_Sig" 
  by eval

lemma P384_SHA256_TV07_Verify: 
  "SEC1_P384_ECDSA_Verify SHA256octets P384_SHA256_TV07_Q P384_SHA256_TV07_Msg P384_SHA256_TV07_Sig" 
  by eval

subsection‹Test Vector: P384_SHA256_TV08›

subsubsection‹Given Test Vector Values›

definition P384_SHA256_TV08_Msg :: octets where
  "P384_SHA256_TV08_Msg = nat_to_octets_len 0xd98b9a7d4fe9d0fd95de5056af164a8b7882cd34ab5bde83a2abb32dc361eb56a479a3a6119db3b91dcad26a42d2206749567f0d97c34a981a91fc734921821a429f6a53401743a5c406ba9d560f956203abc9d1f32f1a13e7d7b290f75c95fdbf857ea597021461c06a3aacfa554ede3d69e4ff03bbbee5b7463ec77de2b3b2 128"

definition P384_SHA256_TV08_d :: nat where
  "P384_SHA256_TV08_d = 0x2e780550984f3a00cb1e412429b33493c6eb6cd86d12f9d80588c247dcf567bd04296d2d4b24b889d9c54954b7f38f57"

definition P384_SHA256_TV08_Qx :: nat where
  "P384_SHA256_TV08_Qx = 0x37dac42ef04663238443ef33e8addee2e78c40d50a1751913a7f5c37d1f23a26c7f86e16055c788b8ca9554f06b2f2ef"

definition P384_SHA256_TV08_Qy :: nat where
  "P384_SHA256_TV08_Qy = 0xbbed1549652904e3d00c39b01cc0460dbaf3185e6190c2705677a9701de1fe56dff4f4d8418ee15059ff8fc36800982d"

definition P384_SHA256_TV08_Q :: "int point" where
  "P384_SHA256_TV08_Q = Point (int P384_SHA256_TV08_Qx) (int P384_SHA256_TV08_Qy)"

definition P384_SHA256_TV08_k :: nat where
  "P384_SHA256_TV08_k = 0xb788ca82811b0d4e4841765c71eafaa1e575378beedcd3860d8b92db3d070ac5aef7c425067860fbee6c50cf0c642bbb"

definition P384_SHA256_TV08_R :: nat where
  "P384_SHA256_TV08_R = 0x7292b3851870daeb2555a8a2fb198ead78739fcfb75327e5c32a82c6b77d58983e5ad548ccb75dcf9411039c9576d9b9"

definition P384_SHA256_TV08_S :: nat where
  "P384_SHA256_TV08_S = 0xa378c61802d9f1dd062b6e18f16416a954018f77df4df95ad1b983570377d5cfce4cc7861759e802c52f81abc4f49aac"

definition P384_SHA256_TV08_Sig :: "nat × nat" where
  "P384_SHA256_TV08_Sig = (P384_SHA256_TV08_R, P384_SHA256_TV08_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA256_TV08_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV08_d" 
  by eval

lemma P384_SHA256_TV08_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA256_TV08_Q" 
  by eval

lemma P384_SHA256_TV08_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA256_TV08_d = P384_SHA256_TV08_Q" 
  by eval

lemma P384_SHA256_TV08_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA256_TV08_d P384_SHA256_TV08_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA256_TV08_dQvalidPair' P384_SHA256_TV08_d_valid) 

lemma P384_SHA256_TV08_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV08_k" 
  by eval

lemma P384_SHA256_TV08_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA256octets P384_SHA256_TV08_d P384_SHA256_TV08_Msg P384_SHA256_TV08_k = Some P384_SHA256_TV08_Sig" 
  by eval

lemma P384_SHA256_TV08_Verify: 
  "SEC1_P384_ECDSA_Verify SHA256octets P384_SHA256_TV08_Q P384_SHA256_TV08_Msg P384_SHA256_TV08_Sig" 
  by eval

subsection‹Test Vector: P384_SHA256_TV09›

subsubsection‹Given Test Vector Values›

definition P384_SHA256_TV09_Msg :: octets where
  "P384_SHA256_TV09_Msg = nat_to_octets_len 0x1b4c754ac1c28dc415a71eac816bde68de7e8db66409af835838c5bb2c605111108a3bf13606ed5d8ade5ed72e50503e0de664416393d178ea4eec834d8d6f15039847b410080fd5529b426e5aadd8451c20ebd92d787921f33e147bcbeb327b104d4aab1157fc1df33e4d768404b5ccb7110055c2508c600f429fd0c21b5784 128"

definition P384_SHA256_TV09_d :: nat where
  "P384_SHA256_TV09_d = 0xa24d0fe90808aecc5d90626d7e6da7c9be5dfd4e1233c7f0f71f1b7c1c6fd318fafe18559c94718f044cf02ed5107cb1"

definition P384_SHA256_TV09_Qx :: nat where
  "P384_SHA256_TV09_Qx = 0xec8ae1fb9bb88589d27d6f27d790392853396f37bc0c381631d85800fc668eea0886bf1c6cff801147df19778d5b1604"

definition P384_SHA256_TV09_Qy :: nat where
  "P384_SHA256_TV09_Qy = 0x1e1a8336c1e2506f8ee388b55cc648ae73b9295ea78467979d2affb364536fad28120f51ec62a67cbb6ce7784780389f"

definition P384_SHA256_TV09_Q :: "int point" where
  "P384_SHA256_TV09_Q = Point (int P384_SHA256_TV09_Qx) (int P384_SHA256_TV09_Qy)"

definition P384_SHA256_TV09_k :: nat where
  "P384_SHA256_TV09_k = 0x755d025509b73cf1ea8817beb772ad150b4c17a52378be187daffe3db0158921e5e552d1ca3c85df28519939f3cb794d"

definition P384_SHA256_TV09_R :: nat where
  "P384_SHA256_TV09_R = 0x23ff2ffa62bbd427d49995d9c9950116e0d5a06ef076a4553448bc109e6482c5e87d4c833bc88de0bc722bc98cae2e61"

definition P384_SHA256_TV09_S :: nat where
  "P384_SHA256_TV09_S = 0x9aea13d487c3ea6917e16374caafcf0321c12a80d28902dd8cd81909bb04b8c439e2491e504756742d0d0bfb15a9c34c"

definition P384_SHA256_TV09_Sig :: "nat × nat" where
  "P384_SHA256_TV09_Sig = (P384_SHA256_TV09_R, P384_SHA256_TV09_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA256_TV09_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV09_d" 
  by eval

lemma P384_SHA256_TV09_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA256_TV09_Q" 
  by eval

lemma P384_SHA256_TV09_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA256_TV09_d = P384_SHA256_TV09_Q" 
  by eval

lemma P384_SHA256_TV09_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA256_TV09_d P384_SHA256_TV09_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA256_TV09_dQvalidPair' P384_SHA256_TV09_d_valid) 

lemma P384_SHA256_TV09_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV09_k" 
  by eval

lemma P384_SHA256_TV09_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA256octets P384_SHA256_TV09_d P384_SHA256_TV09_Msg P384_SHA256_TV09_k = Some P384_SHA256_TV09_Sig" 
  by eval

lemma P384_SHA256_TV09_Verify: 
  "SEC1_P384_ECDSA_Verify SHA256octets P384_SHA256_TV09_Q P384_SHA256_TV09_Msg P384_SHA256_TV09_Sig" 
  by eval

subsection‹Test Vector: P384_SHA256_TV10›

subsubsection‹Given Test Vector Values›

definition P384_SHA256_TV10_Msg :: octets where
  "P384_SHA256_TV10_Msg = nat_to_octets_len 0x3cd8c053741dd9f974c6c5dbf8a1e5728e9b5eafb1cbcfc3452f5fbbda32a8c7564dee157e8d902c52514361da6d972934a56b3276e2a9379e328e24282e0db697c5bc29090fc489ec46b7b188325dd4e96494c250de0f4a89fe2ccf919eaefcfb50c288113e6df92714feb7f46e0822478c796d0f4ff3447a32997e892693ce 128"

definition P384_SHA256_TV10_d :: nat where
  "P384_SHA256_TV10_d = 0x1c172e25732555afee7ded67a496f3f11babc0875898619f4519c29321e201e8ba1149f2c20b48e5efba235d58fea7c3"

definition P384_SHA256_TV10_Qx :: nat where
  "P384_SHA256_TV10_Qx = 0x13e9e2c8bbcfe26e8f5f43c86268c5980ee693236a6b8777f3a7323718baa21005b482d08aafc6fa6e3667d91353544c"

definition P384_SHA256_TV10_Qy :: nat where
  "P384_SHA256_TV10_Qy = 0x9ba181b3ee505be030f87ecd249b00670a791489b42af04976013483ff95b630c91c01e95757e906129f2f9b4ce719a8"

definition P384_SHA256_TV10_Q :: "int point" where
  "P384_SHA256_TV10_Q = Point (int P384_SHA256_TV10_Qx) (int P384_SHA256_TV10_Qy)"

definition P384_SHA256_TV10_k :: nat where
  "P384_SHA256_TV10_k = 0x08aec9a9e58bdc028805eb5dc86073d05fff1f5fb3fd17f510fc08f9272d84ba7aa66b6f77d84fe6360bd538192bf01a"

definition P384_SHA256_TV10_R :: nat where
  "P384_SHA256_TV10_R = 0x2b4337c3dfbc886ffad7858ae2480cb62227e12205a70361c42f1a5ca9e658ee30fc3cf4030d85bd065edad83b99821f"

definition P384_SHA256_TV10_S :: nat where
  "P384_SHA256_TV10_S = 0x2550cef8574bf17fb3d6b0c9d04ab266962bac3621bac233ff2e4989712d2a4a07171c0aebd3040cd6a32c3bd3efb8b5"

definition P384_SHA256_TV10_Sig :: "nat × nat" where
  "P384_SHA256_TV10_Sig = (P384_SHA256_TV10_R, P384_SHA256_TV10_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA256_TV10_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV10_d" 
  by eval

lemma P384_SHA256_TV10_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA256_TV10_Q" 
  by eval

lemma P384_SHA256_TV10_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA256_TV10_d = P384_SHA256_TV10_Q" 
  by eval

lemma P384_SHA256_TV10_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA256_TV10_d P384_SHA256_TV10_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA256_TV10_dQvalidPair' P384_SHA256_TV10_d_valid) 

lemma P384_SHA256_TV10_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV10_k" 
  by eval

lemma P384_SHA256_TV10_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA256octets P384_SHA256_TV10_d P384_SHA256_TV10_Msg P384_SHA256_TV10_k = Some P384_SHA256_TV10_Sig" 
  by eval

lemma P384_SHA256_TV10_Verify: 
  "SEC1_P384_ECDSA_Verify SHA256octets P384_SHA256_TV10_Q P384_SHA256_TV10_Msg P384_SHA256_TV10_Sig" 
  by eval

subsection‹Test Vector: P384_SHA256_TV11›

subsubsection‹Given Test Vector Values›

definition P384_SHA256_TV11_Msg :: octets where
  "P384_SHA256_TV11_Msg = nat_to_octets_len 0xed955dda6d9650124804d3deb6aeef900e520faf98b1ef6f14efcada7ca2433f09329b70897305e59c89024d76e466b28fe02cb2a9b12e2478c66470259d7c282137a19e5a04ffadea55245c0f34a681593fedc42931d8b3321b3d82e9cc102cd00540ad311ec7bd8c9d06db21bea4ca3dc74d98931ae0d40494aefc2345132c 128"

definition P384_SHA256_TV11_d :: nat where
  "P384_SHA256_TV11_d = 0x5b96555dbd602e71d4d5d3aee19fd1ea084ee23d4f55c10937056762bc2015cbded2e898a487f5482ab7e1e971245907"

definition P384_SHA256_TV11_Qx :: nat where
  "P384_SHA256_TV11_Qx = 0x6e14c17bb831b0112d7f3543c5fd17c78379a516c9e0539b03b8b4bfdead2820343fc84b0382807573ded6c4d97b7003"

definition P384_SHA256_TV11_Qy :: nat where
  "P384_SHA256_TV11_Qy = 0x7f60021d2de77546db666721c9aec84c3e2ba8de0ba77443600dc77e6839bbf9316271adb22d4cb47d08f745ecb1dafd"

definition P384_SHA256_TV11_Q :: "int point" where
  "P384_SHA256_TV11_Q = Point (int P384_SHA256_TV11_Qx) (int P384_SHA256_TV11_Qy)"

definition P384_SHA256_TV11_k :: nat where
  "P384_SHA256_TV11_k = 0x7ad6f4ffd2b429ba10c6f112f800cacf1ad508cf8eba880893bb9659c1ddaaec57dcdc093a114500460d457bdde324f2"

definition P384_SHA256_TV11_R :: nat where
  "P384_SHA256_TV11_R = 0xfaea950ca513806bc59028c638d6302ffc86978c3ff1f06db015dd7c4777050186cb8dd871f5e926e1416539c1939c2f"

definition P384_SHA256_TV11_S :: nat where
  "P384_SHA256_TV11_S = 0x2c592240eabb8a1f9878e1b5c9d5d3ced7b3a7ae571f5a86494ed2ca567a36eb72e7bea8934bded29594bccf67ca84bd"

definition P384_SHA256_TV11_Sig :: "nat × nat" where
  "P384_SHA256_TV11_Sig = (P384_SHA256_TV11_R, P384_SHA256_TV11_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA256_TV11_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV11_d" 
  by eval

lemma P384_SHA256_TV11_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA256_TV11_Q" 
  by eval

lemma P384_SHA256_TV11_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA256_TV11_d = P384_SHA256_TV11_Q" 
  by eval

lemma P384_SHA256_TV11_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA256_TV11_d P384_SHA256_TV11_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA256_TV11_dQvalidPair' P384_SHA256_TV11_d_valid) 

lemma P384_SHA256_TV11_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV11_k" 
  by eval

lemma P384_SHA256_TV11_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA256octets P384_SHA256_TV11_d P384_SHA256_TV11_Msg P384_SHA256_TV11_k = Some P384_SHA256_TV11_Sig" 
  by eval

lemma P384_SHA256_TV11_Verify: 
  "SEC1_P384_ECDSA_Verify SHA256octets P384_SHA256_TV11_Q P384_SHA256_TV11_Msg P384_SHA256_TV11_Sig" 
  by eval

subsection‹Test Vector: P384_SHA256_TV12›

subsubsection‹Given Test Vector Values›

definition P384_SHA256_TV12_Msg :: octets where
  "P384_SHA256_TV12_Msg = nat_to_octets_len 0xce395b001da2a58e49691605d44af4206306f62f561bf2394060d2a5591a350277166bed043819035f1e60b5b3fb5ae113ddd0473f8ef6b2b050c472c2a264e1d8b3ca82a4f158c40f2d78d9ce5e5ea6de243f2e1f13f47f6c6f403b270912c81c636be35b396ca58468b3fb60aa83911d61441a0528d973bc31f965d4059080 128"

definition P384_SHA256_TV12_d :: nat where
  "P384_SHA256_TV12_d = 0x8df9c3c710a25192f3dea970910bb3784e3509874cccf4334823eb9f7a8d05b067f2d812d61e878e24b093089a0b8245"

definition P384_SHA256_TV12_Qx :: nat where
  "P384_SHA256_TV12_Qx = 0x92c9e32b20cbe6d4ed0727c6c942cf804a72031d6dfd69078b5e78ebce2d192268f1f5e2abce5aaf1f8d6a35f136837f"

definition P384_SHA256_TV12_Qy :: nat where
  "P384_SHA256_TV12_Qy = 0xd5167905fa7689e03b9fb1487c566f62b36f2bc1c4a2bfb6a836113b5c8d46f7c1ca51b628b14397fbc06ec9a07f4849"

definition P384_SHA256_TV12_Q :: "int point" where
  "P384_SHA256_TV12_Q = Point (int P384_SHA256_TV12_Qx) (int P384_SHA256_TV12_Qy)"

definition P384_SHA256_TV12_k :: nat where
  "P384_SHA256_TV12_k = 0x258dd05919735cd48627c9fe9fac5c252604aa7c2ae0460d7c1149cd96b7bd2ba195ad393bf392a2499f06aead5ba050"

definition P384_SHA256_TV12_R :: nat where
  "P384_SHA256_TV12_R = 0x413793bcce52eda0f5b675a8d687cce86d5c9e1659b38a89e96246b5e05f8b0934d17dbba3b2ea44c838aa5fd87125d1"

definition P384_SHA256_TV12_S :: nat where
  "P384_SHA256_TV12_S = 0xce7309fc2d6e3438818a1a29a997410b025b0403de20795b97c86c46034a6b02afeed279aeb06522d4de941bfdf50469"

definition P384_SHA256_TV12_Sig :: "nat × nat" where
  "P384_SHA256_TV12_Sig = (P384_SHA256_TV12_R, P384_SHA256_TV12_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA256_TV12_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV12_d" 
  by eval

lemma P384_SHA256_TV12_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA256_TV12_Q" 
  by eval

lemma P384_SHA256_TV12_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA256_TV12_d = P384_SHA256_TV12_Q" 
  by eval

lemma P384_SHA256_TV12_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA256_TV12_d P384_SHA256_TV12_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA256_TV12_dQvalidPair' P384_SHA256_TV12_d_valid) 

lemma P384_SHA256_TV12_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV12_k" 
  by eval

lemma P384_SHA256_TV12_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA256octets P384_SHA256_TV12_d P384_SHA256_TV12_Msg P384_SHA256_TV12_k = Some P384_SHA256_TV12_Sig" 
  by eval

lemma P384_SHA256_TV12_Verify: 
  "SEC1_P384_ECDSA_Verify SHA256octets P384_SHA256_TV12_Q P384_SHA256_TV12_Msg P384_SHA256_TV12_Sig" 
  by eval

subsection‹Test Vector: P384_SHA256_TV13›

subsubsection‹Given Test Vector Values›

definition P384_SHA256_TV13_Msg :: octets where
  "P384_SHA256_TV13_Msg = nat_to_octets_len 0xffefe316455ae4ffdb890bb804bf7d31424ea060ecacff419d0f7134ff76ad434063c0ec0f8bb7059584d3a03f3625bb9e9f66ace1a47ac4b8f3e76fc7c420c55edb1427d1fa15b387ad73d02b0595c4e74321be8822752230a0dcfb85d60bfa186da7623a8ec3eb1633f0a294b23ae87216b14ccee9ef56418dcfab9427371e 128"

definition P384_SHA256_TV13_d :: nat where
  "P384_SHA256_TV13_d = 0x6002cb01ad2ce6e7101665d47729c863b6435c3875de57a93f99da834f73e3e6e2b3880e06de3e6bd1d51ea1807ab0d7"

definition P384_SHA256_TV13_Qx :: nat where
  "P384_SHA256_TV13_Qx = 0xe4216e1a20af8e8e3e74653ac016545001066e53e64af679ad1c85841bb475aed3e00ead052ae9955f48d675ff4ace56"

definition P384_SHA256_TV13_Qy :: nat where
  "P384_SHA256_TV13_Qy = 0x8804c17641be21d4c6386902c9c5c888af25d97ca383703ea4a85cf93bbab360c0bbd2993374da499a303778650270b9"

definition P384_SHA256_TV13_Q :: "int point" where
  "P384_SHA256_TV13_Q = Point (int P384_SHA256_TV13_Qx) (int P384_SHA256_TV13_Qy)"

definition P384_SHA256_TV13_k :: nat where
  "P384_SHA256_TV13_k = 0x6b9507fd2844df0949f8b67b6fde986e50173713ac03df2edf65cb339859321cd3a2b9aab8356f95dec62460ab19c822"

definition P384_SHA256_TV13_R :: nat where
  "P384_SHA256_TV13_R = 0x018891f6381ed358b422f79a299cf0789cee783ba388af4d82cbbe17f3709751b7fd9400e9702820c28b9afc62fdf489"

definition P384_SHA256_TV13_S :: nat where
  "P384_SHA256_TV13_S = 0xaef73bd590802b2fd2a65c4f7fec89f9b24ecc199a69254785925f334cd1977c5e1f858bd9830d7d7d243ea707b1af0b"

definition P384_SHA256_TV13_Sig :: "nat × nat" where
  "P384_SHA256_TV13_Sig = (P384_SHA256_TV13_R, P384_SHA256_TV13_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA256_TV13_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV13_d" 
  by eval

lemma P384_SHA256_TV13_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA256_TV13_Q" 
  by eval

lemma P384_SHA256_TV13_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA256_TV13_d = P384_SHA256_TV13_Q" 
  by eval

lemma P384_SHA256_TV13_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA256_TV13_d P384_SHA256_TV13_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA256_TV13_dQvalidPair' P384_SHA256_TV13_d_valid) 

lemma P384_SHA256_TV13_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV13_k" 
  by eval

lemma P384_SHA256_TV13_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA256octets P384_SHA256_TV13_d P384_SHA256_TV13_Msg P384_SHA256_TV13_k = Some P384_SHA256_TV13_Sig" 
  by eval

lemma P384_SHA256_TV13_Verify: 
  "SEC1_P384_ECDSA_Verify SHA256octets P384_SHA256_TV13_Q P384_SHA256_TV13_Msg P384_SHA256_TV13_Sig" 
  by eval

subsection‹Test Vector: P384_SHA256_TV14›

subsubsection‹Given Test Vector Values›

definition P384_SHA256_TV14_Msg :: octets where
  "P384_SHA256_TV14_Msg = nat_to_octets_len 0x304bccb718b3a9e12669913490cc5bcc1979287b56c628fad706c354241e88d10e81445a2853e3fc32ece094ba1abc3fdcab61da27f9a0fca739371049fed462ee6b08fa31cde12720f8144a6f00ce9b1a7a6eadd231f126717074b4efb5c72ce673ca5859000a436f67a338d698759f12c461247c45a361fb6cb661fdbe6714 128"

definition P384_SHA256_TV14_d :: nat where
  "P384_SHA256_TV14_d = 0xd8559c3543afc6f7b3dc037a687bad2630283757ba7862fd23ed14e2151a4cf5fed3d249268f780e0b96b6b46274a2d5"

definition P384_SHA256_TV14_Qx :: nat where
  "P384_SHA256_TV14_Qx = 0x5f94223918f2ec9f0a08342cb99e724881c92453957c59672860f69daac01b660331a0f5845e50f1f27766b219c89e7e"

definition P384_SHA256_TV14_Qy :: nat where
  "P384_SHA256_TV14_Qy = 0xd76d83396130d10d1168d76c7fc83742ffffbe66d9f4da4ca3f95f5ad6dac8cc7bb65d16d317d37aa99fdbf30ec7439c"

definition P384_SHA256_TV14_Q :: "int point" where
  "P384_SHA256_TV14_Q = Point (int P384_SHA256_TV14_Qx) (int P384_SHA256_TV14_Qy)"

definition P384_SHA256_TV14_k :: nat where
  "P384_SHA256_TV14_k = 0x4ad5a92b5b8e170b71c8a7ed419dc624c7680004562b8d16a37b6e639f581ce81d5f0d98cce44d54c4e7136229148340"

definition P384_SHA256_TV14_R :: nat where
  "P384_SHA256_TV14_R = 0xf7baa6a5488ab462ea59aa31a36402b15880c68110b6069f51ede0c3b52a7b1e5bf926fdbe95768931b7d5f87058835c"

definition P384_SHA256_TV14_S :: nat where
  "P384_SHA256_TV14_S = 0x28b1c4ef448a432f7c91b98b0c6471691e888211b6af907369a8930859b8cdb2e94f466a44f4e52f46df9b0d65e35de6"

definition P384_SHA256_TV14_Sig :: "nat × nat" where
  "P384_SHA256_TV14_Sig = (P384_SHA256_TV14_R, P384_SHA256_TV14_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA256_TV14_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV14_d" 
  by eval

lemma P384_SHA256_TV14_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA256_TV14_Q" 
  by eval

lemma P384_SHA256_TV14_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA256_TV14_d = P384_SHA256_TV14_Q" 
  by eval

lemma P384_SHA256_TV14_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA256_TV14_d P384_SHA256_TV14_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA256_TV14_dQvalidPair' P384_SHA256_TV14_d_valid) 

lemma P384_SHA256_TV14_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV14_k" 
  by eval

lemma P384_SHA256_TV14_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA256octets P384_SHA256_TV14_d P384_SHA256_TV14_Msg P384_SHA256_TV14_k = Some P384_SHA256_TV14_Sig" 
  by eval

lemma P384_SHA256_TV14_Verify: 
  "SEC1_P384_ECDSA_Verify SHA256octets P384_SHA256_TV14_Q P384_SHA256_TV14_Msg P384_SHA256_TV14_Sig" 
  by eval

subsection‹Test Vector: P384_SHA256_TV15›

subsubsection‹Given Test Vector Values›

definition P384_SHA256_TV15_Msg :: octets where
  "P384_SHA256_TV15_Msg = nat_to_octets_len 0x64f9f05c2805acf59c047b5f5d2e20c39277b6d6380f70f87b72327a76170b872bfe4b25c451602acfb6a631bb885e2655aee8abe44f69c90fb21ffde03cef2a452c468c6369867dfd8aa26ac24e16aa53b292375a8d8fbf988e302bf00088e4c061aa12c421d8fe3cbd7273b0e8993701df1c59431f436a08b8e15bd123d133 128"

definition P384_SHA256_TV15_d :: nat where
  "P384_SHA256_TV15_d = 0xb9208cbfd186ddfa3efd5b71342ae1efb01a13ebc4c2a992a2cbee7254b7846a4252ece1104b89d13d835911f8511224"

definition P384_SHA256_TV15_Qx :: nat where
  "P384_SHA256_TV15_Qx = 0x166e6d96cb60d916fd19888a2dd945a3306ff0d7b0a5e30729f47d3dac3de2be3fd5cd7437e9a80d6c48cf960d2d36f8"

definition P384_SHA256_TV15_Qy :: nat where
  "P384_SHA256_TV15_Qy = 0xe6b2b70f131092ae210f29cc6bad701318bddb31bddf921695855c6208941100d0cee5d10799f8b835afe3ea510e8229"

definition P384_SHA256_TV15_Q :: "int point" where
  "P384_SHA256_TV15_Q = Point (int P384_SHA256_TV15_Qx) (int P384_SHA256_TV15_Qy)"

definition P384_SHA256_TV15_k :: nat where
  "P384_SHA256_TV15_k = 0xda706ab5f61531f2378b3c0a2b342108cd119eadaa88b859df64923bccfb0ec2393fd312826f65c15a6587d1d460015b"

definition P384_SHA256_TV15_R :: nat where
  "P384_SHA256_TV15_R = 0xd9124c42858080c62400e4d4d8136304e03d910cbe9b9b3487f4d27c7e0540a314d34bef8c850045c8746ca631c11c42"

definition P384_SHA256_TV15_S :: nat where
  "P384_SHA256_TV15_S = 0xbbf6424a3b70166fa799f49e918439d515327039258ef9bd88435a59c9c19659f8ec3c8660720b0c08354ff60e0f5a76"

definition P384_SHA256_TV15_Sig :: "nat × nat" where
  "P384_SHA256_TV15_Sig = (P384_SHA256_TV15_R, P384_SHA256_TV15_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA256_TV15_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV15_d" 
  by eval

lemma P384_SHA256_TV15_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA256_TV15_Q" 
  by eval

lemma P384_SHA256_TV15_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA256_TV15_d = P384_SHA256_TV15_Q" 
  by eval

lemma P384_SHA256_TV15_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA256_TV15_d P384_SHA256_TV15_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA256_TV15_dQvalidPair' P384_SHA256_TV15_d_valid) 

lemma P384_SHA256_TV15_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA256_TV15_k" 
  by eval

lemma P384_SHA256_TV15_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA256octets P384_SHA256_TV15_d P384_SHA256_TV15_Msg P384_SHA256_TV15_k = Some P384_SHA256_TV15_Sig" 
  by eval

lemma P384_SHA256_TV15_Verify: 
  "SEC1_P384_ECDSA_Verify SHA256octets P384_SHA256_TV15_Q P384_SHA256_TV15_Msg P384_SHA256_TV15_Sig" 
  by eval

section‹NIST Test Vectors for ECDSA: Curve P-384, Hash SHA-384›
  
subsection‹Test Vector: P384_SHA384_TV01›

subsubsection‹Given Test Vector Values›

definition P384_SHA384_TV01_Msg :: octets where
  "P384_SHA384_TV01_Msg = nat_to_octets_len 0x6b45d88037392e1371d9fd1cd174e9c1838d11c3d6133dc17e65fa0c485dcca9f52d41b60161246039e42ec784d49400bffdb51459f5de654091301a09378f93464d52118b48d44b30d781eb1dbed09da11fb4c818dbd442d161aba4b9edc79f05e4b7e401651395b53bd8b5bd3f2aaa6a00877fa9b45cadb8e648550b4c6cbe 128"

definition P384_SHA384_TV01_d :: nat where
  "P384_SHA384_TV01_d = 0x201b432d8df14324182d6261db3e4b3f46a8284482d52e370da41e6cbdf45ec2952f5db7ccbce3bc29449f4fb080ac97"

definition P384_SHA384_TV01_Qx :: nat where
  "P384_SHA384_TV01_Qx = 0xc2b47944fb5de342d03285880177ca5f7d0f2fcad7678cce4229d6e1932fcac11bfc3c3e97d942a3c56bf34123013dbf"

definition P384_SHA384_TV01_Qy :: nat where
  "P384_SHA384_TV01_Qy = 0x37257906a8223866eda0743c519616a76a758ae58aee81c5fd35fbf3a855b7754a36d4a0672df95d6c44a81cf7620c2d"

definition P384_SHA384_TV01_Q :: "int point" where
  "P384_SHA384_TV01_Q = Point (int P384_SHA384_TV01_Qx) (int P384_SHA384_TV01_Qy)"

definition P384_SHA384_TV01_k :: nat where
  "P384_SHA384_TV01_k = 0xdcedabf85978e090f733c6e16646fa34df9ded6e5ce28c6676a00f58a25283db8885e16ce5bf97f917c81e1f25c9c771"

definition P384_SHA384_TV01_R :: nat where
  "P384_SHA384_TV01_R = 0x50835a9251bad008106177ef004b091a1e4235cd0da84fff54542b0ed755c1d6f251609d14ecf18f9e1ddfe69b946e32"

definition P384_SHA384_TV01_S :: nat where
  "P384_SHA384_TV01_S = 0x0475f3d30c6463b646e8d3bf2455830314611cbde404be518b14464fdb195fdcc92eb222e61f426a4a592c00a6a89721"

definition P384_SHA384_TV01_Sig :: "nat × nat" where
  "P384_SHA384_TV01_Sig = (P384_SHA384_TV01_R, P384_SHA384_TV01_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA384_TV01_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV01_d" 
  by eval

lemma P384_SHA384_TV01_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA384_TV01_Q" 
  by eval

lemma P384_SHA384_TV01_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA384_TV01_d = P384_SHA384_TV01_Q" 
  by eval

lemma P384_SHA384_TV01_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA384_TV01_d P384_SHA384_TV01_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA384_TV01_dQvalidPair' P384_SHA384_TV01_d_valid) 

lemma P384_SHA384_TV01_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV01_k" 
  by eval

lemma P384_SHA384_TV01_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA384octets P384_SHA384_TV01_d P384_SHA384_TV01_Msg P384_SHA384_TV01_k = Some P384_SHA384_TV01_Sig" 
  by eval

lemma P384_SHA384_TV01_Verify: 
  "SEC1_P384_ECDSA_Verify SHA384octets P384_SHA384_TV01_Q P384_SHA384_TV01_Msg P384_SHA384_TV01_Sig" 
  by eval

subsection‹Test Vector: P384_SHA384_TV02›

subsubsection‹Given Test Vector Values›

definition P384_SHA384_TV02_Msg :: octets where
  "P384_SHA384_TV02_Msg = nat_to_octets_len 0xd768f41e6e8ec2125d6cf5786d1ba96668ac6566c5cdbbe407f7f2051f3ad6b1acdbfe13edf0d0a86fa110f405406b69085219b5a234ebdb93153241f785d45811b3540d1c37424cc7194424787a51b79679266484c787fb1ded6d1a26b9567d5ea68f04be416caf3be9bd2cafa208fe2a9e234d3ae557c65d3fe6da4cb48da4 128"

definition P384_SHA384_TV02_d :: nat where
  "P384_SHA384_TV02_d = 0x23d9f4ea6d87b7d6163d64256e3449255db14786401a51daa7847161bf56d494325ad2ac8ba928394e01061d882c3528"

definition P384_SHA384_TV02_Qx :: nat where
  "P384_SHA384_TV02_Qx = 0x5d42d6301c54a438f65970bae2a098cbc567e98840006e356221966c86d82e8eca515bca850eaa3cd41f175f03a0cbfd"

definition P384_SHA384_TV02_Qy :: nat where
  "P384_SHA384_TV02_Qy = 0x4aef5a0ceece95d382bd70ab5ce1cb77408bae42b51a08816d5e5e1d3da8c18fcc95564a752730b0aabea983ccea4e2e"

definition P384_SHA384_TV02_Q :: "int point" where
  "P384_SHA384_TV02_Q = Point (int P384_SHA384_TV02_Qx) (int P384_SHA384_TV02_Qy)"

definition P384_SHA384_TV02_k :: nat where
  "P384_SHA384_TV02_k = 0x67ba379366049008593eac124f59ab017358892ee0c063d38f3758bb849fd25d867c3561563cac1532a323b228dc0890"

definition P384_SHA384_TV02_R :: nat where
  "P384_SHA384_TV02_R = 0xfb318f4cb1276282bb43f733a7fb7c567ce94f4d02924fc758635ab2d1107108bf159b85db080cdc3b30fbb5400016f3"

definition P384_SHA384_TV02_S :: nat where
  "P384_SHA384_TV02_S = 0x588e3d7af5da03eae255ecb1813100d95edc243476b724b22db8e85377660d7645ddc1c2c2ee4eaea8b683dbe22f86ca"

definition P384_SHA384_TV02_Sig :: "nat × nat" where
  "P384_SHA384_TV02_Sig = (P384_SHA384_TV02_R, P384_SHA384_TV02_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA384_TV02_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV02_d" 
  by eval

lemma P384_SHA384_TV02_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA384_TV02_Q" 
  by eval

lemma P384_SHA384_TV02_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA384_TV02_d = P384_SHA384_TV02_Q" 
  by eval

lemma P384_SHA384_TV02_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA384_TV02_d P384_SHA384_TV02_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA384_TV02_dQvalidPair' P384_SHA384_TV02_d_valid) 

lemma P384_SHA384_TV02_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV02_k" 
  by eval

lemma P384_SHA384_TV02_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA384octets P384_SHA384_TV02_d P384_SHA384_TV02_Msg P384_SHA384_TV02_k = Some P384_SHA384_TV02_Sig" 
  by eval

lemma P384_SHA384_TV02_Verify: 
  "SEC1_P384_ECDSA_Verify SHA384octets P384_SHA384_TV02_Q P384_SHA384_TV02_Msg P384_SHA384_TV02_Sig" 
  by eval

subsection‹Test Vector: P384_SHA384_TV03›

subsubsection‹Given Test Vector Values›

definition P384_SHA384_TV03_Msg :: octets where
  "P384_SHA384_TV03_Msg = nat_to_octets_len 0x6af6652e92a17b7898e40b6776fabaf0d74cf88d8f0ebfa6088309cbe09fac472eeac2aa8ea96b8c12e993d14c93f8ef4e8b547afe7ae5e4f3973170b35deb3239898918c70c1056332c3f894cd643d2d9b93c2561aac069577bbab45803250a31cd62226cab94d8cba7261dce9fe88c210c212b54329d76a273522c8ba91ddf 128"

definition P384_SHA384_TV03_d :: nat where
  "P384_SHA384_TV03_d = 0xb5f670e98d8befc46f6f51fb2997069550c2a52ebfb4e5e25dd905352d9ef89eed5c2ecd16521853aadb1b52b8c42ae6"

definition P384_SHA384_TV03_Qx :: nat where
  "P384_SHA384_TV03_Qx = 0x44ffb2a3a95e12d87c72b5ea0a8a7cb89f56b3bd46342b2303608d7216301c21b5d2921d80b6628dc512ccb84e2fc278"

definition P384_SHA384_TV03_Qy :: nat where
  "P384_SHA384_TV03_Qy = 0xe4c1002f1828abaec768cadcb7cf42fbf93b1709ccae6df5b134c41fae2b9a188bfbe1eccff0bd348517d7227f2071a6"

definition P384_SHA384_TV03_Q :: "int point" where
  "P384_SHA384_TV03_Q = Point (int P384_SHA384_TV03_Qx) (int P384_SHA384_TV03_Qy)"

definition P384_SHA384_TV03_k :: nat where
  "P384_SHA384_TV03_k = 0x229e67638f712f57bea4c2b02279d5ccad1e7c9e201c77f6f01aeb81ea90e62b44b2d2107fd66d35e56608fff65e28e4"

definition P384_SHA384_TV03_R :: nat where
  "P384_SHA384_TV03_R = 0xb11db592e4ebc75b6472b879b1d8ce57452c615aef20f67a280f8bca9b11a30ad4ac9d69541258c7dd5d0b4ab8dd7d49"

definition P384_SHA384_TV03_S :: nat where
  "P384_SHA384_TV03_S = 0x4eb51db8004e46d438359abf060a9444616cb46b4f99c9a05b53ba6df02e914c9c0b6cc3a9791d804d2e4c0984dab1cc"

definition P384_SHA384_TV03_Sig :: "nat × nat" where
  "P384_SHA384_TV03_Sig = (P384_SHA384_TV03_R, P384_SHA384_TV03_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA384_TV03_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV03_d" 
  by eval

lemma P384_SHA384_TV03_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA384_TV03_Q" 
  by eval

lemma P384_SHA384_TV03_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA384_TV03_d = P384_SHA384_TV03_Q" 
  by eval

lemma P384_SHA384_TV03_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA384_TV03_d P384_SHA384_TV03_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA384_TV03_dQvalidPair' P384_SHA384_TV03_d_valid) 

lemma P384_SHA384_TV03_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV03_k" 
  by eval

lemma P384_SHA384_TV03_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA384octets P384_SHA384_TV03_d P384_SHA384_TV03_Msg P384_SHA384_TV03_k = Some P384_SHA384_TV03_Sig" 
  by eval

lemma P384_SHA384_TV03_Verify: 
  "SEC1_P384_ECDSA_Verify SHA384octets P384_SHA384_TV03_Q P384_SHA384_TV03_Msg P384_SHA384_TV03_Sig" 
  by eval

subsection‹Test Vector: P384_SHA384_TV04›

subsubsection‹Given Test Vector Values›

definition P384_SHA384_TV04_Msg :: octets where
  "P384_SHA384_TV04_Msg = nat_to_octets_len 0xb96d74b2265dd895d94e25092fb9262dc4f2f7a328a3c0c3da134b2d0a4e2058ca994e3445c5ff4f812738e1b0c0f7a126486942a12e674a21f22d0886d68df2375f41685d694d487a718024933a7c4306f33f1a4267d469c530b0fed4e7dea520a19dd68bf0203cc87cad652260ed43b7b23f6ed140d3085875190191a0381a 128"

definition P384_SHA384_TV04_d :: nat where
  "P384_SHA384_TV04_d = 0xde5975d8932533f092e76295ed6b23f10fc5fba48bfb82c6cc714826baf0126813247f8bd51d5738503654ab22459976"

definition P384_SHA384_TV04_Qx :: nat where
  "P384_SHA384_TV04_Qx = 0xf1fabafc01fec7e96d982528d9ef3a2a18b7fe8ae0fa0673977341c7ae4ae8d8d3d67420343d013a984f5f61da29ae38"

definition P384_SHA384_TV04_Qy :: nat where
  "P384_SHA384_TV04_Qy = 0x1a31cf902c46343d01b2ebb614bc789c313b5f91f9302ad9418e9c797563e2fa3d44500f47b4e26ad8fdec1a816d1dcf"

definition P384_SHA384_TV04_Q :: "int point" where
  "P384_SHA384_TV04_Q = Point (int P384_SHA384_TV04_Qx) (int P384_SHA384_TV04_Qy)"

definition P384_SHA384_TV04_k :: nat where
  "P384_SHA384_TV04_k = 0xfc5940e661542436f9265c34bce407eff6364bd471aa79b90c906d923e15c9ed96eea4e86f3238ea86161d13b7d9359d"

definition P384_SHA384_TV04_R :: nat where
  "P384_SHA384_TV04_R = 0xc2fbdd6a56789024082173725d797ef9fd6accb6ae664b7260f9e83cb8ab2490428c8b9c52e153612295432fec4d59cd"

definition P384_SHA384_TV04_S :: nat where
  "P384_SHA384_TV04_S = 0x8056c5bb57f41f73082888b234fcda320a33250b5da012ba1fdb4924355ae679012d81d2c08fc0f8634c708a4833232f"

definition P384_SHA384_TV04_Sig :: "nat × nat" where
  "P384_SHA384_TV04_Sig = (P384_SHA384_TV04_R, P384_SHA384_TV04_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA384_TV04_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV04_d" 
  by eval

lemma P384_SHA384_TV04_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA384_TV04_Q" 
  by eval

lemma P384_SHA384_TV04_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA384_TV04_d = P384_SHA384_TV04_Q" 
  by eval

lemma P384_SHA384_TV04_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA384_TV04_d P384_SHA384_TV04_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA384_TV04_dQvalidPair' P384_SHA384_TV04_d_valid) 

lemma P384_SHA384_TV04_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV04_k" 
  by eval

lemma P384_SHA384_TV04_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA384octets P384_SHA384_TV04_d P384_SHA384_TV04_Msg P384_SHA384_TV04_k = Some P384_SHA384_TV04_Sig" 
  by eval

lemma P384_SHA384_TV04_Verify: 
  "SEC1_P384_ECDSA_Verify SHA384octets P384_SHA384_TV04_Q P384_SHA384_TV04_Msg P384_SHA384_TV04_Sig" 
  by eval

subsection‹Test Vector: P384_SHA384_TV05›

subsubsection‹Given Test Vector Values›

definition P384_SHA384_TV05_Msg :: octets where
  "P384_SHA384_TV05_Msg = nat_to_octets_len 0x7cec7480a037ff40c232c1d2d6e8cd4c080bbeecdaf3886fccc9f129bb6d202c316eca76c8ad4e76079afe622f833a16f4907e817260c1fa68b10c7a151a37eb8c036b057ed4652c353db4b4a34b37c9a2b300fb5f5fcfb8aa8adae13db359160f70a9241546140e550af0073468683377e6771b6508327408c245d78911c2cc 128"

definition P384_SHA384_TV05_d :: nat where
  "P384_SHA384_TV05_d = 0x11e0d470dc31fab0f5722f87b74a6c8d7414115e58ceb38bfcdced367beac3adbf1fe9ba5a04f72e978b1eb54597eabc"

definition P384_SHA384_TV05_Qx :: nat where
  "P384_SHA384_TV05_Qx = 0x1950166989164cbfd97968c7e8adb6fbca1873ebef811ea259eb48b7d584627f0e6d6c64defe23cbc95236505a252aa1"

definition P384_SHA384_TV05_Qy :: nat where
  "P384_SHA384_TV05_Qy = 0x41ef424b5cb076d4e32accd9250ea75fcf4ffd81814040c050d58c0a29b06be11edf67c911b403e418b7277417e52906"

definition P384_SHA384_TV05_Q :: "int point" where
  "P384_SHA384_TV05_Q = Point (int P384_SHA384_TV05_Qx) (int P384_SHA384_TV05_Qy)"

definition P384_SHA384_TV05_k :: nat where
  "P384_SHA384_TV05_k = 0xe56904028226eb04f8d071e3f9cefec91075a81ca0fa87b44cae148fe1ce9827b5d1910db2336d0eb9813ddba3e4d7b5"

definition P384_SHA384_TV05_R :: nat where
  "P384_SHA384_TV05_R = 0xc38ef30f55624e8935680c29f8c24824877cf48ffc0ef015e62de1068893353030d1193bf9d34237d7ce6ba92c98b0fe"

definition P384_SHA384_TV05_S :: nat where
  "P384_SHA384_TV05_S = 0x651b8c3d5c9d5b936d300802a06d82ad54f7b1ba4327b2f031c0c5b0cb215ad4354edc7f932d934e877dfa1cf51b13fe"

definition P384_SHA384_TV05_Sig :: "nat × nat" where
  "P384_SHA384_TV05_Sig = (P384_SHA384_TV05_R, P384_SHA384_TV05_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA384_TV05_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV05_d" 
  by eval

lemma P384_SHA384_TV05_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA384_TV05_Q" 
  by eval

lemma P384_SHA384_TV05_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA384_TV05_d = P384_SHA384_TV05_Q" 
  by eval

lemma P384_SHA384_TV05_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA384_TV05_d P384_SHA384_TV05_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA384_TV05_dQvalidPair' P384_SHA384_TV05_d_valid) 

lemma P384_SHA384_TV05_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV05_k" 
  by eval

lemma P384_SHA384_TV05_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA384octets P384_SHA384_TV05_d P384_SHA384_TV05_Msg P384_SHA384_TV05_k = Some P384_SHA384_TV05_Sig" 
  by eval

lemma P384_SHA384_TV05_Verify: 
  "SEC1_P384_ECDSA_Verify SHA384octets P384_SHA384_TV05_Q P384_SHA384_TV05_Msg P384_SHA384_TV05_Sig" 
  by eval

subsection‹Test Vector: P384_SHA384_TV06›

subsubsection‹Given Test Vector Values›

definition P384_SHA384_TV06_Msg :: octets where
  "P384_SHA384_TV06_Msg = nat_to_octets_len 0x00ce978603229710345c9ad7c1c2dba3596b196528eea25bd822d43ca8f76a024e29217703dd0652c8a615284fc3edcc1c5ad1c8d5a8521c8e104c016a24e50c2e25066dcb56596f913b872767e3627aa3e55ec812e9fdac7c2f1beade83aef093e24c9c953982adf431a776880ae4583be158e11cdab1cbca3ad3a66900213d 128"

definition P384_SHA384_TV06_d :: nat where
  "P384_SHA384_TV06_d = 0x5c6bbf9fbcbb7b97c9535f57b431ed1ccae1945b7e8a4f1b032016b07810bd24a9e20055c0e9306650df59ef7e2cd8c2"

definition P384_SHA384_TV06_Qx :: nat where
  "P384_SHA384_TV06_Qx = 0x2e01c5b59e619e00b79060a1e8ef695472e23bf9a511fc3d5ed77a334a242557098e40972713732c5291c97adf9cf2cf"

definition P384_SHA384_TV06_Qy :: nat where
  "P384_SHA384_TV06_Qy = 0x563e3fe4ad807e803b9e961b08da4dde4cea8925649da0d93221ce4cdceabc6a1db7612180a8c6bef3579c65539b97e9"

definition P384_SHA384_TV06_Q :: "int point" where
  "P384_SHA384_TV06_Q = Point (int P384_SHA384_TV06_Qx) (int P384_SHA384_TV06_Qy)"

definition P384_SHA384_TV06_k :: nat where
  "P384_SHA384_TV06_k = 0x03d23f1277b949cb6380211ad9d338e6f76c3eedac95989b91d0243cfb734a54b19bca45a5d13d6a4b9f815d919eea77"

definition P384_SHA384_TV06_R :: nat where
  "P384_SHA384_TV06_R = 0xabab65308f0b79c4f3a9ff28dd490acb0c320434094cef93e75adfe17e5820dc1f77544cfaaacdc8cf9ac8b38e174bef"

definition P384_SHA384_TV06_S :: nat where
  "P384_SHA384_TV06_S = 0x11b783d879a6de054b316af7d56e526c3dce96c85289122e3ad927cfa77bfc50b4a96c97f85b1b8221be2df083ff58fb"

definition P384_SHA384_TV06_Sig :: "nat × nat" where
  "P384_SHA384_TV06_Sig = (P384_SHA384_TV06_R, P384_SHA384_TV06_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA384_TV06_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV06_d" 
  by eval

lemma P384_SHA384_TV06_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA384_TV06_Q" 
  by eval

lemma P384_SHA384_TV06_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA384_TV06_d = P384_SHA384_TV06_Q" 
  by eval

lemma P384_SHA384_TV06_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA384_TV06_d P384_SHA384_TV06_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA384_TV06_dQvalidPair' P384_SHA384_TV06_d_valid) 

lemma P384_SHA384_TV06_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV06_k" 
  by eval

lemma P384_SHA384_TV06_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA384octets P384_SHA384_TV06_d P384_SHA384_TV06_Msg P384_SHA384_TV06_k = Some P384_SHA384_TV06_Sig" 
  by eval

lemma P384_SHA384_TV06_Verify: 
  "SEC1_P384_ECDSA_Verify SHA384octets P384_SHA384_TV06_Q P384_SHA384_TV06_Msg P384_SHA384_TV06_Sig" 
  by eval

subsection‹Test Vector: P384_SHA384_TV07›

subsubsection‹Given Test Vector Values›

definition P384_SHA384_TV07_Msg :: octets where
  "P384_SHA384_TV07_Msg = nat_to_octets_len 0x54a255c18692c6162a46add176a0ae8361dcb8948f092d8d7bac83e160431794d3b9812849bf1994bcdcfba56e8540c8a9ee5b93414548f2a653191b6bb28bda8dc70d45cc1b92a489f58a2d54f85766cb3c90de7dd88e690d8ebc9a79987eee1989df35af5e35522f83d85c48dda89863171c8b0bf4853ae28c2ac45c764416 128"

definition P384_SHA384_TV07_d :: nat where
  "P384_SHA384_TV07_d = 0xffc7dedeff8343721f72046bc3c126626c177b0e48e247f44fd61f8469d4d5f0a74147fabaa334495cc1f986ebc5f0b1"

definition P384_SHA384_TV07_Qx :: nat where
  "P384_SHA384_TV07_Qx = 0x51c78c979452edd53b563f63eb3e854a5b23e87f1b2103942b65f77d024471f75c8ce1cc0dfef83292b368112aa5126e"

definition P384_SHA384_TV07_Qy :: nat where
  "P384_SHA384_TV07_Qy = 0x313e6aaf09caa3ba30f13072b2134878f14a4a01ee86326cccbff3d079b4df097dc57985e8c8c834a10cb9d766169366"

definition P384_SHA384_TV07_Q :: "int point" where
  "P384_SHA384_TV07_Q = Point (int P384_SHA384_TV07_Qx) (int P384_SHA384_TV07_Qy)"

definition P384_SHA384_TV07_k :: nat where
  "P384_SHA384_TV07_k = 0xc3de91dbe4f777698773da70dd610ef1a7efe4dc00d734399c7dd100728006a502822a5a7ff9129ffd8adf6c1fc1211a"

definition P384_SHA384_TV07_R :: nat where
  "P384_SHA384_TV07_R = 0xf4f477855819ad8b1763f53691b76afbc4a31a638b1e08c293f9bcd55decf797f9913ca128d4b45b2e2ea3e82c6cf565"

definition P384_SHA384_TV07_S :: nat where
  "P384_SHA384_TV07_S = 0x7c26be29569ef95480a6d0c1af49dc10a51a0a8931345e48c0c39498bfb94d62962980b56143a7b41a2fddc8794c1b7f"

definition P384_SHA384_TV07_Sig :: "nat × nat" where
  "P384_SHA384_TV07_Sig = (P384_SHA384_TV07_R, P384_SHA384_TV07_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA384_TV07_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV07_d" 
  by eval

lemma P384_SHA384_TV07_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA384_TV07_Q" 
  by eval

lemma P384_SHA384_TV07_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA384_TV07_d = P384_SHA384_TV07_Q" 
  by eval

lemma P384_SHA384_TV07_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA384_TV07_d P384_SHA384_TV07_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA384_TV07_dQvalidPair' P384_SHA384_TV07_d_valid) 

lemma P384_SHA384_TV07_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV07_k" 
  by eval

lemma P384_SHA384_TV07_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA384octets P384_SHA384_TV07_d P384_SHA384_TV07_Msg P384_SHA384_TV07_k = Some P384_SHA384_TV07_Sig" 
  by eval

lemma P384_SHA384_TV07_Verify: 
  "SEC1_P384_ECDSA_Verify SHA384octets P384_SHA384_TV07_Q P384_SHA384_TV07_Msg P384_SHA384_TV07_Sig" 
  by eval

subsection‹Test Vector: P384_SHA384_TV08›

subsubsection‹Given Test Vector Values›

definition P384_SHA384_TV08_Msg :: octets where
  "P384_SHA384_TV08_Msg = nat_to_octets_len 0x692a78f90d4f9d5aee5da536314a78d68c1feabbfe5d1ccea7f6059a66c4b310f8051c411c409ccf6e19a0cbd8b8e100c48317fe8c6d4f8a638b9551ce7ee178020f04f7da3001a0e6855225fb3c9b375e4ed964588a1a41a095f3f476c42d52ffd23ce1702c93b56d4425d3befcf75d0951b6fd5c05b05455bdaf205fe70ca2 128"

definition P384_SHA384_TV08_d :: nat where
  "P384_SHA384_TV08_d = 0xadca364ef144a21df64b163615e8349cf74ee9dbf728104215c532073a7f74e2f67385779f7f74ab344cc3c7da061cf6"

definition P384_SHA384_TV08_Qx :: nat where
  "P384_SHA384_TV08_Qx = 0xef948daae68242330a7358ef73f23b56c07e37126266db3fa6eea233a04a9b3e4915233dd6754427cd4b71b75854077d"

definition P384_SHA384_TV08_Qy :: nat where
  "P384_SHA384_TV08_Qy = 0x009453ef1828eaff9e17c856d4fc1895ab60051312c3e1db1e3766566438b2990cbf9945c2545619e3e0145bc6a79004"

definition P384_SHA384_TV08_Q :: "int point" where
  "P384_SHA384_TV08_Q = Point (int P384_SHA384_TV08_Qx) (int P384_SHA384_TV08_Qy)"

definition P384_SHA384_TV08_k :: nat where
  "P384_SHA384_TV08_k = 0xa2da3fae2e6da3cf11b49861afb34fba357fea89f54b35ce5ed7434ae09103fe53e2be75b93fc579fedf919f6d5e407e"

definition P384_SHA384_TV08_R :: nat where
  "P384_SHA384_TV08_R = 0xdda994b9c428b57e9f8bbaebba0d682e3aac6ed828e3a1e99a7fc4c804bff8df151137f539c7389d80e23d9f3ee497bf"

definition P384_SHA384_TV08_S :: nat where
  "P384_SHA384_TV08_S = 0xa0d6b10ceffd0e1b29cf784476f9173ba6ecd2cfc7929725f2d6e24e0db5a4721683640eaa2bbe151fb57560f9ce594b"

definition P384_SHA384_TV08_Sig :: "nat × nat" where
  "P384_SHA384_TV08_Sig = (P384_SHA384_TV08_R, P384_SHA384_TV08_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA384_TV08_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV08_d" 
  by eval

lemma P384_SHA384_TV08_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA384_TV08_Q" 
  by eval

lemma P384_SHA384_TV08_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA384_TV08_d = P384_SHA384_TV08_Q" 
  by eval

lemma P384_SHA384_TV08_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA384_TV08_d P384_SHA384_TV08_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA384_TV08_dQvalidPair' P384_SHA384_TV08_d_valid) 

lemma P384_SHA384_TV08_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV08_k" 
  by eval

lemma P384_SHA384_TV08_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA384octets P384_SHA384_TV08_d P384_SHA384_TV08_Msg P384_SHA384_TV08_k = Some P384_SHA384_TV08_Sig" 
  by eval

lemma P384_SHA384_TV08_Verify: 
  "SEC1_P384_ECDSA_Verify SHA384octets P384_SHA384_TV08_Q P384_SHA384_TV08_Msg P384_SHA384_TV08_Sig" 
  by eval

subsection‹Test Vector: P384_SHA384_TV09›

subsubsection‹Given Test Vector Values›

definition P384_SHA384_TV09_Msg :: octets where
  "P384_SHA384_TV09_Msg = nat_to_octets_len 0x3b309bb912ab2a51681451ed18ad79e95d968abc35423a67036a02af92f575a0c89f1b668afe22c7037ad1199e757a8f06b281c33e9a40bab69c9874e0bb680b905d909b9dc24a9fe89bb3d7f7d47082b25093c59754f8c19d1f81f30334a8cdd50a3cb72f96d4b3c305e60a439a7e93aeb640dd3c8de37d63c60fb469c2d3ed 128"

definition P384_SHA384_TV09_d :: nat where
  "P384_SHA384_TV09_d = 0x39bea008ec8a217866dcbdb1b93da34d1d3e851d011df9ef44b7828b3453a54aa70f1df9932170804eacd207e4f7e91d"

definition P384_SHA384_TV09_Qx :: nat where
  "P384_SHA384_TV09_Qx = 0x5709ec4305a9c3271c304face6c148142490b827a73a4c17affcfd01fffd7eaa65d2fdedfa2419fc64ed910823513faf"

definition P384_SHA384_TV09_Qy :: nat where
  "P384_SHA384_TV09_Qy = 0xb083cda1cf3be6371b6c06e729ea6299213428db57119347247ec1fcd44204386cc0bca3f452d9d864b39efbfc89d6b2"

definition P384_SHA384_TV09_Q :: "int point" where
  "P384_SHA384_TV09_Q = Point (int P384_SHA384_TV09_Qx) (int P384_SHA384_TV09_Qy)"

definition P384_SHA384_TV09_k :: nat where
  "P384_SHA384_TV09_k = 0x3c90cc7b6984056f570542a51cbe497ce4c11aeae8fc35e8fd6a0d9adeb650e8644f9d1d5e4341b5adc81e27f284c08f"

definition P384_SHA384_TV09_R :: nat where
  "P384_SHA384_TV09_R = 0xd13646895afb1bfd1953551bb922809c95ad65d6abe94eb3719c899aa1f6dba6b01222c7f283900fe98628b7597b6ea6"

definition P384_SHA384_TV09_S :: nat where
  "P384_SHA384_TV09_S = 0x4a9a38afda04c0a6b0058943b679bd02205b14d0f3d49b8f31aac289129780cdb1c555def8c3f9106b478729e0c7efaa"

definition P384_SHA384_TV09_Sig :: "nat × nat" where
  "P384_SHA384_TV09_Sig = (P384_SHA384_TV09_R, P384_SHA384_TV09_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA384_TV09_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV09_d" 
  by eval

lemma P384_SHA384_TV09_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA384_TV09_Q" 
  by eval

lemma P384_SHA384_TV09_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA384_TV09_d = P384_SHA384_TV09_Q" 
  by eval

lemma P384_SHA384_TV09_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA384_TV09_d P384_SHA384_TV09_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA384_TV09_dQvalidPair' P384_SHA384_TV09_d_valid) 

lemma P384_SHA384_TV09_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV09_k" 
  by eval

lemma P384_SHA384_TV09_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA384octets P384_SHA384_TV09_d P384_SHA384_TV09_Msg P384_SHA384_TV09_k = Some P384_SHA384_TV09_Sig" 
  by eval

lemma P384_SHA384_TV09_Verify: 
  "SEC1_P384_ECDSA_Verify SHA384octets P384_SHA384_TV09_Q P384_SHA384_TV09_Msg P384_SHA384_TV09_Sig" 
  by eval

subsection‹Test Vector: P384_SHA384_TV10›

subsubsection‹Given Test Vector Values›

definition P384_SHA384_TV10_Msg :: octets where
  "P384_SHA384_TV10_Msg = nat_to_octets_len 0xf072b72b8783289463da118613c43824d11441dba364c289de03ff5fab3a6f60e85957d8ff211f1cb62fa90216fb727106f692e5ae0844b11b710e5a12c69df3ed895b94e8769ecd15ff433762d6e8e94d8e6a72645b213b0231344e2c968056766c5dd6b5a5df41971858b85e99afbf859400f839b42cd129068efabeea4a26 128"

definition P384_SHA384_TV10_d :: nat where
  "P384_SHA384_TV10_d = 0xe849cf948b241362e3e20c458b52df044f2a72deb0f41c1bb0673e7c04cdd70811215059032b5ca3cc69c345dcce4cf7"

definition P384_SHA384_TV10_Qx :: nat where
  "P384_SHA384_TV10_Qx = 0x06c037a0cbf43fdf335dff33de06d34348405353f9fdf2ce1361efba30fb204aea9dbd2e30da0a10fd2d876188371be6"

definition P384_SHA384_TV10_Qy :: nat where
  "P384_SHA384_TV10_Qy = 0x360d38f3940e34679204b98fbf70b8a4d97f25443e46d0807ab634ed5891ad864dd7703557aa933cd380e26eea662a43"

definition P384_SHA384_TV10_Q :: "int point" where
  "P384_SHA384_TV10_Q = Point (int P384_SHA384_TV10_Qx) (int P384_SHA384_TV10_Qy)"

definition P384_SHA384_TV10_k :: nat where
  "P384_SHA384_TV10_k = 0x32386b2593c85e877b70e5e5495936f65dc49553caef1aa6cc14d9cd370c442a0ccfab4c0da9ec311b67913b1b575a9d"

definition P384_SHA384_TV10_R :: nat where
  "P384_SHA384_TV10_R = 0x5886078d3495767e330c7507b7ca0fa07a50e59912a416d89f0ab1aa4e88153d6eaf00882d1b4aa64153153352d853b5"

definition P384_SHA384_TV10_S :: nat where
  "P384_SHA384_TV10_S = 0x2cc10023bf1bf8ccfd14b06b82cc2114449a352389c8ff9f6f78cdc4e32bde69f3869da0e17f691b329682ae7a36e1aa"

definition P384_SHA384_TV10_Sig :: "nat × nat" where
  "P384_SHA384_TV10_Sig = (P384_SHA384_TV10_R, P384_SHA384_TV10_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA384_TV10_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV10_d" 
  by eval

lemma P384_SHA384_TV10_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA384_TV10_Q" 
  by eval

lemma P384_SHA384_TV10_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA384_TV10_d = P384_SHA384_TV10_Q" 
  by eval

lemma P384_SHA384_TV10_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA384_TV10_d P384_SHA384_TV10_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA384_TV10_dQvalidPair' P384_SHA384_TV10_d_valid) 

lemma P384_SHA384_TV10_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV10_k" 
  by eval

lemma P384_SHA384_TV10_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA384octets P384_SHA384_TV10_d P384_SHA384_TV10_Msg P384_SHA384_TV10_k = Some P384_SHA384_TV10_Sig" 
  by eval

lemma P384_SHA384_TV10_Verify: 
  "SEC1_P384_ECDSA_Verify SHA384octets P384_SHA384_TV10_Q P384_SHA384_TV10_Msg P384_SHA384_TV10_Sig" 
  by eval

subsection‹Test Vector: P384_SHA384_TV11›

subsubsection‹Given Test Vector Values›

definition P384_SHA384_TV11_Msg :: octets where
  "P384_SHA384_TV11_Msg = nat_to_octets_len 0xcf4945350be8133b575c4ad6c9585e0b83ff1ed17989b6cd6c71b41b5264e828b4e115995b1ae77528e7e9002ac1b5669064442645929f9d7dd70927cb93f95edeb73e8624f4bc897ec4c2c7581cb626916f29b2d6e6c2fba8c59a71e30754b459d81b912a12798182bcff4019c7bdfe929cc769bcc2414befe7d2906add4271 128"

definition P384_SHA384_TV11_d :: nat where
  "P384_SHA384_TV11_d = 0xd89607475d509ef23dc9f476eae4280c986de741b63560670fa2bd605f5049f1972792c0413a5b3b4b34e7a38b70b7ca"

definition P384_SHA384_TV11_Qx :: nat where
  "P384_SHA384_TV11_Qx = 0x49a1c631f31cf5c45b2676b1f130cbf9be683d0a50dffae0d147c1e9913ab1090c6529a84f47ddc7cf025921b771355a"

definition P384_SHA384_TV11_Qy :: nat where
  "P384_SHA384_TV11_Qy = 0x1e207eece62f2bcc6bdabc1113158145170be97469a2904eaaa93aad85b86a19719207f3e423051f5b9cbbe2754eefcb"

definition P384_SHA384_TV11_Q :: "int point" where
  "P384_SHA384_TV11_Q = Point (int P384_SHA384_TV11_Qx) (int P384_SHA384_TV11_Qy)"

definition P384_SHA384_TV11_k :: nat where
  "P384_SHA384_TV11_k = 0x78613c570c8d33b7dd1bd1561d87e36282e8cf4843e7c344a2b2bb6a0da94756d670eeaffe434f7ae7c780f7cf05ca08"

definition P384_SHA384_TV11_R :: nat where
  "P384_SHA384_TV11_R = 0x66f92b39aa3f4aeb9e2dc03ac3855406fa3ebbab0a6c88a78d7a03482f0c9868d7b78bc081ede0947c7f37bf193074ba"

definition P384_SHA384_TV11_S :: nat where
  "P384_SHA384_TV11_S = 0xe5c64ed98d7f3701193f25dd237d59c91c0da6e26215e0889d82e6d3e416693f8d58843cf30ab10ab8d0edd9170b53ad"

definition P384_SHA384_TV11_Sig :: "nat × nat" where
  "P384_SHA384_TV11_Sig = (P384_SHA384_TV11_R, P384_SHA384_TV11_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA384_TV11_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV11_d" 
  by eval

lemma P384_SHA384_TV11_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA384_TV11_Q" 
  by eval

lemma P384_SHA384_TV11_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA384_TV11_d = P384_SHA384_TV11_Q" 
  by eval

lemma P384_SHA384_TV11_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA384_TV11_d P384_SHA384_TV11_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA384_TV11_dQvalidPair' P384_SHA384_TV11_d_valid) 

lemma P384_SHA384_TV11_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV11_k" 
  by eval

lemma P384_SHA384_TV11_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA384octets P384_SHA384_TV11_d P384_SHA384_TV11_Msg P384_SHA384_TV11_k = Some P384_SHA384_TV11_Sig" 
  by eval

lemma P384_SHA384_TV11_Verify: 
  "SEC1_P384_ECDSA_Verify SHA384octets P384_SHA384_TV11_Q P384_SHA384_TV11_Msg P384_SHA384_TV11_Sig" 
  by eval

subsection‹Test Vector: P384_SHA384_TV12›

subsubsection‹Given Test Vector Values›

definition P384_SHA384_TV12_Msg :: octets where
  "P384_SHA384_TV12_Msg = nat_to_octets_len 0xd9b5cf0b50416573ff3c63133275a18394dd4326be2041e8d97e6e4e3855a4a177e9d26dfd223fe8aa74564edb49bd72de19916fb6f001f44530d5c18e2c332bce1b7415df5927ece5f3824f34d174b963136b53aef1fb78fb0c06a201a40b2db38e4d8216fc1e392a798c8ab4b3a314496b7f1087804ebfa89bf96e9cdb80c0 128"

definition P384_SHA384_TV12_d :: nat where
  "P384_SHA384_TV12_d = 0x083e7152734adf342520ae377087a223688de2899b10cfcb34a0b36bca500a4dfa530e2343e6a39da7ae1eb0862b4a0d"

definition P384_SHA384_TV12_Qx :: nat where
  "P384_SHA384_TV12_Qx = 0x70a0f16b6c61172659b027ed19b18fd8f57bd28dc0501f207bd6b0bb065b5671cf3dd1ed13d388dcf6ccc766597aa604"

definition P384_SHA384_TV12_Qy :: nat where
  "P384_SHA384_TV12_Qy = 0x4f845bf01c3c3f6126a7368c3454f51425801ee0b72e63fb6799b4420bfdebe3e37c7246db627cc82c09654979c700bb"

definition P384_SHA384_TV12_Q :: "int point" where
  "P384_SHA384_TV12_Q = Point (int P384_SHA384_TV12_Qx) (int P384_SHA384_TV12_Qy)"

definition P384_SHA384_TV12_k :: nat where
  "P384_SHA384_TV12_k = 0x28096ababe29a075fbdf894709a20d0fdedb01ed3eeacb642a33a0da6aed726e13caf6cf206792ec359f0c9f9b567552"

definition P384_SHA384_TV12_R :: nat where
  "P384_SHA384_TV12_R = 0xee2923f9b9999ea05b5e57f505bed5c6ba0420def42c6fa90eef7a6ef770786525546de27cdeb2f8586f8f29fb4ee67c"

definition P384_SHA384_TV12_S :: nat where
  "P384_SHA384_TV12_S = 0x50ef923fb217c4cf65a48b94412fda430fac685f0da7bd574557c6c50f5b22e0c8354d99f2c2f2c2691f252f93c7d84a"

definition P384_SHA384_TV12_Sig :: "nat × nat" where
  "P384_SHA384_TV12_Sig = (P384_SHA384_TV12_R, P384_SHA384_TV12_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA384_TV12_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV12_d" 
  by eval

lemma P384_SHA384_TV12_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA384_TV12_Q" 
  by eval

lemma P384_SHA384_TV12_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA384_TV12_d = P384_SHA384_TV12_Q" 
  by eval

lemma P384_SHA384_TV12_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA384_TV12_d P384_SHA384_TV12_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA384_TV12_dQvalidPair' P384_SHA384_TV12_d_valid) 

lemma P384_SHA384_TV12_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV12_k" 
  by eval

lemma P384_SHA384_TV12_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA384octets P384_SHA384_TV12_d P384_SHA384_TV12_Msg P384_SHA384_TV12_k = Some P384_SHA384_TV12_Sig" 
  by eval

lemma P384_SHA384_TV12_Verify: 
  "SEC1_P384_ECDSA_Verify SHA384octets P384_SHA384_TV12_Q P384_SHA384_TV12_Msg P384_SHA384_TV12_Sig" 
  by eval

subsection‹Test Vector: P384_SHA384_TV13›

subsubsection‹Given Test Vector Values›

definition P384_SHA384_TV13_Msg :: octets where
  "P384_SHA384_TV13_Msg = nat_to_octets_len 0x9e4042d8438a405475b7dab1cd783eb6ce1d1bffa46ac9dfda622b23ac31057b922eced8e2ed7b3241efeafd7c9ab372bf16230f7134647f2956fb793989d3c885a5ae064e85ed971b64f5f561e7ddb79d49aa6ebe727c671c67879b794554c04de0e05d68264855745ef3c9567bd646d5c5f8728b797c181b6b6a876e167663 128"

definition P384_SHA384_TV13_d :: nat where
  "P384_SHA384_TV13_d = 0x63578d416215aff2cc78f9b926d4c7740a77c142944e104aa7422b19a616898262d46a8a942d5e8d5db135ee8b09a368"

definition P384_SHA384_TV13_Qx :: nat where
  "P384_SHA384_TV13_Qx = 0xcadbacef4406099316db2ce3206adc636c2bb0a835847ed7941efb02862472f3150338f13f4860d47f39b7e098f0a390"

definition P384_SHA384_TV13_Qy :: nat where
  "P384_SHA384_TV13_Qy = 0x752ad0f22c9c264336cde11bbc95d1816ed4d1b1500db6b8dce259a42832e613c31178c2c7995206a62e201ba108f570"

definition P384_SHA384_TV13_Q :: "int point" where
  "P384_SHA384_TV13_Q = Point (int P384_SHA384_TV13_Qx) (int P384_SHA384_TV13_Qy)"

definition P384_SHA384_TV13_k :: nat where
  "P384_SHA384_TV13_k = 0x7b69c5d5b4d05c9950dc94c27d58403b4c52c004b80a80418ad3a89aabc5d34f21926729e76afd280cc8ee88c9805a2a"

definition P384_SHA384_TV13_R :: nat where
  "P384_SHA384_TV13_R = 0xdb054addb6161ee49c6ce2e4d646d7670754747b6737ca8516e9d1e87859937c3ef9b1d2663e10d7e4bd00ec85b7a97a"

definition P384_SHA384_TV13_S :: nat where
  "P384_SHA384_TV13_S = 0xfcc504e0f00ef29587e4bc22faada4db30e2cb1ac552680a65785ae87beb666c792513f2be7a3180fc544296841a0e27"

definition P384_SHA384_TV13_Sig :: "nat × nat" where
  "P384_SHA384_TV13_Sig = (P384_SHA384_TV13_R, P384_SHA384_TV13_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA384_TV13_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV13_d" 
  by eval

lemma P384_SHA384_TV13_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA384_TV13_Q" 
  by eval

lemma P384_SHA384_TV13_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA384_TV13_d = P384_SHA384_TV13_Q" 
  by eval

lemma P384_SHA384_TV13_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA384_TV13_d P384_SHA384_TV13_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA384_TV13_dQvalidPair' P384_SHA384_TV13_d_valid) 

lemma P384_SHA384_TV13_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV13_k" 
  by eval

lemma P384_SHA384_TV13_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA384octets P384_SHA384_TV13_d P384_SHA384_TV13_Msg P384_SHA384_TV13_k = Some P384_SHA384_TV13_Sig" 
  by eval

lemma P384_SHA384_TV13_Verify: 
  "SEC1_P384_ECDSA_Verify SHA384octets P384_SHA384_TV13_Q P384_SHA384_TV13_Msg P384_SHA384_TV13_Sig" 
  by eval

subsection‹Test Vector: P384_SHA384_TV14›

subsubsection‹Given Test Vector Values›

definition P384_SHA384_TV14_Msg :: octets where
  "P384_SHA384_TV14_Msg = nat_to_octets_len 0x0b14a7484a40b68a3ce1273b8a48b8fdb65ba900d98541c4bbd07b97e31bcc4c85545a03e9deab3c563f47a036ff60d0361684ba241b5aa68bb46f440da22181ee328a011de98eff34ba235ec10612b07bdfa6b3dc4ccc5e82d3a8d057e1862fef3def5a1804696f84699fda2ec4175a54a4d08bcb4f0406fdac4eddadf5e29b 128"

definition P384_SHA384_TV14_d :: nat where
  "P384_SHA384_TV14_d = 0xed4df19971658b74868800b3b81bc877807743b25c65740f1d6377542afe2c6427612c840ada31a8eb794718f37c7283"

definition P384_SHA384_TV14_Qx :: nat where
  "P384_SHA384_TV14_Qx = 0x33093a0568757e8b58df5b72ea5fe5bf26e6f7aeb541b4c6a8c189c93721749bcaceccf2982a2f0702586a9f812fc66f"

definition P384_SHA384_TV14_Qy :: nat where
  "P384_SHA384_TV14_Qy = 0xebe320d09e1f0662189d50b85a20403b821ac0d000afdbf66a0a33f304726c69e354d81c50b94ba3a5250efc31319cd1"

definition P384_SHA384_TV14_Q :: "int point" where
  "P384_SHA384_TV14_Q = Point (int P384_SHA384_TV14_Qx) (int P384_SHA384_TV14_Qy)"

definition P384_SHA384_TV14_k :: nat where
  "P384_SHA384_TV14_k = 0xd9b4cd1bdfa83e608289634dbfcee643f07315baf743fc91922880b55a2feda3b38ddf6040d3ba10985cd1285fc690d5"

definition P384_SHA384_TV14_R :: nat where
  "P384_SHA384_TV14_R = 0x009c74063e206a4259b53decff5445683a03f44fa67252b76bd3581081c714f882f882df915e97dbeab061fa8b3cc4e7"

definition P384_SHA384_TV14_S :: nat where
  "P384_SHA384_TV14_S = 0xd40e09d3468b46699948007e8f59845766dbf694b9c62066890dd055c0cb9a0caf0aa611fb9f466ad0bbb00dbe29d7eb"

definition P384_SHA384_TV14_Sig :: "nat × nat" where
  "P384_SHA384_TV14_Sig = (P384_SHA384_TV14_R, P384_SHA384_TV14_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA384_TV14_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV14_d" 
  by eval

lemma P384_SHA384_TV14_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA384_TV14_Q" 
  by eval

lemma P384_SHA384_TV14_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA384_TV14_d = P384_SHA384_TV14_Q" 
  by eval

lemma P384_SHA384_TV14_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA384_TV14_d P384_SHA384_TV14_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA384_TV14_dQvalidPair' P384_SHA384_TV14_d_valid) 

lemma P384_SHA384_TV14_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV14_k" 
  by eval

lemma P384_SHA384_TV14_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA384octets P384_SHA384_TV14_d P384_SHA384_TV14_Msg P384_SHA384_TV14_k = Some P384_SHA384_TV14_Sig" 
  by eval

lemma P384_SHA384_TV14_Verify: 
  "SEC1_P384_ECDSA_Verify SHA384octets P384_SHA384_TV14_Q P384_SHA384_TV14_Msg P384_SHA384_TV14_Sig" 
  by eval

subsection‹Test Vector: P384_SHA384_TV15›

subsubsection‹Given Test Vector Values›

definition P384_SHA384_TV15_Msg :: octets where
  "P384_SHA384_TV15_Msg = nat_to_octets_len 0x0e646c6c3cc0f9fdedef934b7195fe3837836a9f6f263968af95ef84cd035750f3cdb649de745c874a6ef66b3dd83b66068b4335bc0a97184182e3965c722b3b1aee488c3620adb835a8140e199f4fc83a88b02881816b366a09316e25685217f9221157fc05b2d8d2bc855372183da7af3f0a14148a09def37a332f8eb40dc9 128"

definition P384_SHA384_TV15_d :: nat where
  "P384_SHA384_TV15_d = 0xe9c7e9a79618d6ff3274da1abd0ff3ed0ec1ae3b54c3a4fd8d68d98fb04326b7633fc637e0b195228d0edba6bb1468fb"

definition P384_SHA384_TV15_Qx :: nat where
  "P384_SHA384_TV15_Qx = 0xa39ac353ca787982c577aff1e8601ce192aa90fd0de4c0ed627f66a8b6f02ae51315543f72ffc1c48a7269b25e7c289a"

definition P384_SHA384_TV15_Qy :: nat where
  "P384_SHA384_TV15_Qy = 0x9064a507b66b340b6e0e0d5ffaa67dd20e6dafc0ea6a6faee1635177af256f9108a22e9edf736ab4ae8e96dc207b1fa9"

definition P384_SHA384_TV15_Q :: "int point" where
  "P384_SHA384_TV15_Q = Point (int P384_SHA384_TV15_Qx) (int P384_SHA384_TV15_Qy)"

definition P384_SHA384_TV15_k :: nat where
  "P384_SHA384_TV15_k = 0xb094cb3a5c1440cfab9dc56d0ec2eff00f2110dea203654c70757254aa5912a7e73972e607459b1f4861e0b08a5cc763"

definition P384_SHA384_TV15_R :: nat where
  "P384_SHA384_TV15_R = 0xee82c0f90501136eb0dc0e459ad17bf3be1b1c8b8d05c60068a9306a346326ff7344776a95f1f7e2e2cf9477130e735c"

definition P384_SHA384_TV15_S :: nat where
  "P384_SHA384_TV15_S = 0xaf10b90f203af23b7500e070536e64629ba19245d6ef39aab57fcdb1b73c4c6bf7070c6263544633d3d358c12a178138"

definition P384_SHA384_TV15_Sig :: "nat × nat" where
  "P384_SHA384_TV15_Sig = (P384_SHA384_TV15_R, P384_SHA384_TV15_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA384_TV15_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV15_d" 
  by eval

lemma P384_SHA384_TV15_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA384_TV15_Q" 
  by eval

lemma P384_SHA384_TV15_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA384_TV15_d = P384_SHA384_TV15_Q" 
  by eval

lemma P384_SHA384_TV15_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA384_TV15_d P384_SHA384_TV15_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA384_TV15_dQvalidPair' P384_SHA384_TV15_d_valid) 

lemma P384_SHA384_TV15_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA384_TV15_k" 
  by eval

lemma P384_SHA384_TV15_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA384octets P384_SHA384_TV15_d P384_SHA384_TV15_Msg P384_SHA384_TV15_k = Some P384_SHA384_TV15_Sig" 
  by eval

lemma P384_SHA384_TV15_Verify: 
  "SEC1_P384_ECDSA_Verify SHA384octets P384_SHA384_TV15_Q P384_SHA384_TV15_Msg P384_SHA384_TV15_Sig" 
  by eval

section‹NIST Test Vectors for ECDSA: Curve P-384, Hash SHA-512›
  
subsection‹Test Vector: P384_SHA512_TV01›

subsubsection‹Given Test Vector Values›

definition P384_SHA512_TV01_Msg :: octets where
  "P384_SHA512_TV01_Msg = nat_to_octets_len 0x67d9eb88f289454d61def4764d1573db49b875cfb11e139d7eacc4b7a79d3db3bf7208191b2b2078cbbcc974ec0da1ed5e0c10ec37f6181bf81c0f32972a125df64e3b3e1d838ec7da8dfe0b7fcc911e43159a79c73df5fa252b98790be511d8a732fcbf011aacc7d45d8027d50a347703d613ceda09f650c6104c9459537c8f 128"

definition P384_SHA512_TV01_d :: nat where
  "P384_SHA512_TV01_d = 0x217afba406d8ab32ee07b0f27eef789fc201d121ffab76c8fbe3c2d352c594909abe591c6f86233992362c9d631baf7c"

definition P384_SHA512_TV01_Qx :: nat where
  "P384_SHA512_TV01_Qx = 0xfb937e4a303617b71b6c1a25f2ac786087328a3e26bdef55e52d46ab5e69e5411bf9fc55f5df9994d2bf82e8f39a153e"

definition P384_SHA512_TV01_Qy :: nat where
  "P384_SHA512_TV01_Qy = 0xa97d9075e92fa5bfe67e6ec18e21cc4d11fde59a68aef72c0e46a28f31a9d60385f41f39da468f4e6c3d3fbac9046765"

definition P384_SHA512_TV01_Q :: "int point" where
  "P384_SHA512_TV01_Q = Point (int P384_SHA512_TV01_Qx) (int P384_SHA512_TV01_Qy)"

definition P384_SHA512_TV01_k :: nat where
  "P384_SHA512_TV01_k = 0x90338a7f6ffce541366ca2987c3b3ca527992d1efcf1dd2723fbd241a24cff19990f2af5fd6419ed2104b4a59b5ae631"

definition P384_SHA512_TV01_R :: nat where
  "P384_SHA512_TV01_R = 0xc269d9c4619aafdf5f4b3100211dddb14693abe25551e04f9499c91152a296d7449c08b36f87d1e16e8e15fee4a7f5c8"

definition P384_SHA512_TV01_S :: nat where
  "P384_SHA512_TV01_S = 0x77ffed5c61665152d52161dc13ac3fbae5786928a3d736f42d34a9e4d6d4a70a02d5af90fa37a23a318902ae2656c071"

definition P384_SHA512_TV01_Sig :: "nat × nat" where
  "P384_SHA512_TV01_Sig = (P384_SHA512_TV01_R, P384_SHA512_TV01_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA512_TV01_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV01_d" 
  by eval

lemma P384_SHA512_TV01_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA512_TV01_Q" 
  by eval

lemma P384_SHA512_TV01_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA512_TV01_d = P384_SHA512_TV01_Q" 
  by eval

lemma P384_SHA512_TV01_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA512_TV01_d P384_SHA512_TV01_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA512_TV01_dQvalidPair' P384_SHA512_TV01_d_valid) 

lemma P384_SHA512_TV01_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV01_k" 
  by eval

lemma P384_SHA512_TV01_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA512octets P384_SHA512_TV01_d P384_SHA512_TV01_Msg P384_SHA512_TV01_k = Some P384_SHA512_TV01_Sig" 
  by eval

lemma P384_SHA512_TV01_Verify: 
  "SEC1_P384_ECDSA_Verify SHA512octets P384_SHA512_TV01_Q P384_SHA512_TV01_Msg P384_SHA512_TV01_Sig" 
  by eval

subsection‹Test Vector: P384_SHA512_TV02›

subsubsection‹Given Test Vector Values›

definition P384_SHA512_TV02_Msg :: octets where
  "P384_SHA512_TV02_Msg = nat_to_octets_len 0x45db86829c363c80160659e3c5c7d7971abb1f6f0d495709bba908d7aa99c9df64b3408a51bd69aba8870e2aaff488ef138f3123cf94391d081f357e21906a4e2f311defe527c55e0231579957c51def507f835cceb466eb2593a509dcbee2f09e0dde6693b2bfe17697c9e86dd672f5797339cbe9ea8a7c6309b061eca7aef5 128"

definition P384_SHA512_TV02_d :: nat where
  "P384_SHA512_TV02_d = 0x0a3f45a28a355381a919372f60320d6610cfb69c3e318eb1607db3cadfc42b728b77a6a9e9e333de9183c58933daf60f"

definition P384_SHA512_TV02_Qx :: nat where
  "P384_SHA512_TV02_Qx = 0x832cbb7061a719a316e73dbad348fa67cd17c33f40b9000a3d3b691a2a2cd821052566717c3ead01089b56086af1366f"

definition P384_SHA512_TV02_Qy :: nat where
  "P384_SHA512_TV02_Qy = 0x1e15a048d1dce642d9ebcbfac7f92b1bcee90fd0240cc79abd29e32e0e655c4ee1fd34fb88178bba92aca100e7794ed0"

definition P384_SHA512_TV02_Q :: "int point" where
  "P384_SHA512_TV02_Q = Point (int P384_SHA512_TV02_Qx) (int P384_SHA512_TV02_Qy)"

definition P384_SHA512_TV02_k :: nat where
  "P384_SHA512_TV02_k = 0x2a78e651623ba604c42cf094fc7d046629306f508853427ba091448800d1092c041bb2323035fc9d19a8d44950f7dcc3"

definition P384_SHA512_TV02_R :: nat where
  "P384_SHA512_TV02_R = 0x0db0cc9a2bda8dd7e565ad36f91b1c5756d78164dc8a72a5bee4b6bc45ea38c7a16b01d05b1893d4e06b62db24c30385"

definition P384_SHA512_TV02_S :: nat where
  "P384_SHA512_TV02_S = 0xabd383edaeda7d0b8de1b54fcd3c28874fed62ab266f1f84c8ba796a7b54e5e0695fdb43ce7fe90ed00fa468d87bca64"

definition P384_SHA512_TV02_Sig :: "nat × nat" where
  "P384_SHA512_TV02_Sig = (P384_SHA512_TV02_R, P384_SHA512_TV02_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA512_TV02_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV02_d" 
  by eval

lemma P384_SHA512_TV02_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA512_TV02_Q" 
  by eval

lemma P384_SHA512_TV02_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA512_TV02_d = P384_SHA512_TV02_Q" 
  by eval

lemma P384_SHA512_TV02_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA512_TV02_d P384_SHA512_TV02_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA512_TV02_dQvalidPair' P384_SHA512_TV02_d_valid) 

lemma P384_SHA512_TV02_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV02_k" 
  by eval

lemma P384_SHA512_TV02_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA512octets P384_SHA512_TV02_d P384_SHA512_TV02_Msg P384_SHA512_TV02_k = Some P384_SHA512_TV02_Sig" 
  by eval

lemma P384_SHA512_TV02_Verify: 
  "SEC1_P384_ECDSA_Verify SHA512octets P384_SHA512_TV02_Q P384_SHA512_TV02_Msg P384_SHA512_TV02_Sig" 
  by eval

subsection‹Test Vector: P384_SHA512_TV03›

subsubsection‹Given Test Vector Values›

definition P384_SHA512_TV03_Msg :: octets where
  "P384_SHA512_TV03_Msg = nat_to_octets_len 0x4672fce0721d37c5be166bffa4b30d753bcf104b9b414db994b3ed33f36af4935ea59a0bb92db66448b3f57dad4fc67cef10ce141bf82c536be604b89a0bc0e8bca605b867880049d97142d30538fc543bd9d4fab7fdbe2f703815cdb6361beb66acff764bc275f910d1662445b07b92830db69a5994857f53657ed5ca282648 128"

definition P384_SHA512_TV03_d :: nat where
  "P384_SHA512_TV03_d = 0x2e408c57921939f0e0fe2e80ce74a4fa4a1b4fa7ab070206298fe894d655be50e2583af9e45544b5d69c73dce8a2c8e7"

definition P384_SHA512_TV03_Qx :: nat where
  "P384_SHA512_TV03_Qx = 0xa2b24a5ad4a2e91f12199ed7699e3f297e27bf8b8ea8fbe7ed28366f3544cd8e680c238450f8a6422b40829d6647b25c"

definition P384_SHA512_TV03_Qy :: nat where
  "P384_SHA512_TV03_Qy = 0x2732be0075536e6519f6a099b975a40f8e0de337fa4d48bd0762b43f41cab8deafdef9cfbb9973e457801e3bf9c93304"

definition P384_SHA512_TV03_Q :: "int point" where
  "P384_SHA512_TV03_Q = Point (int P384_SHA512_TV03_Qx) (int P384_SHA512_TV03_Qy)"

definition P384_SHA512_TV03_k :: nat where
  "P384_SHA512_TV03_k = 0xb10b6258afdde81f9c971cc1526d942e20cafac02f59fee10f98e99b8674636bff1d84a6eaa49c0de8d8cfdc90d8ce84"

definition P384_SHA512_TV03_R :: nat where
  "P384_SHA512_TV03_R = 0xbe428a8de89a364a134719141ee8d776a3a8338f1132b07e01b28573d8eaf3b9008b63304c48821e53638b6141f9660b"

definition P384_SHA512_TV03_S :: nat where
  "P384_SHA512_TV03_S = 0x866181dbef5c147d391bed6adcee408c339982c307adc718c2b9ab9e5642d8dedc36dd6402559a3ab614c99c1e56b529"

definition P384_SHA512_TV03_Sig :: "nat × nat" where
  "P384_SHA512_TV03_Sig = (P384_SHA512_TV03_R, P384_SHA512_TV03_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA512_TV03_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV03_d" 
  by eval

lemma P384_SHA512_TV03_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA512_TV03_Q" 
  by eval

lemma P384_SHA512_TV03_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA512_TV03_d = P384_SHA512_TV03_Q" 
  by eval

lemma P384_SHA512_TV03_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA512_TV03_d P384_SHA512_TV03_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA512_TV03_dQvalidPair' P384_SHA512_TV03_d_valid) 

lemma P384_SHA512_TV03_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV03_k" 
  by eval

lemma P384_SHA512_TV03_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA512octets P384_SHA512_TV03_d P384_SHA512_TV03_Msg P384_SHA512_TV03_k = Some P384_SHA512_TV03_Sig" 
  by eval

lemma P384_SHA512_TV03_Verify: 
  "SEC1_P384_ECDSA_Verify SHA512octets P384_SHA512_TV03_Q P384_SHA512_TV03_Msg P384_SHA512_TV03_Sig" 
  by eval

subsection‹Test Vector: P384_SHA512_TV04›

subsubsection‹Given Test Vector Values›

definition P384_SHA512_TV04_Msg :: octets where
  "P384_SHA512_TV04_Msg = nat_to_octets_len 0x9ae48fdd9bfc5cb0f4d4761e28b2073bda05a3e3fe82c212e66701dc4573cc67a829b0f82d7520b1bf11db0c6d1743822bbe41bb0adbd7222aa5fae70fbd1a31f2d4453a01c81e064d775388468be96f6063f8673b7b8d4455fe1bd4c801ad5e625a015eaa4a1a18da490d2af8642201eaba3c611cbd65f861d8e19ca82a1ee6 128"

definition P384_SHA512_TV04_d :: nat where
  "P384_SHA512_TV04_d = 0x1c285da72a8eb1c3c38faab8d3bb4e68dc95c797082b9a3991a21c1de54759071ecf2265fb1eff504ab24174bc6710cf"

definition P384_SHA512_TV04_Qx :: nat where
  "P384_SHA512_TV04_Qx = 0x11acb1b5cc59a4f1df1913a8d6e91cbdafb8206dc44aff7d9da45906b664fc33194d9935a82aa4d62f39618897c86025"

definition P384_SHA512_TV04_Qy :: nat where
  "P384_SHA512_TV04_Qy = 0x832ed0b9575fff52a3603bfe89f312751b4c396da98324117a61b3f525d27b2266f6cfb22be07e50b6874435e380ed62"

definition P384_SHA512_TV04_Q :: "int point" where
  "P384_SHA512_TV04_Q = Point (int P384_SHA512_TV04_Qx) (int P384_SHA512_TV04_Qy)"

definition P384_SHA512_TV04_k :: nat where
  "P384_SHA512_TV04_k = 0x2513075e02cc7fb3cff7b7adde46da31c5493749b5cf02758bd5b098a838bfd4d5e4c7fb8268bdc37e219c30efebe878"

definition P384_SHA512_TV04_R :: nat where
  "P384_SHA512_TV04_R = 0xb3d638b3be45f14f170da5bdc22d2114deac93ab340a25b3af2b5c18584bb9147e00dc6c67a2274f79aa4838793eb63f"

definition P384_SHA512_TV04_S :: nat where
  "P384_SHA512_TV04_S = 0x876112bdca2c725eb2f6dbd76d07710a31f0c16d38430cb0817f320a25a9ecfec8a66137d0304612ae29a6a484fd3319"

definition P384_SHA512_TV04_Sig :: "nat × nat" where
  "P384_SHA512_TV04_Sig = (P384_SHA512_TV04_R, P384_SHA512_TV04_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA512_TV04_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV04_d" 
  by eval

lemma P384_SHA512_TV04_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA512_TV04_Q" 
  by eval

lemma P384_SHA512_TV04_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA512_TV04_d = P384_SHA512_TV04_Q" 
  by eval

lemma P384_SHA512_TV04_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA512_TV04_d P384_SHA512_TV04_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA512_TV04_dQvalidPair' P384_SHA512_TV04_d_valid) 

lemma P384_SHA512_TV04_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV04_k" 
  by eval

lemma P384_SHA512_TV04_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA512octets P384_SHA512_TV04_d P384_SHA512_TV04_Msg P384_SHA512_TV04_k = Some P384_SHA512_TV04_Sig" 
  by eval

lemma P384_SHA512_TV04_Verify: 
  "SEC1_P384_ECDSA_Verify SHA512octets P384_SHA512_TV04_Q P384_SHA512_TV04_Msg P384_SHA512_TV04_Sig" 
  by eval

subsection‹Test Vector: P384_SHA512_TV05›

subsubsection‹Given Test Vector Values›

definition P384_SHA512_TV05_Msg :: octets where
  "P384_SHA512_TV05_Msg = nat_to_octets_len 0x817d6a110a8fd0ca7b4d565558f68b59a156744d4c5aac5c6610c95451793de2a756f774558c61d21818d3ebeeeb71d132da1c23a02f4b305eccc5cd46bd21dfc173a8a91098354f10ffbb21bf63d9f4c3feb231c736504549a78fd76d39f3ad35c36178f5c233742d2917d5611d2073124845f1e3615b2ef25199a7a547e882 128"

definition P384_SHA512_TV05_d :: nat where
  "P384_SHA512_TV05_d = 0x9da37e104938019fbdcf247e3df879a282c45f8fb57e6655e36b47723af42bec3b820f660436deb3de123a21de0ca37b"

definition P384_SHA512_TV05_Qx :: nat where
  "P384_SHA512_TV05_Qx = 0x722d0ea6891d509b18b85ca56f74deb5c3030d2a30433824123d430d03c99279572c3b28ecf01e747b9db8acc55d0ba3"

definition P384_SHA512_TV05_Qy :: nat where
  "P384_SHA512_TV05_Qy = 0x7e2605ea7092214f366f3639037bffd89fe103c646e990839d3a1ced8d78edb5b9bc60d834fd8e2a3c17e920bdae023a"

definition P384_SHA512_TV05_Q :: "int point" where
  "P384_SHA512_TV05_Q = Point (int P384_SHA512_TV05_Qx) (int P384_SHA512_TV05_Qy)"

definition P384_SHA512_TV05_k :: nat where
  "P384_SHA512_TV05_k = 0xc8c18e53a9aa5915288c33132bd09323638f7995cd89162073984ed84e72e07a37e18c4c023933eace92c35d10e6b1b6"

definition P384_SHA512_TV05_R :: nat where
  "P384_SHA512_TV05_R = 0x6512a8a2be731e301dcf4803764297862bbfa0ac8daed64d8e98b34618ecb20520fc5d3cf890b7783edf86e7ea407541"

definition P384_SHA512_TV05_S :: nat where
  "P384_SHA512_TV05_S = 0x4ff10301f7b4168fae066361376007c1d7aa89a75c87719d0b54711ffef5ef3726f3eef84f7ebc025c110bde511b17f6"

definition P384_SHA512_TV05_Sig :: "nat × nat" where
  "P384_SHA512_TV05_Sig = (P384_SHA512_TV05_R, P384_SHA512_TV05_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA512_TV05_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV05_d" 
  by eval

lemma P384_SHA512_TV05_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA512_TV05_Q" 
  by eval

lemma P384_SHA512_TV05_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA512_TV05_d = P384_SHA512_TV05_Q" 
  by eval

lemma P384_SHA512_TV05_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA512_TV05_d P384_SHA512_TV05_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA512_TV05_dQvalidPair' P384_SHA512_TV05_d_valid) 

lemma P384_SHA512_TV05_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV05_k" 
  by eval

lemma P384_SHA512_TV05_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA512octets P384_SHA512_TV05_d P384_SHA512_TV05_Msg P384_SHA512_TV05_k = Some P384_SHA512_TV05_Sig" 
  by eval

lemma P384_SHA512_TV05_Verify: 
  "SEC1_P384_ECDSA_Verify SHA512octets P384_SHA512_TV05_Q P384_SHA512_TV05_Msg P384_SHA512_TV05_Sig" 
  by eval

subsection‹Test Vector: P384_SHA512_TV06›

subsubsection‹Given Test Vector Values›

definition P384_SHA512_TV06_Msg :: octets where
  "P384_SHA512_TV06_Msg = nat_to_octets_len 0x464f10ec6fb229a51db5fd0e122f2cb8a9a022117e2987f4007bf5565b2c16aba0714e2e3cdd0c100d55ac3017e36fc7501ad8309ab9572aa65424c9eb2e580a119c55777676ec498df53ef6ae78fd8a988130ee0e6082bf1ef71cd4c946021018a8ca7154d13b174c638912613b0bdb9001c302bf7e443ad2124ab2c1cce212 128"

definition P384_SHA512_TV06_d :: nat where
  "P384_SHA512_TV06_d = 0x0661ab3bf9f7bef51bec7dff758de289154557beb9ce18cc4b8cc09a871e8322af259cf188b593dc62f03a19e75f7f69"

definition P384_SHA512_TV06_Qx :: nat where
  "P384_SHA512_TV06_Qx = 0xb4f100558043858efa728082d9b99ad5192b59b0947434f5ba7ff2514508a6d71ba54e7221c31cb0712103272b3f6fa4"

definition P384_SHA512_TV06_Qy :: nat where
  "P384_SHA512_TV06_Qy = 0x34f6df4eeb2da11498044635067c2715ed15ae251c78ffb9030d87909ea8539b66394e93109ca54c0406cf99960c3e93"

definition P384_SHA512_TV06_Q :: "int point" where
  "P384_SHA512_TV06_Q = Point (int P384_SHA512_TV06_Qx) (int P384_SHA512_TV06_Qy)"

definition P384_SHA512_TV06_k :: nat where
  "P384_SHA512_TV06_k = 0x84a87137edb6894f96c5a8e94a3765162034feb84dfea94e1c71411170c285a80321ec7999e25861844143209804882c"

definition P384_SHA512_TV06_R :: nat where
  "P384_SHA512_TV06_R = 0x4dc9d1b949b36e3c3847ac1c7ed114e1bc9cbe76119cf6fcd3f1b69ee6ee54e3255f1bb288fe2f8bd6d4049a21793c27"

definition P384_SHA512_TV06_S :: nat where
  "P384_SHA512_TV06_S = 0x56a561d647b62ccae1e6df818b1a6fbde66c82ef0ff69ee415f183e7daf76be22630c7e02cd3fd729dfa490f26824584"

definition P384_SHA512_TV06_Sig :: "nat × nat" where
  "P384_SHA512_TV06_Sig = (P384_SHA512_TV06_R, P384_SHA512_TV06_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA512_TV06_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV06_d" 
  by eval

lemma P384_SHA512_TV06_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA512_TV06_Q" 
  by eval

lemma P384_SHA512_TV06_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA512_TV06_d = P384_SHA512_TV06_Q" 
  by eval

lemma P384_SHA512_TV06_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA512_TV06_d P384_SHA512_TV06_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA512_TV06_dQvalidPair' P384_SHA512_TV06_d_valid) 

lemma P384_SHA512_TV06_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV06_k" 
  by eval

lemma P384_SHA512_TV06_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA512octets P384_SHA512_TV06_d P384_SHA512_TV06_Msg P384_SHA512_TV06_k = Some P384_SHA512_TV06_Sig" 
  by eval

lemma P384_SHA512_TV06_Verify: 
  "SEC1_P384_ECDSA_Verify SHA512octets P384_SHA512_TV06_Q P384_SHA512_TV06_Msg P384_SHA512_TV06_Sig" 
  by eval

subsection‹Test Vector: P384_SHA512_TV07›

subsubsection‹Given Test Vector Values›

definition P384_SHA512_TV07_Msg :: octets where
  "P384_SHA512_TV07_Msg = nat_to_octets_len 0x4e3e0fb96320ddccde8b463c273654c4f7164920b1d63430921d2e808dee403e6420eedda0a557b911d00736a4f8798dd4ef26673efd6d190988ad4929ec64f8685cfb76070a36cd6a3a4bf2f54fb08a349d44642b6f614043fef9b2813b63457c76537d23da7b37310334f7ba76edf1999dad86f72aa3446445a65952ac4e50 128"

definition P384_SHA512_TV07_d :: nat where
  "P384_SHA512_TV07_d = 0x66e7cfdeb7f264cf786e35210f458c32223c3a12a3bc4b63d53a5776bc9b069928452484f6241caa3781fd1a4109d4db"

definition P384_SHA512_TV07_Qx :: nat where
  "P384_SHA512_TV07_Qx = 0x3c7682de540ab231daf21bf9fc80bda6abf7e17dcc79d476c7b7c3bd4d42d386877fd8ba495c1b0333e04fb5fd2a1505"

definition P384_SHA512_TV07_Qy :: nat where
  "P384_SHA512_TV07_Qy = 0x0a1582e4f4d72abea9d3476aff8369c41261f0c5dddf2ca82e10f7a163f73df09473d9e5e2552187104e4cc7c6d83611"

definition P384_SHA512_TV07_Q :: "int point" where
  "P384_SHA512_TV07_Q = Point (int P384_SHA512_TV07_Qx) (int P384_SHA512_TV07_Qy)"

definition P384_SHA512_TV07_k :: nat where
  "P384_SHA512_TV07_k = 0x2fa266f5cce190eb77614933ca6a55121ad8bae168ff7a9043d96d13b5ca2fe70101ff9fe1e2b2cd7413e6aa8f49abde"

definition P384_SHA512_TV07_R :: nat where
  "P384_SHA512_TV07_R = 0xe7ecda9da0c52d0474a9f70094dc8f061d7d6a22210d3b69a7be8f389aa666f256322099b87d16ad35357ea856574dba"

definition P384_SHA512_TV07_S :: nat where
  "P384_SHA512_TV07_S = 0xba348eb40a2830ec5a1130264ac0a8675420b1ae243e808a778135809ece21f42c0c881166321102b4f02df4c5c7ed9d"

definition P384_SHA512_TV07_Sig :: "nat × nat" where
  "P384_SHA512_TV07_Sig = (P384_SHA512_TV07_R, P384_SHA512_TV07_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA512_TV07_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV07_d" 
  by eval

lemma P384_SHA512_TV07_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA512_TV07_Q" 
  by eval

lemma P384_SHA512_TV07_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA512_TV07_d = P384_SHA512_TV07_Q" 
  by eval

lemma P384_SHA512_TV07_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA512_TV07_d P384_SHA512_TV07_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA512_TV07_dQvalidPair' P384_SHA512_TV07_d_valid) 

lemma P384_SHA512_TV07_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV07_k" 
  by eval

lemma P384_SHA512_TV07_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA512octets P384_SHA512_TV07_d P384_SHA512_TV07_Msg P384_SHA512_TV07_k = Some P384_SHA512_TV07_Sig" 
  by eval

lemma P384_SHA512_TV07_Verify: 
  "SEC1_P384_ECDSA_Verify SHA512octets P384_SHA512_TV07_Q P384_SHA512_TV07_Msg P384_SHA512_TV07_Sig" 
  by eval

subsection‹Test Vector: P384_SHA512_TV08›

subsubsection‹Given Test Vector Values›

definition P384_SHA512_TV08_Msg :: octets where
  "P384_SHA512_TV08_Msg = nat_to_octets_len 0xc466b6b6baf7e6ffa876ec06105e2d43534e0517c07b1c4c9fb67ba81ce09525a7721ec3c290f2b1f65b6463d41598e7a25b2238501629953a5ca955b644354fb6856733a2e5bb8f5bc21a0c803493f5539f9fb83aab3dba2c982989c2270c61ab244b68bfe1b948d00c2ed975e09c29b5f8a7effcad8652a148cc880d503217 128"

definition P384_SHA512_TV08_d :: nat where
  "P384_SHA512_TV08_d = 0x92c2f7ee64af86d003ab484e12b82fcf245fc330761057fec5b7af8f7e0a2d85b468c21d171460fcb829cae7b986316d"

definition P384_SHA512_TV08_Qx :: nat where
  "P384_SHA512_TV08_Qx = 0xca43a306479bf8fb537d4b9ff9d635bbb2a0d60d9e854d5b7e269d09d91f78c6b90b616e4c931629453645a2bb371e14"

definition P384_SHA512_TV08_Qy :: nat where
  "P384_SHA512_TV08_Qy = 0x356c4d7f10e690614eaf7f82ba0f9dc1aad98130c0ad9fe353deec565cc04bef789a0a4242322e0058b46cd02f2de77d"

definition P384_SHA512_TV08_Q :: "int point" where
  "P384_SHA512_TV08_Q = Point (int P384_SHA512_TV08_Qx) (int P384_SHA512_TV08_Qy)"

definition P384_SHA512_TV08_k :: nat where
  "P384_SHA512_TV08_k = 0x6ec81fb74f8725ba225f317264460ee300cfd2f02092000989acbdad4799cf55c244a65c557113328fe20282e6badb55"

definition P384_SHA512_TV08_R :: nat where
  "P384_SHA512_TV08_R = 0xcd7a4309bcebc25a8e10899fe2eda5f8b2dbcf329cd2f3d65befd67393e83fba2f8a67a15c01a6ac8314f9f5e87a9dca"

definition P384_SHA512_TV08_S :: nat where
  "P384_SHA512_TV08_S = 0x6dcfc0426bc148e67e91d4784e3d7e9bc3b7ce3676be62daa7f3f55dfdff6d9dc735b5e3e0bbd0785db1f76f7ac065f3"

definition P384_SHA512_TV08_Sig :: "nat × nat" where
  "P384_SHA512_TV08_Sig = (P384_SHA512_TV08_R, P384_SHA512_TV08_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA512_TV08_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV08_d" 
  by eval

lemma P384_SHA512_TV08_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA512_TV08_Q" 
  by eval

lemma P384_SHA512_TV08_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA512_TV08_d = P384_SHA512_TV08_Q" 
  by eval

lemma P384_SHA512_TV08_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA512_TV08_d P384_SHA512_TV08_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA512_TV08_dQvalidPair' P384_SHA512_TV08_d_valid) 

lemma P384_SHA512_TV08_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV08_k" 
  by eval

lemma P384_SHA512_TV08_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA512octets P384_SHA512_TV08_d P384_SHA512_TV08_Msg P384_SHA512_TV08_k = Some P384_SHA512_TV08_Sig" 
  by eval

lemma P384_SHA512_TV08_Verify: 
  "SEC1_P384_ECDSA_Verify SHA512octets P384_SHA512_TV08_Q P384_SHA512_TV08_Msg P384_SHA512_TV08_Sig" 
  by eval

subsection‹Test Vector: P384_SHA512_TV09›

subsubsection‹Given Test Vector Values›

definition P384_SHA512_TV09_Msg :: octets where
  "P384_SHA512_TV09_Msg = nat_to_octets_len 0xfeac892b7720af80b3c9eede51e923f18d3d0c5de4c31f4aa75e36df7c7c2fd8f41778851a24b69e67dccb65e159dd5c383243bad7cfedcc5e85c8a01c34b0b94ba8e07e4c024c09d279b3731e8b62f9562d3c4f5042567efe42a9d0eaaabab28bc6f11232fc8ceaaf4518d9f3b2bebf020294496b7f6b879e69503f75fecd3d 128"

definition P384_SHA512_TV09_d :: nat where
  "P384_SHA512_TV09_d = 0x15347caaad1067f1848a676bd0a8c52021ae604b79d02775a0459226e0391a3acd26653c916fcfe86149fb0ee0904476"

definition P384_SHA512_TV09_Qx :: nat where
  "P384_SHA512_TV09_Qx = 0xe5a0463163964d984f5bad0072d45bc2059939e60a826ccca36c151460ae360f5d6679f60fe43e999b6da5841c96e48a"

definition P384_SHA512_TV09_Qy :: nat where
  "P384_SHA512_TV09_Qy = 0x30f2dd425a3fa2c95d34124217250b39e3b4a14f3e6e415ae8e5b0409eb72f43f78b64d0ce6f2d49980d6f04cd1391db"

definition P384_SHA512_TV09_Q :: "int point" where
  "P384_SHA512_TV09_Q = Point (int P384_SHA512_TV09_Qx) (int P384_SHA512_TV09_Qy)"

definition P384_SHA512_TV09_k :: nat where
  "P384_SHA512_TV09_k = 0x1a2d224db4bb9c241ca5cab18920fad615fa25c1db0de0f024cb3ace0d11ef72b056885446659f67650fdff692517b1c"

definition P384_SHA512_TV09_R :: nat where
  "P384_SHA512_TV09_R = 0x87b4de0fb21df38dfc9a4b1e350da67547e307f55b5b9dd6615e408afe7c3553a6e02722847367439e636074faa2182b"

definition P384_SHA512_TV09_S :: nat where
  "P384_SHA512_TV09_S = 0x375d965753b9ed6c6c08576726f8308c2f8dbd2737824464e71265d47907e26f615bbeb8203ec617520d4ecd1851dc44"

definition P384_SHA512_TV09_Sig :: "nat × nat" where
  "P384_SHA512_TV09_Sig = (P384_SHA512_TV09_R, P384_SHA512_TV09_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA512_TV09_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV09_d" 
  by eval

lemma P384_SHA512_TV09_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA512_TV09_Q" 
  by eval

lemma P384_SHA512_TV09_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA512_TV09_d = P384_SHA512_TV09_Q" 
  by eval

lemma P384_SHA512_TV09_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA512_TV09_d P384_SHA512_TV09_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA512_TV09_dQvalidPair' P384_SHA512_TV09_d_valid) 

lemma P384_SHA512_TV09_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV09_k" 
  by eval

lemma P384_SHA512_TV09_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA512octets P384_SHA512_TV09_d P384_SHA512_TV09_Msg P384_SHA512_TV09_k = Some P384_SHA512_TV09_Sig" 
  by eval

lemma P384_SHA512_TV09_Verify: 
  "SEC1_P384_ECDSA_Verify SHA512octets P384_SHA512_TV09_Q P384_SHA512_TV09_Msg P384_SHA512_TV09_Sig" 
  by eval

subsection‹Test Vector: P384_SHA512_TV10›

subsubsection‹Given Test Vector Values›

definition P384_SHA512_TV10_Msg :: octets where
  "P384_SHA512_TV10_Msg = nat_to_octets_len 0xcf2982e3bf174ce547741b969403cd11e9553067e6af8177d89511a0eb040db924530bdba65d8b1ff714228db0737c1756f509e1506014a10736e65be2f91980a73891496e90ff2714a3601c7565cdcef5a395e2e0e1652f138d90d61eaa9cba993b823245647f6e07cec9b8b4449cd68a29741cd1579c66e548ca0d0acf33aa 128"

definition P384_SHA512_TV10_d :: nat where
  "P384_SHA512_TV10_d = 0xac1cb5e59bda2eff3413a3bab80308f9fb32c595283c795de4c17fdae8d4647b5f108fd0801aee22adb7db129283b5aa"

definition P384_SHA512_TV10_Qx :: nat where
  "P384_SHA512_TV10_Qx = 0xbc6b1a718284803553c173089c397870aaaecca579bb8e81a8cfa12473cd2057567fa8726a19ed427cc035baeec2c551"

definition P384_SHA512_TV10_Qy :: nat where
  "P384_SHA512_TV10_Qy = 0x14f82997d1129b669f0015350e47ad561b1b13441af4fb44656f15ed0c5706984d66655accc52f2e943eef39cb1cdc21"

definition P384_SHA512_TV10_Q :: "int point" where
  "P384_SHA512_TV10_Q = Point (int P384_SHA512_TV10_Qx) (int P384_SHA512_TV10_Qy)"

definition P384_SHA512_TV10_k :: nat where
  "P384_SHA512_TV10_k = 0x8053a46e875f446056b06d4318fa3e8977622de7207cbf0996bf35b0e9b19aaa507f642bcf0be9f048f1af09806f6946"

definition P384_SHA512_TV10_R :: nat where
  "P384_SHA512_TV10_R = 0xa994eb15b64114ce8a9342d18b5edda96a6d76314a5ac03da723699177d352a4a9f3b7121b11a91e43a6af4025da51d6"

definition P384_SHA512_TV10_S :: nat where
  "P384_SHA512_TV10_S = 0x8183ae33a888e99aa76882da0a6705ad102f2bbd9572fad0d2e4d6d70151970469e00c5220e59c14724d771c1384b302"

definition P384_SHA512_TV10_Sig :: "nat × nat" where
  "P384_SHA512_TV10_Sig = (P384_SHA512_TV10_R, P384_SHA512_TV10_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA512_TV10_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV10_d" 
  by eval

lemma P384_SHA512_TV10_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA512_TV10_Q" 
  by eval

lemma P384_SHA512_TV10_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA512_TV10_d = P384_SHA512_TV10_Q" 
  by eval

lemma P384_SHA512_TV10_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA512_TV10_d P384_SHA512_TV10_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA512_TV10_dQvalidPair' P384_SHA512_TV10_d_valid) 

lemma P384_SHA512_TV10_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV10_k" 
  by eval

lemma P384_SHA512_TV10_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA512octets P384_SHA512_TV10_d P384_SHA512_TV10_Msg P384_SHA512_TV10_k = Some P384_SHA512_TV10_Sig" 
  by eval

lemma P384_SHA512_TV10_Verify: 
  "SEC1_P384_ECDSA_Verify SHA512octets P384_SHA512_TV10_Q P384_SHA512_TV10_Msg P384_SHA512_TV10_Sig" 
  by eval

subsection‹Test Vector: P384_SHA512_TV11›

subsubsection‹Given Test Vector Values›

definition P384_SHA512_TV11_Msg :: octets where
  "P384_SHA512_TV11_Msg = nat_to_octets_len 0xbf9fdd4107ef5a6070108771ac9eee4f0c8043bf0d04db772a47294f4137e2439d94b337114b074e57e0cb78d0ccf352a2833e9788ee2a1a9ffeacd34f38fcefb86653d70c7dadd4cf6548d608e70acdef6c7530974b92c813798add659752a8c72b05e1ad9c65c21834ce6fbe49d8a1426b5a54270794436d284364fac6ec1a 128"

definition P384_SHA512_TV11_d :: nat where
  "P384_SHA512_TV11_d = 0x205f1eb3dfacff2bdd8590e43e613b92512d6a415c5951bda7a6c37db3aae39b9b7ec6edd256609e75373419087fa71f"

definition P384_SHA512_TV11_Qx :: nat where
  "P384_SHA512_TV11_Qx = 0xc9f1f63a18c761b077a1ec35fbb2de635db9b8592c36194a01769b57728c7755d4c79b3d5b97a1a4631e30c86d03f13c"

definition P384_SHA512_TV11_Qy :: nat where
  "P384_SHA512_TV11_Qy = 0xf8c4a38770054d5cc9bb9182e6d4638242c4fd16e869ac22e44c4b9402d594e0c6f5df6a9a7de32a4893d9f6588f1950"

definition P384_SHA512_TV11_Q :: "int point" where
  "P384_SHA512_TV11_Q = Point (int P384_SHA512_TV11_Qx) (int P384_SHA512_TV11_Qy)"

definition P384_SHA512_TV11_k :: nat where
  "P384_SHA512_TV11_k = 0xecd395c5d8b7d6e6b2b19644e0d2e6086c912c6a0f5b8ed4b94b7290b65852c9741ce8eeb08d8751ead8a183e17d76c6"

definition P384_SHA512_TV11_R :: nat where
  "P384_SHA512_TV11_R = 0xe81331d78b438b0b8d98c1be03385ba5d614af182f1677f259126cc3de7eaac6c19b02be955d936b6bf9c27c6796e6f0"

definition P384_SHA512_TV11_S :: nat where
  "P384_SHA512_TV11_S = 0x17c2b7a8e0fc93909762aa9f86f9561e759ecb88f02337b2018363be6095d9e4324a6d3296046686624b5efad6b52878"

definition P384_SHA512_TV11_Sig :: "nat × nat" where
  "P384_SHA512_TV11_Sig = (P384_SHA512_TV11_R, P384_SHA512_TV11_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA512_TV11_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV11_d" 
  by eval

lemma P384_SHA512_TV11_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA512_TV11_Q" 
  by eval

lemma P384_SHA512_TV11_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA512_TV11_d = P384_SHA512_TV11_Q" 
  by eval

lemma P384_SHA512_TV11_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA512_TV11_d P384_SHA512_TV11_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA512_TV11_dQvalidPair' P384_SHA512_TV11_d_valid) 

lemma P384_SHA512_TV11_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV11_k" 
  by eval

lemma P384_SHA512_TV11_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA512octets P384_SHA512_TV11_d P384_SHA512_TV11_Msg P384_SHA512_TV11_k = Some P384_SHA512_TV11_Sig" 
  by eval

lemma P384_SHA512_TV11_Verify: 
  "SEC1_P384_ECDSA_Verify SHA512octets P384_SHA512_TV11_Q P384_SHA512_TV11_Msg P384_SHA512_TV11_Sig" 
  by eval

subsection‹Test Vector: P384_SHA512_TV12›

subsubsection‹Given Test Vector Values›

definition P384_SHA512_TV12_Msg :: octets where
  "P384_SHA512_TV12_Msg = nat_to_octets_len 0x5d634fb39a2239256107dc68db19751540b4badac9ecf2fce644724401d6d632b3ae3b2e6d05746b77ddc0c899878032248c263eda08d3d004d35952ad7a9cfe19343d14b37f9f632245e7b7b5fae3cb31c5231f82b9f1884f2de7578fbf156c430257031ba97bc6579843bc7f59fcb9a6449a4cd942dffa6adb929cf219f0ad 128"

definition P384_SHA512_TV12_d :: nat where
  "P384_SHA512_TV12_d = 0xe21e3a739e7ded418df5d3e7bc2c4ae8da76266a1fc4c89e5b09923db80a72217f1e96158031be42914cf3ee725748c1"

definition P384_SHA512_TV12_Qx :: nat where
  "P384_SHA512_TV12_Qx = 0x0f753171922b5334f3dd2778a64ce2da8295121939beae71ad85e5344e893be0fd03cf14e1f031adec098e0c4409449c"

definition P384_SHA512_TV12_Qy :: nat where
  "P384_SHA512_TV12_Qy = 0x45c10a0ffc0eb2f1cec5c89b698061108313ee7d449ad580efad344f0e7cf35be8a18fca620f112e57bdc746abdace55"

definition P384_SHA512_TV12_Q :: "int point" where
  "P384_SHA512_TV12_Q = Point (int P384_SHA512_TV12_Qx) (int P384_SHA512_TV12_Qy)"

definition P384_SHA512_TV12_k :: nat where
  "P384_SHA512_TV12_k = 0xd06bea06b25e6c30e866b1eb0657b45673e37b709013fb28fd7373afc8277cbc861354f821d0bd1927e52ec083a0f41f"

definition P384_SHA512_TV12_R :: nat where
  "P384_SHA512_TV12_R = 0xe8d4a31dd0e7d2522be62a32608e744c3775ceb606dc897899f0c73f1a40ce9a8be854cd506e65cd81fd7fa2c616cb7b"

definition P384_SHA512_TV12_S :: nat where
  "P384_SHA512_TV12_S = 0x8151b681b6b6046d3c36f332d06d9ba7751e740631cdb759f88c50a25a8e950d5023df8a15c77243743733c4feaf21d5"

definition P384_SHA512_TV12_Sig :: "nat × nat" where
  "P384_SHA512_TV12_Sig = (P384_SHA512_TV12_R, P384_SHA512_TV12_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA512_TV12_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV12_d" 
  by eval

lemma P384_SHA512_TV12_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA512_TV12_Q" 
  by eval

lemma P384_SHA512_TV12_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA512_TV12_d = P384_SHA512_TV12_Q" 
  by eval

lemma P384_SHA512_TV12_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA512_TV12_d P384_SHA512_TV12_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA512_TV12_dQvalidPair' P384_SHA512_TV12_d_valid) 

lemma P384_SHA512_TV12_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV12_k" 
  by eval

lemma P384_SHA512_TV12_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA512octets P384_SHA512_TV12_d P384_SHA512_TV12_Msg P384_SHA512_TV12_k = Some P384_SHA512_TV12_Sig" 
  by eval

lemma P384_SHA512_TV12_Verify: 
  "SEC1_P384_ECDSA_Verify SHA512octets P384_SHA512_TV12_Q P384_SHA512_TV12_Msg P384_SHA512_TV12_Sig" 
  by eval

subsection‹Test Vector: P384_SHA512_TV13›

subsubsection‹Given Test Vector Values›

definition P384_SHA512_TV13_Msg :: octets where
  "P384_SHA512_TV13_Msg = nat_to_octets_len 0xc9b4ff721b3e886f0dc05856ffff0aabb64a8504b1746a47fdd73e6b7ebc068f06ac7ffa44c757e4de207fc3cbfaf0469d3ac6795d40630bcafe8c658627e4bc6b86fd6a2135afbc18ccc8e6d0e1e86016930ca92edc5aa3fbe2c57de136d0ea5f41642b6a5d0ddeb380f2454d76a16639d663687f2a2e29fb9304243900d26d 128"

definition P384_SHA512_TV13_d :: nat where
  "P384_SHA512_TV13_d = 0x93434d3c03ec1da8510b74902c3b3e0cb9e8d7dccad37594d28b93e065b468d9af4892a03763a63eae060c769119c23c"

definition P384_SHA512_TV13_Qx :: nat where
  "P384_SHA512_TV13_Qx = 0xa52c25f2af70e5bc6a992ecef4ea54e831ed5b9453747d28aec5cffb2fcfee05be80c5cbab21606b5507aa23878adee1"

definition P384_SHA512_TV13_Qy :: nat where
  "P384_SHA512_TV13_Qy = 0x2cf2a9afeff83f3041dc8a05f016ccae58aa1a0e0dc6be9d928e97f2598c9ba5e9718d5eb74c9cfb516fd8c09f55f5b9"

definition P384_SHA512_TV13_Q :: "int point" where
  "P384_SHA512_TV13_Q = Point (int P384_SHA512_TV13_Qx) (int P384_SHA512_TV13_Qy)"

definition P384_SHA512_TV13_k :: nat where
  "P384_SHA512_TV13_k = 0x13d047708ae5228d6e3bbada0e385afdb3b735b31123454fdf40afe3c36efed563fd2cce84dcc45c553b0993d9ca9ec3"

definition P384_SHA512_TV13_R :: nat where
  "P384_SHA512_TV13_R = 0xa0203f6f2c456baac03538ed506a182e57a25151802cf4b2557613b2fb615ebd4c50ddc505f87c048a45bad3b2fc371c"

definition P384_SHA512_TV13_S :: nat where
  "P384_SHA512_TV13_S = 0x0eab56457c4080400fa3af124761d5a01fef35f9649edba8b97d22116386f3b8b363e97ef3f82616d5d825df1cf865ef"

definition P384_SHA512_TV13_Sig :: "nat × nat" where
  "P384_SHA512_TV13_Sig = (P384_SHA512_TV13_R, P384_SHA512_TV13_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA512_TV13_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV13_d" 
  by eval

lemma P384_SHA512_TV13_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA512_TV13_Q" 
  by eval

lemma P384_SHA512_TV13_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA512_TV13_d = P384_SHA512_TV13_Q" 
  by eval

lemma P384_SHA512_TV13_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA512_TV13_d P384_SHA512_TV13_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA512_TV13_dQvalidPair' P384_SHA512_TV13_d_valid) 

lemma P384_SHA512_TV13_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV13_k" 
  by eval

lemma P384_SHA512_TV13_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA512octets P384_SHA512_TV13_d P384_SHA512_TV13_Msg P384_SHA512_TV13_k = Some P384_SHA512_TV13_Sig" 
  by eval

lemma P384_SHA512_TV13_Verify: 
  "SEC1_P384_ECDSA_Verify SHA512octets P384_SHA512_TV13_Q P384_SHA512_TV13_Msg P384_SHA512_TV13_Sig" 
  by eval

subsection‹Test Vector: P384_SHA512_TV14›

subsubsection‹Given Test Vector Values›

definition P384_SHA512_TV14_Msg :: octets where
  "P384_SHA512_TV14_Msg = nat_to_octets_len 0xdb2ad659cf21bc9c1f7e6469c5f262b73261d49f7b1755fc137636e8ce0202f929dca4466c422284c10be8f351f36333ebc04b1888cba217c0fec872b2dfc3aa0d544e5e06a9518a8cfe3df5b20fbcb14a9bf218e3bf6a8e024530a17bab50906be34d9f9bba69af0b11d8ed426b9ec75c3bd1f2e5b8756e4a72ff846bc9e498 128"

definition P384_SHA512_TV14_d :: nat where
  "P384_SHA512_TV14_d = 0xe36339ddbe8787062a9bc4e1540690915dd2a2f11b3fe9ee946e281a0a2cbed426df405ed9cb0eca42f85443efd09e0c"

definition P384_SHA512_TV14_Qx :: nat where
  "P384_SHA512_TV14_Qx = 0xa1ffb4b790d1593e907369b69de10b93cddbb02c6131f787422364d9d692768ef8097970306cce16c97f2b10c538efa7"

definition P384_SHA512_TV14_Qy :: nat where
  "P384_SHA512_TV14_Qy = 0xd0692028601ea794d2563ffe9facc7273938fab47dd00b8960be15549a9c2b3f8552583eb4c6cd212fe486c159c79153"

definition P384_SHA512_TV14_Q :: "int point" where
  "P384_SHA512_TV14_Q = Point (int P384_SHA512_TV14_Qx) (int P384_SHA512_TV14_Qy)"

definition P384_SHA512_TV14_k :: nat where
  "P384_SHA512_TV14_k = 0x2226f7329378cecd697f36ae151546643d67760856854661e31d424fae662da910e2157da9bb6dfbe3622296e0b5710c"

definition P384_SHA512_TV14_R :: nat where
  "P384_SHA512_TV14_R = 0x20dcc25b67dd997621f437f65d78347fb57f8295b1b14453b1128203cda892bcfe726a2f107d30975d63172e56f11d76"

definition P384_SHA512_TV14_S :: nat where
  "P384_SHA512_TV14_S = 0x51cff592cbef75ef8321c8fa1e4229c4298b8180e427bee4e91d1e24fc28a729cf296beb728960d2a58cf26773d8e2e2"

definition P384_SHA512_TV14_Sig :: "nat × nat" where
  "P384_SHA512_TV14_Sig = (P384_SHA512_TV14_R, P384_SHA512_TV14_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA512_TV14_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV14_d" 
  by eval

lemma P384_SHA512_TV14_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA512_TV14_Q" 
  by eval

lemma P384_SHA512_TV14_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA512_TV14_d = P384_SHA512_TV14_Q" 
  by eval

lemma P384_SHA512_TV14_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA512_TV14_d P384_SHA512_TV14_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA512_TV14_dQvalidPair' P384_SHA512_TV14_d_valid) 

lemma P384_SHA512_TV14_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV14_k" 
  by eval

lemma P384_SHA512_TV14_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA512octets P384_SHA512_TV14_d P384_SHA512_TV14_Msg P384_SHA512_TV14_k = Some P384_SHA512_TV14_Sig" 
  by eval

lemma P384_SHA512_TV14_Verify: 
  "SEC1_P384_ECDSA_Verify SHA512octets P384_SHA512_TV14_Q P384_SHA512_TV14_Msg P384_SHA512_TV14_Sig" 
  by eval

subsection‹Test Vector: P384_SHA512_TV15›

subsubsection‹Given Test Vector Values›

definition P384_SHA512_TV15_Msg :: octets where
  "P384_SHA512_TV15_Msg = nat_to_octets_len 0xdbd8ddc02771a5ff7359d5216536b2e524a2d0b6ff180fa29a41a8847b6f45f1b1d52344d32aea62a23ea3d8584deaaea38ee92d1314fdb4fbbecdad27ac810f02de0452332939f644aa9fe526d313cea81b9c3f6a8dbbeafc899d0cdaeb1dca05160a8a039662c4c845a3dbb07be2bc8c9150e344103e404411668c48aa7792 128"

definition P384_SHA512_TV15_d :: nat where
  "P384_SHA512_TV15_d = 0x5da87be7af63fdaf40662bd2ba87597f54d7d52fae4b298308956cddbe5664f1e3c48cc6fd3c99291b0ce7a62a99a855"

definition P384_SHA512_TV15_Qx :: nat where
  "P384_SHA512_TV15_Qx = 0x54c79da7f8faeeee6f3a1fdc664e405d5c0fb3b904715f3a9d89d6fda7eabe6cee86ef82c19fca0d1a29e09c1acfcf18"

definition P384_SHA512_TV15_Qy :: nat where
  "P384_SHA512_TV15_Qy = 0x926c17d68778eb066c2078cdb688b17399e54bde5a79ef1852352a58967dff02c17a792d39f95c76d146fdc086fe26b0"

definition P384_SHA512_TV15_Q :: "int point" where
  "P384_SHA512_TV15_Q = Point (int P384_SHA512_TV15_Qx) (int P384_SHA512_TV15_Qy)"

definition P384_SHA512_TV15_k :: nat where
  "P384_SHA512_TV15_k = 0x1b686b45a31b31f6de9ed5362e18a3f8c8feded3d3b251b134835843b7ae8ede57c61dc61a30993123ac7699de4b6eac"

definition P384_SHA512_TV15_R :: nat where
  "P384_SHA512_TV15_R = 0x9dbfa147375767dde81b014f1e3bf579c44dd22486998a9b6f9e0920e53faa11eed29a4e2356e393afd1f5c1b060a958"

definition P384_SHA512_TV15_S :: nat where
  "P384_SHA512_TV15_S = 0xe4d318391f7cbfe70da78908d42db85225c85f4f2ff413ecad50aad5833abe91bdd5f6d64b0cd281398eab19452087dd"

definition P384_SHA512_TV15_Sig :: "nat × nat" where
  "P384_SHA512_TV15_Sig = (P384_SHA512_TV15_R, P384_SHA512_TV15_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P384_SHA512_TV15_d_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV15_d" 
  by eval

lemma P384_SHA512_TV15_QonCurve: 
  "SEC1_P384_on_curve' P384_SHA512_TV15_Q" 
  by eval

lemma P384_SHA512_TV15_dQvalidPair': 
  "SEC1_P384_ECkeyGen P384_SHA512_TV15_d = P384_SHA512_TV15_Q" 
  by eval

lemma P384_SHA512_TV15_dQvalidPair: 
  "SEC1_P384_ECkeyPairValid P384_SHA512_TV15_d P384_SHA512_TV15_Q" 
  by (simp add: SEC1_P384.ECkeyPairEqKeyGen P384_SHA512_TV15_dQvalidPair' P384_SHA512_TV15_d_valid) 

lemma P384_SHA512_TV15_k_valid: 
  "SEC1_P384_ECprivateKeyValid P384_SHA512_TV15_k" 
  by eval

lemma P384_SHA512_TV15_Sign: 
  "SEC1_P384_ECDSA_Sign' SHA512octets P384_SHA512_TV15_d P384_SHA512_TV15_Msg P384_SHA512_TV15_k = Some P384_SHA512_TV15_Sig" 
  by eval

lemma P384_SHA512_TV15_Verify: 
  "SEC1_P384_ECDSA_Verify SHA512octets P384_SHA512_TV15_Q P384_SHA512_TV15_Msg P384_SHA512_TV15_Sig" 
  by eval

section‹NIST Test Vectors for ECDSA: Curve P-521, Hash SHA-224›
  
subsection‹Test Vector: P521_SHA224_TV01›

subsubsection‹Given Test Vector Values›

definition P521_SHA224_TV01_Msg :: octets where
  "P521_SHA224_TV01_Msg = nat_to_octets_len 0x58ec2b2ceb80207ff51b17688bd5850f9388ce0b4a4f7316f5af6f52cfc4dde4192b6dbd97b56f93d1e4073517ac6c6140429b5484e266d07127e28b8e613ddf65888cbd5242b2f0eee4d5754eb11f25dfa5c3f87c790de371856c882731a157083a00d8eae29a57884dbbfcd98922c12cf5d73066daabe3bf3f42cfbdb9d853 128"

definition P521_SHA224_TV01_d :: nat where
  "P521_SHA224_TV01_d = 0x1d7bb864c5b5ecae019296cf9b5c63a166f5f1113942819b1933d889a96d12245777a99428f93de4fc9a18d709bf91889d7f8dddd522b4c364aeae13c983e9fae46"

definition P521_SHA224_TV01_Qx :: nat where
  "P521_SHA224_TV01_Qx = 0x1a7596d38aac7868327ddc1ef5e8178cf052b7ebc512828e8a45955d85bef49494d15278198bbcc5454358c12a2af9a3874e7002e1a2f02fcb36ff3e3b4bc0c69e7"

definition P521_SHA224_TV01_Qy :: nat where
  "P521_SHA224_TV01_Qy = 0x184902e515982bb225b8c84f245e61b327c08e94d41c07d0b4101a963e02fe52f6a9f33e8b1de2394e0cb74c40790b4e489b5500e6804cabed0fe8c192443d4027b"

definition P521_SHA224_TV01_Q :: "int point" where
  "P521_SHA224_TV01_Q = Point (int P521_SHA224_TV01_Qx) (int P521_SHA224_TV01_Qy)"

definition P521_SHA224_TV01_k :: nat where
  "P521_SHA224_TV01_k = 0x141f679033b27ec29219afd8aa123d5e535c227badbe2c86ff6eafa5116e9778000f538579a80ca4739b1675b8ff8b6245347852aa524fe9aad781f9b672e0bb3ff"

definition P521_SHA224_TV01_R :: nat where
  "P521_SHA224_TV01_R = 0x06b973a638bde22d8c1c0d804d94e40538526093705f92c0c4dac2c72e7db013a9c89ffc5b12a396886305ddf0cbaa7f10cdd4cd8866334c8abfc800e5cca365391"

definition P521_SHA224_TV01_S :: nat where
  "P521_SHA224_TV01_S = 0x0b0a01eca07a3964dd27d9ba6f3750615ea36434979dc73e153cd8ed1dbcde2885ead5757ebcabba117a64fcff9b5085d848f107f0c9ecc83dfa2fa09ada3503028"

definition P521_SHA224_TV01_Sig :: "nat × nat" where
  "P521_SHA224_TV01_Sig = (P521_SHA224_TV01_R, P521_SHA224_TV01_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA224_TV01_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV01_d" 
  by eval

lemma P521_SHA224_TV01_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA224_TV01_Q" 
  by eval

lemma P521_SHA224_TV01_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA224_TV01_d = P521_SHA224_TV01_Q" 
  by eval

lemma P521_SHA224_TV01_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA224_TV01_d P521_SHA224_TV01_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA224_TV01_dQvalidPair' P521_SHA224_TV01_d_valid) 

lemma P521_SHA224_TV01_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV01_k" 
  by eval

lemma P521_SHA224_TV01_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA224octets P521_SHA224_TV01_d P521_SHA224_TV01_Msg P521_SHA224_TV01_k = Some P521_SHA224_TV01_Sig" 
  by eval

lemma P521_SHA224_TV01_Verify: 
  "SEC1_P521_ECDSA_Verify SHA224octets P521_SHA224_TV01_Q P521_SHA224_TV01_Msg P521_SHA224_TV01_Sig" 
  by eval

subsection‹Test Vector: P521_SHA224_TV02›

subsubsection‹Given Test Vector Values›

definition P521_SHA224_TV02_Msg :: octets where
  "P521_SHA224_TV02_Msg = nat_to_octets_len 0x2449a53e0581f1b56d1e463b1c1686d33b3491efe1f3cc0443ba05d65694597cc7a2595bda9cae939166eb03cec624a788c9bbab69a39fb6554649131a56b26295683d8ac1aea969040413df405325425146c1e3a138d2f4f772ae2ed917cc36465acd66150058622440d7e77b3ad621e1c43a3f277da88d850d608079d9b911 128"

definition P521_SHA224_TV02_d :: nat where
  "P521_SHA224_TV02_d = 0x17e49b8ea8f9d1b7c0378e378a7a42e68e12cf78779ed41dcd29a090ae7e0f883b0d0f2cbc8f0473c0ad6732bea40d371a7f363bc6537d075bd1a4c23e558b0bc73"

definition P521_SHA224_TV02_Qx :: nat where
  "P521_SHA224_TV02_Qx = 0x0156cd2c485012ea5d5aadad724fb87558637de37b34485c4cf7c8cbc3e4f106cb1efd3e64f0adf99ddb51e3ac991bdd90785172386cdaf2c582cc46d6c99b0fed1"

definition P521_SHA224_TV02_Qy :: nat where
  "P521_SHA224_TV02_Qy = 0x1edeeda717554252b9f1e13553d4af028ec9e158dbe12332684fc1676dc731f39138a5d301376505a9ab04d562cc1659b0be9cb2b5e03bad8b412f2699c245b0ba2"

definition P521_SHA224_TV02_Q :: "int point" where
  "P521_SHA224_TV02_Q = Point (int P521_SHA224_TV02_Qx) (int P521_SHA224_TV02_Qy)"

definition P521_SHA224_TV02_k :: nat where
  "P521_SHA224_TV02_k = 0x1dc3e60a788caa5f62cb079f332d7e5c918974643dca3ab3566a599642cd84964fbef43ce94290041fe3d2c8c26104d9c73a57a7d4724613242531083b49e255f33"

definition P521_SHA224_TV02_R :: nat where
  "P521_SHA224_TV02_R = 0x12592c0be6cce18efb2b972cd193d036dcb850f2390fa8b9b86b2f876548bc424fb3bc13c1e5c415fa09d0ecfcae5bf76fb23e8322d7eecb264a2ae6d20ef50d405"

definition P521_SHA224_TV02_S :: nat where
  "P521_SHA224_TV02_S = 0x11bc9713be88e3b9912a3e5f5d7b56f20573e979b1a75d04ce339f724bddffa4665d25995fe24d32507d8a07c5e10169f5338ef2827737f7b0291752b21237217e3"

definition P521_SHA224_TV02_Sig :: "nat × nat" where
  "P521_SHA224_TV02_Sig = (P521_SHA224_TV02_R, P521_SHA224_TV02_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA224_TV02_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV02_d" 
  by eval

lemma P521_SHA224_TV02_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA224_TV02_Q" 
  by eval

lemma P521_SHA224_TV02_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA224_TV02_d = P521_SHA224_TV02_Q" 
  by eval

lemma P521_SHA224_TV02_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA224_TV02_d P521_SHA224_TV02_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA224_TV02_dQvalidPair' P521_SHA224_TV02_d_valid) 

lemma P521_SHA224_TV02_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV02_k" 
  by eval

lemma P521_SHA224_TV02_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA224octets P521_SHA224_TV02_d P521_SHA224_TV02_Msg P521_SHA224_TV02_k = Some P521_SHA224_TV02_Sig" 
  by eval

lemma P521_SHA224_TV02_Verify: 
  "SEC1_P521_ECDSA_Verify SHA224octets P521_SHA224_TV02_Q P521_SHA224_TV02_Msg P521_SHA224_TV02_Sig" 
  by eval

subsection‹Test Vector: P521_SHA224_TV03›

subsubsection‹Given Test Vector Values›

definition P521_SHA224_TV03_Msg :: octets where
  "P521_SHA224_TV03_Msg = nat_to_octets_len 0x7ba05797b5b67e1adfafb7fae20c0c0abe1543c94cee92d5021e1abc57720a6107999c70eacf3d4a79702cd4e6885fa1b7155398ac729d1ed6b45e51fe114c46caf444b20b406ad9cde6b9b2687aa645b46b51ab790b67047219e7290df1a797f35949aaf912a0a8556bb21018e7f70427c0fc018e461755378b981d0d9df3a9 128"

definition P521_SHA224_TV03_d :: nat where
  "P521_SHA224_TV03_d = 0x135ea346852f837d10c1b2dfb8012ae8215801a7e85d4446dadd993c68d1e9206e1d8651b7ed763b95f707a52410eeef4f21ae9429828289eaea1fd9caadf826ace"

definition P521_SHA224_TV03_Qx :: nat where
  "P521_SHA224_TV03_Qx = 0x18d40cc4573892b3e467d314c39c95615ee0510e3e4dbc9fa28f6cd1f73e7acde15ad7c8c5339df9a7774f8155130e7d1f8de9139ddd6dfe1841c1e64c38ea98243"

definition P521_SHA224_TV03_Qy :: nat where
  "P521_SHA224_TV03_Qy = 0x17021782d33dc513716c83afe7ba5e7abef9cb25b31f483661115b8d6b5ae469aaf6f3d54baa3b658a9af9b6249fd4d5ea7a07cb8b600f1df72b81dac614cfc384a"

definition P521_SHA224_TV03_Q :: "int point" where
  "P521_SHA224_TV03_Q = Point (int P521_SHA224_TV03_Qx) (int P521_SHA224_TV03_Qy)"

definition P521_SHA224_TV03_k :: nat where
  "P521_SHA224_TV03_k = 0x0c24acc1edb3777212e5b0bac744eadf4eda11fa150753b355bf96b189e6f57fc02284bb22d8b3cd8bba7a09aae9f4ea955b382063425a6f8da2f99b9647b147172"

definition P521_SHA224_TV03_R :: nat where
  "P521_SHA224_TV03_R = 0x183da7b8a9f9d5f08903359c1a2435b085fcf26a2ed09ab71357bb7634054acc569535e6fe81d28233e4703005fc4bf83ce794d9463d575795aa0f03398e854cefd"

definition P521_SHA224_TV03_S :: nat where
  "P521_SHA224_TV03_S = 0x0b3621145b9866ab7809139795cc30cd0404127a7f0fafa793660491009f6c53724fdb0b1ffbf0fd51c131180b8a957fe66e76d2970247c024261c768dee9abbfb9"

definition P521_SHA224_TV03_Sig :: "nat × nat" where
  "P521_SHA224_TV03_Sig = (P521_SHA224_TV03_R, P521_SHA224_TV03_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA224_TV03_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV03_d" 
  by eval

lemma P521_SHA224_TV03_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA224_TV03_Q" 
  by eval

lemma P521_SHA224_TV03_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA224_TV03_d = P521_SHA224_TV03_Q" 
  by eval

lemma P521_SHA224_TV03_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA224_TV03_d P521_SHA224_TV03_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA224_TV03_dQvalidPair' P521_SHA224_TV03_d_valid) 

lemma P521_SHA224_TV03_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV03_k" 
  by eval

lemma P521_SHA224_TV03_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA224octets P521_SHA224_TV03_d P521_SHA224_TV03_Msg P521_SHA224_TV03_k = Some P521_SHA224_TV03_Sig" 
  by eval

lemma P521_SHA224_TV03_Verify: 
  "SEC1_P521_ECDSA_Verify SHA224octets P521_SHA224_TV03_Q P521_SHA224_TV03_Msg P521_SHA224_TV03_Sig" 
  by eval

subsection‹Test Vector: P521_SHA224_TV04›

subsubsection‹Given Test Vector Values›

definition P521_SHA224_TV04_Msg :: octets where
  "P521_SHA224_TV04_Msg = nat_to_octets_len 0x716dabdb22a1c854ec60420249905a1d7ca68dd573efaff7542e76f0eae54a1828db69a39a1206cd05e10e681f24881b131e042ed9e19f5995c253840e937b809dfb8027fed71d541860f318691c13a2eb514daa5889410f256305f3b5b47cc16f7a7dad6359589b5f4568de4c4aae2357a8ea5e0ebaa5b89063eb3aa44eb952 128"

definition P521_SHA224_TV04_d :: nat where
  "P521_SHA224_TV04_d = 0x1393cb1ee9bfd7f7b9c057ecc66b43e807e12515f66ed7e9c9210ba1514693965988e567fbad7c3f17231aacee0e9b9a4b1940504b1cd4fd5edfaa62ba4e3e476fc"

definition P521_SHA224_TV04_Qx :: nat where
  "P521_SHA224_TV04_Qx = 0x1e855c935139c8092092cfa733db1292530506eeb2bbb1687f9602c36d97a6714e998892d5d3b842d1896a6ece9d549e9792881a256256137b3dff180c96cc5d07b"

definition P521_SHA224_TV04_Qy :: nat where
  "P521_SHA224_TV04_Qy = 0x18d83b6e93cd287311f7bf7c1d7f9eeabcf0b69c12f2d8f40e333e81e956d968532a37a4c04d761874df293b484cd7053b03fdbc2fdcd3b4c412d6f272fb7c93fe6"

definition P521_SHA224_TV04_Q :: "int point" where
  "P521_SHA224_TV04_Q = Point (int P521_SHA224_TV04_Qx) (int P521_SHA224_TV04_Qy)"

definition P521_SHA224_TV04_k :: nat where
  "P521_SHA224_TV04_k = 0x1d98619bdc04735d30c222fc67da82c069aea5f449af5e8c4db10c1786c0cb9e6f2cc0bb66fa6be18c485570d648dafcd0a973c43d5c94e9a9dacbd3170e53fa2a0"

definition P521_SHA224_TV04_R :: nat where
  "P521_SHA224_TV04_R = 0x0bf47fabe107ce0ec03e2ad60a79b058e1bebb18568b6a8cdbe86032e71aa30c15766105b2ea952cfa79bcab046df601159f96e179bbcf252dc68ac73d31481fdae"

definition P521_SHA224_TV04_S :: nat where
  "P521_SHA224_TV04_S = 0x1f918fec69cd07d90f9d892b7117e7519c3224947f4262f1fd97077dd5386a6c78aeddff3ee97e59ea353f06029f1336f0d6ef5c0f4b17ca59343a55319b7bfc3db"

definition P521_SHA224_TV04_Sig :: "nat × nat" where
  "P521_SHA224_TV04_Sig = (P521_SHA224_TV04_R, P521_SHA224_TV04_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA224_TV04_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV04_d" 
  by eval

lemma P521_SHA224_TV04_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA224_TV04_Q" 
  by eval

lemma P521_SHA224_TV04_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA224_TV04_d = P521_SHA224_TV04_Q" 
  by eval

lemma P521_SHA224_TV04_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA224_TV04_d P521_SHA224_TV04_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA224_TV04_dQvalidPair' P521_SHA224_TV04_d_valid) 

lemma P521_SHA224_TV04_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV04_k" 
  by eval

lemma P521_SHA224_TV04_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA224octets P521_SHA224_TV04_d P521_SHA224_TV04_Msg P521_SHA224_TV04_k = Some P521_SHA224_TV04_Sig" 
  by eval

lemma P521_SHA224_TV04_Verify: 
  "SEC1_P521_ECDSA_Verify SHA224octets P521_SHA224_TV04_Q P521_SHA224_TV04_Msg P521_SHA224_TV04_Sig" 
  by eval

subsection‹Test Vector: P521_SHA224_TV05›

subsubsection‹Given Test Vector Values›

definition P521_SHA224_TV05_Msg :: octets where
  "P521_SHA224_TV05_Msg = nat_to_octets_len 0x9cc9c2f131fe3ac7ea91ae6d832c7788cbbf34f68e839269c336ceef7bef6f20c0a62ea8cc340a333a3002145d07eba4cf4026a0c4b26b0217a0046701de92d573d7c87a386a1ea68dc80525b7dcc9be41b451ad9f3d16819e2a0a0b5a0c56736da3709e64761f97cae2399de2a4022dc4c3d73c7a1735c36dbde86c4bc5b6f7 128"

definition P521_SHA224_TV05_d :: nat where
  "P521_SHA224_TV05_d = 0x179fa164e051c5851e8a37d82c181e809a05fea9a3f083299b22684f59aa27e40dc5a33b3f7949338764d46bfe1f355134750518b856d98d9167ef07aac3092c549"

definition P521_SHA224_TV05_Qx :: nat where
  "P521_SHA224_TV05_Qx = 0x1857cc7bbed20e87b3fd9a104956aa20c6502192910e0e7598410526ebfe1c99397b85189612a60c51fb8f4dd5cb08a8cd2e702563062dcb043410715c5323a0046"

definition P521_SHA224_TV05_Qy :: nat where
  "P521_SHA224_TV05_Qy = 0x1fce8d135284310d2f38c216030634b32cd223222f0d9d8d2b7c55477c4b8b74fc6c96a6092f34b05ca44d3633a5037c2166c479a032bb4f949f89fc1ba5236d07d"

definition P521_SHA224_TV05_Q :: "int point" where
  "P521_SHA224_TV05_Q = Point (int P521_SHA224_TV05_Qx) (int P521_SHA224_TV05_Qy)"

definition P521_SHA224_TV05_k :: nat where
  "P521_SHA224_TV05_k = 0x16d9704c0cee791f2938bb2a8a595752a3635c2f557efeecefd719414b5f2aaf846080f582c76eae7a8fddf81859b49d0131c212524d55defa67dca1a9a28ca400f"

definition P521_SHA224_TV05_R :: nat where
  "P521_SHA224_TV05_R = 0x1c9a4e51774384e8362876a87c572e6463a54413c7c6252c552ebb182f83e45ace436ade4ca373d8a7216e83efb62c8b41c4d5132a0afa65078f16d189baca39187"

definition P521_SHA224_TV05_S :: nat where
  "P521_SHA224_TV05_S = 0x1e92a7dd5fea29a666398e1df5775cbb5664fe6943fe4c1d2bba516b7543c84df584458e53919c4ffab579a26fb3c892a5d1a77b0a07428c89350f8b559e627b014"

definition P521_SHA224_TV05_Sig :: "nat × nat" where
  "P521_SHA224_TV05_Sig = (P521_SHA224_TV05_R, P521_SHA224_TV05_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA224_TV05_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV05_d" 
  by eval

lemma P521_SHA224_TV05_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA224_TV05_Q" 
  by eval

lemma P521_SHA224_TV05_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA224_TV05_d = P521_SHA224_TV05_Q" 
  by eval

lemma P521_SHA224_TV05_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA224_TV05_d P521_SHA224_TV05_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA224_TV05_dQvalidPair' P521_SHA224_TV05_d_valid) 

lemma P521_SHA224_TV05_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV05_k" 
  by eval

lemma P521_SHA224_TV05_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA224octets P521_SHA224_TV05_d P521_SHA224_TV05_Msg P521_SHA224_TV05_k = Some P521_SHA224_TV05_Sig" 
  by eval

lemma P521_SHA224_TV05_Verify: 
  "SEC1_P521_ECDSA_Verify SHA224octets P521_SHA224_TV05_Q P521_SHA224_TV05_Msg P521_SHA224_TV05_Sig" 
  by eval

subsection‹Test Vector: P521_SHA224_TV06›

subsubsection‹Given Test Vector Values›

definition P521_SHA224_TV06_Msg :: octets where
  "P521_SHA224_TV06_Msg = nat_to_octets_len 0x14c69f8d660f7a6b37b13a6d9788eff16311b67598ab8368039ea1d9146e54f55a83b3d13d7ac9652135933c68fafd993a582253be0deea282d86046c2fb6fd3a7b2c80874ced28d8bed791bd4134c796bb7baf195bdd0dc6fa03fdb7f98755ca063fb1349e56fd0375cf94774df4203b34495404ebb86f1c7875b85174c574c 128"

definition P521_SHA224_TV06_d :: nat where
  "P521_SHA224_TV06_d = 0x13dabca37130ba278eae2b3d106b5407711b0d3b437fbf1c952f0773571570764d2c7cb8896a8815f3f1975b21adc6697898e5c0a4242092fc1b80db819a4702df4"

definition P521_SHA224_TV06_Qx :: nat where
  "P521_SHA224_TV06_Qx = 0x0bc2aebf40cd435bc37d73c09d05f2fd71321111a767c2b0d446f90dd4a186839c694ceb734e027e7ee948f0f63e4d3f1656d3d543df23c342a599306909b347109"

definition P521_SHA224_TV06_Qy :: nat where
  "P521_SHA224_TV06_Qy = 0x1f4c98ac03f0718e58d5d1762c920445b11dbdd60ec7f60095809204e14965a4ecb0be6fea06adbac8ba431d6f144c75c199225df2a619a34be99897125b3a10af8"

definition P521_SHA224_TV06_Q :: "int point" where
  "P521_SHA224_TV06_Q = Point (int P521_SHA224_TV06_Qx) (int P521_SHA224_TV06_Qy)"

definition P521_SHA224_TV06_k :: nat where
  "P521_SHA224_TV06_k = 0x0401187c8b89945a1e48cda9ee52167789f4121e67482a7ac797899f5d3d2e623aed31e4adae08a8d43e69028fa074d2650317cbc765f6ed191cf0317b4bae57881"

definition P521_SHA224_TV06_R :: nat where
  "P521_SHA224_TV06_R = 0x1e572afed754016fba43fc33e352932c4db65efcb84e2bd159b40fc5925893b161effc40240be28d8c07154d2615f605c6f0451b976522d95afd37f46602df7a12a"

definition P521_SHA224_TV06_S :: nat where
  "P521_SHA224_TV06_S = 0x030370c1c5352c2b663ac1858b42f69545b2f58ed5b2c007f303726977d3c756b5d644ec6788f94c886f78269aa190a3d8d1ae10e4fd24d937c4556fb9e1953fd6d"

definition P521_SHA224_TV06_Sig :: "nat × nat" where
  "P521_SHA224_TV06_Sig = (P521_SHA224_TV06_R, P521_SHA224_TV06_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA224_TV06_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV06_d" 
  by eval

lemma P521_SHA224_TV06_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA224_TV06_Q" 
  by eval

lemma P521_SHA224_TV06_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA224_TV06_d = P521_SHA224_TV06_Q" 
  by eval

lemma P521_SHA224_TV06_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA224_TV06_d P521_SHA224_TV06_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA224_TV06_dQvalidPair' P521_SHA224_TV06_d_valid) 

lemma P521_SHA224_TV06_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV06_k" 
  by eval

lemma P521_SHA224_TV06_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA224octets P521_SHA224_TV06_d P521_SHA224_TV06_Msg P521_SHA224_TV06_k = Some P521_SHA224_TV06_Sig" 
  by eval

lemma P521_SHA224_TV06_Verify: 
  "SEC1_P521_ECDSA_Verify SHA224octets P521_SHA224_TV06_Q P521_SHA224_TV06_Msg P521_SHA224_TV06_Sig" 
  by eval

subsection‹Test Vector: P521_SHA224_TV07›

subsubsection‹Given Test Vector Values›

definition P521_SHA224_TV07_Msg :: octets where
  "P521_SHA224_TV07_Msg = nat_to_octets_len 0x8d8e75df200c177dbfe61be61567b82177ea5ec58e2781168d2277d2fd42668f01248ca3eb29ffa2689b12ae40f9c429532b6d2e1f15891322b825a0a072a1c68fa09e78cfdef3e95ed6fdf7233a43cb68236560d49a3278f0b3f47cb08f475bd9ab2f60755ea4a1767de9313b71a1b9ea87ef33f34682efbda263b0f8cc2f52 128"

definition P521_SHA224_TV07_d :: nat where
  "P521_SHA224_TV07_d = 0x198681adbde7840d7ccd9cf1fb82056433fb4dd26bddf909af7b3b99da1ca2c05c8d4560ecd80ba68f376f8b487897e374e99a9288ed7e3645cc0d00a478aae8d16"

definition P521_SHA224_TV07_Qx :: nat where
  "P521_SHA224_TV07_Qx = 0x057ce3777af7032f1f82308682e71fe09f88bf29dacd5018a725e1caa4b1e2bfdd894fe618f9266f31ba089856dc9c1b70e4a2faa08b4b744d1aafcd5ae99e2c736"

definition P521_SHA224_TV07_Qy :: nat where
  "P521_SHA224_TV07_Qy = 0x199bcfef2021bc5890d7d39ec5dc0c26956801e84cae742cf6c50386eb289b6e97754dd25a94abf81f1cb1b36935b5eb29f4b32a6516d2ff6a7d23064a0daec94b3"

definition P521_SHA224_TV07_Q :: "int point" where
  "P521_SHA224_TV07_Q = Point (int P521_SHA224_TV07_Qx) (int P521_SHA224_TV07_Qy)"

definition P521_SHA224_TV07_k :: nat where
  "P521_SHA224_TV07_k = 0x19d2d74ad8ee2d85048f386998a71899ef6c960b4ab324e5fd1c0a076c5a632fd0009500076522e052c5c9806eef7056da48df6b16eb71cdf0f1838b0e21715fce0"

definition P521_SHA224_TV07_R :: nat where
  "P521_SHA224_TV07_R = 0x18ecacbcffd5414bbb96728e5f2d4c90178e27733d13617e134ec788022db124374bbaa11e2c77fe3f38d1af6e998e1b0266b77380984c423e80ffa6ff2bcafd57a"

definition P521_SHA224_TV07_S :: nat where
  "P521_SHA224_TV07_S = 0x1c727f34b6a378f3087721a54e9796499b597ecf6666b8f18312d67e1190a8a66e878efc2367b551267494e0245979ef4deed6d2cbf2c3711af6d82ccfeb101a377"

definition P521_SHA224_TV07_Sig :: "nat × nat" where
  "P521_SHA224_TV07_Sig = (P521_SHA224_TV07_R, P521_SHA224_TV07_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA224_TV07_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV07_d" 
  by eval

lemma P521_SHA224_TV07_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA224_TV07_Q" 
  by eval

lemma P521_SHA224_TV07_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA224_TV07_d = P521_SHA224_TV07_Q" 
  by eval

lemma P521_SHA224_TV07_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA224_TV07_d P521_SHA224_TV07_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA224_TV07_dQvalidPair' P521_SHA224_TV07_d_valid) 

lemma P521_SHA224_TV07_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV07_k" 
  by eval

lemma P521_SHA224_TV07_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA224octets P521_SHA224_TV07_d P521_SHA224_TV07_Msg P521_SHA224_TV07_k = Some P521_SHA224_TV07_Sig" 
  by eval

lemma P521_SHA224_TV07_Verify: 
  "SEC1_P521_ECDSA_Verify SHA224octets P521_SHA224_TV07_Q P521_SHA224_TV07_Msg P521_SHA224_TV07_Sig" 
  by eval

subsection‹Test Vector: P521_SHA224_TV08›

subsubsection‹Given Test Vector Values›

definition P521_SHA224_TV08_Msg :: octets where
  "P521_SHA224_TV08_Msg = nat_to_octets_len 0x10631c3d438870f311c905e569a58e56d20a2a560e857f0f9bac2bb7233ec40c79de145294da0937e6b5e5c34fff4e6270823e5c8553c07d4adf25f614845b2eac731c5773ebbd716ab45698d156d043859945de57473389954d223522fbafecf560b07ef9ba861bcc1df9a7a89cdd6debf4cd9bf2cf28c193393569ccbd0398 128"

definition P521_SHA224_TV08_d :: nat where
  "P521_SHA224_TV08_d = 0x08c4c0fd9696d86e99a6c1c32349a89a0b0c8384f2829d1281730d4e9af1df1ad5a0bcfccc6a03a703b210defd5d49a6fb82536f88b885776f0f7861c6fc010ef37"

definition P521_SHA224_TV08_Qx :: nat where
  "P521_SHA224_TV08_Qx = 0x164ac88ed9afe137f648dd89cdd9956682830cac5f7c1a06d19a1b19f82bb1d22dfeefea30d35c11202fed93fd5ce64835d27c6564d6e181287fa04a2d20994986b"

definition P521_SHA224_TV08_Qy :: nat where
  "P521_SHA224_TV08_Qy = 0x05cb83669265f5380ccefe6b4f85fdf0049e6703f6f378a0b2e52ed0fbbcf300afebb722f4ed48e3819cb976c1d60e2ba05646b478f6dfecfbae730e9644c297f00"

definition P521_SHA224_TV08_Q :: "int point" where
  "P521_SHA224_TV08_Q = Point (int P521_SHA224_TV08_Qx) (int P521_SHA224_TV08_Qy)"

definition P521_SHA224_TV08_k :: nat where
  "P521_SHA224_TV08_k = 0x189801432cba9bf8c0763d43b6ec3b8636e62324587a4e27905b09a58e4aa66d07d096dbce87824e837be1c243dd741f983c535a5dd2f077aac8beee9918258d3cb"

definition P521_SHA224_TV08_R :: nat where
  "P521_SHA224_TV08_R = 0x0917723f7241e8dc7cd746b699ab621d068dd3a90e906aaf0a4862744b96fd4e5ccdb9c7796c27f7196e693d06ec209464c3ea60ad6313e9b77cceaa14767e6651c"

definition P521_SHA224_TV08_S :: nat where
  "P521_SHA224_TV08_S = 0x0957b0ecdc3668f6efa5d0957615bcfffd6419c5e57579b74f960f65ae3fb9e8284322ff710b066f7e0959ac926d3cf9a594bdb70bbec756c96910b26a2486dee9e"

definition P521_SHA224_TV08_Sig :: "nat × nat" where
  "P521_SHA224_TV08_Sig = (P521_SHA224_TV08_R, P521_SHA224_TV08_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA224_TV08_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV08_d" 
  by eval

lemma P521_SHA224_TV08_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA224_TV08_Q" 
  by eval

lemma P521_SHA224_TV08_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA224_TV08_d = P521_SHA224_TV08_Q" 
  by eval

lemma P521_SHA224_TV08_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA224_TV08_d P521_SHA224_TV08_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA224_TV08_dQvalidPair' P521_SHA224_TV08_d_valid) 

lemma P521_SHA224_TV08_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV08_k" 
  by eval

lemma P521_SHA224_TV08_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA224octets P521_SHA224_TV08_d P521_SHA224_TV08_Msg P521_SHA224_TV08_k = Some P521_SHA224_TV08_Sig" 
  by eval

lemma P521_SHA224_TV08_Verify: 
  "SEC1_P521_ECDSA_Verify SHA224octets P521_SHA224_TV08_Q P521_SHA224_TV08_Msg P521_SHA224_TV08_Sig" 
  by eval

subsection‹Test Vector: P521_SHA224_TV09›

subsubsection‹Given Test Vector Values›

definition P521_SHA224_TV09_Msg :: octets where
  "P521_SHA224_TV09_Msg = nat_to_octets_len 0x80aad6d696cbe654faa0d0a24d2f50d46e4f00a1b488ea1a98ed06c44d1d0c568beb4ab3674fc2b1d2d3da1053f28940e89ba1244899e8515cabdd66e99a77df31e90d93e37a8a240e803a998209988fc829e239150da058a300489e33bf3dcdaf7d06069e74569fee77f4e3875d0a713ccd2b7e9d7be62b34b6e375e84209ef 128"

definition P521_SHA224_TV09_d :: nat where
  "P521_SHA224_TV09_d = 0x1466d14f8fbe25544b209c5e6a000b771ef107867e28ed489a42015119d1aa64bff51d6b7a0ac88673bbc3618c917561cff4a41cdb7c2833dab5ebb9d0ddf2ca256"

definition P521_SHA224_TV09_Qx :: nat where
  "P521_SHA224_TV09_Qx = 0x1dc8b71d55700573a26af6698b92b66180cf43e153edadb720780321dbb4e71d28e0a488e4201d207fc4848fe9dd10dcabec44492656a3ff7a665fe932445c82d0b"

definition P521_SHA224_TV09_Qy :: nat where
  "P521_SHA224_TV09_Qy = 0x1920b16331b7abeb3db883a31288ef66f80b7728b008b3cc33e03a68f68d9e653a86e3177bbc00014fa5ea4c1608c0d455c2e2ac7bd8ab8519ebf19955edf1baf8d"

definition P521_SHA224_TV09_Q :: "int point" where
  "P521_SHA224_TV09_Q = Point (int P521_SHA224_TV09_Qx) (int P521_SHA224_TV09_Qy)"

definition P521_SHA224_TV09_k :: nat where
  "P521_SHA224_TV09_k = 0x160d04420e0d31b0df476f83393b1f9aff68389cc3299e42ef348d97646f7531a722b66ddfb9501bbb5c4a41d84c78be7233b11489bceb817d23060e6017433fab8"

definition P521_SHA224_TV09_R :: nat where
  "P521_SHA224_TV09_R = 0x08077aabd0a342f03f912007c586cfedfc63f93d1118f720d5b62b3ce141a60f86f111dfd8fc2e31a6778981f1a5e28f29a7369bd7897bb41240c8d3a9c170e0ee0"

definition P521_SHA224_TV09_S :: nat where
  "P521_SHA224_TV09_S = 0x00abc75fc154b93840579457820957e89d1260fee0a4b9bb1946f61ca1e71afd76bb5e1077b3e38ceb39d1fac5ef8b217c4110617b3ad118e02b3fcc2a39ef38613"

definition P521_SHA224_TV09_Sig :: "nat × nat" where
  "P521_SHA224_TV09_Sig = (P521_SHA224_TV09_R, P521_SHA224_TV09_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA224_TV09_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV09_d" 
  by eval

lemma P521_SHA224_TV09_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA224_TV09_Q" 
  by eval

lemma P521_SHA224_TV09_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA224_TV09_d = P521_SHA224_TV09_Q" 
  by eval

lemma P521_SHA224_TV09_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA224_TV09_d P521_SHA224_TV09_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA224_TV09_dQvalidPair' P521_SHA224_TV09_d_valid) 

lemma P521_SHA224_TV09_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV09_k" 
  by eval

lemma P521_SHA224_TV09_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA224octets P521_SHA224_TV09_d P521_SHA224_TV09_Msg P521_SHA224_TV09_k = Some P521_SHA224_TV09_Sig" 
  by eval

lemma P521_SHA224_TV09_Verify: 
  "SEC1_P521_ECDSA_Verify SHA224octets P521_SHA224_TV09_Q P521_SHA224_TV09_Msg P521_SHA224_TV09_Sig" 
  by eval

subsection‹Test Vector: P521_SHA224_TV10›

subsubsection‹Given Test Vector Values›

definition P521_SHA224_TV10_Msg :: octets where
  "P521_SHA224_TV10_Msg = nat_to_octets_len 0x8a7792a2870d2dd341cd9c4a2a9ec2da753dcb0f692b70b64cef2e22071389c70b3b188dea5f409fb435cbd09082f59de6bc2ff9e65f91b7acc51e6e7f8e513148cb3c7c4664f227d5c704626b0fda447aa87b9d47cd99789b88628eb642ed250312de5ba6b25f3d5342a3cbb7ebd69b0044ee2b4c9ba5e3f5195afb6bea823d 128"

definition P521_SHA224_TV10_d :: nat where
  "P521_SHA224_TV10_d = 0x01a99fcf54c9b85010f20dc4e48199266c70767e18b2c618044542cd0e23733817776a1a45dbd74a8e8244a313d96c779f723013cd88886cb7a08ef7ee8fdd862e7"

definition P521_SHA224_TV10_Qx :: nat where
  "P521_SHA224_TV10_Qx = 0x1912d33b01d51e2f777bdbd1ada23f2b1a9faf2be2f2a3b152547db9b149b697dd71824ca96547462e347bc4ef9530e7466318c25338c7e04323b1ba5fd25ea7162"

definition P521_SHA224_TV10_Qy :: nat where
  "P521_SHA224_TV10_Qy = 0x0bbe9b1e3a84accd69b76b253f556c63e3f374e3de0d1f5e3600fc19215533b2e40d6b32c3af33314d223ea2366a51d1a337af858f69326389276f91be5c466e649"

definition P521_SHA224_TV10_Q :: "int point" where
  "P521_SHA224_TV10_Q = Point (int P521_SHA224_TV10_Qx) (int P521_SHA224_TV10_Qy)"

definition P521_SHA224_TV10_k :: nat where
  "P521_SHA224_TV10_k = 0x14fafd60cb026f50c23481867772411bb426ec6b97054e025b35db74fe8ea8f74faa2d36e7d40b4652d1f61794878510b49b7b4fe4349afccd24fc45fec2fd9e9e7"

definition P521_SHA224_TV10_R :: nat where
  "P521_SHA224_TV10_R = 0x18b1df1b6d7030a23a154cacce4a2e3761cc6251ff8bf6c9f6c89d0a15123baef9b338ada59728349ce685c03109fcde512ed01a40afd2ca34e1bc02ecf2871d45c"

definition P521_SHA224_TV10_S :: nat where
  "P521_SHA224_TV10_S = 0x0a399f9b9e21aeddf450429fec2dc5749e4a4c7e4f94cee736004dcc089c47635da22845992cd076a4f0a01d2cc1b0af6e17b81a802361699b862157ad6cad8bd1d"

definition P521_SHA224_TV10_Sig :: "nat × nat" where
  "P521_SHA224_TV10_Sig = (P521_SHA224_TV10_R, P521_SHA224_TV10_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA224_TV10_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV10_d" 
  by eval

lemma P521_SHA224_TV10_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA224_TV10_Q" 
  by eval

lemma P521_SHA224_TV10_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA224_TV10_d = P521_SHA224_TV10_Q" 
  by eval

lemma P521_SHA224_TV10_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA224_TV10_d P521_SHA224_TV10_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA224_TV10_dQvalidPair' P521_SHA224_TV10_d_valid) 

lemma P521_SHA224_TV10_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV10_k" 
  by eval

lemma P521_SHA224_TV10_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA224octets P521_SHA224_TV10_d P521_SHA224_TV10_Msg P521_SHA224_TV10_k = Some P521_SHA224_TV10_Sig" 
  by eval

lemma P521_SHA224_TV10_Verify: 
  "SEC1_P521_ECDSA_Verify SHA224octets P521_SHA224_TV10_Q P521_SHA224_TV10_Msg P521_SHA224_TV10_Sig" 
  by eval

subsection‹Test Vector: P521_SHA224_TV11›

subsubsection‹Given Test Vector Values›

definition P521_SHA224_TV11_Msg :: octets where
  "P521_SHA224_TV11_Msg = nat_to_octets_len 0xf971bcd396efb8392207b5ca72ac62649b47732fba8feaa8e84f7fb36b3edb5d7b5333fbfa39a4f882cb42fe57cd1ace43d06aaad33d0603741a18bc261caa14f29ead389f7c20536d406e9d39c34079812ba26b39baedf5feb1ef1f79990496dd019c87e38c38c486ec1c251da2a8a9a57854b80fcd513285e8dee8c43a9890 128"

definition P521_SHA224_TV11_d :: nat where
  "P521_SHA224_TV11_d = 0x1b6015d898611fbaf0b66a344fa18d1d488564352bf1c2da40f52cd997952f8ccb436b693851f9ccb69c519d8a033cf27035c27233324f10e9969a3b384e1c1dc73"

definition P521_SHA224_TV11_Qx :: nat where
  "P521_SHA224_TV11_Qx = 0x110c6177ceb44b0aec814063f297c0c890671220413dbd900e4f037a67d87583eaf4b6a9a1d2092472c17641362313c6a96f19829bb982e76e3a993932b848c7a97"

definition P521_SHA224_TV11_Qy :: nat where
  "P521_SHA224_TV11_Qy = 0x0f6e566c4e49b2ee70a900dc53295640f3a4a66732df80b29f497f4ae2fa61d0949f7f4b12556967bb92201a4f5d1384d741120c95b617b99c47a61e11c93a482d6"

definition P521_SHA224_TV11_Q :: "int point" where
  "P521_SHA224_TV11_Q = Point (int P521_SHA224_TV11_Qx) (int P521_SHA224_TV11_Qy)"

definition P521_SHA224_TV11_k :: nat where
  "P521_SHA224_TV11_k = 0x1a88667b9bdfe72fb87a6999a59b8b139e18ef9273261549bc394d884db5aa64a0bc7c7d38a8ef17333478d2119d826e2540560d65f52b9a6dc91be1340cfd8f8f8"

definition P521_SHA224_TV11_R :: nat where
  "P521_SHA224_TV11_R = 0x015f73def52ea47ddb03e0a5d154999642202e06e6734ac930c1dc84756c67bbb1cca9f21f92d61bfdb2052c5dd2833349610f68139393d77250a7662ef7bd17cbe"

definition P521_SHA224_TV11_S :: nat where
  "P521_SHA224_TV11_S = 0x155c744a729f83b27d1f325a91e63a0d564fe96ff91eaa1bad3bff17d2abffa065d14a1d20a04dd993f6ed3260b60bcc6401e31f6bc75aaafe03e8c1a9cd14d2708"

definition P521_SHA224_TV11_Sig :: "nat × nat" where
  "P521_SHA224_TV11_Sig = (P521_SHA224_TV11_R, P521_SHA224_TV11_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA224_TV11_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV11_d" 
  by eval

lemma P521_SHA224_TV11_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA224_TV11_Q" 
  by eval

lemma P521_SHA224_TV11_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA224_TV11_d = P521_SHA224_TV11_Q" 
  by eval

lemma P521_SHA224_TV11_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA224_TV11_d P521_SHA224_TV11_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA224_TV11_dQvalidPair' P521_SHA224_TV11_d_valid) 

lemma P521_SHA224_TV11_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV11_k" 
  by eval

lemma P521_SHA224_TV11_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA224octets P521_SHA224_TV11_d P521_SHA224_TV11_Msg P521_SHA224_TV11_k = Some P521_SHA224_TV11_Sig" 
  by eval

lemma P521_SHA224_TV11_Verify: 
  "SEC1_P521_ECDSA_Verify SHA224octets P521_SHA224_TV11_Q P521_SHA224_TV11_Msg P521_SHA224_TV11_Sig" 
  by eval

subsection‹Test Vector: P521_SHA224_TV12›

subsubsection‹Given Test Vector Values›

definition P521_SHA224_TV12_Msg :: octets where
  "P521_SHA224_TV12_Msg = nat_to_octets_len 0xec0d468447222506b4ead04ea1a17e2aa96eeb3e5f066367975dbaea426104f2111c45e206752896e5fa7594d74ed184493598783cb8079e0e915b638d5c317fa978d9011b44a76b28d752462adf305bde321431f7f34b017c9a35bae8786755a62e746480fa3524d398a6ff5fdc6cec54c07221cce61e46fd0a1af932fa8a33 128"

definition P521_SHA224_TV12_d :: nat where
  "P521_SHA224_TV12_d = 0x05e0d47bf37f83bcc9cd834245c42420b68751ac552f8a4aae8c24b6064ae3d33508ecd2c17ec391558ec79c8440117ad80e5e22770dac7f2017b755255000c853c"

definition P521_SHA224_TV12_Qx :: nat where
  "P521_SHA224_TV12_Qx = 0x1a6effc96a7f23a44bf9988f64e5cfafdae23fa14e4bee530af35d7a4ddf6b80dcd0d937be9dd2db3adcda2f5216fecbce867ee67e7e3773082f255156e31358c2f"

definition P521_SHA224_TV12_Qy :: nat where
  "P521_SHA224_TV12_Qy = 0x1e7760190dfbe07ec2df87067597087de262c1e0a12355456faba91b2e7277050d73b924e14c0e93b8457a8b3e1f4207ce6e754274f88ad75c000d1b2977edc9c1a"

definition P521_SHA224_TV12_Q :: "int point" where
  "P521_SHA224_TV12_Q = Point (int P521_SHA224_TV12_Qx) (int P521_SHA224_TV12_Qy)"

definition P521_SHA224_TV12_k :: nat where
  "P521_SHA224_TV12_k = 0x18afea9a6a408db1e7a7bb1437a3d276f231eacfc57678bfa229d78681cbe4e800e6065332a3128db65d3aa446bb35b517dca26b02e106e1311881a95b0302d15e8"

definition P521_SHA224_TV12_R :: nat where
  "P521_SHA224_TV12_R = 0x01c49b3c1d21f1678bdbe1ac12167e95e06617190bdee1a729c1c649210da19e2e210f6689e1310513bfe2ac6c0f4ee5f324f344b31b18df341eaadb826d07adc9b"

definition P521_SHA224_TV12_S :: nat where
  "P521_SHA224_TV12_S = 0x129d4931ba457443012f6ffecd002f2abc3a4b65a58fee8457917ebcf24b29a1d3055b7fc62939a74ebb0c3582172ee7c3c75e0b2fa2367c6e04df63a7a91d593ad"

definition P521_SHA224_TV12_Sig :: "nat × nat" where
  "P521_SHA224_TV12_Sig = (P521_SHA224_TV12_R, P521_SHA224_TV12_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA224_TV12_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV12_d" 
  by eval

lemma P521_SHA224_TV12_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA224_TV12_Q" 
  by eval

lemma P521_SHA224_TV12_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA224_TV12_d = P521_SHA224_TV12_Q" 
  by eval

lemma P521_SHA224_TV12_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA224_TV12_d P521_SHA224_TV12_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA224_TV12_dQvalidPair' P521_SHA224_TV12_d_valid) 

lemma P521_SHA224_TV12_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV12_k" 
  by eval

lemma P521_SHA224_TV12_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA224octets P521_SHA224_TV12_d P521_SHA224_TV12_Msg P521_SHA224_TV12_k = Some P521_SHA224_TV12_Sig" 
  by eval

lemma P521_SHA224_TV12_Verify: 
  "SEC1_P521_ECDSA_Verify SHA224octets P521_SHA224_TV12_Q P521_SHA224_TV12_Msg P521_SHA224_TV12_Sig" 
  by eval

subsection‹Test Vector: P521_SHA224_TV13›

subsubsection‹Given Test Vector Values›

definition P521_SHA224_TV13_Msg :: octets where
  "P521_SHA224_TV13_Msg = nat_to_octets_len 0xd891da97d2b612fa6483ee7870e0f10fc12a89f9e33d636f587f72e0049f5888782ccde3ea737e2abca41492bac291e20de5b84157a43c5ea900aef761006a4471072ab6ae6d515ffe227695d3ff2341355b8398f72a723ae947f9618237c4b6642a36974860b452c0c6202688bc0814710cbbff4b8e0d1395e8671ae67ada01 128"

definition P521_SHA224_TV13_d :: nat where
  "P521_SHA224_TV13_d = 0x1804ab8f90ff518b58019a0b30c9ed8e00326d42671b71b067e6f815ac6752fa35016bd33455ab51ad4550424034419db8314a91362c28e29a80fbd193670f56ace"

definition P521_SHA224_TV13_Qx :: nat where
  "P521_SHA224_TV13_Qx = 0x0a79529d23a832412825c3c2ad5f121c436af0f29990347ecfa586ce2e57fd3c7e0624d8db1f099c53473dbc2578f85416ad2ac958a162051014fb96bf07f9e1d17"

definition P521_SHA224_TV13_Qy :: nat where
  "P521_SHA224_TV13_Qy = 0x17c0750f26df0c621d2d243c6c99f195f0086947b1bf0f43731555f5d677e2d4a082fb5fe8da87e1592a5fa31777da3299cede5a6f756edf81c85b77853388bb3ab"

definition P521_SHA224_TV13_Q :: "int point" where
  "P521_SHA224_TV13_Q = Point (int P521_SHA224_TV13_Qx) (int P521_SHA224_TV13_Qy)"

definition P521_SHA224_TV13_k :: nat where
  "P521_SHA224_TV13_k = 0x042d7c36fec0415bc875deb0fab0c64548554062e618aee3aa6670ffd68ab579fe620d3a9316357267fd3111c0ed567dca663acd94b646d2ba0771953cd9690ef42"

definition P521_SHA224_TV13_R :: nat where
  "P521_SHA224_TV13_R = 0x0d01dfbef126febbdfa03ef43603fd73bc7d2296dce052216e965fed7bb8cbbc24142bfcddb60c2e0bef185833a225daa0c91a2d9665176d4ad9986da785f4bfcf0"

definition P521_SHA224_TV13_S :: nat where
  "P521_SHA224_TV13_S = 0x16627e2614dbcd371693c10bbf579c90c31a46c8d88adf59912c0c529047b053a7c7715142f64dcf5945dbc69ff5b706c4b0f5448d04dd1f0b5a4c3765148bf253d"

definition P521_SHA224_TV13_Sig :: "nat × nat" where
  "P521_SHA224_TV13_Sig = (P521_SHA224_TV13_R, P521_SHA224_TV13_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA224_TV13_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV13_d" 
  by eval

lemma P521_SHA224_TV13_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA224_TV13_Q" 
  by eval

lemma P521_SHA224_TV13_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA224_TV13_d = P521_SHA224_TV13_Q" 
  by eval

lemma P521_SHA224_TV13_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA224_TV13_d P521_SHA224_TV13_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA224_TV13_dQvalidPair' P521_SHA224_TV13_d_valid) 

lemma P521_SHA224_TV13_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV13_k" 
  by eval

lemma P521_SHA224_TV13_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA224octets P521_SHA224_TV13_d P521_SHA224_TV13_Msg P521_SHA224_TV13_k = Some P521_SHA224_TV13_Sig" 
  by eval

lemma P521_SHA224_TV13_Verify: 
  "SEC1_P521_ECDSA_Verify SHA224octets P521_SHA224_TV13_Q P521_SHA224_TV13_Msg P521_SHA224_TV13_Sig" 
  by eval

subsection‹Test Vector: P521_SHA224_TV14›

subsubsection‹Given Test Vector Values›

definition P521_SHA224_TV14_Msg :: octets where
  "P521_SHA224_TV14_Msg = nat_to_octets_len 0x924e4afc979d1fd1ec8ab17e02b69964a1f025882611d9ba57c772175926944e42c68422d15f9326285538a348f9301e593e02c35a9817b160c05e21003d202473db69df695191be22db05615561951867f8425f88c29ba8997a41a2f96b5cee791307369671543373ea91d5ed9d6a34794d33305db8975b061864e6b0fe775f 128"

definition P521_SHA224_TV14_d :: nat where
  "P521_SHA224_TV14_d = 0x0159bff3a4e42b133e20148950452d99681de6649a56b904ee3358d6dd01fb6c76ea05345cb9ea216e5f5db9ecec201880bdff0ed02ac28a6891c164036c538b8a8"

definition P521_SHA224_TV14_Qx :: nat where
  "P521_SHA224_TV14_Qx = 0x12d7f260e570cf548743d0557077139d65245c7b854ca58c85920ac2b290f2abfeccd3bb4217ee4a29b92513ddce3b5cbf7488fb65180bb74aeb7575f8682337ef5"

definition P521_SHA224_TV14_Qy :: nat where
  "P521_SHA224_TV14_Qy = 0x17560186230c7e8bff0bffce1272afcd37534f317b453b40716436a44e4731a3ec90a8f17c53357bc54e6ff22fc5b4ca892321aa7891252d140ece88e25258b63d5"

definition P521_SHA224_TV14_Q :: "int point" where
  "P521_SHA224_TV14_Q = Point (int P521_SHA224_TV14_Qx) (int P521_SHA224_TV14_Qy)"

definition P521_SHA224_TV14_k :: nat where
  "P521_SHA224_TV14_k = 0x14b8a30f988cefdc0edec59537264edb0b697d8c4f9e8507cf72bc01c761304bd2019da1d67e577b84c1c43dd034b7569f16635a771542b0399737025b8d817e1c3"

definition P521_SHA224_TV14_R :: nat where
  "P521_SHA224_TV14_R = 0x0fc50939ebca4f4daa83e7eaf6907cb08f330c01d6ea497b86becda43dfcad47cb5c48f5eb2cc924228628070bcd144088c449a7873242ba86badf796097dbecd6d"

definition P521_SHA224_TV14_S :: nat where
  "P521_SHA224_TV14_S = 0x0ccb6463c4301ba5c043e47ed508d57dd908fd0d533af89fd3b11e76343a1cf2954ce90b0eb18cbc36acd6d76b3906612d8a0feec6ebed13d88650ed9c708b28a11"

definition P521_SHA224_TV14_Sig :: "nat × nat" where
  "P521_SHA224_TV14_Sig = (P521_SHA224_TV14_R, P521_SHA224_TV14_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA224_TV14_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV14_d" 
  by eval

lemma P521_SHA224_TV14_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA224_TV14_Q" 
  by eval

lemma P521_SHA224_TV14_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA224_TV14_d = P521_SHA224_TV14_Q" 
  by eval

lemma P521_SHA224_TV14_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA224_TV14_d P521_SHA224_TV14_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA224_TV14_dQvalidPair' P521_SHA224_TV14_d_valid) 

lemma P521_SHA224_TV14_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV14_k" 
  by eval

lemma P521_SHA224_TV14_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA224octets P521_SHA224_TV14_d P521_SHA224_TV14_Msg P521_SHA224_TV14_k = Some P521_SHA224_TV14_Sig" 
  by eval

lemma P521_SHA224_TV14_Verify: 
  "SEC1_P521_ECDSA_Verify SHA224octets P521_SHA224_TV14_Q P521_SHA224_TV14_Msg P521_SHA224_TV14_Sig" 
  by eval

subsection‹Test Vector: P521_SHA224_TV15›

subsubsection‹Given Test Vector Values›

definition P521_SHA224_TV15_Msg :: octets where
  "P521_SHA224_TV15_Msg = nat_to_octets_len 0xc64319c8aa1c1ae676630045ae488aedebca19d753704182c4bf3b306b75db98e9be438234233c2f14e3b97c2f55236950629885ac1e0bd015db0f912913ffb6f1361c4cc25c3cd434583b0f7a5a9e1a549aa523614268037973b65eb59c0c16a19a49bfaa13d507b29d5c7a146cd8da2917665100ac9de2d75fa48cb708ac79 128"

definition P521_SHA224_TV15_d :: nat where
  "P521_SHA224_TV15_d = 0x17418dfc0fc3d38f02aa06b7df6afa9e0d08540fc40da2b459c727cff052eb0827bdb3d53f61eb3033eb083c224086e48e3eea7e85e31428ffe517328e253f166ad"

definition P521_SHA224_TV15_Qx :: nat where
  "P521_SHA224_TV15_Qx = 0x00188366b9419a900ab0ed9633426d51e25e8dc03f4f0e7549904243981ec469c8d6d938f6714ee620e63bb0ec536376a73d24d40e58ad9eb44d1e6063f2eb4c51d"

definition P521_SHA224_TV15_Qy :: nat where
  "P521_SHA224_TV15_Qy = 0x09889b9203d52b9243fd515294a674afd6b81df4637ffdddc43a7414741eda78d8aa862c9cbbb618acec55bb9a29aac59616fc804a52a97a9fc4d03254f4469effe"

definition P521_SHA224_TV15_Q :: "int point" where
  "P521_SHA224_TV15_Q = Point (int P521_SHA224_TV15_Qx) (int P521_SHA224_TV15_Qy)"

definition P521_SHA224_TV15_k :: nat where
  "P521_SHA224_TV15_k = 0x1211c8824dcbfa0e1e15a04779c9068aed2431daeac298260795e6a80401f11f6d52d36bcee3cfa36627989c49d11475163aa201d2cd4c5394144a6bb500bbaf02b"

definition P521_SHA224_TV15_R :: nat where
  "P521_SHA224_TV15_R = 0x1d59401b8ac438855d545a699991142685077a409de2418c7ccfe01a4771b3870e76287a9654c209b58a12b0f51e8dc568e33140a6b630324f7ef17caa64bf4c139"

definition P521_SHA224_TV15_S :: nat where
  "P521_SHA224_TV15_S = 0x143af360b7971095b3b50679a13cd49217189eaee4713f4201720175216573c68f7ac6f688bfe6eb940a2d971809bf36c0a77decc553b025ed41935a3898685183b"

definition P521_SHA224_TV15_Sig :: "nat × nat" where
  "P521_SHA224_TV15_Sig = (P521_SHA224_TV15_R, P521_SHA224_TV15_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA224_TV15_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV15_d" 
  by eval

lemma P521_SHA224_TV15_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA224_TV15_Q" 
  by eval

lemma P521_SHA224_TV15_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA224_TV15_d = P521_SHA224_TV15_Q" 
  by eval

lemma P521_SHA224_TV15_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA224_TV15_d P521_SHA224_TV15_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA224_TV15_dQvalidPair' P521_SHA224_TV15_d_valid) 

lemma P521_SHA224_TV15_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA224_TV15_k" 
  by eval

lemma P521_SHA224_TV15_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA224octets P521_SHA224_TV15_d P521_SHA224_TV15_Msg P521_SHA224_TV15_k = Some P521_SHA224_TV15_Sig" 
  by eval

lemma P521_SHA224_TV15_Verify: 
  "SEC1_P521_ECDSA_Verify SHA224octets P521_SHA224_TV15_Q P521_SHA224_TV15_Msg P521_SHA224_TV15_Sig" 
  by eval

section‹NIST Test Vectors for ECDSA: Curve P-521, Hash SHA-256›
  
subsection‹Test Vector: P521_SHA256_TV01›

subsubsection‹Given Test Vector Values›

definition P521_SHA256_TV01_Msg :: octets where
  "P521_SHA256_TV01_Msg = nat_to_octets_len 0x8ab8176b16278db54f84328ae0b75ef8f0cd18afdf40c04ad0927ed0f6d9e47470396c8e87cde7a9be2ffbfe6c9658c88b7de4d582111119c433b2e4a504493f0a1166e3a3ea0d7b93358f4a297d63f65a5e752f94e2ee7f49ebcc742fa3eb03a617d00c574245b77a20033854d82964b2949e2247637239ab00baf4d170d97c 128"

definition P521_SHA256_TV01_d :: nat where
  "P521_SHA256_TV01_d = 0x1e8c05996b85e6f3f875712a09c1b40672b5e7a78d5852de01585c5fb990bf3812c3245534a714389ae9014d677a449efd658254e610da8e6cad33414b9d33e0d7a"

definition P521_SHA256_TV01_Qx :: nat where
  "P521_SHA256_TV01_Qx = 0x07d042ca19408524e68b981f1419351e3b84736c77fe58fee7d11317df2e850d960c7dd10d10ba714c8a609d163502b79d682e8bbecd4f52591d2748533e45a867a"

definition P521_SHA256_TV01_Qy :: nat where
  "P521_SHA256_TV01_Qy = 0x197ac6416111ccf987d290459ebc8ad9ec56e49059c992155539a36a626631f4a2d89164b985154f2dddc0281ee5b5178271f3a76a0914c3fcd1f97be8e8376efb3"

definition P521_SHA256_TV01_Q :: "int point" where
  "P521_SHA256_TV01_Q = Point (int P521_SHA256_TV01_Qx) (int P521_SHA256_TV01_Qy)"

definition P521_SHA256_TV01_k :: nat where
  "P521_SHA256_TV01_k = 0x0dc8daaacddb8fd2ff5c34a5ce183a42261ad3c64dbfc095e58924364dc47ea1c05e2599aae917c2c95f47d6bb37da008af9f55730ddbe4d8ded24f9e8daa46db6a"

definition P521_SHA256_TV01_R :: nat where
  "P521_SHA256_TV01_R = 0x09dd1f2a716843eedec7a6645ac834d4336e7b18e35701f06cae9d6b290d41491424735f3b57e829ad5de055eaeef1778f051c1ee152bf2131a081e53df2a567a8a"

definition P521_SHA256_TV01_S :: nat where
  "P521_SHA256_TV01_S = 0x02148e8428d70a72bc9fa986c38c2c97deda0420f222f9dc99d32c0acba699dc7ba0a2b79ce5999ff61bd0b233c744a893bc105bca5c235423e531612da65d72e62"

definition P521_SHA256_TV01_Sig :: "nat × nat" where
  "P521_SHA256_TV01_Sig = (P521_SHA256_TV01_R, P521_SHA256_TV01_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA256_TV01_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV01_d" 
  by eval

lemma P521_SHA256_TV01_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA256_TV01_Q" 
  by eval

lemma P521_SHA256_TV01_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA256_TV01_d = P521_SHA256_TV01_Q" 
  by eval

lemma P521_SHA256_TV01_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA256_TV01_d P521_SHA256_TV01_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA256_TV01_dQvalidPair' P521_SHA256_TV01_d_valid) 

lemma P521_SHA256_TV01_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV01_k" 
  by eval

lemma P521_SHA256_TV01_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA256octets P521_SHA256_TV01_d P521_SHA256_TV01_Msg P521_SHA256_TV01_k = Some P521_SHA256_TV01_Sig" 
  by eval

lemma P521_SHA256_TV01_Verify: 
  "SEC1_P521_ECDSA_Verify SHA256octets P521_SHA256_TV01_Q P521_SHA256_TV01_Msg P521_SHA256_TV01_Sig" 
  by eval

subsection‹Test Vector: P521_SHA256_TV02›

subsubsection‹Given Test Vector Values›

definition P521_SHA256_TV02_Msg :: octets where
  "P521_SHA256_TV02_Msg = nat_to_octets_len 0xc4bc2cec829036469e55acdd277745034e4e3cc4fcd2f50ec8bd89055c19795a1e051ccf9aa178e12f9beab6a016a7257e391faa536eaa5c969396d4e1ade36795a82ebc709d9422de8497e5b68e7292538d4ccdc6dd66d27a3ece6a2844962b77db073df9489c9710585ba03d53fa430dbc6626dc03b61d53fc180b9af5dea6 128"

definition P521_SHA256_TV02_d :: nat where
  "P521_SHA256_TV02_d = 0x0b65bf33b2f27d52cbfabcadce741e691bf4762089afd37964de1a0deda98331bf8c74020a14b52d44d26e2f6fa7bcddbe83be7db17a0c8a1b376469cf92c6da27c"

definition P521_SHA256_TV02_Qx :: nat where
  "P521_SHA256_TV02_Qx = 0x10038bb9a7aea626de68c14c64243150e72c69e2f8a1ab922bfbdaa6f33d24fb4542c0324357b0dd640bbcd07632ecd253f64ca2bfbfbf3de9b24fffd0568ab82da"

definition P521_SHA256_TV02_Qy :: nat where
  "P521_SHA256_TV02_Qy = 0x0faf867d95308cc36d6f46844a0f535dc70f9768eed011a2464d2f308fa1d8e72c3616aec7e70516908183ffce7fdd36984a15f73efaa3858c2edf16a784d40e6c2"

definition P521_SHA256_TV02_Q :: "int point" where
  "P521_SHA256_TV02_Q = Point (int P521_SHA256_TV02_Qx) (int P521_SHA256_TV02_Qy)"

definition P521_SHA256_TV02_k :: nat where
  "P521_SHA256_TV02_k = 0x14aeb96c57d99677a1f5e4588064215e7e9af4027bfb8f31ff6126dbf341b8e6f719465e4273e91ba32670feca802549808322b7ee108bb20653cf20f93284d365f"

definition P521_SHA256_TV02_R :: nat where
  "P521_SHA256_TV02_R = 0x075ead62edf7d86c5d1bc2443d1aeb5dc034fd999e6ea012cef7499d9d050cd97d262095884e9fc89a42e15bd3dee80fe3c1ba10f4caabc4aabb86347023028b663"

definition P521_SHA256_TV02_S :: nat where
  "P521_SHA256_TV02_S = 0x129a992a6ff66d41948d11fa680f732b1a74315b804c982805190ed9d2fae223f2b149980b9241998cdea0c5672595a8a49d5186a0ef7a46c0a376f925bdda81726"

definition P521_SHA256_TV02_Sig :: "nat × nat" where
  "P521_SHA256_TV02_Sig = (P521_SHA256_TV02_R, P521_SHA256_TV02_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA256_TV02_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV02_d" 
  by eval

lemma P521_SHA256_TV02_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA256_TV02_Q" 
  by eval

lemma P521_SHA256_TV02_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA256_TV02_d = P521_SHA256_TV02_Q" 
  by eval

lemma P521_SHA256_TV02_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA256_TV02_d P521_SHA256_TV02_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA256_TV02_dQvalidPair' P521_SHA256_TV02_d_valid) 

lemma P521_SHA256_TV02_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV02_k" 
  by eval

lemma P521_SHA256_TV02_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA256octets P521_SHA256_TV02_d P521_SHA256_TV02_Msg P521_SHA256_TV02_k = Some P521_SHA256_TV02_Sig" 
  by eval

lemma P521_SHA256_TV02_Verify: 
  "SEC1_P521_ECDSA_Verify SHA256octets P521_SHA256_TV02_Q P521_SHA256_TV02_Msg P521_SHA256_TV02_Sig" 
  by eval

subsection‹Test Vector: P521_SHA256_TV03›

subsubsection‹Given Test Vector Values›

definition P521_SHA256_TV03_Msg :: octets where
  "P521_SHA256_TV03_Msg = nat_to_octets_len 0x1c1b641d0511a0625a4b33e7639d7a057e27f3a7f818e67f593286c8a4c827bb1f3e4f399027e57f18a45403a310c785b50e5a03517c72b45ef8c242a57b162debf2e80c1cf6c7b90237aede5f4ab1fcaf8187be3beb524c223cc0ceff24429eb181a5eea364a748c713214880d976c2cd497fd65ab3854ad0d6c2c1913d3a06 128"

definition P521_SHA256_TV03_d :: nat where
  "P521_SHA256_TV03_d = 0x02c4e660609e99becd61c14d043e8b419a663010cc1d8f9469897d7d0a4f076a619a7214a2a9d07957b028f7d8539ba7430d0b9a7de08beeeae8452d7bb0eac669d"

definition P521_SHA256_TV03_Qx :: nat where
  "P521_SHA256_TV03_Qx = 0x0fb3868238ca840dbb36ecc6cf04f5f773ea0ab8e8b0fdcf779dc4039a8d7146a417504e953c0cb5e7f4e599cc2c168deda8b7f16084b5582f89f2ece4cae5167f7"

definition P521_SHA256_TV03_Qy :: nat where
  "P521_SHA256_TV03_Qy = 0x1f90b5c15eeda48e747cf3ee8183166a49dbfac6161cbd09d29d40a6854f4c495e88a435892a920cdaad20d41985890b648badd4f0a858ffcbd9afdfc23134ede18"

definition P521_SHA256_TV03_Q :: "int point" where
  "P521_SHA256_TV03_Q = Point (int P521_SHA256_TV03_Qx) (int P521_SHA256_TV03_Qy)"

definition P521_SHA256_TV03_k :: nat where
  "P521_SHA256_TV03_k = 0x1f875bbf882cd6dd034a87916c7b3ba54b41b2ea2ce84ebaf4e393fcf7291fee09dec2b5bb8b6490997c9e62f077c34f0947fe14cec99b906dd6bf0b5d301e75ca1"

definition P521_SHA256_TV03_R :: nat where
  "P521_SHA256_TV03_R = 0x07aa70425697736b298233249f5d0cf25c99e640c9ff88035ef1804820e1bfe7d043755f02d7a079494f7fa6dc26740c4e6b7b430c63f29c67bbd3a5c88d2f0e8d1"

definition P521_SHA256_TV03_S :: nat where
  "P521_SHA256_TV03_S = 0x0e0d42e4ff11cf5be37a9fda348514d5097a662f214687cbfb28ff42d635b13029871ca4f464bb1fbce02d5da4d5fb61b2a071844259fc863d136197bec3a61e7c7"

definition P521_SHA256_TV03_Sig :: "nat × nat" where
  "P521_SHA256_TV03_Sig = (P521_SHA256_TV03_R, P521_SHA256_TV03_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA256_TV03_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV03_d" 
  by eval

lemma P521_SHA256_TV03_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA256_TV03_Q" 
  by eval

lemma P521_SHA256_TV03_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA256_TV03_d = P521_SHA256_TV03_Q" 
  by eval

lemma P521_SHA256_TV03_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA256_TV03_d P521_SHA256_TV03_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA256_TV03_dQvalidPair' P521_SHA256_TV03_d_valid) 

lemma P521_SHA256_TV03_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV03_k" 
  by eval

lemma P521_SHA256_TV03_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA256octets P521_SHA256_TV03_d P521_SHA256_TV03_Msg P521_SHA256_TV03_k = Some P521_SHA256_TV03_Sig" 
  by eval

lemma P521_SHA256_TV03_Verify: 
  "SEC1_P521_ECDSA_Verify SHA256octets P521_SHA256_TV03_Q P521_SHA256_TV03_Msg P521_SHA256_TV03_Sig" 
  by eval

subsection‹Test Vector: P521_SHA256_TV04›

subsubsection‹Given Test Vector Values›

definition P521_SHA256_TV04_Msg :: octets where
  "P521_SHA256_TV04_Msg = nat_to_octets_len 0xadb5f069b2b501a3ebb83d4f1808eb07710ac4a7b12532996855a20bcc54b2f76812915f632163c3654ff13d187d007152617cf859200194b59c5e81fc6cc9eb1ceb75d654050f260caa79c265254089270ccd02607fdcf3246119738c496dc3a4bd5d3be15789fc3d29a08d6d921febe2f40aef286d5d4330b07198c7f4588e 128"

definition P521_SHA256_TV04_d :: nat where
  "P521_SHA256_TV04_d = 0x17c3522007a90357ff0bda7d3a36e66df88ca9721fb80e8f63f50255d47ee819068d018f14c6dd7c6ad176f69a4500e6f63caf5cf780531004f85009c69b9c1230c"

definition P521_SHA256_TV04_Qx :: nat where
  "P521_SHA256_TV04_Qx = 0x13a4bea0eed80c66ea973a9d3d4a90b6abbb5dee57d8affaf93390a8783a20982eba644d2e2809f66530adeeee7f9a1da7515447e9ba118999f76f170c375f621f7"

definition P521_SHA256_TV04_Qy :: nat where
  "P521_SHA256_TV04_Qy = 0x12f9dfaee40a75d8442b39b37a5c19ea124b464236e9b9a31bae6780cfd50f7ea4a700154b5ea0feeb64e9b35a1b0e33e46900cca1f34d13bb17e5017769841af27"

definition P521_SHA256_TV04_Q :: "int point" where
  "P521_SHA256_TV04_Q = Point (int P521_SHA256_TV04_Qx) (int P521_SHA256_TV04_Qy)"

definition P521_SHA256_TV04_k :: nat where
  "P521_SHA256_TV04_k = 0x18388a49caeda35859ef02702c1fd45ff26991998bd9d5e189c12c36cdae3f642ddd4a79561bd1d3e1cd9359de8f5c9e1604a312d207a27b08a6033f2741794ced5"

definition P521_SHA256_TV04_R :: nat where
  "P521_SHA256_TV04_R = 0x15c6264795837dfea19f91876455f564f073c5c84a3c9d76e67872ae0447ba0d4850d8721302b25bec7ebfedd2721de140b2f3dead547042b24b0876117e7093cc1"

definition P521_SHA256_TV04_S :: nat where
  "P521_SHA256_TV04_S = 0x060eb74236c189a28ed20bd0822eb22d75f7d97c9043a3c8e3f6d4c90bc8ca02ac4d37c1171c799a1c7dfd2fcbf83406b5e48c051e0fbf0fd937bfe6c3db4e18154"

definition P521_SHA256_TV04_Sig :: "nat × nat" where
  "P521_SHA256_TV04_Sig = (P521_SHA256_TV04_R, P521_SHA256_TV04_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA256_TV04_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV04_d" 
  by eval

lemma P521_SHA256_TV04_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA256_TV04_Q" 
  by eval

lemma P521_SHA256_TV04_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA256_TV04_d = P521_SHA256_TV04_Q" 
  by eval

lemma P521_SHA256_TV04_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA256_TV04_d P521_SHA256_TV04_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA256_TV04_dQvalidPair' P521_SHA256_TV04_d_valid) 

lemma P521_SHA256_TV04_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV04_k" 
  by eval

lemma P521_SHA256_TV04_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA256octets P521_SHA256_TV04_d P521_SHA256_TV04_Msg P521_SHA256_TV04_k = Some P521_SHA256_TV04_Sig" 
  by eval

lemma P521_SHA256_TV04_Verify: 
  "SEC1_P521_ECDSA_Verify SHA256octets P521_SHA256_TV04_Q P521_SHA256_TV04_Msg P521_SHA256_TV04_Sig" 
  by eval

subsection‹Test Vector: P521_SHA256_TV05›

subsubsection‹Given Test Vector Values›

definition P521_SHA256_TV05_Msg :: octets where
  "P521_SHA256_TV05_Msg = nat_to_octets_len 0xf253484d121d1ce8a88def6a3e9e78c47f4025ead6f73285bf90647102645b0c32d4d86742a50b8b7a42d5f6156a6faf588212b7dc72c3ffd13973bdba732b554d8bffc57d04f8167aef21ee941ee6ffb6cce0f49445bd707da8deb35dca650aaf761c3aa66a5ebccddd15aee21293f63061a7f4bfc3787c2cd62c806a1a9985 128"

definition P521_SHA256_TV05_d :: nat where
  "P521_SHA256_TV05_d = 0x0c4dad55871d3bd65b016d143ddd7a195cc868b3048c8bbcb1435622036bdb5e0dec7178ca0138c610238e0365968f6ddd191bbfacc91948088044d9966f652ff25"

definition P521_SHA256_TV05_Qx :: nat where
  "P521_SHA256_TV05_Qx = 0x014858a3b9bd426b678fdcf93fc53d17e7a9e8fe022442aaaba65399d12fd3a6a381958fb0f07ac6088f4e490506ec0f1ab4d0dbd461126f7eb46ff69cfa8bd88af"

definition P521_SHA256_TV05_Qy :: nat where
  "P521_SHA256_TV05_Qy = 0x18c18ce29ecc6d79d26a2de0cd31c4b32e84b5e90f6ba748f86c5afbd89618aceb9079460cbd1a8261ed5476973e61bf1d17ea78b022387443800c9247d21dde550"

definition P521_SHA256_TV05_Q :: "int point" where
  "P521_SHA256_TV05_Q = Point (int P521_SHA256_TV05_Qx) (int P521_SHA256_TV05_Qy)"

definition P521_SHA256_TV05_k :: nat where
  "P521_SHA256_TV05_k = 0x05577108f4187a173e5c29e927a8fc8f5ffd37e184254a6e381ff1018955aec91a35f30085e8cee6a7555c10f9efdce26d62f2b4b52dfdbaeafc3a30983e2d50d5b"

definition P521_SHA256_TV05_R :: nat where
  "P521_SHA256_TV05_R = 0x0344375ae7c804cbe32ced7a20976efae5d9c19eb88b6e24514d1d0cfb728b0f4601098b18b2e98f42b5222dd5237d4d87767007bf5acb185c5526d72047e2cb1a1"

definition P521_SHA256_TV05_S :: nat where
  "P521_SHA256_TV05_S = 0x02de4cfa908c73c1102d6fb7062baf54a056a9517701e036c9c51e09899d60051612d59348945f845dffebec5aa395b2fac7229929033615788777306ccad96d0a3"

definition P521_SHA256_TV05_Sig :: "nat × nat" where
  "P521_SHA256_TV05_Sig = (P521_SHA256_TV05_R, P521_SHA256_TV05_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA256_TV05_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV05_d" 
  by eval

lemma P521_SHA256_TV05_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA256_TV05_Q" 
  by eval

lemma P521_SHA256_TV05_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA256_TV05_d = P521_SHA256_TV05_Q" 
  by eval

lemma P521_SHA256_TV05_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA256_TV05_d P521_SHA256_TV05_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA256_TV05_dQvalidPair' P521_SHA256_TV05_d_valid) 

lemma P521_SHA256_TV05_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV05_k" 
  by eval

lemma P521_SHA256_TV05_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA256octets P521_SHA256_TV05_d P521_SHA256_TV05_Msg P521_SHA256_TV05_k = Some P521_SHA256_TV05_Sig" 
  by eval

lemma P521_SHA256_TV05_Verify: 
  "SEC1_P521_ECDSA_Verify SHA256octets P521_SHA256_TV05_Q P521_SHA256_TV05_Msg P521_SHA256_TV05_Sig" 
  by eval

subsection‹Test Vector: P521_SHA256_TV06›

subsubsection‹Given Test Vector Values›

definition P521_SHA256_TV06_Msg :: octets where
  "P521_SHA256_TV06_Msg = nat_to_octets_len 0x33bab1c369c495db1610965bc0b0546a216e8dd00cd0e602a605d40bc8812bbf1ffa67143f896c436b8f7cf0bed308054f1e1ff77f4d0a13c1e831efbd0e2fcfb3eadab9f755f070ba9aeaceb0a5110f2f8b0c1f7b1aa96a7f2d038a1b72e26400819b1f73d925ea4e34d6acaf59d0a461a34ce5d65c9c937a80e844e323a16d 128"

definition P521_SHA256_TV06_d :: nat where
  "P521_SHA256_TV06_d = 0x03d4749fadcc2008f098de70545a669133c548ce0e32eec1276ff531bcff53533144555728ad8906d17f091cc0514571691107350b6561858e90dbe19633aaf31bf"

definition P521_SHA256_TV06_Qx :: nat where
  "P521_SHA256_TV06_Qx = 0x10fe5986b65f6e65d13c88c4d2aed781a91026904f82129d46779bdadaf6b733c845a934e941ab4a285efdea9c96ecc9dc784d87e4d937b42c337b3a9cb111a9600"

definition P521_SHA256_TV06_Qy :: nat where
  "P521_SHA256_TV06_Qy = 0x077853768a2a4d6f596f57414e57ec60b76d3cd5ece8351cd1f335ebcb8801a3d91fb82c65caaeb5c31eea9918367bb5906863ff3ccaf7a6cee415e0d75c15ac2e0"

definition P521_SHA256_TV06_Q :: "int point" where
  "P521_SHA256_TV06_Q = Point (int P521_SHA256_TV06_Qx) (int P521_SHA256_TV06_Qy)"

definition P521_SHA256_TV06_k :: nat where
  "P521_SHA256_TV06_k = 0x1fbb4de337b09e935a6dc6215ffcfcb85d236cc490585e73251a8b8bac37cfa36c5d1df5f4536d33659be1e7a442529a783452f7efda74a4f661b6a127f9248aaf7"

definition P521_SHA256_TV06_R :: nat where
  "P521_SHA256_TV06_R = 0x09d8f10eeff6178594c89d6e8184f9502117384813243ddf9ccf3c8eac5dc6502c472dfc1487a5caffc569f7dedd14a8ebcb310e9bacdb79fb6655aba026cdf87f2"

definition P521_SHA256_TV06_S :: nat where
  "P521_SHA256_TV06_S = 0x0f74236c7915d638708d17c9f10e39dda358faf9bbb821d8dcda0d151aac143bfb165ad0a23a65cd3de532e32cad928728f5ae1c16f58fc16577f3ca8e36f9e708b"

definition P521_SHA256_TV06_Sig :: "nat × nat" where
  "P521_SHA256_TV06_Sig = (P521_SHA256_TV06_R, P521_SHA256_TV06_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA256_TV06_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV06_d" 
  by eval

lemma P521_SHA256_TV06_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA256_TV06_Q" 
  by eval

lemma P521_SHA256_TV06_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA256_TV06_d = P521_SHA256_TV06_Q" 
  by eval

lemma P521_SHA256_TV06_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA256_TV06_d P521_SHA256_TV06_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA256_TV06_dQvalidPair' P521_SHA256_TV06_d_valid) 

lemma P521_SHA256_TV06_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV06_k" 
  by eval

lemma P521_SHA256_TV06_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA256octets P521_SHA256_TV06_d P521_SHA256_TV06_Msg P521_SHA256_TV06_k = Some P521_SHA256_TV06_Sig" 
  by eval

lemma P521_SHA256_TV06_Verify: 
  "SEC1_P521_ECDSA_Verify SHA256octets P521_SHA256_TV06_Q P521_SHA256_TV06_Msg P521_SHA256_TV06_Sig" 
  by eval

subsection‹Test Vector: P521_SHA256_TV07›

subsubsection‹Given Test Vector Values›

definition P521_SHA256_TV07_Msg :: octets where
  "P521_SHA256_TV07_Msg = nat_to_octets_len 0x08c8b7faaac8e1154042d162dca1df0f66e0001b3c5ecf49b6a4334ce4e8a754a1a8e4daf8ec09cf1e521c96547aed5172ef852e82c03cddd851a9f992183ac5199594f288dbcc53a9bb6128561ff3236a7b4b0dce8eaf7d45e64e782955ee1b690ce6a73ece47dc4409b690de6b7928cbe60c42fc6a5ddf1d729faf1cc3885e 128"

definition P521_SHA256_TV07_d :: nat where
  "P521_SHA256_TV07_d = 0x096a77b591bba65023ba92f8a51029725b555caf6eff129879d28f6400e760439d6e69ce662f6f1aecf3869f7b6057b530a3c6ff8ed9e86d5944f583ee0b3fbb570"

definition P521_SHA256_TV07_Qx :: nat where
  "P521_SHA256_TV07_Qx = 0x0fdf6aed933dba73913142ef8bdcd4b760db8500831cd11d7707ab852a6372c05d112a1e7fbc7b514c42142c7370d9f4129493cd75cc6f2daf83747078f15229db6"

definition P521_SHA256_TV07_Qy :: nat where
  "P521_SHA256_TV07_Qy = 0x0ef91dffb3c43080a59534b95ca585ee87f6145f6a0199b2b82c89f456d8bd8e6ac71c78039c08177184484eb2ebd372f189db3a58fab961a75a18afec1ee32764a"

definition P521_SHA256_TV07_Q :: "int point" where
  "P521_SHA256_TV07_Q = Point (int P521_SHA256_TV07_Qx) (int P521_SHA256_TV07_Qy)"

definition P521_SHA256_TV07_k :: nat where
  "P521_SHA256_TV07_k = 0x13aa7b0471317a2a139c2f90df1c40d75e5a8a830fbaf87030fffdb2ef6f2c93d1310c9ed7fe9d7bcd4fe46537ff2495bc9c4f0aaff11461f5e4bebbfbce9a8740a"

definition P521_SHA256_TV07_R :: nat where
  "P521_SHA256_TV07_R = 0x1c7a21800962c91d4651553633b18612d931bb88bff8b743ed595b4e869437e50f8e84fbf334c99061db123a1c40b73b07e203790561a37df65a660355ba2017d78"

definition P521_SHA256_TV07_S :: nat where
  "P521_SHA256_TV07_S = 0x1301e1782559a38f1ca0eebe9bed0f5c7c33103d506a24f8a688f500ee1fe37f97b6685319279e82e6fe43cfd823ccbc123309974cffa76c4f8d41ec02a3cbc45f1"

definition P521_SHA256_TV07_Sig :: "nat × nat" where
  "P521_SHA256_TV07_Sig = (P521_SHA256_TV07_R, P521_SHA256_TV07_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA256_TV07_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV07_d" 
  by eval

lemma P521_SHA256_TV07_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA256_TV07_Q" 
  by eval

lemma P521_SHA256_TV07_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA256_TV07_d = P521_SHA256_TV07_Q" 
  by eval

lemma P521_SHA256_TV07_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA256_TV07_d P521_SHA256_TV07_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA256_TV07_dQvalidPair' P521_SHA256_TV07_d_valid) 

lemma P521_SHA256_TV07_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV07_k" 
  by eval

lemma P521_SHA256_TV07_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA256octets P521_SHA256_TV07_d P521_SHA256_TV07_Msg P521_SHA256_TV07_k = Some P521_SHA256_TV07_Sig" 
  by eval

lemma P521_SHA256_TV07_Verify: 
  "SEC1_P521_ECDSA_Verify SHA256octets P521_SHA256_TV07_Q P521_SHA256_TV07_Msg P521_SHA256_TV07_Sig" 
  by eval

subsection‹Test Vector: P521_SHA256_TV08›

subsubsection‹Given Test Vector Values›

definition P521_SHA256_TV08_Msg :: octets where
  "P521_SHA256_TV08_Msg = nat_to_octets_len 0xba74eed74282811631bd2069e862381e4e2a1e4e9a357b1c159a9ce69786f864b60fe90eeb32d8b72b099986fc594965a33285f7185b415df58fead7b8b50fc60d073680881d7435609ad1d22fd21e789b6730e232b0d2e888889fb82d6ad0337ab909308676164d4f47df44b21190eca8ba0f94995e60ad9bb02938461eee61 128"

definition P521_SHA256_TV08_d :: nat where
  "P521_SHA256_TV08_d = 0x015152382bfd4f7932a8668026e705e9e73daa8bade21e80ea62cf91bd2448ebc4487b508ca2bdaaf072e3706ba87252d64761c6885a65dcafa64c5573c224ae9e6"

definition P521_SHA256_TV08_Qx :: nat where
  "P521_SHA256_TV08_Qx = 0x00b8c7c0186a77dc6e9addd2018188a6a40c3e2ba396f30bbd9293dba2841d57d60866b37f587432719b544d8bf7eb06d90a8c0dc9c93b0c53d53b2f667077228ca"

definition P521_SHA256_TV08_Qy :: nat where
  "P521_SHA256_TV08_Qy = 0x1dd2e5c73ab908ae34f701689f1cd3cf5186d3a2bc941e208bf3ef970e5e429ee9b154d73286b2e5da423e75b7c7b78c7bdf915da92279db43265a0cdefca51f86a"

definition P521_SHA256_TV08_Q :: "int point" where
  "P521_SHA256_TV08_Q = Point (int P521_SHA256_TV08_Qx) (int P521_SHA256_TV08_Qy)"

definition P521_SHA256_TV08_k :: nat where
  "P521_SHA256_TV08_k = 0x0d03506999f5cc9ec3304072984a20a9c64a22ad9b418495ca904f4bbddc96e76d34672cb52763339d3f3bc5b1701c00a675b972797e3a086314da1a8d338436566"

definition P521_SHA256_TV08_R :: nat where
  "P521_SHA256_TV08_R = 0x085406c0ff5ec91f598bb579ad8714ad718c3e133d5dcc2e67c5d2339c146b69919cac07f3bc2bda218f4c7c8be04855e2ca6fff7fbdc4fc0fda87c8c3081cad4f5"

definition P521_SHA256_TV08_S :: nat where
  "P521_SHA256_TV08_S = 0x1b45f2066e583636215ae135afc202b8bf3f301eccff2e1c0198b9aeddf695fa8179488e7b622fc307f601e2f6551815117cc836bb09ef888f8e64a45d9c84ad30c"

definition P521_SHA256_TV08_Sig :: "nat × nat" where
  "P521_SHA256_TV08_Sig = (P521_SHA256_TV08_R, P521_SHA256_TV08_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA256_TV08_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV08_d" 
  by eval

lemma P521_SHA256_TV08_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA256_TV08_Q" 
  by eval

lemma P521_SHA256_TV08_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA256_TV08_d = P521_SHA256_TV08_Q" 
  by eval

lemma P521_SHA256_TV08_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA256_TV08_d P521_SHA256_TV08_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA256_TV08_dQvalidPair' P521_SHA256_TV08_d_valid) 

lemma P521_SHA256_TV08_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV08_k" 
  by eval

lemma P521_SHA256_TV08_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA256octets P521_SHA256_TV08_d P521_SHA256_TV08_Msg P521_SHA256_TV08_k = Some P521_SHA256_TV08_Sig" 
  by eval

lemma P521_SHA256_TV08_Verify: 
  "SEC1_P521_ECDSA_Verify SHA256octets P521_SHA256_TV08_Q P521_SHA256_TV08_Msg P521_SHA256_TV08_Sig" 
  by eval

subsection‹Test Vector: P521_SHA256_TV09›

subsubsection‹Given Test Vector Values›

definition P521_SHA256_TV09_Msg :: octets where
  "P521_SHA256_TV09_Msg = nat_to_octets_len 0xdc71f171a28bdc30968c39f08f999b88dc04c550e261ecf1124d67f05edeae7e87fe9b8135a96fe2bc3996a4f47213d9d191184a76bd6310e1ee5cb67ea7fc3ef6f641a0ba165198040fa668192b75a4754fc02c224bd4a74aade5a8c814adf151c2bfeda65165a04ef359e39847c84e312afb66d4cd1db50d41ef3fe5f31296 128"

definition P521_SHA256_TV09_d :: nat where
  "P521_SHA256_TV09_d = 0x1750ff0ca0c166560b2034bc5760fe0b3915340bc43216e9de0c1d4a76550e8b2036e8b874230f8d29354aed43e183610f24fd4abd4b0be2f111dae942bd7a121f7"

definition P521_SHA256_TV09_Qx :: nat where
  "P521_SHA256_TV09_Qx = 0x1b4b8947192a7c0166c0e0b2791e217370836283e805f3ee11cfb78445aba3c5bc39fe594e01916617ad59e7c8e740d8f2d07d88905d3f33bd5e51aafd4943c5dc6"

definition P521_SHA256_TV09_Qy :: nat where
  "P521_SHA256_TV09_Qy = 0x1175d117232836c28e717ce2a55e59f4ec550effde30d18e3d99e42c6aa2283c7b3e7f2f6ff1fca605dde78c3a5bffa689347b4c93f51ba59a1787bb7d5e43861dc"

definition P521_SHA256_TV09_Q :: "int point" where
  "P521_SHA256_TV09_Q = Point (int P521_SHA256_TV09_Qx) (int P521_SHA256_TV09_Qy)"

definition P521_SHA256_TV09_k :: nat where
  "P521_SHA256_TV09_k = 0x023645023d6bdf20652cdce1185c4ef225c66d54f18632d99ccf743bf554d04c214c88ce52a4f71ec75c899ad1b3c07c34112ca20b55c217ff1d72c9528e2774ce8"

definition P521_SHA256_TV09_R :: nat where
  "P521_SHA256_TV09_R = 0x1e933f68ce0f8403cb16822b8e0564b1d39a35f27b53e4ae0bcdff3e051759464afbc34998ba7c8a7ee34ef6c1aaa722cffe48356fd0b738058358d4c768b3186c1"

definition P521_SHA256_TV09_S :: nat where
  "P521_SHA256_TV09_S = 0x0a67368a305508ce6d25d29c84f552a4a513998990fef4936244f891a2909c30d5fdc9e8a267ecbf3c597138f4a08f7e92bee57d5420eadd700fee864bf78b2614b"

definition P521_SHA256_TV09_Sig :: "nat × nat" where
  "P521_SHA256_TV09_Sig = (P521_SHA256_TV09_R, P521_SHA256_TV09_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA256_TV09_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV09_d" 
  by eval

lemma P521_SHA256_TV09_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA256_TV09_Q" 
  by eval

lemma P521_SHA256_TV09_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA256_TV09_d = P521_SHA256_TV09_Q" 
  by eval

lemma P521_SHA256_TV09_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA256_TV09_d P521_SHA256_TV09_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA256_TV09_dQvalidPair' P521_SHA256_TV09_d_valid) 

lemma P521_SHA256_TV09_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV09_k" 
  by eval

lemma P521_SHA256_TV09_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA256octets P521_SHA256_TV09_d P521_SHA256_TV09_Msg P521_SHA256_TV09_k = Some P521_SHA256_TV09_Sig" 
  by eval

lemma P521_SHA256_TV09_Verify: 
  "SEC1_P521_ECDSA_Verify SHA256octets P521_SHA256_TV09_Q P521_SHA256_TV09_Msg P521_SHA256_TV09_Sig" 
  by eval

subsection‹Test Vector: P521_SHA256_TV10›

subsubsection‹Given Test Vector Values›

definition P521_SHA256_TV10_Msg :: octets where
  "P521_SHA256_TV10_Msg = nat_to_octets_len 0xb895788d7828aaeace4f6b61a072ffa344d8ea324962ba6dab5efda93f65bf64a0f2ac6d5721d03ee70e2aef21cdba69fd29040199160e3a293b772ffb961ed694a8dc82800dab79367a4809a864e4aff6bc837aaa868e952b771b76591c0bb82249034e3208e593d85973d3fea753a95b16e221b2561644535c0131fe834ae7 128"

definition P521_SHA256_TV10_d :: nat where
  "P521_SHA256_TV10_d = 0x023048bc16e00e58c4a4c7cc62ee80ea57f745bda35715510ed0fc29f62359ff60b0cf85b673383b87a6e1a792d93ab8549281515850fa24d6a2d93a20a2fff3d6e"

definition P521_SHA256_TV10_Qx :: nat where
  "P521_SHA256_TV10_Qx = 0x0ba3dc98326a15999351a2ec6c59e221d7d9e7ee7152a6f71686c9797f3f330d3150123620d547813ba9d7cc6c6d35cc9a087d07dff780e4821e74ad05f3762efd6"

definition P521_SHA256_TV10_Qy :: nat where
  "P521_SHA256_TV10_Qy = 0x18b051af9824b5f614d23ecadd591e38edbfe910ad6cbebc3e8a6bec11ea90691c17deb3bc5f34a4a3acd90b7b10f521f6ee7b3cfbfdc03b72d5a8783a4a77c3e4c"

definition P521_SHA256_TV10_Q :: "int point" where
  "P521_SHA256_TV10_Q = Point (int P521_SHA256_TV10_Qx) (int P521_SHA256_TV10_Qy)"

definition P521_SHA256_TV10_k :: nat where
  "P521_SHA256_TV10_k = 0x06099d2667f06c58798757632d07d8b3efbe9c1323efb0c244be6b12b3b163ba1b7cf5246c98dcc0771665a66696d687af5f28ed664fd87d5093df6427523d4db84"

definition P521_SHA256_TV10_R :: nat where
  "P521_SHA256_TV10_R = 0x10dc80ea853064a2ba5a781f108aca3785c5ec0aa45aa05ba31d4de671170797589e863d54a3a986aadf6f670277f50355713dfb27d4ec7e348f787910b3cd668cd"

definition P521_SHA256_TV10_S :: nat where
  "P521_SHA256_TV10_S = 0x018572bfad4f62e3694d1f2e6ffd432faed2e2b9d7e3611a07138212f1e79e6c394839f7cfae96bc368422630016fb9346681eadc5f9699e7331c3b5fde6d65e4c6"

definition P521_SHA256_TV10_Sig :: "nat × nat" where
  "P521_SHA256_TV10_Sig = (P521_SHA256_TV10_R, P521_SHA256_TV10_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA256_TV10_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV10_d" 
  by eval

lemma P521_SHA256_TV10_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA256_TV10_Q" 
  by eval

lemma P521_SHA256_TV10_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA256_TV10_d = P521_SHA256_TV10_Q" 
  by eval

lemma P521_SHA256_TV10_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA256_TV10_d P521_SHA256_TV10_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA256_TV10_dQvalidPair' P521_SHA256_TV10_d_valid) 

lemma P521_SHA256_TV10_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV10_k" 
  by eval

lemma P521_SHA256_TV10_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA256octets P521_SHA256_TV10_d P521_SHA256_TV10_Msg P521_SHA256_TV10_k = Some P521_SHA256_TV10_Sig" 
  by eval

lemma P521_SHA256_TV10_Verify: 
  "SEC1_P521_ECDSA_Verify SHA256octets P521_SHA256_TV10_Q P521_SHA256_TV10_Msg P521_SHA256_TV10_Sig" 
  by eval

subsection‹Test Vector: P521_SHA256_TV11›

subsubsection‹Given Test Vector Values›

definition P521_SHA256_TV11_Msg :: octets where
  "P521_SHA256_TV11_Msg = nat_to_octets_len 0x2c5bd848c476e34b427cfe5676692e588e1957957db7b5704492bd02104a38216535607f5d092dc40020130c04a3aaf0f1c52409834926d69a05d3f3188187a71d402a10ba34eac8629b4c6359b1095f30f710219298bf06b9f19bfc299981d7e251ca232a0a85338a7e02464731d1b25d4a1f68baf97064516590644820c998 128"

definition P521_SHA256_TV11_d :: nat where
  "P521_SHA256_TV11_d = 0x02b8b866ce4503bb40ffc2c3c990465c72473f901d6ebe6a119ca49fcec8221b3b4fa7ec4e8e9a10dbd90c739065ad6a3a0dd98d1d6f6dcb0720f25a99357a40938"

definition P521_SHA256_TV11_Qx :: nat where
  "P521_SHA256_TV11_Qx = 0x1b8c7a169d5455f16bfe5df1ba5d6ec9c76e4bad9968d4f5f96be5878a7b6f71d74bfac0076dd278bc4630629f3294646f17d6b6c712b0087e2c4d576039cfdc8b9"

definition P521_SHA256_TV11_Qy :: nat where
  "P521_SHA256_TV11_Qy = 0x18faffd5422dfd1b61432fa77b9a288b2b7d546656c0dcca3032179e6f45ee3cf61d6a447fc51731cb54457343a41569fcf78cef42895f4da5efcb14ea1fc065f8d"

definition P521_SHA256_TV11_Q :: "int point" where
  "P521_SHA256_TV11_Q = Point (int P521_SHA256_TV11_Qx) (int P521_SHA256_TV11_Qy)"

definition P521_SHA256_TV11_k :: nat where
  "P521_SHA256_TV11_k = 0x0ac89e813f94042292aa1e77c73773c85cf881a9343b3f50711f13fa17b50f4e5cb04ac5f6fc3106a6ef4c9732016c4e08e301eefac19199459129a41a7589e0628"

definition P521_SHA256_TV11_R :: nat where
  "P521_SHA256_TV11_R = 0x05bc7a253a028ee8b7253979b8d689d41d8df6fae7736341f22e28b6faf0cbbdebbd2ef4d73e56d2021af2c646dc15539a7c1e1c4dc9c7674808bd7968d8a66f947"

definition P521_SHA256_TV11_S :: nat where
  "P521_SHA256_TV11_S = 0x0fd71575837a43a4cf1c47d0485cfd503c2cf36ebcea0fdef946ad29acb7fb2e7c6daf6b4eb741eb211081aed6207d02569f1518988f275ad94c7fd4735cb18a92e"

definition P521_SHA256_TV11_Sig :: "nat × nat" where
  "P521_SHA256_TV11_Sig = (P521_SHA256_TV11_R, P521_SHA256_TV11_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA256_TV11_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV11_d" 
  by eval

lemma P521_SHA256_TV11_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA256_TV11_Q" 
  by eval

lemma P521_SHA256_TV11_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA256_TV11_d = P521_SHA256_TV11_Q" 
  by eval

lemma P521_SHA256_TV11_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA256_TV11_d P521_SHA256_TV11_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA256_TV11_dQvalidPair' P521_SHA256_TV11_d_valid) 

lemma P521_SHA256_TV11_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV11_k" 
  by eval

lemma P521_SHA256_TV11_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA256octets P521_SHA256_TV11_d P521_SHA256_TV11_Msg P521_SHA256_TV11_k = Some P521_SHA256_TV11_Sig" 
  by eval

lemma P521_SHA256_TV11_Verify: 
  "SEC1_P521_ECDSA_Verify SHA256octets P521_SHA256_TV11_Q P521_SHA256_TV11_Msg P521_SHA256_TV11_Sig" 
  by eval

subsection‹Test Vector: P521_SHA256_TV12›

subsubsection‹Given Test Vector Values›

definition P521_SHA256_TV12_Msg :: octets where
  "P521_SHA256_TV12_Msg = nat_to_octets_len 0x65a0b97048067a0c9040acbb5d7f6e2e6ac462e1e0064a8ce5b5bbf8e57059e25a3ef8c80fc9037ae08f63e63f5bdb9378c322ad9b2daf839fad7a75b1027abb6f70f110247da7e971c7c52914e5a4f7761854432fa16b2a521e7bcaee2c735a87cad20c535bf6d04a87340c229bf9af8647eedca9e2dc0b5aa90f7fea3cdc0a 128"

definition P521_SHA256_TV12_d :: nat where
  "P521_SHA256_TV12_d = 0x0a43b32ad7327ec92c0a67279f417c8ada6f40d6282fe79d6dc23b8702147a31162e646291e8df460d39d7cdbdd7b2e7c6c89509b7ed3071b68d4a518ba48e63662"

definition P521_SHA256_TV12_Qx :: nat where
  "P521_SHA256_TV12_Qx = 0x172fb25a3e22c2a88975d7a814f3e02d5bb74cfb0aaa082c5af580019b429fddd8c7f9e09b6938f62e8c31019b25571aaceef3c0d479079db9a9b533ee8e1670abd"

definition P521_SHA256_TV12_Qy :: nat where
  "P521_SHA256_TV12_Qy = 0x0ff5516223b6cc7c711705f15b91db559014e96d3839249c5c849f2aced228a8998177a1e91177abbb24b57a8ea84d944e0c95da860ae0925f1b40c0e1b7c9e0a46"

definition P521_SHA256_TV12_Q :: "int point" where
  "P521_SHA256_TV12_Q = Point (int P521_SHA256_TV12_Qx) (int P521_SHA256_TV12_Qy)"

definition P521_SHA256_TV12_k :: nat where
  "P521_SHA256_TV12_k = 0x0383eda042e06c0297fbd279a2ad40559c5c12ad458f73458eebcc92b308d3c4fcec20a5b59f698e16fa6ea02dba8661b6955f67c052f67b0a56460869f24cfdf7d"

definition P521_SHA256_TV12_R :: nat where
  "P521_SHA256_TV12_R = 0x1b9c35356b9d068f33aa22a61370dae44a6cb030497a34fb52af23c6b684677370268f06bb4433be6795a71de570088aec17ce0c9933d2f76c7edce7f406f62fedd"

definition P521_SHA256_TV12_S :: nat where
  "P521_SHA256_TV12_S = 0x06f07ea453cfa20ad604ba855332f62834657b0b795684d50c1562a675456e37f4dae45f0df47d8e27e47bc9ce9c9cbba1554c5b94b0b17401b73c8d0c0902c6cc4"

definition P521_SHA256_TV12_Sig :: "nat × nat" where
  "P521_SHA256_TV12_Sig = (P521_SHA256_TV12_R, P521_SHA256_TV12_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA256_TV12_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV12_d" 
  by eval

lemma P521_SHA256_TV12_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA256_TV12_Q" 
  by eval

lemma P521_SHA256_TV12_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA256_TV12_d = P521_SHA256_TV12_Q" 
  by eval

lemma P521_SHA256_TV12_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA256_TV12_d P521_SHA256_TV12_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA256_TV12_dQvalidPair' P521_SHA256_TV12_d_valid) 

lemma P521_SHA256_TV12_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV12_k" 
  by eval

lemma P521_SHA256_TV12_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA256octets P521_SHA256_TV12_d P521_SHA256_TV12_Msg P521_SHA256_TV12_k = Some P521_SHA256_TV12_Sig" 
  by eval

lemma P521_SHA256_TV12_Verify: 
  "SEC1_P521_ECDSA_Verify SHA256octets P521_SHA256_TV12_Q P521_SHA256_TV12_Msg P521_SHA256_TV12_Sig" 
  by eval

subsection‹Test Vector: P521_SHA256_TV13›

subsubsection‹Given Test Vector Values›

definition P521_SHA256_TV13_Msg :: octets where
  "P521_SHA256_TV13_Msg = nat_to_octets_len 0xd6e366a87808eea5d39fe77cac4b8c754e865a796062e2ec89f72165cd41fe04c48148068c570e0d29afe9011e7e7a2461f4d9897d8c1fa14b4ff88cab40059d17ab724f4039244e97fcecb07f9ffeec2fb9d6b1896700fe374104a8c44af01a10e93b268d25367bf2bef488b8abcc1ef0e14c3e6e1621b2d58753f21e28b86f 128"

definition P521_SHA256_TV13_d :: nat where
  "P521_SHA256_TV13_d = 0x03c08fdccb089faee91dac3f56f556654a153cebb32f238488d925afd4c7027707118a372f2a2db132516e12ec25f1664953f123ac2ac8f12e0dcbbb61ff40fb721"

definition P521_SHA256_TV13_Qx :: nat where
  "P521_SHA256_TV13_Qx = 0x193301fc0791996ca29e2350723bd9aa0991ddbb4a78348ee72bdcd9ed63ce110ba3496f2ce0331b5c00d4d674c1b70114e17ce44a73c3e16bab14ed1ee924202e4"

definition P521_SHA256_TV13_Qy :: nat where
  "P521_SHA256_TV13_Qy = 0x0aea9b288cfb2933ec0a40efa8e2108774e09b3863b3193d0dac6cc16ccaa5bd5f9ce133aec5cd3b62cbaeec04703e4b61b19572705db38cfaa1907c3d7c785b0cd"

definition P521_SHA256_TV13_Q :: "int point" where
  "P521_SHA256_TV13_Q = Point (int P521_SHA256_TV13_Qx) (int P521_SHA256_TV13_Qy)"

definition P521_SHA256_TV13_k :: nat where
  "P521_SHA256_TV13_k = 0x0d0e90d5ee7b5036655ad5c8f6a112c4b21c9449ca91c5c78421e364a2160bbac4428303657bc11ea69f59fb0fe85a41b8f155a362343094456fd2a39f2a79e4804"

definition P521_SHA256_TV13_R :: nat where
  "P521_SHA256_TV13_R = 0x1a8c23a2965d365a4c2ffd0802ae8b3a69c6b84a1ba77fd8a5f2f61e8ec3a1dcb336f136e2a997252eaa94caf9b5ad6c9ecff5bf33abf547ca84985bb89908a11d7"

definition P521_SHA256_TV13_S :: nat where
  "P521_SHA256_TV13_S = 0x1cc42a2dd97aa42b9df5ea430e0d4cb13106dd6da6e8c9315c96ed7b052db365bbde6960c9a965954a4398c18ea7db9593bbfc3c3b6b3466ff806fccac3de6424ab"

definition P521_SHA256_TV13_Sig :: "nat × nat" where
  "P521_SHA256_TV13_Sig = (P521_SHA256_TV13_R, P521_SHA256_TV13_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA256_TV13_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV13_d" 
  by eval

lemma P521_SHA256_TV13_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA256_TV13_Q" 
  by eval

lemma P521_SHA256_TV13_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA256_TV13_d = P521_SHA256_TV13_Q" 
  by eval

lemma P521_SHA256_TV13_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA256_TV13_d P521_SHA256_TV13_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA256_TV13_dQvalidPair' P521_SHA256_TV13_d_valid) 

lemma P521_SHA256_TV13_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV13_k" 
  by eval

lemma P521_SHA256_TV13_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA256octets P521_SHA256_TV13_d P521_SHA256_TV13_Msg P521_SHA256_TV13_k = Some P521_SHA256_TV13_Sig" 
  by eval

lemma P521_SHA256_TV13_Verify: 
  "SEC1_P521_ECDSA_Verify SHA256octets P521_SHA256_TV13_Q P521_SHA256_TV13_Msg P521_SHA256_TV13_Sig" 
  by eval

subsection‹Test Vector: P521_SHA256_TV14›

subsubsection‹Given Test Vector Values›

definition P521_SHA256_TV14_Msg :: octets where
  "P521_SHA256_TV14_Msg = nat_to_octets_len 0xf99e1d272d0f5fb9c4f986e873d070ec638422bc04b47c715595e2cf1a701cdf88bc6c4b20085b357bad12ccba67cac8a5ca07f31ba432f9154ff1fadefd487a83a9c37e49fb70a2f170e58889cab0552e0a3806ccfa2a60d96e346851d84b7de6d1a4b8cf37567dc161a84f13421e3412457d4bc27f6213453c8519a2d7daa2 128"

definition P521_SHA256_TV14_d :: nat where
  "P521_SHA256_TV14_d = 0x0969b515f356f8bb605ee131e80e8831e340902f3c6257270f7dedb2ba9d876a2ae55b4a17f5d9acd46c1b26366c7e4e4e90a0ee5cff69ed9b278e5b1156a435f7e"

definition P521_SHA256_TV14_Qx :: nat where
  "P521_SHA256_TV14_Qx = 0x0fc7ae62b05ed6c34077cbcbb869629528a1656e2e6d403884e79a21f5f612e91fc83c3a8ac1478d58852f0e8ba120d5855983afd1a719949afa8a21aec407516c3"

definition P521_SHA256_TV14_Qy :: nat where
  "P521_SHA256_TV14_Qy = 0x0aa705da6459a90eaa2c057f2e6614fb72fc730d6fdebe70e968c93dbc9858534768ea2666553cd01db132331441823950a17e8d2345a3cab039c22b21bfe7bd3b9"

definition P521_SHA256_TV14_Q :: "int point" where
  "P521_SHA256_TV14_Q = Point (int P521_SHA256_TV14_Qx) (int P521_SHA256_TV14_Qy)"

definition P521_SHA256_TV14_k :: nat where
  "P521_SHA256_TV14_k = 0x19029260f88e19360b70c11107a92f06faa64524cfbd9f70fecf02bd5a94f390582a7f4c92c5313bb91dc881596768d86f75a0d6f452094adbe11d6643d1a0b2135"

definition P521_SHA256_TV14_R :: nat where
  "P521_SHA256_TV14_R = 0x07f2158e9b9fa995199608263969498923cf918fdc736427c72ce27ce4a3540dce2e8e5e63a8fc7ba46f7fa42480efbf79c6ed39521f6e6ec056079e453e80a89d9"

definition P521_SHA256_TV14_S :: nat where
  "P521_SHA256_TV14_S = 0x08e349eed6f1e28b0dbf0a8aeb1d67e59a95b54a699f083db885f50d702f3c6a4069591afaa5b80b3c75efb1674ebd32c7ead0040d115945f9a52ee3a51806cad45"

definition P521_SHA256_TV14_Sig :: "nat × nat" where
  "P521_SHA256_TV14_Sig = (P521_SHA256_TV14_R, P521_SHA256_TV14_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA256_TV14_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV14_d" 
  by eval

lemma P521_SHA256_TV14_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA256_TV14_Q" 
  by eval

lemma P521_SHA256_TV14_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA256_TV14_d = P521_SHA256_TV14_Q" 
  by eval

lemma P521_SHA256_TV14_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA256_TV14_d P521_SHA256_TV14_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA256_TV14_dQvalidPair' P521_SHA256_TV14_d_valid) 

lemma P521_SHA256_TV14_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV14_k" 
  by eval

lemma P521_SHA256_TV14_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA256octets P521_SHA256_TV14_d P521_SHA256_TV14_Msg P521_SHA256_TV14_k = Some P521_SHA256_TV14_Sig" 
  by eval

lemma P521_SHA256_TV14_Verify: 
  "SEC1_P521_ECDSA_Verify SHA256octets P521_SHA256_TV14_Q P521_SHA256_TV14_Msg P521_SHA256_TV14_Sig" 
  by eval

subsection‹Test Vector: P521_SHA256_TV15›

subsubsection‹Given Test Vector Values›

definition P521_SHA256_TV15_Msg :: octets where
  "P521_SHA256_TV15_Msg = nat_to_octets_len 0x91f1ca8ce6681f4e1f117b918ae787a888798a9df3afc9d0e922f51cdd6e7f7e55da996f7e3615f1d41e4292479859a44fa18a5a006662610f1aaa2884f843c2e73d441753e0ead51dffc366250616c706f07128940dd6312ff3eda6f0e2b4e441b3d74c592b97d9cd910f979d7f39767b379e7f36a7519f2a4a251ef5e8aae1 128"

definition P521_SHA256_TV15_d :: nat where
  "P521_SHA256_TV15_d = 0x013be0bf0cb060dbba02e90e43c6ba6022f201de35160192d33574a67f3f79df969d3ae87850071aac346b5f386fc645ed1977bea2e8446e0c5890784e369124418"

definition P521_SHA256_TV15_Qx :: nat where
  "P521_SHA256_TV15_Qx = 0x167d8b8308259c730931db828a5f69697ec0773a79bdedbaaf15114a4937011c5ae36ab0503957373fee6b1c4650f91a3b0c92c2d604a3559dd2e856a9a84f551d9"

definition P521_SHA256_TV15_Qy :: nat where
  "P521_SHA256_TV15_Qy = 0x19d2c1346aadaa3090b5981f5353243300a4ff0ab961c4ee530f4133fe85e6aab5bad42e747eee0298c2b8051c8be7049109ad3e1b572dda1cac4a03010f99f206e"

definition P521_SHA256_TV15_Q :: "int point" where
  "P521_SHA256_TV15_Q = Point (int P521_SHA256_TV15_Qx) (int P521_SHA256_TV15_Qy)"

definition P521_SHA256_TV15_k :: nat where
  "P521_SHA256_TV15_k = 0x1a363a344996aac9a3ac040066a65856edfb36f10bb687d4821a2e0299b329c6b60e3547dde03bdbd1afa98b0b75d79cf5aac0ef7a3116266cadf3dfbd46f8a4bfc"

definition P521_SHA256_TV15_R :: nat where
  "P521_SHA256_TV15_R = 0x1ff097485faf32ce9e0c557ee064587c12c4834e7f0988cf181d07ba9ee15ae85a8208b61850080fc4bbedbd82536181d43973459f0d696ac5e6b8f2330b179d180"

definition P521_SHA256_TV15_S :: nat where
  "P521_SHA256_TV15_S = 0x0306dc3c382af13c99d44db7a84ed813c8719c6ed3bbe751ead0d487b5a4aa018129862b7d282cce0bc2059a56d7722f4b226f9deb85da12d5b40648bf6ec568128"

definition P521_SHA256_TV15_Sig :: "nat × nat" where
  "P521_SHA256_TV15_Sig = (P521_SHA256_TV15_R, P521_SHA256_TV15_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA256_TV15_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV15_d" 
  by eval

lemma P521_SHA256_TV15_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA256_TV15_Q" 
  by eval

lemma P521_SHA256_TV15_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA256_TV15_d = P521_SHA256_TV15_Q" 
  by eval

lemma P521_SHA256_TV15_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA256_TV15_d P521_SHA256_TV15_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA256_TV15_dQvalidPair' P521_SHA256_TV15_d_valid) 

lemma P521_SHA256_TV15_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA256_TV15_k" 
  by eval

lemma P521_SHA256_TV15_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA256octets P521_SHA256_TV15_d P521_SHA256_TV15_Msg P521_SHA256_TV15_k = Some P521_SHA256_TV15_Sig" 
  by eval

lemma P521_SHA256_TV15_Verify: 
  "SEC1_P521_ECDSA_Verify SHA256octets P521_SHA256_TV15_Q P521_SHA256_TV15_Msg P521_SHA256_TV15_Sig" 
  by eval

section‹NIST Test Vectors for ECDSA: Curve P-521, Hash SHA-384›
  
subsection‹Test Vector: P521_SHA384_TV01›

subsubsection‹Given Test Vector Values›

definition P521_SHA384_TV01_Msg :: octets where
  "P521_SHA384_TV01_Msg = nat_to_octets_len 0xdbc094402c5b559d53168c6f0c550d827499c6fb2186ae2db15b89b4e6f46220386d6f01bebde91b6ceb3ec7b4696e2cbfd14894dd0b7d656d23396ce920044f9ca514bf115cf98ecaa55b950a9e49365c2f3a05be5020e93db92c37437513044973e792af814d0ffad2c8ecc89ae4b35ccb19318f0b988a7d33ec5a4fe85dfe 128"

definition P521_SHA384_TV01_d :: nat where
  "P521_SHA384_TV01_d = 0x095976d387d814e68aeb09abecdbf4228db7232cd3229569ade537f33e07ed0da0abdee84ab057c9a00049f45250e2719d1ecaccf91c0e6fcdd4016b75bdd98a950"

definition P521_SHA384_TV01_Qx :: nat where
  "P521_SHA384_TV01_Qx = 0x13b4ab7bc1ddf7fd74ca6f75ac560c94169f435361e74eba1f8e759ac70ab3af138d8807aca3d8e73b5c2eb787f6dcca2718122bd94f08943a686b115d869d3f406"

definition P521_SHA384_TV01_Qy :: nat where
  "P521_SHA384_TV01_Qy = 0x0f293c1d627b44e7954d0546270665888144a94d437679d074787959d0d944d8223b9d4b5d068b4fbbd1176a004b476810475cd2a200b83eccd226d08b444a71e71"

definition P521_SHA384_TV01_Q :: "int point" where
  "P521_SHA384_TV01_Q = Point (int P521_SHA384_TV01_Qx) (int P521_SHA384_TV01_Qy)"

definition P521_SHA384_TV01_k :: nat where
  "P521_SHA384_TV01_k = 0x0a8d90686bd1104627836afe698effe22c51aa3b651737a940f2b0f9cd72c594575e550adb142e467a3f631f4429514df8296d8f5144df86faa9e3a8f13939ad5b3"

definition P521_SHA384_TV01_R :: nat where
  "P521_SHA384_TV01_R = 0x02128f77df66d16a604ffcd1a515e039d49bf6b91a215b814b2a1c88d32039521fbd142f717817b838450229025670d99c1fd5ab18bd965f093cae7accff0675aae"

definition P521_SHA384_TV01_S :: nat where
  "P521_SHA384_TV01_S = 0x008dc65a243700a84619dce14e44ea8557e36631db1a55de15865497dbfd66e76a7471f78e510c04e613ced332aa563432a1017da8b81c146059ccc7930153103a6"

definition P521_SHA384_TV01_Sig :: "nat × nat" where
  "P521_SHA384_TV01_Sig = (P521_SHA384_TV01_R, P521_SHA384_TV01_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA384_TV01_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV01_d" 
  by eval

lemma P521_SHA384_TV01_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA384_TV01_Q" 
  by eval

lemma P521_SHA384_TV01_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA384_TV01_d = P521_SHA384_TV01_Q" 
  by eval

lemma P521_SHA384_TV01_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA384_TV01_d P521_SHA384_TV01_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA384_TV01_dQvalidPair' P521_SHA384_TV01_d_valid) 

lemma P521_SHA384_TV01_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV01_k" 
  by eval

lemma P521_SHA384_TV01_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA384octets P521_SHA384_TV01_d P521_SHA384_TV01_Msg P521_SHA384_TV01_k = Some P521_SHA384_TV01_Sig" 
  by eval

lemma P521_SHA384_TV01_Verify: 
  "SEC1_P521_ECDSA_Verify SHA384octets P521_SHA384_TV01_Q P521_SHA384_TV01_Msg P521_SHA384_TV01_Sig" 
  by eval

subsection‹Test Vector: P521_SHA384_TV02›

subsubsection‹Given Test Vector Values›

definition P521_SHA384_TV02_Msg :: octets where
  "P521_SHA384_TV02_Msg = nat_to_octets_len 0x114187efd1f6d6c46473fed0c1922987c79be2144439c6f61183caf2045bfb419f8cddc82267d14540624975f27232117729ccfeacccc7ecd5b71473c69d128152931865a60e6a104b67afe5ed443bdbcdc45372f1a85012bbc4614d4c0c534aacd9ab78664dda9b1f1e255878e8ac59e23c56a686f567e4b15c66f0e7c0931e 128"

definition P521_SHA384_TV02_d :: nat where
  "P521_SHA384_TV02_d = 0x04ceb9896da32f2df630580de979515d698fbf1dd96bea889b98fc0efd0751ed35e6bcf75bc5d99172b0960ffd3d8b683fbffd4174b379fbdecd7b138bb9025574b"

definition P521_SHA384_TV02_Qx :: nat where
  "P521_SHA384_TV02_Qx = 0x0e7a3d30d5bd443549d50e9b297aaa87bc80b5c9e94169602d9d43d6d0c490c0bed8cc2170288b106bdbf4c9f1ce53fd699af0b4c64b494b08520e57dc01ab9a8b0"

definition P521_SHA384_TV02_Qy :: nat where
  "P521_SHA384_TV02_Qy = 0x1d81056d37aec8a75d588f6d05977416e6f24ad0117a7f4450036d695612e7bc2771caed80e580314eebc88c8fc51c453f066e752481f212b57165d67f8a44f375a"

definition P521_SHA384_TV02_Q :: "int point" where
  "P521_SHA384_TV02_Q = Point (int P521_SHA384_TV02_Qx) (int P521_SHA384_TV02_Qy)"

definition P521_SHA384_TV02_k :: nat where
  "P521_SHA384_TV02_k = 0x046639c5a3ec15afae5e4a7a418ac760846512d880c359bc2c751b199ce43b10887e861b14127809754dbea47f6cc0140d2817e3f5b9a80ce01abd81f81b748433a"

definition P521_SHA384_TV02_R :: nat where
  "P521_SHA384_TV02_R = 0x0f913de91e19bd8f943d542ae357bacc942a0967abc9be6c06239a379db8cc733fa50013e0b0f088bce9d630262feaa33b30d84f91bcf5ce9976e4e740fcb112f84"

definition P521_SHA384_TV02_S :: nat where
  "P521_SHA384_TV02_S = 0x08a73a5c9c24235e0d9cecaac653f68ce5a6fb186ce67fa058d6ddbbd4d0a8c4d194e571148e8ad6c8882b4e33d2f60fb23dd7d07a1ae60864e8277918f592b3dc6"

definition P521_SHA384_TV02_Sig :: "nat × nat" where
  "P521_SHA384_TV02_Sig = (P521_SHA384_TV02_R, P521_SHA384_TV02_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA384_TV02_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV02_d" 
  by eval

lemma P521_SHA384_TV02_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA384_TV02_Q" 
  by eval

lemma P521_SHA384_TV02_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA384_TV02_d = P521_SHA384_TV02_Q" 
  by eval

lemma P521_SHA384_TV02_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA384_TV02_d P521_SHA384_TV02_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA384_TV02_dQvalidPair' P521_SHA384_TV02_d_valid) 

lemma P521_SHA384_TV02_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV02_k" 
  by eval

lemma P521_SHA384_TV02_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA384octets P521_SHA384_TV02_d P521_SHA384_TV02_Msg P521_SHA384_TV02_k = Some P521_SHA384_TV02_Sig" 
  by eval

lemma P521_SHA384_TV02_Verify: 
  "SEC1_P521_ECDSA_Verify SHA384octets P521_SHA384_TV02_Q P521_SHA384_TV02_Msg P521_SHA384_TV02_Sig" 
  by eval

subsection‹Test Vector: P521_SHA384_TV03›

subsubsection‹Given Test Vector Values›

definition P521_SHA384_TV03_Msg :: octets where
  "P521_SHA384_TV03_Msg = nat_to_octets_len 0x6744b69fc2420fe00f2352399bd58719e4ecdd6d602e2c80f194d607e58b27a0854745bfd6d504de2eb30b04cee0f44af710dd77e2f816ac3ac5692fad2d1d417893bb0edba2707a4c146a486f8728ca696d35cc52e9c7187c82d4bdb92eb954794e5ad15133f6bfea1f025da32ada710a3014cf11095b3ff69a94d087f17753 128"

definition P521_SHA384_TV03_d :: nat where
  "P521_SHA384_TV03_d = 0x00a8db566bd771a9689ea5188c63d586b9c8b576dbe74c06d618576f61365e90b843d00347fdd084fec4ba229fe671ccdd5d9a3afee821a84af9560cd455ed72e8f"

definition P521_SHA384_TV03_Qx :: nat where
  "P521_SHA384_TV03_Qx = 0x04f5b790cbe2984b71d41af5efed6c6893d15e13f31816d55a9c2926a104eee66f1ada83115d1388551218773b8b9d1138e3e3f027bb4392c90c14fd232580b4a11"

definition P521_SHA384_TV03_Qy :: nat where
  "P521_SHA384_TV03_Qy = 0x0660eb160e9bfc8c5619e70e948e238c6fd37739bc1bb657b8e8436e63628f91992be7e63d9a7359623a1340642777b22026feb51116a6c50c54c3589b9bd39b6cb"

definition P521_SHA384_TV03_Q :: "int point" where
  "P521_SHA384_TV03_Q = Point (int P521_SHA384_TV03_Qx) (int P521_SHA384_TV03_Qy)"

definition P521_SHA384_TV03_k :: nat where
  "P521_SHA384_TV03_k = 0x1e7b5e53571a24bd102dd7ad44a4b8d8a4e60e5957bc3c4e5d3c73109f55233f072e572c7892f425ba5e64d3cb7966096bb34a47e26cd5b3e3b44108b310d9f681b"

definition P521_SHA384_TV03_R :: nat where
  "P521_SHA384_TV03_R = 0x1a88bcd7e2bdff6e497d943dde432fb3f855a7177c466319cb53b701230c299db030276269685857d1e3f28110e690f2f529c8d18115eb381f313bc891d92ad278e"

definition P521_SHA384_TV03_S :: nat where
  "P521_SHA384_TV03_S = 0x146f1984ea879274dfd5e86ad92e564a4de081523ddbb1c397b8f9595911ef2e6501bc081584d5340f7aa47e1af036234ac6f27a5ac31f78dd3b0ff1a62693c630d"

definition P521_SHA384_TV03_Sig :: "nat × nat" where
  "P521_SHA384_TV03_Sig = (P521_SHA384_TV03_R, P521_SHA384_TV03_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA384_TV03_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV03_d" 
  by eval

lemma P521_SHA384_TV03_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA384_TV03_Q" 
  by eval

lemma P521_SHA384_TV03_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA384_TV03_d = P521_SHA384_TV03_Q" 
  by eval

lemma P521_SHA384_TV03_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA384_TV03_d P521_SHA384_TV03_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA384_TV03_dQvalidPair' P521_SHA384_TV03_d_valid) 

lemma P521_SHA384_TV03_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV03_k" 
  by eval

lemma P521_SHA384_TV03_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA384octets P521_SHA384_TV03_d P521_SHA384_TV03_Msg P521_SHA384_TV03_k = Some P521_SHA384_TV03_Sig" 
  by eval

lemma P521_SHA384_TV03_Verify: 
  "SEC1_P521_ECDSA_Verify SHA384octets P521_SHA384_TV03_Q P521_SHA384_TV03_Msg P521_SHA384_TV03_Sig" 
  by eval

subsection‹Test Vector: P521_SHA384_TV04›

subsubsection‹Given Test Vector Values›

definition P521_SHA384_TV04_Msg :: octets where
  "P521_SHA384_TV04_Msg = nat_to_octets_len 0x16001f4dcf9e76aa134b12b867f252735144e523e40fba9b4811b07448a24ef4ccf3e81fe9d7f8097ae1d216a51b6eefc83880885e5b14a5eeee025c4232319c4b8bce26807d1b386ad6a964deb3bdca30ee196cfdd717facfad5c77d9b1d05fdd96875e9675e85029ecbf4f94c524624746b7c42870c14a9a1454acf3354474 128"

definition P521_SHA384_TV04_d :: nat where
  "P521_SHA384_TV04_d = 0x1a300b8bf028449344d0e736145d9dd7c4075a783cb749e1ec7988d60440a07021a25a3de74ea5e3d7bd4ab774d8ad6163adae31877ef0b2bd50e26e9e4be8a7b66"

definition P521_SHA384_TV04_Qx :: nat where
  "P521_SHA384_TV04_Qx = 0x05055b9ad726ba8a48219b0ecbfffb89f8428de895b231f676705b7de9f2022d9ff4e0114ebb52dea342f9bf76b2fb060c020e29d92074ebb1fbfe5290a58c8bc10"

definition P521_SHA384_TV04_Qy :: nat where
  "P521_SHA384_TV04_Qy = 0x0415af7f20a6e945315adbf757316bb486c80780a0a3a15b4b9609f126d7341053a2b726ab63cb46feee527b0bf532b32b477e5671aea23d9b3c3e604b9029954b5"

definition P521_SHA384_TV04_Q :: "int point" where
  "P521_SHA384_TV04_Q = Point (int P521_SHA384_TV04_Qx) (int P521_SHA384_TV04_Qy)"

definition P521_SHA384_TV04_k :: nat where
  "P521_SHA384_TV04_k = 0x05a2e92717bb4dab3ee76724d4d9c2d58a32b873e491e36127985f0c9960c610962ca1c4510dba75c98d83beebdc58b1d8678e054640951d11db1bd2d8a4ab8476b"

definition P521_SHA384_TV04_R :: nat where
  "P521_SHA384_TV04_R = 0x104a78ce94f878822daaf00ee527fbdbf6cceb3cbb23a2caa485e4109466de8910252f92379ab292cac8d1eda164f880c0067696e733fc8588a27703a3e1f5b8f1f"

definition P521_SHA384_TV04_S :: nat where
  "P521_SHA384_TV04_S = 0x1ffe23e8ab5a31668a81161a234ea14879771fe9866f8872eb6edb672e0fe91d2bb75c9767a2dfbac7c15c802211236b22ea41ecd055a0b8b311ffc4255f86d5c67"

definition P521_SHA384_TV04_Sig :: "nat × nat" where
  "P521_SHA384_TV04_Sig = (P521_SHA384_TV04_R, P521_SHA384_TV04_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA384_TV04_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV04_d" 
  by eval

lemma P521_SHA384_TV04_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA384_TV04_Q" 
  by eval

lemma P521_SHA384_TV04_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA384_TV04_d = P521_SHA384_TV04_Q" 
  by eval

lemma P521_SHA384_TV04_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA384_TV04_d P521_SHA384_TV04_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA384_TV04_dQvalidPair' P521_SHA384_TV04_d_valid) 

lemma P521_SHA384_TV04_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV04_k" 
  by eval

lemma P521_SHA384_TV04_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA384octets P521_SHA384_TV04_d P521_SHA384_TV04_Msg P521_SHA384_TV04_k = Some P521_SHA384_TV04_Sig" 
  by eval

lemma P521_SHA384_TV04_Verify: 
  "SEC1_P521_ECDSA_Verify SHA384octets P521_SHA384_TV04_Q P521_SHA384_TV04_Msg P521_SHA384_TV04_Sig" 
  by eval

subsection‹Test Vector: P521_SHA384_TV05›

subsubsection‹Given Test Vector Values›

definition P521_SHA384_TV05_Msg :: octets where
  "P521_SHA384_TV05_Msg = nat_to_octets_len 0xa9824a7b810aa16690083a00d422842971baf400c3563baa789c5653fc13416111c0236c67c68e95a13cec0df50324dcc9ae780ce4232607cb57dd9b2c61b382f0fa51fd4e283e2c55ffe272597651659fbd88cd03bfa9652cd54b01a7034c83a602709879e1325c77969bebfd93932ce09a23eae607374602201614ff84b141 128"

definition P521_SHA384_TV05_d :: nat where
  "P521_SHA384_TV05_d = 0x06a253acd79912a74270fc0703ed6507ab20a970f2bc2277f782062092cf0e60ae1ca1bb44dec003169bc25ef6e7123dd04692f77b181a6d7e692e66b09d35a540c"

definition P521_SHA384_TV05_Qx :: nat where
  "P521_SHA384_TV05_Qx = 0x1f15c6b1df156fdd8381cd7446e039435e445f8f36f0247475058da0e371bf72753f6e39f98066bc79370b038c39687ba18e16cb118fe6538b7568c5403c251f6b7"

definition P521_SHA384_TV05_Qy :: nat where
  "P521_SHA384_TV05_Qy = 0x12d2b4f46b854eeae75f1c63f55b76bf0c604d47f870c28a50ecdeb52bba1dd9a0ff12e680804ff864111207652da7dd10b49edf66bb86be00bc06672de91982457"

definition P521_SHA384_TV05_Q :: "int point" where
  "P521_SHA384_TV05_Q = Point (int P521_SHA384_TV05_Qx) (int P521_SHA384_TV05_Qy)"

definition P521_SHA384_TV05_k :: nat where
  "P521_SHA384_TV05_k = 0x165faf3727e42fd61345cfa7b93e55fb4bf583b24bdc14ce635b6c99dbd788012f14da9a210b677c44acdd851e672f1a48188d6b8946c0efeebfe8a597ba0090a2c"

definition P521_SHA384_TV05_R :: nat where
  "P521_SHA384_TV05_R = 0x1ad9463d2759abd568626548578deefdcd8b2d050ce6d9c7ed05feca20167484b86e89bdcc936fd647e0f8aedd7b6add2b8cf13ff6ff013c2b5540c6c56fda97a0c"

definition P521_SHA384_TV05_S :: nat where
  "P521_SHA384_TV05_S = 0x1645a7d0e11015256cfb034adca198695eea6aedd44d9fbf496850ccfed950f43fffd8dbf41e113f2d3837d8a5dd62b2ed580112ff05800b1f73196e5576810e15b"

definition P521_SHA384_TV05_Sig :: "nat × nat" where
  "P521_SHA384_TV05_Sig = (P521_SHA384_TV05_R, P521_SHA384_TV05_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA384_TV05_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV05_d" 
  by eval

lemma P521_SHA384_TV05_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA384_TV05_Q" 
  by eval

lemma P521_SHA384_TV05_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA384_TV05_d = P521_SHA384_TV05_Q" 
  by eval

lemma P521_SHA384_TV05_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA384_TV05_d P521_SHA384_TV05_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA384_TV05_dQvalidPair' P521_SHA384_TV05_d_valid) 

lemma P521_SHA384_TV05_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV05_k" 
  by eval

lemma P521_SHA384_TV05_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA384octets P521_SHA384_TV05_d P521_SHA384_TV05_Msg P521_SHA384_TV05_k = Some P521_SHA384_TV05_Sig" 
  by eval

lemma P521_SHA384_TV05_Verify: 
  "SEC1_P521_ECDSA_Verify SHA384octets P521_SHA384_TV05_Q P521_SHA384_TV05_Msg P521_SHA384_TV05_Sig" 
  by eval

subsection‹Test Vector: P521_SHA384_TV06›

subsubsection‹Given Test Vector Values›

definition P521_SHA384_TV06_Msg :: octets where
  "P521_SHA384_TV06_Msg = nat_to_octets_len 0x90d8bbf714fd2120d2144022bf29520842d9fbd2dc8bb734b3e892ba0285c6a342d6e1e37cc11a62083566e45b039cc65506d20a7d8b51d763d25f0d9eaf3d38601af612c5798a8a2c712d968592b6ed689b88bbab95259ad34da26af9dda80f2f8a02960370bdb7e7595c0a4fffb465d7ad0c4665b5ec0e7d50c6a8238c7f53 128"

definition P521_SHA384_TV06_d :: nat where
  "P521_SHA384_TV06_d = 0x0d5a5d3ddfd2170f9d2653b91967efc8a5157f8720d740dd974e272aab000cc1a4e6c630348754ab923cafb5056fc584b3706628051c557fce67744ee58ba7a56d0"

definition P521_SHA384_TV06_Qx :: nat where
  "P521_SHA384_TV06_Qx = 0x128a4da5fc995678e457ceb3929adee93c280f851abe900fa21f4f809dafad4e33b381e0cd49ce8dd50e2e281cea162bfd60a1d6a1c0ee2228e6a011e171b559ab8"

definition P521_SHA384_TV06_Qy :: nat where
  "P521_SHA384_TV06_Qy = 0x06eb0917cd72256992c49ea527f6bb0315f13d8047794a0f1da1e93737703b1c2a74a00441ef3b47b6a2ff789c49ae32d91cabe7b29247aeec44f6c40a76597a2ca"

definition P521_SHA384_TV06_Q :: "int point" where
  "P521_SHA384_TV06_Q = Point (int P521_SHA384_TV06_Qx) (int P521_SHA384_TV06_Qy)"

definition P521_SHA384_TV06_k :: nat where
  "P521_SHA384_TV06_k = 0x03269983a5c2bcc98e9476f5abf82424566b1f08b17204d29e310ece88f99eb677a537f86fe2529e409cfef2c12929644100099e0de2f27c0f0ac11105a4dca935b"

definition P521_SHA384_TV06_R :: nat where
  "P521_SHA384_TV06_R = 0x1a5257ae1e8187ba954f535b86ff9b8d6a181a3b95c250d090cb4e9c3bfbd03aa64696a76c569728ef67780d6338d70ce46da40b87a3e49bfe154b93930890dfa93"

definition P521_SHA384_TV06_S :: nat where
  "P521_SHA384_TV06_S = 0x05b6ccdfd5c63c7db76d3a0478064a2a376e0e050cb093be795a72a549247c2e4adba9183145c63d46479dbbdcf09986a6f64c09c7e16abc4853f6376c9558b014a"

definition P521_SHA384_TV06_Sig :: "nat × nat" where
  "P521_SHA384_TV06_Sig = (P521_SHA384_TV06_R, P521_SHA384_TV06_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA384_TV06_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV06_d" 
  by eval

lemma P521_SHA384_TV06_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA384_TV06_Q" 
  by eval

lemma P521_SHA384_TV06_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA384_TV06_d = P521_SHA384_TV06_Q" 
  by eval

lemma P521_SHA384_TV06_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA384_TV06_d P521_SHA384_TV06_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA384_TV06_dQvalidPair' P521_SHA384_TV06_d_valid) 

lemma P521_SHA384_TV06_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV06_k" 
  by eval

lemma P521_SHA384_TV06_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA384octets P521_SHA384_TV06_d P521_SHA384_TV06_Msg P521_SHA384_TV06_k = Some P521_SHA384_TV06_Sig" 
  by eval

lemma P521_SHA384_TV06_Verify: 
  "SEC1_P521_ECDSA_Verify SHA384octets P521_SHA384_TV06_Q P521_SHA384_TV06_Msg P521_SHA384_TV06_Sig" 
  by eval

subsection‹Test Vector: P521_SHA384_TV07›

subsubsection‹Given Test Vector Values›

definition P521_SHA384_TV07_Msg :: octets where
  "P521_SHA384_TV07_Msg = nat_to_octets_len 0x09952b1e09995e95bf0022e911c6ab1a463b0a1fdd0eec69117b34af1103c720b57600217de7cd178fef92de5391e550af72a8dcf7badf25b06dd039417f9a7d0f5be88fcd4e9655931d5b605452a667c9d1bae91d3476e7d51cff4108f116a49966fb3a7cff8df1c09734ce5620faf2dccb3dc5d94e7e9ac812da31f6d07a38 128"

definition P521_SHA384_TV07_d :: nat where
  "P521_SHA384_TV07_d = 0x1bcedf920fa148361671b43c64e3186e1937eb1bd4b28cbd84c421472394552889bc05509aa732ef69d732b21b750523fdfd811f36467690fe94e01e64c9d5cbbe9"

definition P521_SHA384_TV07_Qx :: nat where
  "P521_SHA384_TV07_Qx = 0x0d33c151d202a5d4d831348e940b027ee32e4b0b9b48d823a05c67ff3bdaee0189fc6680565f352c062e99968afc643208b4f9c7af185b861658a88c4ad0fcc8ba2"

definition P521_SHA384_TV07_Qy :: nat where
  "P521_SHA384_TV07_Qy = 0x0e4441ddb546468ad8ffa6074f137edfbb81e82e0e7d8f05c4c54598aa996a9cde54cb371f642bfdd4ae7eca5b769696030027129a4183da93567ad142a2dff5183"

definition P521_SHA384_TV07_Q :: "int point" where
  "P521_SHA384_TV07_Q = Point (int P521_SHA384_TV07_Qx) (int P521_SHA384_TV07_Qy)"

definition P521_SHA384_TV07_k :: nat where
  "P521_SHA384_TV07_k = 0x046e619b83aac868b26d0b3cbfab55e630e0b55c461985b5d00f94ff3a5ce90ff412cebf46bbd84550d2031d573ca27d924624428360708c8d8491c29eb01d30f2e"

definition P521_SHA384_TV07_R :: nat where
  "P521_SHA384_TV07_R = 0x08427c0f0ac0263472cd423c0fb554bf3c851b9c775c566ab0f6878717bd57665830767b05b7789c5c0b078195bd943dc737325552d32877ecb04a7c41bd07cd80c"

definition P521_SHA384_TV07_S :: nat where
  "P521_SHA384_TV07_S = 0x10bb6652d6a624c40a7dd06828f15774130d02369ceb1a7d03b553e16e17b7fa5b5401f15885d5e4fc2e55c0c7a1b97871ab02f76386b93a16aa6e7eb65debac6dd"

definition P521_SHA384_TV07_Sig :: "nat × nat" where
  "P521_SHA384_TV07_Sig = (P521_SHA384_TV07_R, P521_SHA384_TV07_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA384_TV07_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV07_d" 
  by eval

lemma P521_SHA384_TV07_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA384_TV07_Q" 
  by eval

lemma P521_SHA384_TV07_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA384_TV07_d = P521_SHA384_TV07_Q" 
  by eval

lemma P521_SHA384_TV07_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA384_TV07_d P521_SHA384_TV07_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA384_TV07_dQvalidPair' P521_SHA384_TV07_d_valid) 

lemma P521_SHA384_TV07_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV07_k" 
  by eval

lemma P521_SHA384_TV07_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA384octets P521_SHA384_TV07_d P521_SHA384_TV07_Msg P521_SHA384_TV07_k = Some P521_SHA384_TV07_Sig" 
  by eval

lemma P521_SHA384_TV07_Verify: 
  "SEC1_P521_ECDSA_Verify SHA384octets P521_SHA384_TV07_Q P521_SHA384_TV07_Msg P521_SHA384_TV07_Sig" 
  by eval

subsection‹Test Vector: P521_SHA384_TV08›

subsubsection‹Given Test Vector Values›

definition P521_SHA384_TV08_Msg :: octets where
  "P521_SHA384_TV08_Msg = nat_to_octets_len 0x0bb0f80cff309c65ff7729c59c517d50fc0ed5be405ef70cb910c3f62c328c90853d4473530b654dda6156e149bc2222a8a7f9be665240e2fbe9d03f78a2356af0bacd1edb84c4801adc8293a8a0bd6123d1cf6ba216aca807a7eb4dca76b493eb6e3dbb69d36f0f00f856222f24d9b93ec34c3b261be2fca0451c00571928e5 128"

definition P521_SHA384_TV08_d :: nat where
  "P521_SHA384_TV08_d = 0x03789e04b3a2a0254ade3380172c150d2fad033885e02ea8bea5b92db3f4adbab190ae423080a1154dfedec694c25eab46ce638be3db4e4cba67bc39f62d6e7db2d"

definition P521_SHA384_TV08_Qx :: nat where
  "P521_SHA384_TV08_Qx = 0x1dbc2cf19627bdccf02432b1761f296275230c150cdde823ce3141ec315d7d05e16b2c29e2a67491078d5316883e933d85b4b10d4f64c477d3c4e0442dc928983a2"

definition P521_SHA384_TV08_Qy :: nat where
  "P521_SHA384_TV08_Qy = 0x07562e720807dd118d3d8b265b3abc61a71fce43e3dce0e7b5ae18b7a4cb01ecc00d39c1f22e150a9a8728997e502144f5b3f6fa9b4cb8a4136212b082ca394e3f6"

definition P521_SHA384_TV08_Q :: "int point" where
  "P521_SHA384_TV08_Q = Point (int P521_SHA384_TV08_Qx) (int P521_SHA384_TV08_Qy)"

definition P521_SHA384_TV08_k :: nat where
  "P521_SHA384_TV08_k = 0x0fbccd8d7804bdd1d1d721b5ec74d4ba37603bc306f9fce2ec241853d8e07334e6b4b12c4ecca0c54bd71193dd7146507933a20737c5f3e15085830fab9b30ca57b"

definition P521_SHA384_TV08_R :: nat where
  "P521_SHA384_TV08_R = 0x181915a3998d8fa214f9715f4ca928d09c36de168dc15c6970a8a062b5cea2dc969b2437ca17b684f78a1fd583aad8e6c762c8f4ab0c91b86a497145e3ca440d307"

definition P521_SHA384_TV08_S :: nat where
  "P521_SHA384_TV08_S = 0x15a6c18c5c77f5470b27d061eafdc26b78561941a3b2ab0f5c81d40899fc053c3d9ed12d7d61e298abbae470009c7b2157731c58d7b16a66fa5abaf5e8a1b8ed394"

definition P521_SHA384_TV08_Sig :: "nat × nat" where
  "P521_SHA384_TV08_Sig = (P521_SHA384_TV08_R, P521_SHA384_TV08_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA384_TV08_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV08_d" 
  by eval

lemma P521_SHA384_TV08_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA384_TV08_Q" 
  by eval

lemma P521_SHA384_TV08_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA384_TV08_d = P521_SHA384_TV08_Q" 
  by eval

lemma P521_SHA384_TV08_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA384_TV08_d P521_SHA384_TV08_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA384_TV08_dQvalidPair' P521_SHA384_TV08_d_valid) 

lemma P521_SHA384_TV08_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV08_k" 
  by eval

lemma P521_SHA384_TV08_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA384octets P521_SHA384_TV08_d P521_SHA384_TV08_Msg P521_SHA384_TV08_k = Some P521_SHA384_TV08_Sig" 
  by eval

lemma P521_SHA384_TV08_Verify: 
  "SEC1_P521_ECDSA_Verify SHA384octets P521_SHA384_TV08_Q P521_SHA384_TV08_Msg P521_SHA384_TV08_Sig" 
  by eval

subsection‹Test Vector: P521_SHA384_TV09›

subsubsection‹Given Test Vector Values›

definition P521_SHA384_TV09_Msg :: octets where
  "P521_SHA384_TV09_Msg = nat_to_octets_len 0x7efacf213382ce30804e78b7256854d759147dba9729c51b2759465715bf2c421034c23dc651c13d6cce95f71fe6a84dfbee5768163ac5789ac0474c5ddf4115684683c5f7c204b33b8bcc0c03ac58f66cef2f53b721fe2fac91ad841126101a88f512a7c2ded38549d9f050d4b7961dda48a1489f026c5d111701762418cfe3 128"

definition P521_SHA384_TV09_d :: nat where
  "P521_SHA384_TV09_d = 0x124700aa9186353e298edefc57bec0c7d0201cca10c1d80dd408d5d71040592b0ac59facdadfa8712445f5977ef8d4854022720c3f02d60e0732dbb2f171fcf1490"

definition P521_SHA384_TV09_Qx :: nat where
  "P521_SHA384_TV09_Qx = 0x0c80fc4cecae5d53348524ddba6a160b735c75b22fdb39af17e2a613d09246e3bb0fd3f2978577f6db5d2118e05c7898024808f8eb8e021d7969cdcf7fc981200bb"

definition P521_SHA384_TV09_Qy :: nat where
  "P521_SHA384_TV09_Qy = 0x1a880c93943fd446d4b3923b574d2221c1bb7b645fb5534dda60e827b497666ff586b77921f7e7f605147947194cffd2fef0678880b89cc0bc7fb74fa96d4b112d7"

definition P521_SHA384_TV09_Q :: "int point" where
  "P521_SHA384_TV09_Q = Point (int P521_SHA384_TV09_Qx) (int P521_SHA384_TV09_Qy)"

definition P521_SHA384_TV09_k :: nat where
  "P521_SHA384_TV09_k = 0x01a05238d595ded5c61d3bf6fde257dbf13095af8a5cb3a2e579e8e4c550fe31d12b71cc2dbcb295e6c4fd0fb8c22d1b741c097cc59d826ced1a8771f09983143c4"

definition P521_SHA384_TV09_R :: nat where
  "P521_SHA384_TV09_R = 0x132762bc81e9922a8d642e3a9d0218affa21fa2331cfcb9e452545c5981c64a8f7e4cc8e68056023b2aa78bead59061d19c7f646c931163a91e544b106b3be8de9e"

definition P521_SHA384_TV09_S :: nat where
  "P521_SHA384_TV09_S = 0x0c3a1b0b000c3169984132add51d611e2cb7069a262a6983d2ae72b459c36e6469509bdb0f473600b8686700b08910779dee9ba83f82e755d4a4ef5f124eb09397f"

definition P521_SHA384_TV09_Sig :: "nat × nat" where
  "P521_SHA384_TV09_Sig = (P521_SHA384_TV09_R, P521_SHA384_TV09_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA384_TV09_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV09_d" 
  by eval

lemma P521_SHA384_TV09_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA384_TV09_Q" 
  by eval

lemma P521_SHA384_TV09_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA384_TV09_d = P521_SHA384_TV09_Q" 
  by eval

lemma P521_SHA384_TV09_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA384_TV09_d P521_SHA384_TV09_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA384_TV09_dQvalidPair' P521_SHA384_TV09_d_valid) 

lemma P521_SHA384_TV09_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV09_k" 
  by eval

lemma P521_SHA384_TV09_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA384octets P521_SHA384_TV09_d P521_SHA384_TV09_Msg P521_SHA384_TV09_k = Some P521_SHA384_TV09_Sig" 
  by eval

lemma P521_SHA384_TV09_Verify: 
  "SEC1_P521_ECDSA_Verify SHA384octets P521_SHA384_TV09_Q P521_SHA384_TV09_Msg P521_SHA384_TV09_Sig" 
  by eval

subsection‹Test Vector: P521_SHA384_TV10›

subsubsection‹Given Test Vector Values›

definition P521_SHA384_TV10_Msg :: octets where
  "P521_SHA384_TV10_Msg = nat_to_octets_len 0x28edff8b9d85f5f58499cc11f492abdfab25e8945975bbaeee910afa2b8fc1295ec61406309ce4e09f4ab4f462959fc2a2786802466eb26d3b01be6919893ae75d0fdc2dc8a82e662550f9fce9627dd364188aaba5c6faa1b2d8a2235adfa5ad0dc140f88a2b2f103f5690e877d07fe8fd30d02d2b2729bd3d8eb5b23a21f54c 128"

definition P521_SHA384_TV10_d :: nat where
  "P521_SHA384_TV10_d = 0x1f532d01af885cb4ad5c329ca5d421c5c021883bd5404c798d617679bb8b094cbb7e15c832fb436325c5302313ce5e496f9513455e7021ffad75777a19b226acfa1"

definition P521_SHA384_TV10_Qx :: nat where
  "P521_SHA384_TV10_Qx = 0x0c0bd76b0027b85bdd879052220da1494d503f6a4bb972105a48ae98e7dda8c2d9fd9336f5646385b961ef68e8464e3a95b00f96614b1a408ceaa2c87b077b6a8fb"

definition P521_SHA384_TV10_Qy :: nat where
  "P521_SHA384_TV10_Qy = 0x17eb7eb5c78db7819af92e8537d110d9f05a5e24f954f4dde21c224d4040f059ec99e051702f390413d2708d18f84d82998c61847475250fb844b20082cbe651a6b"

definition P521_SHA384_TV10_Q :: "int point" where
  "P521_SHA384_TV10_Q = Point (int P521_SHA384_TV10_Qx) (int P521_SHA384_TV10_Qy)"

definition P521_SHA384_TV10_k :: nat where
  "P521_SHA384_TV10_k = 0x14e66853e0f7cd3300ebcae06048532e19cbb95bee140edc1c867ce7310637651445b6dfeb1d99d2e32f2ffb787ebe3fe35032277f185d3dad84f95806924550abe"

definition P521_SHA384_TV10_R :: nat where
  "P521_SHA384_TV10_R = 0x0c5b3a57161098e2e8e16e0a5ae8ecf4a14df14927eea18ed4925d11dc429dda145159323ba970174b194b9b4608a8fa2373b7a825c5e8bd80574e49698285c2c82"

definition P521_SHA384_TV10_S :: nat where
  "P521_SHA384_TV10_S = 0x1a0c038a51796158b42eb5b0dac37aff9ab93b903a47e06ebbdd15946e4bcc9a3b3875b18cf6294c33fc6c3693cef04ed1a43d08951e664c760e2cf3fb4e47490d2"

definition P521_SHA384_TV10_Sig :: "nat × nat" where
  "P521_SHA384_TV10_Sig = (P521_SHA384_TV10_R, P521_SHA384_TV10_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA384_TV10_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV10_d" 
  by eval

lemma P521_SHA384_TV10_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA384_TV10_Q" 
  by eval

lemma P521_SHA384_TV10_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA384_TV10_d = P521_SHA384_TV10_Q" 
  by eval

lemma P521_SHA384_TV10_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA384_TV10_d P521_SHA384_TV10_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA384_TV10_dQvalidPair' P521_SHA384_TV10_d_valid) 

lemma P521_SHA384_TV10_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV10_k" 
  by eval

lemma P521_SHA384_TV10_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA384octets P521_SHA384_TV10_d P521_SHA384_TV10_Msg P521_SHA384_TV10_k = Some P521_SHA384_TV10_Sig" 
  by eval

lemma P521_SHA384_TV10_Verify: 
  "SEC1_P521_ECDSA_Verify SHA384octets P521_SHA384_TV10_Q P521_SHA384_TV10_Msg P521_SHA384_TV10_Sig" 
  by eval

subsection‹Test Vector: P521_SHA384_TV11›

subsubsection‹Given Test Vector Values›

definition P521_SHA384_TV11_Msg :: octets where
  "P521_SHA384_TV11_Msg = nat_to_octets_len 0xbae2a8897c742fd99fbf813351cd009d3f2e18d825ca22e115276484bce8f82f8c7c0c21dd2af208404d8ef45bb5a6c41693912b630897d5246801bf0775aa9bbac8be98cb861d172c3563dc59e78a58ed13c66dea496471b3ad0eeae8995293e4ab97373edc1837ffc95ff1cc0c1e90e64ea8680b2ca5f1e09bf86b99b343b6 128"

definition P521_SHA384_TV11_d :: nat where
  "P521_SHA384_TV11_d = 0x11abf508bca68a85a54bc0659e77efad3c86112c9db04db2883e76144aa446918bb4bb0784b0b6a0e9aa47399fe3de5aaecfd8894a0d130bb0c366c40d9d5050745"

definition P521_SHA384_TV11_Qx :: nat where
  "P521_SHA384_TV11_Qx = 0x05c0ea363a3a12633ea39d564587ebdd3a22a175ef32b9ebfc7311304b19cb3a62b5adc36f6afb6a6f7fabbf810ee89fdb72854fefd613e7798e9b9ff5938ea54c6"

definition P521_SHA384_TV11_Qy :: nat where
  "P521_SHA384_TV11_Qy = 0x0bd06a85e47b885c08124b55a3fcc07ca61647cda6efbfdbd21b24d1ea7a4c7300d46cd798e76063aa979adef6f0698b15e5b7ae8a2ab39ab4f50b2d20614db6317"

definition P521_SHA384_TV11_Q :: "int point" where
  "P521_SHA384_TV11_Q = Point (int P521_SHA384_TV11_Qx) (int P521_SHA384_TV11_Qy)"

definition P521_SHA384_TV11_k :: nat where
  "P521_SHA384_TV11_k = 0x19cadb8c7eb10565aa4567e0709873918720f0e4b42b4817afb0b0547c70cd1100229deae97a276b9c98ea58b01d4839fee86336d749d123b03e8b1a31166acc110"

definition P521_SHA384_TV11_R :: nat where
  "P521_SHA384_TV11_R = 0x0667448a8bbef1c810d40646977dc22f3dfb52a4d80928ded5e976e199cbed02fbd5a08546756ece14548d721a6eb380d0e1a71ad0660dbcac6163c776eedd3e249"

definition P521_SHA384_TV11_S :: nat where
  "P521_SHA384_TV11_S = 0x0ae7f0a238daaddb7fb4a1707fe5132daf653f8e19f732347134c96f1dd798f867c479a4a4609a568a15b61afed70790adbde13ac5f68c468d0230852c1a2c22581"

definition P521_SHA384_TV11_Sig :: "nat × nat" where
  "P521_SHA384_TV11_Sig = (P521_SHA384_TV11_R, P521_SHA384_TV11_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA384_TV11_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV11_d" 
  by eval

lemma P521_SHA384_TV11_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA384_TV11_Q" 
  by eval

lemma P521_SHA384_TV11_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA384_TV11_d = P521_SHA384_TV11_Q" 
  by eval

lemma P521_SHA384_TV11_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA384_TV11_d P521_SHA384_TV11_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA384_TV11_dQvalidPair' P521_SHA384_TV11_d_valid) 

lemma P521_SHA384_TV11_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV11_k" 
  by eval

lemma P521_SHA384_TV11_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA384octets P521_SHA384_TV11_d P521_SHA384_TV11_Msg P521_SHA384_TV11_k = Some P521_SHA384_TV11_Sig" 
  by eval

lemma P521_SHA384_TV11_Verify: 
  "SEC1_P521_ECDSA_Verify SHA384octets P521_SHA384_TV11_Q P521_SHA384_TV11_Msg P521_SHA384_TV11_Sig" 
  by eval

subsection‹Test Vector: P521_SHA384_TV12›

subsubsection‹Given Test Vector Values›

definition P521_SHA384_TV12_Msg :: octets where
  "P521_SHA384_TV12_Msg = nat_to_octets_len 0xd57a26a9593e72bfc87322524639bcaae5f2252d18b99cdaa03b14445b0b8a4dd53928f66a2e4f202fb25b19cad0eb2f1bfda2ab9b0eb668cdcd0fe72f5d9ef2e45e0218590f7ab9d2c9342202610c698bc786cce108a7d4a6730a13e9ea1b470e781f1237d3f84f44abde808516975546bd89075ef9a9732bfd7ee33b6f4399 128"

definition P521_SHA384_TV12_d :: nat where
  "P521_SHA384_TV12_d = 0x18dbf520d58177e4b7a0627674d220137983f486dd2fd3639f19751804e80df0655db6afd829cdf75238de525e1a7a9f048049b593dd64b4b96cc013f970c05ea1f"

definition P521_SHA384_TV12_Qx :: nat where
  "P521_SHA384_TV12_Qx = 0x18b872690c37995be324ddb5c2bd5462841bb062f8e63da248a853de79c3d6bb9a2eb1e6933afda0998ca43491cc807b08ace2d5336a43d0ab50563a2d3d98755f0"

definition P521_SHA384_TV12_Qy :: nat where
  "P521_SHA384_TV12_Qy = 0x002ff31221aa32aa6546f35e8fe5b9361f938362a5e89e77ae130ba8bce3729e912dfac35a2fd21efe84b45b8be2a340850e4b574e1885b35c2afbe196b57c6cf4c"

definition P521_SHA384_TV12_Q :: "int point" where
  "P521_SHA384_TV12_Q = Point (int P521_SHA384_TV12_Qx) (int P521_SHA384_TV12_Qy)"

definition P521_SHA384_TV12_k :: nat where
  "P521_SHA384_TV12_k = 0x098faeb73054639cb2e4442cd68e7b3a13f4b3f397a7b26f303afa40789f8ddd3d918f1ce4f0be53c8cb69c380744e2297d7fc01e2b3daef4ce64dd3a2644234753"

definition P521_SHA384_TV12_R :: nat where
  "P521_SHA384_TV12_R = 0x09c0e7649f814f70a8416cb78bc4601472a363fe97f5c587305778169677860dd97f87b5ab07c3a953bc4615fc34634509d6a25621bdded33ed42446d059509c190"

definition P521_SHA384_TV12_S :: nat where
  "P521_SHA384_TV12_S = 0x120b90e1cfb8a1b5e530df7b17d1128bc051ca4f1a65dd9c9d9d3c59d2f00c7c1e994c52b8671d40294b4d574d2c04475d5bebeacd3a0d3870a54dc7a4805614f40"

definition P521_SHA384_TV12_Sig :: "nat × nat" where
  "P521_SHA384_TV12_Sig = (P521_SHA384_TV12_R, P521_SHA384_TV12_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA384_TV12_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV12_d" 
  by eval

lemma P521_SHA384_TV12_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA384_TV12_Q" 
  by eval

lemma P521_SHA384_TV12_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA384_TV12_d = P521_SHA384_TV12_Q" 
  by eval

lemma P521_SHA384_TV12_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA384_TV12_d P521_SHA384_TV12_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA384_TV12_dQvalidPair' P521_SHA384_TV12_d_valid) 

lemma P521_SHA384_TV12_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV12_k" 
  by eval

lemma P521_SHA384_TV12_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA384octets P521_SHA384_TV12_d P521_SHA384_TV12_Msg P521_SHA384_TV12_k = Some P521_SHA384_TV12_Sig" 
  by eval

lemma P521_SHA384_TV12_Verify: 
  "SEC1_P521_ECDSA_Verify SHA384octets P521_SHA384_TV12_Q P521_SHA384_TV12_Msg P521_SHA384_TV12_Sig" 
  by eval

subsection‹Test Vector: P521_SHA384_TV13›

subsubsection‹Given Test Vector Values›

definition P521_SHA384_TV13_Msg :: octets where
  "P521_SHA384_TV13_Msg = nat_to_octets_len 0x8fdcf5084b12cfc043dd3416b46274e021bbed95d341d3c500c102a5609d3a34de29f8fa9f0adb611a1f47a97ad981f8129d718fc0d6c709eab1a3490db8d550f34eb905b9e00663543afc5bc155e368e0bc919a8b8c9fa42093603537a5614927efa6be819ed42ececbf1a80a61e6e0a7f9b5bc43b9238e62d5df0571fea152 128"

definition P521_SHA384_TV13_d :: nat where
  "P521_SHA384_TV13_d = 0x002764f5696aa813cd55d30948585f86288ae05aeb264ca157cd09e1d09a10515a849b0791b755ccc656a34707be9e52f5762d290a7d2bcd6de52c600ff862eaf4e"

definition P521_SHA384_TV13_Qx :: nat where
  "P521_SHA384_TV13_Qx = 0x127279c88719dc614db387f102e55104ea1c704ac7f57f3bca936f728439b76556730dd7cde2ac1ad0a4c2c2f036ab6f00cf34cb87ea36113571f300713044106d2"

definition P521_SHA384_TV13_Qy :: nat where
  "P521_SHA384_TV13_Qy = 0x134a0786c31f5f2291b83c50fb579ae4c620b95e5a8bdc0c7e1ee6b996c89d764f1b20403e7faa203f397425ada297045dd8ba0e4b155d4900da249e934faab7991"

definition P521_SHA384_TV13_Q :: "int point" where
  "P521_SHA384_TV13_Q = Point (int P521_SHA384_TV13_Qx) (int P521_SHA384_TV13_Qy)"

definition P521_SHA384_TV13_k :: nat where
  "P521_SHA384_TV13_k = 0x08bffb0778cbb06466cecc114b9e89ca243a2b2b5e2597db920bc73a8bbcbe3f57144ad33409ef7faaab430e13f4c42d304d11347360c84972ca20b1539cce3a288"

definition P521_SHA384_TV13_R :: nat where
  "P521_SHA384_TV13_R = 0x1f8f504e64a502e51e7c129517931c3b71f0d8a63b19cfe01ff7c951c6525249608b3ef5d00061d77eb6b3d69581adeaa3732c773bbb9b919c3e7c71fdc09f44d06"

definition P521_SHA384_TV13_S :: nat where
  "P521_SHA384_TV13_S = 0x058044fc64b340604ffd02a5b2918d76fd6fb59ea895feab7aa218e6f1e8c8f226eb9ee345ef8140183a69272582005077b008006aab11597e808d7ff1e8382c924"

definition P521_SHA384_TV13_Sig :: "nat × nat" where
  "P521_SHA384_TV13_Sig = (P521_SHA384_TV13_R, P521_SHA384_TV13_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA384_TV13_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV13_d" 
  by eval

lemma P521_SHA384_TV13_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA384_TV13_Q" 
  by eval

lemma P521_SHA384_TV13_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA384_TV13_d = P521_SHA384_TV13_Q" 
  by eval

lemma P521_SHA384_TV13_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA384_TV13_d P521_SHA384_TV13_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA384_TV13_dQvalidPair' P521_SHA384_TV13_d_valid) 

lemma P521_SHA384_TV13_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV13_k" 
  by eval

lemma P521_SHA384_TV13_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA384octets P521_SHA384_TV13_d P521_SHA384_TV13_Msg P521_SHA384_TV13_k = Some P521_SHA384_TV13_Sig" 
  by eval

lemma P521_SHA384_TV13_Verify: 
  "SEC1_P521_ECDSA_Verify SHA384octets P521_SHA384_TV13_Q P521_SHA384_TV13_Msg P521_SHA384_TV13_Sig" 
  by eval

subsection‹Test Vector: P521_SHA384_TV14›

subsubsection‹Given Test Vector Values›

definition P521_SHA384_TV14_Msg :: octets where
  "P521_SHA384_TV14_Msg = nat_to_octets_len 0x00669f433934992257bed55861df679804107d7fa491672574a7624949c60049b0533383c88d6896c8de860704c3e6a6aefce83efa57c4d57e9ab253da5d15e1f53ab6dce218b592772ab0bc01fee8e63368e85c0639301456fe2d44cd5396a7f2b22761cd03b80eba7883eede8249a2f5db2183bf00550c5c002f45a5e4fb31 128"

definition P521_SHA384_TV14_d :: nat where
  "P521_SHA384_TV14_d = 0x1b0c9acd3eeb618b4b0de4db402206f0f29adc69d7ad324b6db6601b351f723ac8fe949eeacd34228649bf0126276e5aceb0137d00c30dd858aef2d6b6449de2e89"

definition P521_SHA384_TV14_Qx :: nat where
  "P521_SHA384_TV14_Qx = 0x1811c8884486aaa083ddee1c51cb6e861cb830bd5eaa929f72efadbbd1286566ae7e7ba7fde7e02529900d35ee64591652d28798bfc1bed0d192602a9cf5a7d22e3"

definition P521_SHA384_TV14_Qy :: nat where
  "P521_SHA384_TV14_Qy = 0x06d7fc9dd494816cfd29613d4689af67f7d0a2e6fbad5d4d6e0130189172a1ab601c5ca71deaa8bfcb5a190d49da191672ff6fc048e146cb902acec5eae6d87e60a"

definition P521_SHA384_TV14_Q :: "int point" where
  "P521_SHA384_TV14_Q = Point (int P521_SHA384_TV14_Qx) (int P521_SHA384_TV14_Qy)"

definition P521_SHA384_TV14_k :: nat where
  "P521_SHA384_TV14_k = 0x1fdc4f108070af3c66c9ba7b6c1f2603a19ceb4760399df81228cfc7eafde1082b5a0716a3ff82fbe84726f14dd0db3376ca184a78c3c60679bab6cd45f77f9b9ce"

definition P521_SHA384_TV14_R :: nat where
  "P521_SHA384_TV14_R = 0x1ec310339ff056faeb341c4499c43782078b04be1725ae9a6cdcb6011c46d1a4eb3d75c358225e4ec142fd1cd344186f5eb597f7ba559ddfa954824365d5b6edaec"

definition P521_SHA384_TV14_S :: nat where
  "P521_SHA384_TV14_S = 0x005b679a33fdb7e04834f071cd0ac514c04add9f2614ab9bbd9b407b1420fed3f3e02a108e7e279899e43dcf64ae4083c289a87cd7d2103bdc036a95d36800ac7c6"

definition P521_SHA384_TV14_Sig :: "nat × nat" where
  "P521_SHA384_TV14_Sig = (P521_SHA384_TV14_R, P521_SHA384_TV14_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA384_TV14_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV14_d" 
  by eval

lemma P521_SHA384_TV14_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA384_TV14_Q" 
  by eval

lemma P521_SHA384_TV14_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA384_TV14_d = P521_SHA384_TV14_Q" 
  by eval

lemma P521_SHA384_TV14_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA384_TV14_d P521_SHA384_TV14_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA384_TV14_dQvalidPair' P521_SHA384_TV14_d_valid) 

lemma P521_SHA384_TV14_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV14_k" 
  by eval

lemma P521_SHA384_TV14_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA384octets P521_SHA384_TV14_d P521_SHA384_TV14_Msg P521_SHA384_TV14_k = Some P521_SHA384_TV14_Sig" 
  by eval

lemma P521_SHA384_TV14_Verify: 
  "SEC1_P521_ECDSA_Verify SHA384octets P521_SHA384_TV14_Q P521_SHA384_TV14_Msg P521_SHA384_TV14_Sig" 
  by eval

subsection‹Test Vector: P521_SHA384_TV15›

subsubsection‹Given Test Vector Values›

definition P521_SHA384_TV15_Msg :: octets where
  "P521_SHA384_TV15_Msg = nat_to_octets_len 0x4be81dcfab39a64d6f00c0d7fff94dabdf3473dc49f0e12900df328d6584b854fbaebaf3194c433e9e21743342e2dd056b445c8aa7d30a38504b366a8fa889dc8ecec35b3130070787e7bf0f22fab5bea54a07d3a75368605397ba74dbf2923ef20c37a0d9c64caebcc93157456b57b98d4becb13fecb7cc7f3740a6057af287 128"

definition P521_SHA384_TV15_d :: nat where
  "P521_SHA384_TV15_d = 0x181e1037bbec7ca2f271343e5f6e9125162c8a8a46ae8baa7ca7296602ae9d56c994b3b94d359f2b3b3a01deb7a123f07d9e0c2e729d37cc5abdec0f5281931308a"

definition P521_SHA384_TV15_Qx :: nat where
  "P521_SHA384_TV15_Qx = 0x0cfa5a8a3f15eb8c419095673f1d0bd63b396ff9813c18dfe5aa31f40b50b82481f9ed2edd47ae5ea6a48ea01f7e0ad0000edf7b66f8909ee94f141d5a07efe315c"

definition P521_SHA384_TV15_Qy :: nat where
  "P521_SHA384_TV15_Qy = 0x18af728f7318b96d57f19c1104415c8d5989565465e429bc30cf65ced12a1c5856ac86fca02388bc151cf89959a4f048597a9e728f3034aa39259b59870946187bf"

definition P521_SHA384_TV15_Q :: "int point" where
  "P521_SHA384_TV15_Q = Point (int P521_SHA384_TV15_Qx) (int P521_SHA384_TV15_Qy)"

definition P521_SHA384_TV15_k :: nat where
  "P521_SHA384_TV15_k = 0x09078beaba465ba7a8b3624e644ac1e97c654533a58ac755e90bd606e2214f11a48cb51f9007865a0f569d967ea0370801421846a89f3d09eb0a481289270919f14"

definition P521_SHA384_TV15_R :: nat where
  "P521_SHA384_TV15_R = 0x19cf91a38cc20b9269e7467857b1fc7eabb8cea915a3135f727d471e5bfcfb66d321fabe283a2cf38d4c5a6ecb6e8cbee1030474373bb87fcdfcc95cf857a8d25d0"

definition P521_SHA384_TV15_S :: nat where
  "P521_SHA384_TV15_S = 0x1cf9acd9449c57589c950f287842f9e2487c5610955b2b5035f6aacfd2402f511998a1a942b39c307fc2bcab2c8d0dae94b5547ddccfb1012ca985b3edf42bbba8b"

definition P521_SHA384_TV15_Sig :: "nat × nat" where
  "P521_SHA384_TV15_Sig = (P521_SHA384_TV15_R, P521_SHA384_TV15_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA384_TV15_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV15_d" 
  by eval

lemma P521_SHA384_TV15_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA384_TV15_Q" 
  by eval

lemma P521_SHA384_TV15_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA384_TV15_d = P521_SHA384_TV15_Q" 
  by eval

lemma P521_SHA384_TV15_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA384_TV15_d P521_SHA384_TV15_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA384_TV15_dQvalidPair' P521_SHA384_TV15_d_valid) 

lemma P521_SHA384_TV15_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA384_TV15_k" 
  by eval

lemma P521_SHA384_TV15_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA384octets P521_SHA384_TV15_d P521_SHA384_TV15_Msg P521_SHA384_TV15_k = Some P521_SHA384_TV15_Sig" 
  by eval

lemma P521_SHA384_TV15_Verify: 
  "SEC1_P521_ECDSA_Verify SHA384octets P521_SHA384_TV15_Q P521_SHA384_TV15_Msg P521_SHA384_TV15_Sig" 
  by eval

section‹NIST Test Vectors for ECDSA: Curve P-521, Hash SHA-512›
  
subsection‹Test Vector: P521_SHA512_TV01›

subsubsection‹Given Test Vector Values›

definition P521_SHA512_TV01_Msg :: octets where
  "P521_SHA512_TV01_Msg = nat_to_octets_len 0x9ecd500c60e701404922e58ab20cc002651fdee7cbc9336adda33e4c1088fab1964ecb7904dc6856865d6c8e15041ccf2d5ac302e99d346ff2f686531d25521678d4fd3f76bbf2c893d246cb4d7693792fe18172108146853103a51f824acc621cb7311d2463c3361ea707254f2b052bc22cb8012873dcbb95bf1a5cc53ab89f 128"

definition P521_SHA512_TV01_d :: nat where
  "P521_SHA512_TV01_d = 0x0f749d32704bc533ca82cef0acf103d8f4fba67f08d2678e515ed7db886267ffaf02fab0080dca2359b72f574ccc29a0f218c8655c0cccf9fee6c5e567aa14cb926"

definition P521_SHA512_TV01_Qx :: nat where
  "P521_SHA512_TV01_Qx = 0x061387fd6b95914e885f912edfbb5fb274655027f216c4091ca83e19336740fd81aedfe047f51b42bdf68161121013e0d55b117a14e4303f926c8debb77a7fdaad1"

definition P521_SHA512_TV01_Qy :: nat where
  "P521_SHA512_TV01_Qy = 0x0e7d0c75c38626e895ca21526b9f9fdf84dcecb93f2b233390550d2b1463b7ee3f58df7346435ff0434199583c97c665a97f12f706f2357da4b40288def888e59e6"

definition P521_SHA512_TV01_Q :: "int point" where
  "P521_SHA512_TV01_Q = Point (int P521_SHA512_TV01_Qx) (int P521_SHA512_TV01_Qy)"

definition P521_SHA512_TV01_k :: nat where
  "P521_SHA512_TV01_k = 0x03af5ab6caa29a6de86a5bab9aa83c3b16a17ffcd52b5c60c769be3053cdddeac60812d12fecf46cfe1f3db9ac9dcf881fcec3f0aa733d4ecbb83c7593e864c6df1"

definition P521_SHA512_TV01_R :: nat where
  "P521_SHA512_TV01_R = 0x04de826ea704ad10bc0f7538af8a3843f284f55c8b946af9235af5af74f2b76e099e4bc72fd79d28a380f8d4b4c919ac290d248c37983ba05aea42e2dd79fdd33e8"

definition P521_SHA512_TV01_S :: nat where
  "P521_SHA512_TV01_S = 0x087488c859a96fea266ea13bf6d114c429b163be97a57559086edb64aed4a18594b46fb9efc7fd25d8b2de8f09ca0587f54bd287299f47b2ff124aac566e8ee3b43"

definition P521_SHA512_TV01_Sig :: "nat × nat" where
  "P521_SHA512_TV01_Sig = (P521_SHA512_TV01_R, P521_SHA512_TV01_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA512_TV01_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV01_d" 
  by eval

lemma P521_SHA512_TV01_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA512_TV01_Q" 
  by eval

lemma P521_SHA512_TV01_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA512_TV01_d = P521_SHA512_TV01_Q" 
  by eval

lemma P521_SHA512_TV01_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA512_TV01_d P521_SHA512_TV01_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA512_TV01_dQvalidPair' P521_SHA512_TV01_d_valid) 

lemma P521_SHA512_TV01_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV01_k" 
  by eval

lemma P521_SHA512_TV01_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA512octets P521_SHA512_TV01_d P521_SHA512_TV01_Msg P521_SHA512_TV01_k = Some P521_SHA512_TV01_Sig" 
  by eval

lemma P521_SHA512_TV01_Verify: 
  "SEC1_P521_ECDSA_Verify SHA512octets P521_SHA512_TV01_Q P521_SHA512_TV01_Msg P521_SHA512_TV01_Sig" 
  by eval

subsection‹Test Vector: P521_SHA512_TV02›

subsubsection‹Given Test Vector Values›

definition P521_SHA512_TV02_Msg :: octets where
  "P521_SHA512_TV02_Msg = nat_to_octets_len 0xb3c63e5f5a21c4bfe3dbc644354d9a949186d6a9e1dd873828782aa6a0f1df2f64114a430b1c13fe8a2e09099e1ed05ef70de698161039ded73bcb50b312673bb073f8a792ac140a78a8b7f3586dffb1fc8be4f54516d57418ccc9945025ce3acf1eb84f69ceee5e9bd10c18c251dbc481562cd3aae54b54ab618cb1eeda33cf 128"

definition P521_SHA512_TV02_d :: nat where
  "P521_SHA512_TV02_d = 0x1a4d2623a7d59c55f408331ba8d1523b94d6bf8ac83375ceb57a2b395a5bcf977cfc16234d4a97d6f6ee25a99aa5bff15ff535891bcb7ae849a583e01ac49e0e9b6"

definition P521_SHA512_TV02_Qx :: nat where
  "P521_SHA512_TV02_Qx = 0x04d5c8afee038984d2ea96681ec0dccb6b52dfa4ee2e2a77a23c8cf43ef19905a34d6f5d8c5cf0981ed804d89d175b17d1a63522ceb1e785c0f5a1d2f3d15e51352"

definition P521_SHA512_TV02_Qy :: nat where
  "P521_SHA512_TV02_Qy = 0x014368b8e746807b2b68f3615cd78d761a464ddd7918fc8df51d225962fdf1e3dc243e265100ff0ec133359e332e44dd49afd8e5f38fe86133573432d33c02fa0a3"

definition P521_SHA512_TV02_Q :: "int point" where
  "P521_SHA512_TV02_Q = Point (int P521_SHA512_TV02_Qx) (int P521_SHA512_TV02_Qy)"

definition P521_SHA512_TV02_k :: nat where
  "P521_SHA512_TV02_k = 0x0bc2c0f37155859303de6fa539a39714e195c37c6ea826e224c8218584ae09cd0d1cc14d94d93f2d83c96e4ef68517fdb3f383da5404e5a426bfc5d424e253c181b"

definition P521_SHA512_TV02_R :: nat where
  "P521_SHA512_TV02_R = 0x1a3c4a6386c4fb614fba2cb9e74201e1aaa0001aa931a2a939c92e04b8344535a20f53c6e3c69c75c2e5d2fe3549ed27e6713cb0f4a9a94f6189eb33bff7d453fce"

definition P521_SHA512_TV02_S :: nat where
  "P521_SHA512_TV02_S = 0x16a997f81aa0bea2e1469c8c1dab7df02a8b2086ba482c43af04f2174831f2b1761658795adfbdd44190a9b06fe10e578987369f3a2eced147cff89d8c2818f7471"

definition P521_SHA512_TV02_Sig :: "nat × nat" where
  "P521_SHA512_TV02_Sig = (P521_SHA512_TV02_R, P521_SHA512_TV02_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA512_TV02_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV02_d" 
  by eval

lemma P521_SHA512_TV02_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA512_TV02_Q" 
  by eval

lemma P521_SHA512_TV02_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA512_TV02_d = P521_SHA512_TV02_Q" 
  by eval

lemma P521_SHA512_TV02_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA512_TV02_d P521_SHA512_TV02_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA512_TV02_dQvalidPair' P521_SHA512_TV02_d_valid) 

lemma P521_SHA512_TV02_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV02_k" 
  by eval

lemma P521_SHA512_TV02_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA512octets P521_SHA512_TV02_d P521_SHA512_TV02_Msg P521_SHA512_TV02_k = Some P521_SHA512_TV02_Sig" 
  by eval

lemma P521_SHA512_TV02_Verify: 
  "SEC1_P521_ECDSA_Verify SHA512octets P521_SHA512_TV02_Q P521_SHA512_TV02_Msg P521_SHA512_TV02_Sig" 
  by eval

subsection‹Test Vector: P521_SHA512_TV03›

subsubsection‹Given Test Vector Values›

definition P521_SHA512_TV03_Msg :: octets where
  "P521_SHA512_TV03_Msg = nat_to_octets_len 0x6e0f96d56505ffd2d005d5677dbf926345f0ff0a5da456bbcbcfdc2d33c8d878b0bc8511401c73168d161c23a88b04d7a9629a7a6fbcff241071b0d212248fcc2c94fa5c086909adb8f4b9772b4293b4acf5215ea2fc72f8cec57b5a13792d7859b6d40348fc3ba3f5e7062a19075a9edb713ddcd391aefc90f46bbd81e2557b 128"

definition P521_SHA512_TV03_d :: nat where
  "P521_SHA512_TV03_d = 0x14787f95fb1057a2f3867b8407e54abb91740c097dac5024be92d5d65666bb16e4879f3d3904d6eab269cf5e7b632ab3c5f342108d1d4230c30165fba3a1bf1c66f"

definition P521_SHA512_TV03_Qx :: nat where
  "P521_SHA512_TV03_Qx = 0x0c2d540a7557f4530de35bbd94da8a6defbff783f54a65292f8f76341c996cea38795805a1b97174a9147a8644282e0d7040a6f83423ef2a0453248156393a1782e"

definition P521_SHA512_TV03_Qy :: nat where
  "P521_SHA512_TV03_Qy = 0x119f746c5df8cec24e4849ac1870d0d8594c799d2ceb6c3bdf891dfbd2242e7ea24d6aec3166214734acc4cbf4da8f71e2429c5c187b2b3a048527c861f58a9b97f"

definition P521_SHA512_TV03_Q :: "int point" where
  "P521_SHA512_TV03_Q = Point (int P521_SHA512_TV03_Qx) (int P521_SHA512_TV03_Qy)"

definition P521_SHA512_TV03_k :: nat where
  "P521_SHA512_TV03_k = 0x186cd803e6e0c9925022e41cb68671adba3ead5548c2b1cd09348ab19612b7af3820fd14da5fe1d7b550ed1a3c8d2f30592cd7745a3c09ee7b5dcfa9ed31bdd0f1f"

definition P521_SHA512_TV03_R :: nat where
  "P521_SHA512_TV03_R = 0x10ed3ab6d07a15dc3376494501c27ce5f78c8a2b30cc809d3f9c3bf1aef437e590ef66abae4e49065ead1af5f752ec145acfa98329f17bca9991a199579c41f9229"

definition P521_SHA512_TV03_S :: nat where
  "P521_SHA512_TV03_S = 0x08c3457fe1f93d635bb52df9218bf3b49a7a345b8a8a988ac0a254340546752cddf02e6ce47eee58ea398fdc9130e55a4c09f5ae548c715f5bcd539f07a34034d78"

definition P521_SHA512_TV03_Sig :: "nat × nat" where
  "P521_SHA512_TV03_Sig = (P521_SHA512_TV03_R, P521_SHA512_TV03_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA512_TV03_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV03_d" 
  by eval

lemma P521_SHA512_TV03_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA512_TV03_Q" 
  by eval

lemma P521_SHA512_TV03_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA512_TV03_d = P521_SHA512_TV03_Q" 
  by eval

lemma P521_SHA512_TV03_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA512_TV03_d P521_SHA512_TV03_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA512_TV03_dQvalidPair' P521_SHA512_TV03_d_valid) 

lemma P521_SHA512_TV03_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV03_k" 
  by eval

lemma P521_SHA512_TV03_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA512octets P521_SHA512_TV03_d P521_SHA512_TV03_Msg P521_SHA512_TV03_k = Some P521_SHA512_TV03_Sig" 
  by eval

lemma P521_SHA512_TV03_Verify: 
  "SEC1_P521_ECDSA_Verify SHA512octets P521_SHA512_TV03_Q P521_SHA512_TV03_Msg P521_SHA512_TV03_Sig" 
  by eval

subsection‹Test Vector: P521_SHA512_TV04›

subsubsection‹Given Test Vector Values›

definition P521_SHA512_TV04_Msg :: octets where
  "P521_SHA512_TV04_Msg = nat_to_octets_len 0x3f12ab17af3c3680aad22196337cedb0a9dba22387a7c555b46e84176a6f8418004552386ada4deec59fdabb0d25e1c6668a96f100b352f8dabd24b2262bd2a3d0f825602d54150bdc4bcbd5b8e0ca52bc8d2c70ff2af9b03e20730d6bd9ec1d091a3e5c877259bcff4fd2c17a12bfc4b08117ec39fe4762be128d0883a37e9d 128"

definition P521_SHA512_TV04_d :: nat where
  "P521_SHA512_TV04_d = 0x15807c101099c8d1d3f24b212af2c0ce525432d7779262eed0709275de9a1d8a8eeeadf2f909cf08b4720815bc1205a23ad1f825618cb78bde747acad8049ca9742"

definition P521_SHA512_TV04_Qx :: nat where
  "P521_SHA512_TV04_Qx = 0x160d7ea2e128ab3fabd1a3ad5455cb45e2f977c2354a1345d4ae0c7ce4e492fb9ff958eddc2aa61735e5c1971fa6c99beda0f424a20c3ce969380aaa52ef5f5daa8"

definition P521_SHA512_TV04_Qy :: nat where
  "P521_SHA512_TV04_Qy = 0x14e4c83f90d196945fb4fe1e41913488aa53e24c1d2142d35a1eed69fed784c0ef44d71bc21afe0a0065b3b87069217a5abab4355cf8f4ceae5657cd4b9c8008f1f"

definition P521_SHA512_TV04_Q :: "int point" where
  "P521_SHA512_TV04_Q = Point (int P521_SHA512_TV04_Qx) (int P521_SHA512_TV04_Qy)"

definition P521_SHA512_TV04_k :: nat where
  "P521_SHA512_TV04_k = 0x096731f8c52e72ffcc095dd2ee4eec3da13c628f570dba169b4a7460ab471149abdede0b63e4f96faf57eab809c7d2f203fd5ab406c7bd79869b7fae9c62f97c794"

definition P521_SHA512_TV04_R :: nat where
  "P521_SHA512_TV04_R = 0x1e2bf98d1186d7bd3509f517c220de51c9200981e9b344b9fb0d36f34d969026c80311e7e73bb13789a99e0d59e82ebe0e9595d9747204c5f5550c30d934aa30c05"

definition P521_SHA512_TV04_S :: nat where
  "P521_SHA512_TV04_S = 0x12fed45cc874dc3ed3a11dd70f7d5c61451fbea497dd63e226e10364e0718d3722c27c7b4e5027051d54b8f2a57fc58bc070a55b1a5877b0f388d768837ef2e9cec"

definition P521_SHA512_TV04_Sig :: "nat × nat" where
  "P521_SHA512_TV04_Sig = (P521_SHA512_TV04_R, P521_SHA512_TV04_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA512_TV04_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV04_d" 
  by eval

lemma P521_SHA512_TV04_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA512_TV04_Q" 
  by eval

lemma P521_SHA512_TV04_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA512_TV04_d = P521_SHA512_TV04_Q" 
  by eval

lemma P521_SHA512_TV04_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA512_TV04_d P521_SHA512_TV04_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA512_TV04_dQvalidPair' P521_SHA512_TV04_d_valid) 

lemma P521_SHA512_TV04_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV04_k" 
  by eval

lemma P521_SHA512_TV04_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA512octets P521_SHA512_TV04_d P521_SHA512_TV04_Msg P521_SHA512_TV04_k = Some P521_SHA512_TV04_Sig" 
  by eval

lemma P521_SHA512_TV04_Verify: 
  "SEC1_P521_ECDSA_Verify SHA512octets P521_SHA512_TV04_Q P521_SHA512_TV04_Msg P521_SHA512_TV04_Sig" 
  by eval

subsection‹Test Vector: P521_SHA512_TV05›

subsubsection‹Given Test Vector Values›

definition P521_SHA512_TV05_Msg :: octets where
  "P521_SHA512_TV05_Msg = nat_to_octets_len 0xa1eed24b3b7c33296c2491d6ee092ec6124f85cf566bb5bc35bffb5c734e34547242e57593e962fb76aee9e800eed2d702cc301499060b76406b347f3d1c86456978950737703c8159001e6778f69c734a56e5ce5938bd0e0de0877d55adeee48b0d8dfa4ac65fd2d3ce3e12878bac5c7014f9284d161b2a3e7d5c88569a45f6 128"

definition P521_SHA512_TV05_d :: nat where
  "P521_SHA512_TV05_d = 0x18692def0b516edcdd362f42669999cf27a65482f9358fcab312c6869e22ac469b82ca9036fe123935b8b9ed064acb347227a6e377fb156ec833dab9f170c2ac697"

definition P521_SHA512_TV05_Qx :: nat where
  "P521_SHA512_TV05_Qx = 0x1ceee0be3293d8c0fc3e38a78df55e85e6b4bbce0b9995251f0ac55234140f82ae0a434b2bb41dc0aa5ecf950d4628f82c7f4f67651b804d55d844a02c1da6606f7"

definition P521_SHA512_TV05_Qy :: nat where
  "P521_SHA512_TV05_Qy = 0x1f775eb6b3c5e43fc754052d1f7fc5b99137afc15d231a0199a702fc065c917e628a54e038cbfebe05c90988b65183b368a2061e5b5c1b025bbf2b748fae00ba297"

definition P521_SHA512_TV05_Q :: "int point" where
  "P521_SHA512_TV05_Q = Point (int P521_SHA512_TV05_Qx) (int P521_SHA512_TV05_Qy)"

definition P521_SHA512_TV05_k :: nat where
  "P521_SHA512_TV05_k = 0x161cf5d37953e09e12dc0091dc35d5fb3754c5c874e474d2b4a4f1a90b870dff6d99fb156498516e25b9a6a0763170702bb8507fdba4a6131c7258f6ffc3add81fd"

definition P521_SHA512_TV05_R :: nat where
  "P521_SHA512_TV05_R = 0x14dfa43046302b81fd9a34a454dea25ccb594ace8df4f9d98556ca5076bcd44b2a9775dfaca50282b2c8988868e5a31d9eb08e794016996942088d43ad3379eb9a1"

definition P521_SHA512_TV05_S :: nat where
  "P521_SHA512_TV05_S = 0x120be63bd97691f6258b5e78817f2dd6bf5a7bf79d01b8b1c3382860c4b00f89894c72f93a69f3119cb74c90b03e9ede27bd298b357b9616a7282d176f3899aaa24"

definition P521_SHA512_TV05_Sig :: "nat × nat" where
  "P521_SHA512_TV05_Sig = (P521_SHA512_TV05_R, P521_SHA512_TV05_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA512_TV05_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV05_d" 
  by eval

lemma P521_SHA512_TV05_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA512_TV05_Q" 
  by eval

lemma P521_SHA512_TV05_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA512_TV05_d = P521_SHA512_TV05_Q" 
  by eval

lemma P521_SHA512_TV05_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA512_TV05_d P521_SHA512_TV05_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA512_TV05_dQvalidPair' P521_SHA512_TV05_d_valid) 

lemma P521_SHA512_TV05_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV05_k" 
  by eval

lemma P521_SHA512_TV05_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA512octets P521_SHA512_TV05_d P521_SHA512_TV05_Msg P521_SHA512_TV05_k = Some P521_SHA512_TV05_Sig" 
  by eval

lemma P521_SHA512_TV05_Verify: 
  "SEC1_P521_ECDSA_Verify SHA512octets P521_SHA512_TV05_Q P521_SHA512_TV05_Msg P521_SHA512_TV05_Sig" 
  by eval

subsection‹Test Vector: P521_SHA512_TV06›

subsubsection‹Given Test Vector Values›

definition P521_SHA512_TV06_Msg :: octets where
  "P521_SHA512_TV06_Msg = nat_to_octets_len 0x9aace26837695e6596007a54e4bccdd5ffb16dc6844140e2eeeb584b15acb2bbffd203c74440b6ee8db676fd200b4186a8c3e957c19e74d4d865ada83f80655323dfa3570907ed3ce853b6e8cc375ed2d758a2f5ad265dd3b47650517a49b3d02df9e0c60c21576378c2b3a08481eec129b2a75608e13e6420127a3a63c8a3f1 128"

definition P521_SHA512_TV06_d :: nat where
  "P521_SHA512_TV06_d = 0x0a63f9cdefbccdd0d5c9630b309027fa139c31e39ca26686d76c22d4093a2a5e5ec4e2308ce43eb8e563187b5bd811cc6b626eace4063047ac0420c3fdcff5bdc04"

definition P521_SHA512_TV06_Qx :: nat where
  "P521_SHA512_TV06_Qx = 0x14cab9759d4487987b8a00afd16d7199585b730fb0bfe63796272dde9135e7cb9e27cec51207c876d9214214b8c76f82e7363f5086902a577e1c50b4fbf35ce9966"

definition P521_SHA512_TV06_Qy :: nat where
  "P521_SHA512_TV06_Qy = 0x1a83f0caa01ca2166e1206292342f47f358009e8b891d3cb817aec290e0cf2f47e7fc637e39dca03949391839684f76b94d34e5abc7bb750cb44486cce525eb0093"

definition P521_SHA512_TV06_Q :: "int point" where
  "P521_SHA512_TV06_Q = Point (int P521_SHA512_TV06_Qx) (int P521_SHA512_TV06_Qy)"

definition P521_SHA512_TV06_k :: nat where
  "P521_SHA512_TV06_k = 0x01e51fd877dbbcd2ab138fd215d508879298d10c7fcbdcc918802407088eb6ca0f18976a13f2c0a57867b0298512fc85515b209c4435e9ef30ab01ba649838bc7a0"

definition P521_SHA512_TV06_R :: nat where
  "P521_SHA512_TV06_R = 0x11a1323f6132d85482d9b0f73be838d8f9e78647934f2570fededca7c234cc46aa1b97da5ac1b27b714f7a171dc4209cbb0d90e4f793c4c192dc039c31310d6d99b"

definition P521_SHA512_TV06_S :: nat where
  "P521_SHA512_TV06_S = 0x0386a5a0fc55d36ca7231a9537fee6b9e51c2255363d9c9e7cb7185669b302660e23133eb21eb56d305d36e69a79f5b6fa25b46ec61b7f699e1e9e927fb0bceca06"

definition P521_SHA512_TV06_Sig :: "nat × nat" where
  "P521_SHA512_TV06_Sig = (P521_SHA512_TV06_R, P521_SHA512_TV06_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA512_TV06_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV06_d" 
  by eval

lemma P521_SHA512_TV06_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA512_TV06_Q" 
  by eval

lemma P521_SHA512_TV06_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA512_TV06_d = P521_SHA512_TV06_Q" 
  by eval

lemma P521_SHA512_TV06_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA512_TV06_d P521_SHA512_TV06_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA512_TV06_dQvalidPair' P521_SHA512_TV06_d_valid) 

lemma P521_SHA512_TV06_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV06_k" 
  by eval

lemma P521_SHA512_TV06_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA512octets P521_SHA512_TV06_d P521_SHA512_TV06_Msg P521_SHA512_TV06_k = Some P521_SHA512_TV06_Sig" 
  by eval

lemma P521_SHA512_TV06_Verify: 
  "SEC1_P521_ECDSA_Verify SHA512octets P521_SHA512_TV06_Q P521_SHA512_TV06_Msg P521_SHA512_TV06_Sig" 
  by eval

subsection‹Test Vector: P521_SHA512_TV07›

subsubsection‹Given Test Vector Values›

definition P521_SHA512_TV07_Msg :: octets where
  "P521_SHA512_TV07_Msg = nat_to_octets_len 0xac2175940545d4fbab6e2e651c6830aba562e0c11c919e797c43eff9f187a68a9e5a128e3e2a330b955a3f4577d3f826529ad1b03d7b60f7ad678f005053b41dc0f8d267f3685c6abe1a0e9a733c44b2f3ca48b90806f935141c842e3a6c06a58f5343d75e3585971a734f4ae1074ce5b54f74bd9342f4bbca738d260393f43e 128"

definition P521_SHA512_TV07_d :: nat where
  "P521_SHA512_TV07_d = 0x024f7d67dfc0d43a26cc7c19cb511d30a097a1e27e5efe29e9e76e43849af170fd9ad57d5b22b1c8840b59ebf562371871e12d2c1baefc1abaedc872ed5d2666ad6"

definition P521_SHA512_TV07_Qx :: nat where
  "P521_SHA512_TV07_Qx = 0x09da1536154b46e3169265ccba2b4da9b4b06a7462a067c6909f6c0dd8e19a7bc2ac1a47763ec4be06c1bec57d28c55ee936cb19588cc1398fe4ea3bd07e6676b7f"

definition P521_SHA512_TV07_Qy :: nat where
  "P521_SHA512_TV07_Qy = 0x14150cdf25da0925926422e1fd4dcfcffb05bdf8682c54d67a9bd438d21de5af43a15d979b320a847683b6d12ac1383a7183095e9da491c3b4a7c28874625e70f87"

definition P521_SHA512_TV07_Q :: "int point" where
  "P521_SHA512_TV07_Q = Point (int P521_SHA512_TV07_Qx) (int P521_SHA512_TV07_Qy)"

definition P521_SHA512_TV07_k :: nat where
  "P521_SHA512_TV07_k = 0x1c1308f31716d85294b3b5f1dc87d616093b7654907f55289499b419f38ceeb906d2c9fe4cc3d80c5a38c53f9739311b0b198111fede72ebde3b0d2bc4c2ef090d2"

definition P521_SHA512_TV07_R :: nat where
  "P521_SHA512_TV07_R = 0x00dbf787ce07c453c6c6a67b0bf6850c8d6ca693a3e9818d7453487844c9048a7a2e48ff982b64eb9712461b26b5127c4dc57f9a6ad1e15d8cd56d4fd6da7186429"

definition P521_SHA512_TV07_S :: nat where
  "P521_SHA512_TV07_S = 0x0c6f1c7774caf198fc189beb7e21ca92ceccc3f9875f0e2d07dc1d15bcc8f210b6dd376bf65bb6a454bf563d7f563c1041d62d6078828a57538b25ba54723170665"

definition P521_SHA512_TV07_Sig :: "nat × nat" where
  "P521_SHA512_TV07_Sig = (P521_SHA512_TV07_R, P521_SHA512_TV07_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA512_TV07_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV07_d" 
  by eval

lemma P521_SHA512_TV07_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA512_TV07_Q" 
  by eval

lemma P521_SHA512_TV07_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA512_TV07_d = P521_SHA512_TV07_Q" 
  by eval

lemma P521_SHA512_TV07_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA512_TV07_d P521_SHA512_TV07_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA512_TV07_dQvalidPair' P521_SHA512_TV07_d_valid) 

lemma P521_SHA512_TV07_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV07_k" 
  by eval

lemma P521_SHA512_TV07_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA512octets P521_SHA512_TV07_d P521_SHA512_TV07_Msg P521_SHA512_TV07_k = Some P521_SHA512_TV07_Sig" 
  by eval

lemma P521_SHA512_TV07_Verify: 
  "SEC1_P521_ECDSA_Verify SHA512octets P521_SHA512_TV07_Q P521_SHA512_TV07_Msg P521_SHA512_TV07_Sig" 
  by eval

subsection‹Test Vector: P521_SHA512_TV08›

subsubsection‹Given Test Vector Values›

definition P521_SHA512_TV08_Msg :: octets where
  "P521_SHA512_TV08_Msg = nat_to_octets_len 0x6266f09710e2434cb3da3b15396556765db2ddcd221dce257eab7399c7c490135925112932716af1434053b8b9fe340563e57a0b9776f9ac92cbb5fba18b05c0a2fafbed7240b3f93cd1780c980ff5fe92610e36c0177cabe82367c84cee9020cf26c1d74ae3eb9b9b512cb8b3cb3d81b17cf20dc76591b2b394ef1c62ac12ee 128"

definition P521_SHA512_TV08_d :: nat where
  "P521_SHA512_TV08_d = 0x0349471460c205d836aa37dcd6c7322809e4e8ef81501e5da87284b267d843897746b33016f50a7b702964910361ed51d0afd9d8559a47f0b7c25b2bc952ce8ed9e"

definition P521_SHA512_TV08_Qx :: nat where
  "P521_SHA512_TV08_Qx = 0x00bbd4e8a016b0c254e754f68f0f4ed081320d529ecdc7899cfb5a67dd04bc85b3aa6891a3ed2c9861ae76c3847d81780c23ad84153ea2042d7fd5d517a26ff3ce4"

definition P521_SHA512_TV08_Qy :: nat where
  "P521_SHA512_TV08_Qy = 0x0645953afc3c1b3b74fdf503e7d3f982d7ee17611d60f8eb42a4bddbec2b67db1f09b54440c30b44e8071d404658285cb571462001218fc8c5e5b98b9fae28272e6"

definition P521_SHA512_TV08_Q :: "int point" where
  "P521_SHA512_TV08_Q = Point (int P521_SHA512_TV08_Qx) (int P521_SHA512_TV08_Qy)"

definition P521_SHA512_TV08_k :: nat where
  "P521_SHA512_TV08_k = 0x00eb2bd8bb56b9d2e97c51247baf734cc655c39e0bfda35375f0ac2fe82fad699bf1989577e24afb33c3868f91111e24fefe7dec802f3323ac013bec6c048fe5568"

definition P521_SHA512_TV08_R :: nat where
  "P521_SHA512_TV08_R = 0x14bf63bdbc014aa352544bd1e83ede484807ed760619fa6bc38c4f8640840195e1f2f149b29903ca4b6934404fb1f7de5e39b1ea04dba42819c75dbef6a93ebe269"

definition P521_SHA512_TV08_S :: nat where
  "P521_SHA512_TV08_S = 0x05d1bcf2295240ce4415042306abd494b4bda7cf36f2ee2931518d2454faa01c606be120b057062f2f3a174cb09c14f57ab6ef41cb3802140da22074d0e46f908d4"

definition P521_SHA512_TV08_Sig :: "nat × nat" where
  "P521_SHA512_TV08_Sig = (P521_SHA512_TV08_R, P521_SHA512_TV08_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA512_TV08_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV08_d" 
  by eval

lemma P521_SHA512_TV08_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA512_TV08_Q" 
  by eval

lemma P521_SHA512_TV08_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA512_TV08_d = P521_SHA512_TV08_Q" 
  by eval

lemma P521_SHA512_TV08_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA512_TV08_d P521_SHA512_TV08_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA512_TV08_dQvalidPair' P521_SHA512_TV08_d_valid) 

lemma P521_SHA512_TV08_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV08_k" 
  by eval

lemma P521_SHA512_TV08_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA512octets P521_SHA512_TV08_d P521_SHA512_TV08_Msg P521_SHA512_TV08_k = Some P521_SHA512_TV08_Sig" 
  by eval

lemma P521_SHA512_TV08_Verify: 
  "SEC1_P521_ECDSA_Verify SHA512octets P521_SHA512_TV08_Q P521_SHA512_TV08_Msg P521_SHA512_TV08_Sig" 
  by eval

subsection‹Test Vector: P521_SHA512_TV09›

subsubsection‹Given Test Vector Values›

definition P521_SHA512_TV09_Msg :: octets where
  "P521_SHA512_TV09_Msg = nat_to_octets_len 0x3de9e617a6868dca1a1432d503f923535da3f9b34426b2a4822174399c73b1c1ee67311410a58c17202ac767844b2024d8aa21a205707d93865693ac25a24fc87034fa3a7a7e27c3344cb03b87602c15180a5fe6a9dd90cd11af4a0f150207bf2d83f55b12c088adae99aa8cfa659311b3a25beb99056643760d6a282126b9b2 128"

definition P521_SHA512_TV09_d :: nat where
  "P521_SHA512_TV09_d = 0x07788d34758b20efc330c67483be3999d1d1a16fd0da81ed28895ebb35ee21093d37ea1ac808946c275c44454a216195eb3eb3aea1b53a329eca4eb82dd48c784f5"

definition P521_SHA512_TV09_Qx :: nat where
  "P521_SHA512_TV09_Qx = 0x0157d80bd426f6c3cee903c24b73faa02e758607c3e102d6e643b7269c299684fdaba1acddb83ee686a60acca53cddb2fe976149205c8b8ab6ad1458bc00993cc43"

definition P521_SHA512_TV09_Qy :: nat where
  "P521_SHA512_TV09_Qy = 0x16e33cbed05721b284dacc8c8fbe2d118c347fc2e2670e691d5d53daf6ef2dfec464a5fbf46f8efce81ac226915e11d43c11c8229fca2327815e1f8da5fe95021fc"

definition P521_SHA512_TV09_Q :: "int point" where
  "P521_SHA512_TV09_Q = Point (int P521_SHA512_TV09_Qx) (int P521_SHA512_TV09_Qy)"

definition P521_SHA512_TV09_k :: nat where
  "P521_SHA512_TV09_k = 0x0a73477264a9cc69d359464abb1ac098a18c0fb3ea35e4f2e6e1b060dab05bef1255d9f9c9b9fbb89712e5afe13745ae6fd5917a9aedb0f2860d03a0d8f113ea10c"

definition P521_SHA512_TV09_R :: nat where
  "P521_SHA512_TV09_R = 0x07e315d8d958b8ce27eaf4f3782294341d2a46fb1457a60eb9fe93a9ae86f3764716c4f5f124bd6b114781ed59c3f24e18aa35c903211b2f2039d85862932987d68"

definition P521_SHA512_TV09_S :: nat where
  "P521_SHA512_TV09_S = 0x1bcc1d211ebc120a97d465b603a1bb1e470109e0a55d2f1b5c597803931bd6d7718f010d7d289b31533e9fcef3d141974e5955bc7f0ee342b9cad05e29a3dded30e"

definition P521_SHA512_TV09_Sig :: "nat × nat" where
  "P521_SHA512_TV09_Sig = (P521_SHA512_TV09_R, P521_SHA512_TV09_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA512_TV09_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV09_d" 
  by eval

lemma P521_SHA512_TV09_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA512_TV09_Q" 
  by eval

lemma P521_SHA512_TV09_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA512_TV09_d = P521_SHA512_TV09_Q" 
  by eval

lemma P521_SHA512_TV09_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA512_TV09_d P521_SHA512_TV09_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA512_TV09_dQvalidPair' P521_SHA512_TV09_d_valid) 

lemma P521_SHA512_TV09_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV09_k" 
  by eval

lemma P521_SHA512_TV09_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA512octets P521_SHA512_TV09_d P521_SHA512_TV09_Msg P521_SHA512_TV09_k = Some P521_SHA512_TV09_Sig" 
  by eval

lemma P521_SHA512_TV09_Verify: 
  "SEC1_P521_ECDSA_Verify SHA512octets P521_SHA512_TV09_Q P521_SHA512_TV09_Msg P521_SHA512_TV09_Sig" 
  by eval

subsection‹Test Vector: P521_SHA512_TV10›

subsubsection‹Given Test Vector Values›

definition P521_SHA512_TV10_Msg :: octets where
  "P521_SHA512_TV10_Msg = nat_to_octets_len 0xaa48851af7ef17abe233163b7185130f4646203c205e22bcc2a5a3697bcab998c73a9ffe1d3ea0b7978ce7df937a72586eb5ca60b0d939a7d1c115c820171c89c8116b7e2c7b98cf0f14e4c4df3cb2f319ad3ab0ea25ff14526ddc037469f000bf82100acd4cdf94feb4eba4ea1726f0569336604a473aee67d71afebb569209 128"

definition P521_SHA512_TV10_d :: nat where
  "P521_SHA512_TV10_d = 0x1f98696772221e6cccd5569ed8aed3c435ee86a04689c7a64d20c30f6fe1c59cc10c6d2910261d30c3b96117a669e19cfe5b696b68feeacf61f6a3dea55e6e5837a"

definition P521_SHA512_TV10_Qx :: nat where
  "P521_SHA512_TV10_Qx = 0x07002872c200e16d57e8e53f7bce6e9a7832c387f6f9c29c6b75526262c57bc2b56d63e9558c5761c1d62708357f586d3aab41c6a7ca3bf6c32d9c3ca40f9a2796a"

definition P521_SHA512_TV10_Qy :: nat where
  "P521_SHA512_TV10_Qy = 0x1fe3e52472ef224fb38d5a0a14875b52c2f50b82b99eea98d826c77e6a9ccf798de5ffa92a0d65965f740c702a3027be66b9c844f1b2e96c134eb3fdf3edddcf11c"

definition P521_SHA512_TV10_Q :: "int point" where
  "P521_SHA512_TV10_Q = Point (int P521_SHA512_TV10_Qx) (int P521_SHA512_TV10_Qy)"

definition P521_SHA512_TV10_k :: nat where
  "P521_SHA512_TV10_k = 0x1a277cf0414c6adb621d1cc0311ec908401ce040c6687ed45a0cdf2910c42c9f1954a4572d8e659733d5e26cbd35e3260be40017b2f5d38ec42315f5c0b056c596d"

definition P521_SHA512_TV10_R :: nat where
  "P521_SHA512_TV10_R = 0x0d732ba8b3e9c9e0a495249e152e5bee69d94e9ff012d001b140d4b5d082aa9df77e10b65f115a594a50114722db42fa5fbe457c5bd05e7ac7ee510aa68fe7b1e7f"

definition P521_SHA512_TV10_S :: nat where
  "P521_SHA512_TV10_S = 0x134ac5e1ee339727df80c35ff5b2891596dd14d6cfd137bafd50ab98e2c1ab4008a0bd03552618d217912a9ec502a902f2353e757c3b5776309f7f2cfebf913e9cd"

definition P521_SHA512_TV10_Sig :: "nat × nat" where
  "P521_SHA512_TV10_Sig = (P521_SHA512_TV10_R, P521_SHA512_TV10_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA512_TV10_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV10_d" 
  by eval

lemma P521_SHA512_TV10_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA512_TV10_Q" 
  by eval

lemma P521_SHA512_TV10_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA512_TV10_d = P521_SHA512_TV10_Q" 
  by eval

lemma P521_SHA512_TV10_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA512_TV10_d P521_SHA512_TV10_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA512_TV10_dQvalidPair' P521_SHA512_TV10_d_valid) 

lemma P521_SHA512_TV10_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV10_k" 
  by eval

lemma P521_SHA512_TV10_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA512octets P521_SHA512_TV10_d P521_SHA512_TV10_Msg P521_SHA512_TV10_k = Some P521_SHA512_TV10_Sig" 
  by eval

lemma P521_SHA512_TV10_Verify: 
  "SEC1_P521_ECDSA_Verify SHA512octets P521_SHA512_TV10_Q P521_SHA512_TV10_Msg P521_SHA512_TV10_Sig" 
  by eval

subsection‹Test Vector: P521_SHA512_TV11›

subsubsection‹Given Test Vector Values›

definition P521_SHA512_TV11_Msg :: octets where
  "P521_SHA512_TV11_Msg = nat_to_octets_len 0xb0d5d52259af364eb2d1a5027e5f7d0afe4b999cc5dd2268cfe76f51d2f17b541bdd7867e23a1bb897705153d9432a24012108979c6a2c9e2567c9531d012f9e4be764419491a52eae2e127430b0ab58cb8e216515a821b3db206447c235bf44ee304201b483b2a88844abaa18bca0147dfff7e502397dd62e15524f67eb2df2 128"

definition P521_SHA512_TV11_d :: nat where
  "P521_SHA512_TV11_d = 0x13c3852a6bc8825b45fd7da1754078913d77f4e586216a6eb08b6f03adce7464f5dbc2bea0eb7b12d103870ef045f53d67e3600d7eba07aac5db03f71b64db1cceb"

definition P521_SHA512_TV11_Qx :: nat where
  "P521_SHA512_TV11_Qx = 0x0c97a4ebcbbe701c9f7be127e87079edf479b76d3c14bfbee693e1638e5bff8d4705ac0c14597529dbe13356ca85eb03a418edfe144ce6cbf3533016d4efc29dbd4"

definition P521_SHA512_TV11_Qy :: nat where
  "P521_SHA512_TV11_Qy = 0x11c75b7a8894ef64109ac2dea972e7fd5f79b75dab1bf9441a5b8b86f1dc1324426fa6cf4e7b973b44e3d0576c52e5c9edf8ce2fc18cb3c28742d44419f044667f8"

definition P521_SHA512_TV11_Q :: "int point" where
  "P521_SHA512_TV11_Q = Point (int P521_SHA512_TV11_Qx) (int P521_SHA512_TV11_Qy)"

definition P521_SHA512_TV11_k :: nat where
  "P521_SHA512_TV11_k = 0x1e25b86db041f21c2503d547e2b1b655f0b99d5b6c0e1cf2bdbd8a8c6a053f5d79d78c55b4ef75bff764a74edc920b35536e3c470b6f6b8fd53898f3bbc467539ef"

definition P521_SHA512_TV11_R :: nat where
  "P521_SHA512_TV11_R = 0x1dce45ea592b34d016497882c48dc0c7afb1c8e0f81a051800d7ab8da9d237efd892207bc9401f1d30650f66af8d5349fc5b19727756270722d5a8adb0a49b72d0a"

definition P521_SHA512_TV11_S :: nat where
  "P521_SHA512_TV11_S = 0x0b79ffcdc33e028b1ab894cb751ec792a69e3011b201a76f3b878655bc31efd1c0bf3b98aea2b14f262c19d142e008b98e890ebbf464d3b025764dd2f73c4251b1a"

definition P521_SHA512_TV11_Sig :: "nat × nat" where
  "P521_SHA512_TV11_Sig = (P521_SHA512_TV11_R, P521_SHA512_TV11_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA512_TV11_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV11_d" 
  by eval

lemma P521_SHA512_TV11_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA512_TV11_Q" 
  by eval

lemma P521_SHA512_TV11_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA512_TV11_d = P521_SHA512_TV11_Q" 
  by eval

lemma P521_SHA512_TV11_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA512_TV11_d P521_SHA512_TV11_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA512_TV11_dQvalidPair' P521_SHA512_TV11_d_valid) 

lemma P521_SHA512_TV11_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV11_k" 
  by eval

lemma P521_SHA512_TV11_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA512octets P521_SHA512_TV11_d P521_SHA512_TV11_Msg P521_SHA512_TV11_k = Some P521_SHA512_TV11_Sig" 
  by eval

lemma P521_SHA512_TV11_Verify: 
  "SEC1_P521_ECDSA_Verify SHA512octets P521_SHA512_TV11_Q P521_SHA512_TV11_Msg P521_SHA512_TV11_Sig" 
  by eval

subsection‹Test Vector: P521_SHA512_TV12›

subsubsection‹Given Test Vector Values›

definition P521_SHA512_TV12_Msg :: octets where
  "P521_SHA512_TV12_Msg = nat_to_octets_len 0x9599788344976779383a7a0812a096943a1f771ee484d586af1a06207478e4c0be9c200d42460fe837e24b266c8852d80d3c53cc52ffb1913fc3261145fc6da575611efd16c026059a2e64f802517ffd1b6b34de10ad2909c65c2155e8d939b8115400c1d793d23955b15f5d1c13c962ff92b4a815cee0e10f8e14e1f6e6cd38 128"

definition P521_SHA512_TV12_d :: nat where
  "P521_SHA512_TV12_d = 0x1654eaa1f6eec7159ee2d36fb24d15d6d33a128f36c52e2437f7d1b5a44ea4fa965c0a26d0066f92c8b82bd136491e929686c8bde61b7c704daab54ed1e1bdf6b77"

definition P521_SHA512_TV12_Qx :: nat where
  "P521_SHA512_TV12_Qx = 0x1f269692c47a55242bb08731ff920f4915bfcecf4d4431a8b487c90d08565272c52ca90c47397f7604bc643982e34d05178e979c2cff7ea1b9eaec18d69ca7382de"

definition P521_SHA512_TV12_Qy :: nat where
  "P521_SHA512_TV12_Qy = 0x0750bdd866fba3e92c29599c002ac6f9e2bf39af8521b7b133f70510e9918a94d3c279edec97ab75ecda95e3dd7861af84c543371c055dc74eeeff7061726818327"

definition P521_SHA512_TV12_Q :: "int point" where
  "P521_SHA512_TV12_Q = Point (int P521_SHA512_TV12_Qx) (int P521_SHA512_TV12_Qy)"

definition P521_SHA512_TV12_k :: nat where
  "P521_SHA512_TV12_k = 0x1b7519becd00d750459d63a72f13318b6ac61b8c8e7077cf9415c9b4b924f35514c9c28a0fae43d06e31c670a873716156aa7bc744577d62476e038b116576a9e53"

definition P521_SHA512_TV12_R :: nat where
  "P521_SHA512_TV12_R = 0x183bddb46c249e868ef231a1ebd85d0773bf8105a092ab7d884d677a1e9b7d6014d6358c09538a99d9dca8f36f163ac1827df420c3f9360cc66900a9737a7f756f3"

definition P521_SHA512_TV12_S :: nat where
  "P521_SHA512_TV12_S = 0x0d05ee3e64bac4e56d9d8bd511c8a43941e953cba4e5d83c0553acb87091ff54f3aad4d69d9f15e520a2551cc14f2c86bb45513fef0295e381a7635486bd3917b50"

definition P521_SHA512_TV12_Sig :: "nat × nat" where
  "P521_SHA512_TV12_Sig = (P521_SHA512_TV12_R, P521_SHA512_TV12_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA512_TV12_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV12_d" 
  by eval

lemma P521_SHA512_TV12_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA512_TV12_Q" 
  by eval

lemma P521_SHA512_TV12_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA512_TV12_d = P521_SHA512_TV12_Q" 
  by eval

lemma P521_SHA512_TV12_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA512_TV12_d P521_SHA512_TV12_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA512_TV12_dQvalidPair' P521_SHA512_TV12_d_valid) 

lemma P521_SHA512_TV12_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV12_k" 
  by eval

lemma P521_SHA512_TV12_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA512octets P521_SHA512_TV12_d P521_SHA512_TV12_Msg P521_SHA512_TV12_k = Some P521_SHA512_TV12_Sig" 
  by eval

lemma P521_SHA512_TV12_Verify: 
  "SEC1_P521_ECDSA_Verify SHA512octets P521_SHA512_TV12_Q P521_SHA512_TV12_Msg P521_SHA512_TV12_Sig" 
  by eval

subsection‹Test Vector: P521_SHA512_TV13›

subsubsection‹Given Test Vector Values›

definition P521_SHA512_TV13_Msg :: octets where
  "P521_SHA512_TV13_Msg = nat_to_octets_len 0xfdde51acfd04eb0ad892ce9d6c0f90eb91ce765cbe3ce9d3f2defe8f691324d26b968b8b90e77706b068585f2a3ee7bf3e910528f7403c5af745a6f9d7ba6c53abd885c3b1be583415b128f4d3f224daf8563476bd9aa61e9c8518c144335f8f879c03696bddbe3ac37a8fbede29861611feaa87e325e2f60278b4893ed57fb0 128"

definition P521_SHA512_TV13_d :: nat where
  "P521_SHA512_TV13_d = 0x1cba5d561bf18656991eba9a1dde8bde547885ea1f0abe7f2837e569ca52f53df5e64e4a547c4f26458b5d9626ed6d702e5ab1dd585cf36a0c84f768fac946cfd4c"

definition P521_SHA512_TV13_Qx :: nat where
  "P521_SHA512_TV13_Qx = 0x12857c2244fa04db3b73db4847927db63cce2fa6cb22724466d3e20bc950a9250a15eafd99f236a801e5271e8f90d9e8a97f37c12f7da65bce8a2c93bcd25526205"

definition P521_SHA512_TV13_Qy :: nat where
  "P521_SHA512_TV13_Qy = 0x0f394e37c17d5b8e35b488fa05a607dbc74264965043a1fb60e92edc212296ae72d7d6fe2e3457e67be853664e1da64f57e44bd259076b3bb2b06a2c604fea1be9d"

definition P521_SHA512_TV13_Q :: "int point" where
  "P521_SHA512_TV13_Q = Point (int P521_SHA512_TV13_Qx) (int P521_SHA512_TV13_Qy)"

definition P521_SHA512_TV13_k :: nat where
  "P521_SHA512_TV13_k = 0x0e790238796fee7b5885dc0784c7041a4cc7ca4ba757d9f7906ad1fcbab5667e3734bc2309a48047442535ff89144b518f730ff55c0c67eeb4c880c2dfd2fb60d69"

definition P521_SHA512_TV13_R :: nat where
  "P521_SHA512_TV13_R = 0x1d7ce382295a2a109064ea03f0ad8761dd60eefb9c207a20e3c5551e82ac6d2ee5922b3e9655a65ba6c359dcbf8fa843fbe87239a5c3e3eaecec0407d2fcdb687c2"

definition P521_SHA512_TV13_S :: nat where
  "P521_SHA512_TV13_S = 0x161963a6237b8955a8a756d8df5dbd303140bb90143b1da5f07b32f9cb64733dc6316080924733f1e2c81ade9d0be71b5b95b55666026a035a93ab3004d0bc0b19f"

definition P521_SHA512_TV13_Sig :: "nat × nat" where
  "P521_SHA512_TV13_Sig = (P521_SHA512_TV13_R, P521_SHA512_TV13_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA512_TV13_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV13_d" 
  by eval

lemma P521_SHA512_TV13_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA512_TV13_Q" 
  by eval

lemma P521_SHA512_TV13_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA512_TV13_d = P521_SHA512_TV13_Q" 
  by eval

lemma P521_SHA512_TV13_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA512_TV13_d P521_SHA512_TV13_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA512_TV13_dQvalidPair' P521_SHA512_TV13_d_valid) 

lemma P521_SHA512_TV13_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV13_k" 
  by eval

lemma P521_SHA512_TV13_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA512octets P521_SHA512_TV13_d P521_SHA512_TV13_Msg P521_SHA512_TV13_k = Some P521_SHA512_TV13_Sig" 
  by eval

lemma P521_SHA512_TV13_Verify: 
  "SEC1_P521_ECDSA_Verify SHA512octets P521_SHA512_TV13_Q P521_SHA512_TV13_Msg P521_SHA512_TV13_Sig" 
  by eval

subsection‹Test Vector: P521_SHA512_TV14›

subsubsection‹Given Test Vector Values›

definition P521_SHA512_TV14_Msg :: octets where
  "P521_SHA512_TV14_Msg = nat_to_octets_len 0xbeb34c997f905c77451ac392f7957a0ab8b23325bd5c63ca31c109ac8f655a1e3094240cb8a99284f8091de2ab9a7db2504d16251980b86be89ec3a3f41162698bab51848880633e0b71a38f8896335853d8e836a2454ecab2acdcc052c8f659be1d703b13ae1b090334ac50ab0137ddb5e8b924c0e3d2e5789daaef2fdd4a1e 128"

definition P521_SHA512_TV14_d :: nat where
  "P521_SHA512_TV14_d = 0x0972e7ff25adf8a032535e5b19463cfe306b90803bf27fabc6046ae0807d2312fbab85d1da61b80b2d5d48f4e5886f27fca050b84563aee1926ae6b2564cd756d63"

definition P521_SHA512_TV14_Qx :: nat where
  "P521_SHA512_TV14_Qx = 0x1d7f1e9e610619daa9d2efa563610a371677fe8b58048fdc55a98a49970f6afa6649c516f9c72085ca3722aa595f45f2803402b01c832d28aac63d9941f1a25dfea"

definition P521_SHA512_TV14_Qy :: nat where
  "P521_SHA512_TV14_Qy = 0x1571facce3fcfe733a8eef4e8305dfe99103a370f82b3f8d75085414f2592ad44969a2ef8196c8b9809f0eca2f7ddc71c47879e3f37a40b9fecf97992b97af29721"

definition P521_SHA512_TV14_Q :: "int point" where
  "P521_SHA512_TV14_Q = Point (int P521_SHA512_TV14_Qx) (int P521_SHA512_TV14_Qy)"

definition P521_SHA512_TV14_k :: nat where
  "P521_SHA512_TV14_k = 0x0517f6e4002479dc89e8cbb55b7c426d128776ca82cf81be8c1da9557178783f40e3d047db7e77867f1af030a51de470ee3128c22e9c2d642d71e4904ab5a76edfa"

definition P521_SHA512_TV14_R :: nat where
  "P521_SHA512_TV14_R = 0x1c3262a3a3fb74fa5124b71a6c7f7b7e6d56738eabaf7666b372b299b0c99ee8a16be3df88dd955de093fc8c049f76ee83a4138cee41e5fe94755d27a52ee44032f"

definition P521_SHA512_TV14_S :: nat where
  "P521_SHA512_TV14_S = 0x072fd88bb1684c4ca9531748dfce4c161037fcd6ae5c2803b7117fb60d3db5df7df380591aaf3073a3031306b76f062dcc547ded23f6690293c34a710e7e9a226c3"

definition P521_SHA512_TV14_Sig :: "nat × nat" where
  "P521_SHA512_TV14_Sig = (P521_SHA512_TV14_R, P521_SHA512_TV14_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA512_TV14_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV14_d" 
  by eval

lemma P521_SHA512_TV14_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA512_TV14_Q" 
  by eval

lemma P521_SHA512_TV14_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA512_TV14_d = P521_SHA512_TV14_Q" 
  by eval

lemma P521_SHA512_TV14_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA512_TV14_d P521_SHA512_TV14_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA512_TV14_dQvalidPair' P521_SHA512_TV14_d_valid) 

lemma P521_SHA512_TV14_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV14_k" 
  by eval

lemma P521_SHA512_TV14_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA512octets P521_SHA512_TV14_d P521_SHA512_TV14_Msg P521_SHA512_TV14_k = Some P521_SHA512_TV14_Sig" 
  by eval

lemma P521_SHA512_TV14_Verify: 
  "SEC1_P521_ECDSA_Verify SHA512octets P521_SHA512_TV14_Q P521_SHA512_TV14_Msg P521_SHA512_TV14_Sig" 
  by eval

subsection‹Test Vector: P521_SHA512_TV15›

subsubsection‹Given Test Vector Values›

definition P521_SHA512_TV15_Msg :: octets where
  "P521_SHA512_TV15_Msg = nat_to_octets_len 0x543c374af90c34f50ee195006d5f9d8dd986d09ad182fcbefa085567275eee1e742bfe0af3d058675adeb5b9f87f248b00a9fbd2aa779129123a5b983f2f26fc3caf2ea34277550c22fe8c814c739b46972d50232993cddd63a3c99e20f5c5067d9b57e2d5db94317a5a16b5c12b5c4cafbc79cbc2f9940f074bbc7d0dc71e90 128"

definition P521_SHA512_TV15_d :: nat where
  "P521_SHA512_TV15_d = 0x1f0ec8da29295394f2f072672db014861be33bfd9f91349dad5566ff396bea055e53b1d61c8c4e5c9f6e129ed75a49f91cce1d5530ad4e78c2b793a63195eb9f0da"

definition P521_SHA512_TV15_Qx :: nat where
  "P521_SHA512_TV15_Qx = 0x09ec1a3761fe3958073b9647f34202c5e8ca2428d056facc4f3fedc7077fa87f1d1eb30cc74f6e3ff3d3f82df2641cea1eb3ff1529e8a3866ae2055aacec0bf68c4"

definition P521_SHA512_TV15_Qy :: nat where
  "P521_SHA512_TV15_Qy = 0x0bed0261b91f664c3ff53e337d8321cb988c3edc03b46754680097e5a8585245d80d0b7045c75a9c5be7f599d3b5eea08d828acb6294ae515a3df57a37f903ef62e"

definition P521_SHA512_TV15_Q :: "int point" where
  "P521_SHA512_TV15_Q = Point (int P521_SHA512_TV15_Qx) (int P521_SHA512_TV15_Qy)"

definition P521_SHA512_TV15_k :: nat where
  "P521_SHA512_TV15_k = 0x0ac3b6d61ebda99e23301fa198d686a13c0832af594b289c9a55669ce6d62011384769013748b68465527a597ed6858a06a99d50493562b3a7dbcee975ad34657d8"

definition P521_SHA512_TV15_R :: nat where
  "P521_SHA512_TV15_R = 0x0cef3f4babe6f9875e5db28c27d6a197d607c3641a90f10c2cc2cb302ba658aa151dc76c507488b99f4b3c8bb404fb5c852f959273f412cbdd5e713c5e3f0e67f94"

definition P521_SHA512_TV15_S :: nat where
  "P521_SHA512_TV15_S = 0x0097ed9e005416fc944e26bcc3661a09b35c128fcccdc2742739c8a301a338dd77d9d13571612a3b9524a6164b09fe73643bbc31447ee31ef44a490843e4e7db23f"

definition P521_SHA512_TV15_Sig :: "nat × nat" where
  "P521_SHA512_TV15_Sig = (P521_SHA512_TV15_R, P521_SHA512_TV15_S)"

subsubsection‹Checks›
text‹Check that (d,Q) form a valid key pair.  Check that k is a valid private key.
Check that signing Msg with private key d and ephemeral signing key k gives the signature
Sig = (R,S).  Finally check that the verification routine on Msg, Sig = (R,S) and the public 
key Q returns True.  Note that this last check is redundant: since (d,Q) is a valid key pair 
and k is a valid private key, and given that when we sign Msg using d and k we get Sig, 
then we know that the verification routine on Q, Msg, and Sig will return True.›

lemma P521_SHA512_TV15_d_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV15_d" 
  by eval

lemma P521_SHA512_TV15_QonCurve: 
  "SEC1_P521_on_curve' P521_SHA512_TV15_Q" 
  by eval

lemma P521_SHA512_TV15_dQvalidPair': 
  "SEC1_P521_ECkeyGen P521_SHA512_TV15_d = P521_SHA512_TV15_Q" 
  by eval

lemma P521_SHA512_TV15_dQvalidPair: 
  "SEC1_P521_ECkeyPairValid P521_SHA512_TV15_d P521_SHA512_TV15_Q" 
  by (simp add: SEC1_P521.ECkeyPairEqKeyGen P521_SHA512_TV15_dQvalidPair' P521_SHA512_TV15_d_valid) 

lemma P521_SHA512_TV15_k_valid: 
  "SEC1_P521_ECprivateKeyValid P521_SHA512_TV15_k" 
  by eval

lemma P521_SHA512_TV15_Sign: 
  "SEC1_P521_ECDSA_Sign' SHA512octets P521_SHA512_TV15_d P521_SHA512_TV15_Msg P521_SHA512_TV15_k = Some P521_SHA512_TV15_Sig" 
  by eval

lemma P521_SHA512_TV15_Verify: 
  "SEC1_P521_ECDSA_Verify SHA512octets P521_SHA512_TV15_Q P521_SHA512_TV15_Msg P521_SHA512_TV15_Sig" 
  by eval

section‹NIST Test Vectors for ECDH: Curve P-192›
  
subsection‹Test Vector: SEC1_P192_0›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_0_QCAVSx :: nat where 
  "SEC1_P192_0_QCAVSx = 0x42ea6dd9969dd2a61fea1aac7f8e98edcc896c6e55857cc0"

definition SEC1_P192_0_QCAVSy :: nat where 
  "SEC1_P192_0_QCAVSy = 0xdfbe5d7c61fac88b11811bde328e8a0d12bf01a9d204b523"

definition SEC1_P192_0_QCAVS :: "int point" where
  "SEC1_P192_0_QCAVS = Point (int SEC1_P192_0_QCAVSx) (int SEC1_P192_0_QCAVSy)"

definition SEC1_P192_0_dIUT :: nat where 
  "SEC1_P192_0_dIUT = 0xf17d3fea367b74d340851ca4270dcb24c271f445bed9d527"

definition SEC1_P192_0_QIUTx :: nat where 
  "SEC1_P192_0_QIUTx = 0xb15053401f57285637ec324c1cd2139e3a67de3739234b37"

definition SEC1_P192_0_QIUTy :: nat where 
  "SEC1_P192_0_QIUTy = 0xf269c158637482aad644cd692dd1d3ef2c8a7c49e389f7f6"

definition SEC1_P192_0_QIUT :: "int point" where
  "SEC1_P192_0_QIUT = Point (int SEC1_P192_0_QIUTx) (int SEC1_P192_0_QIUTy)"

definition SEC1_P192_0_ZIUT :: nat where 
  "SEC1_P192_0_ZIUT = 0x803d8ab2e5b6e6fca715737c3a82f7ce3c783124f6d51cd0"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_0_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_0_dIUT"
  by eval

lemma SEC1_P192_0_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_0_dIUT = SEC1_P192_0_QIUT" 
  by eval

lemma SEC1_P192_0_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_0_QCAVS" 
  by eval

lemma SEC1_P192_0_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_0_QCAVS" 
  using SEC1_P192_0_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_0_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_0_dIUT SEC1_P192_0_QCAVS = Some SEC1_P192_0_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_1›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_1_QCAVSx :: nat where 
  "SEC1_P192_1_QCAVSx = 0xdeb5712fa027ac8d2f22c455ccb73a91e17b6512b5e030e7"

definition SEC1_P192_1_QCAVSy :: nat where 
  "SEC1_P192_1_QCAVSy = 0x7e2690a02cc9b28708431a29fb54b87b1f0c14e011ac2125"

definition SEC1_P192_1_QCAVS :: "int point" where
  "SEC1_P192_1_QCAVS = Point (int SEC1_P192_1_QCAVSx) (int SEC1_P192_1_QCAVSy)"

definition SEC1_P192_1_dIUT :: nat where 
  "SEC1_P192_1_dIUT = 0x56e853349d96fe4c442448dacb7cf92bb7a95dcf574a9bd5"

definition SEC1_P192_1_QIUTx :: nat where 
  "SEC1_P192_1_QIUTx = 0xc00d435716ffea53fd8c162792414c37665187e582716539"

definition SEC1_P192_1_QIUTy :: nat where 
  "SEC1_P192_1_QIUTy = 0xab711c62aa71a5a18e8a3c48f89dc6fa52fac0108e52a8a0"

definition SEC1_P192_1_QIUT :: "int point" where
  "SEC1_P192_1_QIUT = Point (int SEC1_P192_1_QIUTx) (int SEC1_P192_1_QIUTy)"

definition SEC1_P192_1_ZIUT :: nat where 
  "SEC1_P192_1_ZIUT = 0xc208847568b98835d7312cef1f97f7aa298283152313c29d"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_1_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_1_dIUT"
  by eval

lemma SEC1_P192_1_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_1_dIUT = SEC1_P192_1_QIUT" 
  by eval

lemma SEC1_P192_1_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_1_QCAVS" 
  by eval

lemma SEC1_P192_1_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_1_QCAVS" 
  using SEC1_P192_1_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_1_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_1_dIUT SEC1_P192_1_QCAVS = Some SEC1_P192_1_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_2›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_2_QCAVSx :: nat where 
  "SEC1_P192_2_QCAVSx = 0x4edaa8efc5a0f40f843663ec5815e7762dddc008e663c20f"

definition SEC1_P192_2_QCAVSy :: nat where 
  "SEC1_P192_2_QCAVSy = 0x0a9f8dc67a3e60ef6d64b522185d03df1fc0adfd42478279"

definition SEC1_P192_2_QCAVS :: "int point" where
  "SEC1_P192_2_QCAVS = Point (int SEC1_P192_2_QCAVSx) (int SEC1_P192_2_QCAVSy)"

definition SEC1_P192_2_dIUT :: nat where 
  "SEC1_P192_2_dIUT = 0xc6ef61fe12e80bf56f2d3f7d0bb757394519906d55500949"

definition SEC1_P192_2_QIUTx :: nat where 
  "SEC1_P192_2_QIUTx = 0xe184bc182482f3403c8787b83842477467fcd011db0f6c64"

definition SEC1_P192_2_QIUTy :: nat where 
  "SEC1_P192_2_QIUTy = 0xf9d1c14142f40de8639db97d51a63d2cce1007ccf773cdcb"

definition SEC1_P192_2_QIUT :: "int point" where
  "SEC1_P192_2_QIUT = Point (int SEC1_P192_2_QIUTx) (int SEC1_P192_2_QIUTy)"

definition SEC1_P192_2_ZIUT :: nat where 
  "SEC1_P192_2_ZIUT = 0x87229107047a3b611920d6e3b2c0c89bea4f49412260b8dd"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_2_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_2_dIUT"
  by eval

lemma SEC1_P192_2_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_2_dIUT = SEC1_P192_2_QIUT" 
  by eval

lemma SEC1_P192_2_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_2_QCAVS" 
  by eval

lemma SEC1_P192_2_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_2_QCAVS" 
  using SEC1_P192_2_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_2_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_2_dIUT SEC1_P192_2_QCAVS = Some SEC1_P192_2_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_3›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_3_QCAVSx :: nat where 
  "SEC1_P192_3_QCAVSx = 0x8887c276edeed3e9e866b46d58d895c73fbd80b63e382e88"

definition SEC1_P192_3_QCAVSy :: nat where 
  "SEC1_P192_3_QCAVSy = 0x04c5097ba6645e16206cfb70f7052655947dd44a17f1f9d5"

definition SEC1_P192_3_QCAVS :: "int point" where
  "SEC1_P192_3_QCAVS = Point (int SEC1_P192_3_QCAVSx) (int SEC1_P192_3_QCAVSy)"

definition SEC1_P192_3_dIUT :: nat where 
  "SEC1_P192_3_dIUT = 0xe6747b9c23ba7044f38ff7e62c35e4038920f5a0163d3cda"

definition SEC1_P192_3_QIUTx :: nat where 
  "SEC1_P192_3_QIUTx = 0x2b838dbe73735f37a39a78d3195783d26991e86ff4d92d1a"

definition SEC1_P192_3_QIUTy :: nat where 
  "SEC1_P192_3_QIUTy = 0x60d344942274489f98903b2e7f93f8d197fc9ae60a0ed53a"

definition SEC1_P192_3_QIUT :: "int point" where
  "SEC1_P192_3_QIUT = Point (int SEC1_P192_3_QIUTx) (int SEC1_P192_3_QIUTy)"

definition SEC1_P192_3_ZIUT :: nat where 
  "SEC1_P192_3_ZIUT = 0xeec0bed8fc55e1feddc82158fd6dc0d48a4d796aaf47d46c"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_3_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_3_dIUT"
  by eval

lemma SEC1_P192_3_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_3_dIUT = SEC1_P192_3_QIUT" 
  by eval

lemma SEC1_P192_3_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_3_QCAVS" 
  by eval

lemma SEC1_P192_3_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_3_QCAVS" 
  using SEC1_P192_3_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_3_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_3_dIUT SEC1_P192_3_QCAVS = Some SEC1_P192_3_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_4›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_4_QCAVSx :: nat where 
  "SEC1_P192_4_QCAVSx = 0x0d045f30254adc1fcefa8a5b1f31bf4e739dd327cd18d594"

definition SEC1_P192_4_QCAVSy :: nat where 
  "SEC1_P192_4_QCAVSy = 0x542c314e41427c08278a08ce8d7305f3b5b849c72d8aff73"

definition SEC1_P192_4_QCAVS :: "int point" where
  "SEC1_P192_4_QCAVS = Point (int SEC1_P192_4_QCAVSx) (int SEC1_P192_4_QCAVSy)"

definition SEC1_P192_4_dIUT :: nat where 
  "SEC1_P192_4_dIUT = 0xbeabedd0154a1afcfc85d52181c10f5eb47adc51f655047d"

definition SEC1_P192_4_QIUTx :: nat where 
  "SEC1_P192_4_QIUTx = 0x1f65cf6e8978e1c1bc10bb61a7db311de310088c8cf9768b"

definition SEC1_P192_4_QIUTy :: nat where 
  "SEC1_P192_4_QIUTy = 0xf7d438168e7f42ab14b16af53a7a2f646ff40b53d74cbcc7"

definition SEC1_P192_4_QIUT :: "int point" where
  "SEC1_P192_4_QIUT = Point (int SEC1_P192_4_QIUTx) (int SEC1_P192_4_QIUTy)"

definition SEC1_P192_4_ZIUT :: nat where 
  "SEC1_P192_4_ZIUT = 0x716e743b1b37a2cd8479f0a3d5a74c10ba2599be18d7e2f4"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_4_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_4_dIUT"
  by eval

lemma SEC1_P192_4_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_4_dIUT = SEC1_P192_4_QIUT" 
  by eval

lemma SEC1_P192_4_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_4_QCAVS" 
  by eval

lemma SEC1_P192_4_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_4_QCAVS" 
  using SEC1_P192_4_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_4_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_4_dIUT SEC1_P192_4_QCAVS = Some SEC1_P192_4_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_5›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_5_QCAVSx :: nat where 
  "SEC1_P192_5_QCAVSx = 0xfb35ca20d2e96665c51b98e8f6eb3d79113508d8bccd4516"

definition SEC1_P192_5_QCAVSy :: nat where 
  "SEC1_P192_5_QCAVSy = 0x368eec0d5bfb847721df6aaff0e5d48c444f74bf9cd8a5a7"

definition SEC1_P192_5_QCAVS :: "int point" where
  "SEC1_P192_5_QCAVS = Point (int SEC1_P192_5_QCAVSx) (int SEC1_P192_5_QCAVSy)"

definition SEC1_P192_5_dIUT :: nat where 
  "SEC1_P192_5_dIUT = 0xcf70354226667321d6e2baf40999e2fd74c7a0f793fa8699"

definition SEC1_P192_5_QIUTx :: nat where 
  "SEC1_P192_5_QIUTx = 0x5f4844ffcce61005d24f737db98675e92f7b6543aeb6106c"

definition SEC1_P192_5_QIUTy :: nat where 
  "SEC1_P192_5_QIUTy = 0x5424f598139215d389b6b12b86d58014857f2ddadb540f51"

definition SEC1_P192_5_QIUT :: "int point" where
  "SEC1_P192_5_QIUT = Point (int SEC1_P192_5_QIUTx) (int SEC1_P192_5_QIUTy)"

definition SEC1_P192_5_ZIUT :: nat where 
  "SEC1_P192_5_ZIUT = 0xf67053b934459985a315cb017bf0302891798d45d0e19508"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_5_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_5_dIUT"
  by eval

lemma SEC1_P192_5_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_5_dIUT = SEC1_P192_5_QIUT" 
  by eval

lemma SEC1_P192_5_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_5_QCAVS" 
  by eval

lemma SEC1_P192_5_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_5_QCAVS" 
  using SEC1_P192_5_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_5_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_5_dIUT SEC1_P192_5_QCAVS = Some SEC1_P192_5_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_6›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_6_QCAVSx :: nat where 
  "SEC1_P192_6_QCAVSx = 0x824752960c1307e5f13a83da21c7998ca8b5b00b9549f6d0"

definition SEC1_P192_6_QCAVSy :: nat where 
  "SEC1_P192_6_QCAVSy = 0xbc52d91e234363bc32ee0b6778f25cd8c1847510f4348b94"

definition SEC1_P192_6_QCAVS :: "int point" where
  "SEC1_P192_6_QCAVS = Point (int SEC1_P192_6_QCAVSx) (int SEC1_P192_6_QCAVSy)"

definition SEC1_P192_6_dIUT :: nat where 
  "SEC1_P192_6_dIUT = 0xfe942515237fffdd7b4eb5c64909eee4856a076cdf12bae2"

definition SEC1_P192_6_QIUTx :: nat where 
  "SEC1_P192_6_QIUTx = 0xe6369df79b207b8b8679f7c869cfc264859d1ab55aa401e8"

definition SEC1_P192_6_QIUTy :: nat where 
  "SEC1_P192_6_QIUTy = 0x1f99c71f801a30b52f74da6e5e6dbb62ee4c5da1090cc020"

definition SEC1_P192_6_QIUT :: "int point" where
  "SEC1_P192_6_QIUT = Point (int SEC1_P192_6_QIUTx) (int SEC1_P192_6_QIUTy)"

definition SEC1_P192_6_ZIUT :: nat where 
  "SEC1_P192_6_ZIUT = 0x75822971193edd472bf30151a782619c55ad0b279c9303dd"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_6_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_6_dIUT"
  by eval

lemma SEC1_P192_6_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_6_dIUT = SEC1_P192_6_QIUT" 
  by eval

lemma SEC1_P192_6_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_6_QCAVS" 
  by eval

lemma SEC1_P192_6_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_6_QCAVS" 
  using SEC1_P192_6_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_6_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_6_dIUT SEC1_P192_6_QCAVS = Some SEC1_P192_6_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_7›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_7_QCAVSx :: nat where 
  "SEC1_P192_7_QCAVSx = 0x10bb57020291141981f833b4749e5611034b308e84011d21"

definition SEC1_P192_7_QCAVSy :: nat where 
  "SEC1_P192_7_QCAVSy = 0xe1cacd6b7bd17ed8ddb50b6aee0654c35f2d0eddc1cffcf6"

definition SEC1_P192_7_QCAVS :: "int point" where
  "SEC1_P192_7_QCAVS = Point (int SEC1_P192_7_QCAVSx) (int SEC1_P192_7_QCAVSy)"

definition SEC1_P192_7_dIUT :: nat where 
  "SEC1_P192_7_dIUT = 0x33fed10492afa5bea0333c0af12cac940c4d222455bcd0fe"

definition SEC1_P192_7_QIUTx :: nat where 
  "SEC1_P192_7_QIUTx = 0xef0b28afc41637d737f42e4c8aaceadc84ba2e0b849ca18c"

definition SEC1_P192_7_QIUTy :: nat where 
  "SEC1_P192_7_QIUTy = 0x57797942e552173bba17f73278e029f42335068bd770ddf2"

definition SEC1_P192_7_QIUT :: "int point" where
  "SEC1_P192_7_QIUT = Point (int SEC1_P192_7_QIUTx) (int SEC1_P192_7_QIUTy)"

definition SEC1_P192_7_ZIUT :: nat where 
  "SEC1_P192_7_ZIUT = 0x67cba2cbb69ee78bf1abafb0e6fbe33fa2094c128d59652d"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_7_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_7_dIUT"
  by eval

lemma SEC1_P192_7_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_7_dIUT = SEC1_P192_7_QIUT" 
  by eval

lemma SEC1_P192_7_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_7_QCAVS" 
  by eval

lemma SEC1_P192_7_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_7_QCAVS" 
  using SEC1_P192_7_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_7_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_7_dIUT SEC1_P192_7_QCAVS = Some SEC1_P192_7_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_8›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_8_QCAVSx :: nat where 
  "SEC1_P192_8_QCAVSx = 0x5192fce4185a7758ea1bc56e0e4f4e8b2dce32348d0dced1"

definition SEC1_P192_8_QCAVSy :: nat where 
  "SEC1_P192_8_QCAVSy = 0x20989981beaaf0006d88a96e7971a2fa3a33ba46047fc7ba"

definition SEC1_P192_8_QCAVS :: "int point" where
  "SEC1_P192_8_QCAVS = Point (int SEC1_P192_8_QCAVSx) (int SEC1_P192_8_QCAVSy)"

definition SEC1_P192_8_dIUT :: nat where 
  "SEC1_P192_8_dIUT = 0xf3557c5d70b4c7954960c33568776adbe8e43619abe26b13"

definition SEC1_P192_8_QIUTx :: nat where 
  "SEC1_P192_8_QIUTx = 0xd70112c5f0f0844386494ac1ad99dce2214134176ebfb9af"

definition SEC1_P192_8_QIUTy :: nat where 
  "SEC1_P192_8_QIUTy = 0xd3c187a038510ab31d459e2b7af1a380dd7576af06267548"

definition SEC1_P192_8_QIUT :: "int point" where
  "SEC1_P192_8_QIUT = Point (int SEC1_P192_8_QIUTx) (int SEC1_P192_8_QIUTy)"

definition SEC1_P192_8_ZIUT :: nat where 
  "SEC1_P192_8_ZIUT = 0xcf99a2770a386ca0137d1eca0a226e484297ac3c513f3631"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_8_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_8_dIUT"
  by eval

lemma SEC1_P192_8_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_8_dIUT = SEC1_P192_8_QIUT" 
  by eval

lemma SEC1_P192_8_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_8_QCAVS" 
  by eval

lemma SEC1_P192_8_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_8_QCAVS" 
  using SEC1_P192_8_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_8_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_8_dIUT SEC1_P192_8_QCAVS = Some SEC1_P192_8_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_9›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_9_QCAVSx :: nat where 
  "SEC1_P192_9_QCAVSx = 0x26d019dbe279ead01eed143a91601ada26e2f42225b1c62b"

definition SEC1_P192_9_QCAVSy :: nat where 
  "SEC1_P192_9_QCAVSy = 0x6ca653f08272e0386fc9421fbd580093d7ae6301bca94476"

definition SEC1_P192_9_QCAVS :: "int point" where
  "SEC1_P192_9_QCAVS = Point (int SEC1_P192_9_QCAVSx) (int SEC1_P192_9_QCAVSy)"

definition SEC1_P192_9_dIUT :: nat where 
  "SEC1_P192_9_dIUT = 0x586cfba1c6e81766ed52828f177b1be14ebbc5b83348c311"

definition SEC1_P192_9_QIUTx :: nat where 
  "SEC1_P192_9_QIUTx = 0x58b3c63e56bec9d696bf9a88df2873738391f76368aa2b49"

definition SEC1_P192_9_QIUTy :: nat where 
  "SEC1_P192_9_QIUTy = 0x5776773b261faf7ba2fdc4fe43b92c0b1c7a2fd054a43650"

definition SEC1_P192_9_QIUT :: "int point" where
  "SEC1_P192_9_QIUT = Point (int SEC1_P192_9_QIUTx) (int SEC1_P192_9_QIUTy)"

definition SEC1_P192_9_ZIUT :: nat where 
  "SEC1_P192_9_ZIUT = 0x576331e2b4fb38a112810e1529834de8307fb0a0d2756877"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_9_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_9_dIUT"
  by eval

lemma SEC1_P192_9_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_9_dIUT = SEC1_P192_9_QIUT" 
  by eval

lemma SEC1_P192_9_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_9_QCAVS" 
  by eval

lemma SEC1_P192_9_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_9_QCAVS" 
  using SEC1_P192_9_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_9_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_9_dIUT SEC1_P192_9_QCAVS = Some SEC1_P192_9_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_10›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_10_QCAVSx :: nat where 
  "SEC1_P192_10_QCAVSx = 0x539bc40fe20a0fb267888b647b03eaaf6ec20c02a1e1f8c8"

definition SEC1_P192_10_QCAVSy :: nat where 
  "SEC1_P192_10_QCAVSy = 0x69095e5bb7b4d44c3278a7ee6beca397c45246da9a34c8be"

definition SEC1_P192_10_QCAVS :: "int point" where
  "SEC1_P192_10_QCAVS = Point (int SEC1_P192_10_QCAVSx) (int SEC1_P192_10_QCAVSy)"

definition SEC1_P192_10_dIUT :: nat where 
  "SEC1_P192_10_dIUT = 0xcad8100603a4f65be08d8fc8a1b7e884c5ff65deb3c96d99"

definition SEC1_P192_10_QIUTx :: nat where 
  "SEC1_P192_10_QIUTx = 0xb7fcc0f52c7a411edbed39e10bf02b6ae0f26614c6b325a2"

definition SEC1_P192_10_QIUTy :: nat where 
  "SEC1_P192_10_QIUTy = 0x47483b26eb67776de2b93ab7119d5447573739e3d55e72fb"

definition SEC1_P192_10_QIUT :: "int point" where
  "SEC1_P192_10_QIUT = Point (int SEC1_P192_10_QIUTx) (int SEC1_P192_10_QIUTy)"

definition SEC1_P192_10_ZIUT :: nat where 
  "SEC1_P192_10_ZIUT = 0x902f4501916a0dd945554c3a37b3d780d375a6da713197c4"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_10_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_10_dIUT"
  by eval

lemma SEC1_P192_10_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_10_dIUT = SEC1_P192_10_QIUT" 
  by eval

lemma SEC1_P192_10_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_10_QCAVS" 
  by eval

lemma SEC1_P192_10_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_10_QCAVS" 
  using SEC1_P192_10_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_10_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_10_dIUT SEC1_P192_10_QCAVS = Some SEC1_P192_10_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_11›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_11_QCAVSx :: nat where 
  "SEC1_P192_11_QCAVSx = 0x5d343ddb96318fb4794d10f6c573f99fee5d0d57b996250f"

definition SEC1_P192_11_QCAVSy :: nat where 
  "SEC1_P192_11_QCAVSy = 0x99fbdf9d97dd88ad410235dac36e5b92ce2824b8e587a82c"

definition SEC1_P192_11_QCAVS :: "int point" where
  "SEC1_P192_11_QCAVS = Point (int SEC1_P192_11_QCAVSx) (int SEC1_P192_11_QCAVSy)"

definition SEC1_P192_11_dIUT :: nat where 
  "SEC1_P192_11_dIUT = 0x1edd879cc5c79619cae6c73a691bd5a0395c0ef3b356fcd2"

definition SEC1_P192_11_QIUTx :: nat where 
  "SEC1_P192_11_QIUTx = 0x6ce6adb2c30808f590048c33dffad4524ebf7a5fd39b747b"

definition SEC1_P192_11_QIUTy :: nat where 
  "SEC1_P192_11_QIUTy = 0x4966bd2f3d00569b4d4c0409fbd7a2db752f6d09bca8c25f"

definition SEC1_P192_11_QIUT :: "int point" where
  "SEC1_P192_11_QIUT = Point (int SEC1_P192_11_QIUTx) (int SEC1_P192_11_QIUTy)"

definition SEC1_P192_11_ZIUT :: nat where 
  "SEC1_P192_11_ZIUT = 0x46e4de335054d429863218ae33636fc9b89c628b64b506c7"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_11_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_11_dIUT"
  by eval

lemma SEC1_P192_11_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_11_dIUT = SEC1_P192_11_QIUT" 
  by eval

lemma SEC1_P192_11_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_11_QCAVS" 
  by eval

lemma SEC1_P192_11_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_11_QCAVS" 
  using SEC1_P192_11_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_11_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_11_dIUT SEC1_P192_11_QCAVS = Some SEC1_P192_11_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_12›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_12_QCAVSx :: nat where 
  "SEC1_P192_12_QCAVSx = 0x8d3db9bdce137ffbfb891388c37df6c0cbc90aa5e5376220"

definition SEC1_P192_12_QCAVSy :: nat where 
  "SEC1_P192_12_QCAVSy = 0x135d30b5cb660eef8764ffc744f15c1b5d6dc06ba4416d37"

definition SEC1_P192_12_QCAVS :: "int point" where
  "SEC1_P192_12_QCAVS = Point (int SEC1_P192_12_QCAVSx) (int SEC1_P192_12_QCAVSy)"

definition SEC1_P192_12_dIUT :: nat where 
  "SEC1_P192_12_dIUT = 0x460e452273fe1827602187ad3bebee65cb84423bb4f47537"

definition SEC1_P192_12_QIUTx :: nat where 
  "SEC1_P192_12_QIUTx = 0xd1bd3a3efabf4767fe6380bdf0dbf49d52d4cf0cbb89404c"

definition SEC1_P192_12_QIUTy :: nat where 
  "SEC1_P192_12_QIUTy = 0xc150c2b4c8b3aa35f765f847e4f7f8fd8704d241a181ee99"

definition SEC1_P192_12_QIUT :: "int point" where
  "SEC1_P192_12_QIUT = Point (int SEC1_P192_12_QIUTx) (int SEC1_P192_12_QIUTy)"

definition SEC1_P192_12_ZIUT :: nat where 
  "SEC1_P192_12_ZIUT = 0x1bfe9e5a20ac7a38d8f605b425bb9030be31ef97c101c76c"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_12_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_12_dIUT"
  by eval

lemma SEC1_P192_12_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_12_dIUT = SEC1_P192_12_QIUT" 
  by eval

lemma SEC1_P192_12_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_12_QCAVS" 
  by eval

lemma SEC1_P192_12_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_12_QCAVS" 
  using SEC1_P192_12_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_12_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_12_dIUT SEC1_P192_12_QCAVS = Some SEC1_P192_12_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_13›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_13_QCAVSx :: nat where 
  "SEC1_P192_13_QCAVSx = 0x9e0a6949519c7f5be68c0433c5fdf13064aa13fb29483dc3"

definition SEC1_P192_13_QCAVSy :: nat where 
  "SEC1_P192_13_QCAVSy = 0xe1c8ba63e1f471db23185f50d9c871edea21255b3a63b4b7"

definition SEC1_P192_13_QCAVS :: "int point" where
  "SEC1_P192_13_QCAVS = Point (int SEC1_P192_13_QCAVSx) (int SEC1_P192_13_QCAVSy)"

definition SEC1_P192_13_dIUT :: nat where 
  "SEC1_P192_13_dIUT = 0xb970365008456f8758ecc5a3b33cf3ae6a8d568107a52167"

definition SEC1_P192_13_QIUTx :: nat where 
  "SEC1_P192_13_QIUTx = 0xc1b8610c8c63f8d4abda093b9a11a566044bf65c6faa8999"

definition SEC1_P192_13_QIUTy :: nat where 
  "SEC1_P192_13_QIUTy = 0xa5bc4b3ca095382e9738aee95fe9479b17879b3ad5295559"

definition SEC1_P192_13_QIUT :: "int point" where
  "SEC1_P192_13_QIUT = Point (int SEC1_P192_13_QIUTx) (int SEC1_P192_13_QIUTy)"

definition SEC1_P192_13_ZIUT :: nat where 
  "SEC1_P192_13_ZIUT = 0x0e8c493a4adc445dc9288a3b9b272599224054592d7265b3"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_13_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_13_dIUT"
  by eval

lemma SEC1_P192_13_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_13_dIUT = SEC1_P192_13_QIUT" 
  by eval

lemma SEC1_P192_13_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_13_QCAVS" 
  by eval

lemma SEC1_P192_13_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_13_QCAVS" 
  using SEC1_P192_13_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_13_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_13_dIUT SEC1_P192_13_QCAVS = Some SEC1_P192_13_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_14›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_14_QCAVSx :: nat where 
  "SEC1_P192_14_QCAVSx = 0xbe088238902e9939b3d054eeeb8492daf4bdcf09a2ab77f1"

definition SEC1_P192_14_QCAVSy :: nat where 
  "SEC1_P192_14_QCAVSy = 0x58d6749a3a923dc80440f2661fd35b651617e65294b46375"

definition SEC1_P192_14_QCAVS :: "int point" where
  "SEC1_P192_14_QCAVS = Point (int SEC1_P192_14_QCAVSx) (int SEC1_P192_14_QCAVSy)"

definition SEC1_P192_14_dIUT :: nat where 
  "SEC1_P192_14_dIUT = 0x59c15b8a2464e41dfe4371c7f7dadf470ae425544f8113bd"

definition SEC1_P192_14_QIUTx :: nat where 
  "SEC1_P192_14_QIUTx = 0x1fe776f73567b6ac0b0d6764164de6c5be751ba8d1ff455e"

definition SEC1_P192_14_QIUTy :: nat where 
  "SEC1_P192_14_QIUTy = 0x4c160bf38afb2b71f684261664115ce874553e8b059432d2"

definition SEC1_P192_14_QIUT :: "int point" where
  "SEC1_P192_14_QIUT = Point (int SEC1_P192_14_QIUTx) (int SEC1_P192_14_QIUTy)"

definition SEC1_P192_14_ZIUT :: nat where 
  "SEC1_P192_14_ZIUT = 0x0f1991086b455ded6a1c4146f7bf59fe9b495de566ebc6bf"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_14_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_14_dIUT"
  by eval

lemma SEC1_P192_14_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_14_dIUT = SEC1_P192_14_QIUT" 
  by eval

lemma SEC1_P192_14_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_14_QCAVS" 
  by eval

lemma SEC1_P192_14_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_14_QCAVS" 
  using SEC1_P192_14_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_14_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_14_dIUT SEC1_P192_14_QCAVS = Some SEC1_P192_14_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_15›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_15_QCAVSx :: nat where 
  "SEC1_P192_15_QCAVSx = 0xbf5ae05025e1be617e666d87a4168363873d5761b376b503"

definition SEC1_P192_15_QCAVSy :: nat where 
  "SEC1_P192_15_QCAVSy = 0xe1e6e38b372b6bee0ff5b3502d83735e3b2c26825e4f0fcc"

definition SEC1_P192_15_QCAVS :: "int point" where
  "SEC1_P192_15_QCAVS = Point (int SEC1_P192_15_QCAVSx) (int SEC1_P192_15_QCAVSy)"

definition SEC1_P192_15_dIUT :: nat where 
  "SEC1_P192_15_dIUT = 0xa6e9b885c66b959d1fc2708d591b6d3228e49eb98f726d61"

definition SEC1_P192_15_QIUTx :: nat where 
  "SEC1_P192_15_QIUTx = 0x632bb7651dbf49dde9dd125d13fb234e06617723beed3d1b"

definition SEC1_P192_15_QIUTy :: nat where 
  "SEC1_P192_15_QIUTy = 0xf4ad5209638488397c5f44f994dd7479807e79f4887d2e71"

definition SEC1_P192_15_QIUT :: "int point" where
  "SEC1_P192_15_QIUT = Point (int SEC1_P192_15_QIUTx) (int SEC1_P192_15_QIUTy)"

definition SEC1_P192_15_ZIUT :: nat where 
  "SEC1_P192_15_ZIUT = 0xb30f2127c34df35aaa91dbf0bbe15798e799a03ed11698c1"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_15_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_15_dIUT"
  by eval

lemma SEC1_P192_15_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_15_dIUT = SEC1_P192_15_QIUT" 
  by eval

lemma SEC1_P192_15_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_15_QCAVS" 
  by eval

lemma SEC1_P192_15_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_15_QCAVS" 
  using SEC1_P192_15_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_15_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_15_dIUT SEC1_P192_15_QCAVS = Some SEC1_P192_15_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_16›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_16_QCAVSx :: nat where 
  "SEC1_P192_16_QCAVSx = 0x6cc4feed84c7ab0d09005d660ed34de6955a9461c4138d11"

definition SEC1_P192_16_QCAVSy :: nat where 
  "SEC1_P192_16_QCAVSy = 0x31225f33864ed48da06fa45a913b46cf42557742e35085e6"

definition SEC1_P192_16_QCAVS :: "int point" where
  "SEC1_P192_16_QCAVS = Point (int SEC1_P192_16_QCAVSx) (int SEC1_P192_16_QCAVSy)"

definition SEC1_P192_16_dIUT :: nat where 
  "SEC1_P192_16_dIUT = 0xbdb754096ffbfbd8b0f3cb046ccb7ca149c4e7192067a3ee"

definition SEC1_P192_16_QIUTx :: nat where 
  "SEC1_P192_16_QIUTx = 0xd9c098d421d741f6faab116f3e4731d28c5558e19fe112a1"

definition SEC1_P192_16_QIUTy :: nat where 
  "SEC1_P192_16_QIUTy = 0x38d4dc48ccdb1d3ed8d31fd06784a4f87a68aec1cbd5b08f"

definition SEC1_P192_16_QIUT :: "int point" where
  "SEC1_P192_16_QIUT = Point (int SEC1_P192_16_QIUTx) (int SEC1_P192_16_QIUTy)"

definition SEC1_P192_16_ZIUT :: nat where 
  "SEC1_P192_16_ZIUT = 0x64a5c246599d3e8177a2402a1110eb81e6c456ab4edb5127"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_16_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_16_dIUT"
  by eval

lemma SEC1_P192_16_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_16_dIUT = SEC1_P192_16_QIUT" 
  by eval

lemma SEC1_P192_16_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_16_QCAVS" 
  by eval

lemma SEC1_P192_16_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_16_QCAVS" 
  using SEC1_P192_16_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_16_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_16_dIUT SEC1_P192_16_QCAVS = Some SEC1_P192_16_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_17›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_17_QCAVSx :: nat where 
  "SEC1_P192_17_QCAVSx = 0x36157315bee7afedded58c4e8ba14d3421c401e51135bcc9"

definition SEC1_P192_17_QCAVSy :: nat where 
  "SEC1_P192_17_QCAVSy = 0x37c297ca703f77c52bb062d8ce971db84097ba0c753a418f"

definition SEC1_P192_17_QCAVS :: "int point" where
  "SEC1_P192_17_QCAVS = Point (int SEC1_P192_17_QCAVSx) (int SEC1_P192_17_QCAVSy)"

definition SEC1_P192_17_dIUT :: nat where 
  "SEC1_P192_17_dIUT = 0xd5bcf2534dafc3d99964c7bd63ab7bd15999fe56dd969c42"

definition SEC1_P192_17_QIUTx :: nat where 
  "SEC1_P192_17_QIUTx = 0xfda1d5d28d6fe0e7909d6a8bafa7824db5572ab92ffe7de6"

definition SEC1_P192_17_QIUTy :: nat where 
  "SEC1_P192_17_QIUTy = 0x134a297c1d9c8bbab249abacd951ed11e5a99f92e7991572"

definition SEC1_P192_17_QIUT :: "int point" where
  "SEC1_P192_17_QIUT = Point (int SEC1_P192_17_QIUTx) (int SEC1_P192_17_QIUTy)"

definition SEC1_P192_17_ZIUT :: nat where 
  "SEC1_P192_17_ZIUT = 0x017b8ca53c82fab163da2ab783966a39e061b32c8cfa334d"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_17_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_17_dIUT"
  by eval

lemma SEC1_P192_17_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_17_dIUT = SEC1_P192_17_QIUT" 
  by eval

lemma SEC1_P192_17_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_17_QCAVS" 
  by eval

lemma SEC1_P192_17_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_17_QCAVS" 
  using SEC1_P192_17_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_17_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_17_dIUT SEC1_P192_17_QCAVS = Some SEC1_P192_17_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_18›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_18_QCAVSx :: nat where 
  "SEC1_P192_18_QCAVSx = 0x98464d47f0256f8292e027e8c92582ea77cf9051f5ce8e5d"

definition SEC1_P192_18_QCAVSy :: nat where 
  "SEC1_P192_18_QCAVSy = 0x449552ef7578be96236fe5ed9d0643c0bb6c5a9134b0108d"

definition SEC1_P192_18_QCAVS :: "int point" where
  "SEC1_P192_18_QCAVS = Point (int SEC1_P192_18_QCAVSx) (int SEC1_P192_18_QCAVSy)"

definition SEC1_P192_18_dIUT :: nat where 
  "SEC1_P192_18_dIUT = 0x43d4b9df1053be5b4268104c02244d3bf9594b010b46a8b2"

definition SEC1_P192_18_QIUTx :: nat where 
  "SEC1_P192_18_QIUTx = 0xc3020b7091463d788f1f1d76f7cfeec82ecdb3b7d99c345c"

definition SEC1_P192_18_QIUTy :: nat where 
  "SEC1_P192_18_QIUTy = 0x9a7710d5179591d8f3df0aa122301768ae7db7eee2d7f583"

definition SEC1_P192_18_QIUT :: "int point" where
  "SEC1_P192_18_QIUT = Point (int SEC1_P192_18_QIUTx) (int SEC1_P192_18_QIUTy)"

definition SEC1_P192_18_ZIUT :: nat where 
  "SEC1_P192_18_ZIUT = 0x340ef3db3dbebdd91c62c3d4e1a3da2c7c52a3338b865259"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_18_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_18_dIUT"
  by eval

lemma SEC1_P192_18_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_18_dIUT = SEC1_P192_18_QIUT" 
  by eval

lemma SEC1_P192_18_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_18_QCAVS" 
  by eval

lemma SEC1_P192_18_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_18_QCAVS" 
  using SEC1_P192_18_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_18_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_18_dIUT SEC1_P192_18_QCAVS = Some SEC1_P192_18_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_19›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_19_QCAVSx :: nat where 
  "SEC1_P192_19_QCAVSx = 0x563eb66c334cf6f123bf04c7803b48a3110214237e983bf5"

definition SEC1_P192_19_QCAVSy :: nat where 
  "SEC1_P192_19_QCAVSy = 0x0f351104819199ef07c9a6051d20758f3af79027ea66a53f"

definition SEC1_P192_19_QCAVS :: "int point" where
  "SEC1_P192_19_QCAVS = Point (int SEC1_P192_19_QCAVSx) (int SEC1_P192_19_QCAVSy)"

definition SEC1_P192_19_dIUT :: nat where 
  "SEC1_P192_19_dIUT = 0x94cac2c2ca714746401670d94edbf3f677867b5a03bee7ad"

definition SEC1_P192_19_QIUTx :: nat where 
  "SEC1_P192_19_QIUTx = 0xb18554a2e743ef0aa2f040987c4c451004e096df3d80ddae"

definition SEC1_P192_19_QIUTy :: nat where 
  "SEC1_P192_19_QIUTy = 0x6e3e2c618f896e36ba620077684b70a05ffb79bf5e6c7640"

definition SEC1_P192_19_QIUT :: "int point" where
  "SEC1_P192_19_QIUT = Point (int SEC1_P192_19_QIUTx) (int SEC1_P192_19_QIUTy)"

definition SEC1_P192_19_ZIUT :: nat where 
  "SEC1_P192_19_ZIUT = 0x2162144921df5103d0e6a650fb13fd246f4738d0896ce92f"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_19_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_19_dIUT"
  by eval

lemma SEC1_P192_19_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_19_dIUT = SEC1_P192_19_QIUT" 
  by eval

lemma SEC1_P192_19_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_19_QCAVS" 
  by eval

lemma SEC1_P192_19_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_19_QCAVS" 
  using SEC1_P192_19_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_19_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_19_dIUT SEC1_P192_19_QCAVS = Some SEC1_P192_19_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_20›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_20_QCAVSx :: nat where 
  "SEC1_P192_20_QCAVSx = 0x86828c4ac92b5507618aec7873a1d4fc6543c5be33cf3078"

definition SEC1_P192_20_QCAVSy :: nat where 
  "SEC1_P192_20_QCAVSy = 0xb22ca72437545e10d6d4f052422eb898b737a4b8543ee550"

definition SEC1_P192_20_QCAVS :: "int point" where
  "SEC1_P192_20_QCAVS = Point (int SEC1_P192_20_QCAVSx) (int SEC1_P192_20_QCAVSy)"

definition SEC1_P192_20_dIUT :: nat where 
  "SEC1_P192_20_dIUT = 0x2a3a9e33c8cc3107a9f9265c3bdea1206570e86f92ac7014"

definition SEC1_P192_20_QIUTx :: nat where 
  "SEC1_P192_20_QIUTx = 0xa7ba38be1bc669dd23ccfcee0645b1f0db8cf942deafaeb6"

definition SEC1_P192_20_QIUTy :: nat where 
  "SEC1_P192_20_QIUTy = 0xb82db79d80cd0e37f28d4163adc389dee8fc7797b5c9831b"

definition SEC1_P192_20_QIUT :: "int point" where
  "SEC1_P192_20_QIUT = Point (int SEC1_P192_20_QIUTx) (int SEC1_P192_20_QIUTy)"

definition SEC1_P192_20_ZIUT :: nat where 
  "SEC1_P192_20_ZIUT = 0x4c69e7feed4b11159adfc16a6047a92572ea44e0740b23af"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_20_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_20_dIUT"
  by eval

lemma SEC1_P192_20_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_20_dIUT = SEC1_P192_20_QIUT" 
  by eval

lemma SEC1_P192_20_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_20_QCAVS" 
  by eval

lemma SEC1_P192_20_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_20_QCAVS" 
  using SEC1_P192_20_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_20_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_20_dIUT SEC1_P192_20_QCAVS = Some SEC1_P192_20_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_21›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_21_QCAVSx :: nat where 
  "SEC1_P192_21_QCAVSx = 0x6700a102437781a9581da2bc25ced5abf419da91d3c803df"

definition SEC1_P192_21_QCAVSy :: nat where 
  "SEC1_P192_21_QCAVSy = 0x71396c9cf08bcd91854e3e6e42d8c657ce0f27ab77a9dc4b"

definition SEC1_P192_21_QCAVS :: "int point" where
  "SEC1_P192_21_QCAVS = Point (int SEC1_P192_21_QCAVSx) (int SEC1_P192_21_QCAVSy)"

definition SEC1_P192_21_dIUT :: nat where 
  "SEC1_P192_21_dIUT = 0x4a6b78a98ac98fa8e99a8ece08ec0251125f85c6fd0e289b"

definition SEC1_P192_21_QIUTx :: nat where 
  "SEC1_P192_21_QIUTx = 0xe769dbbcd5ce2d83514b768d3d2d5aa0bcd8f66af15f5500"

definition SEC1_P192_21_QIUTy :: nat where 
  "SEC1_P192_21_QIUTy = 0x2fc6d0b039e0f28f74fbeffe9e883d4dd72296e4e95cae71"

definition SEC1_P192_21_QIUT :: "int point" where
  "SEC1_P192_21_QIUT = Point (int SEC1_P192_21_QIUTx) (int SEC1_P192_21_QIUTy)"

definition SEC1_P192_21_ZIUT :: nat where 
  "SEC1_P192_21_ZIUT = 0x46072acefd67bff50de355ca7a31fa6be59f26e467587259"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_21_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_21_dIUT"
  by eval

lemma SEC1_P192_21_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_21_dIUT = SEC1_P192_21_QIUT" 
  by eval

lemma SEC1_P192_21_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_21_QCAVS" 
  by eval

lemma SEC1_P192_21_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_21_QCAVS" 
  using SEC1_P192_21_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_21_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_21_dIUT SEC1_P192_21_QCAVS = Some SEC1_P192_21_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_22›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_22_QCAVSx :: nat where 
  "SEC1_P192_22_QCAVSx = 0xa82f354cf97bee5d22dc6c079f2902ead44d96a8f614f178"

definition SEC1_P192_22_QCAVSy :: nat where 
  "SEC1_P192_22_QCAVSy = 0xa654a9aa8a1a0802f2ce0ee8a0f4ebe96dee1b37464b1ff2"

definition SEC1_P192_22_QCAVS :: "int point" where
  "SEC1_P192_22_QCAVS = Point (int SEC1_P192_22_QCAVSx) (int SEC1_P192_22_QCAVSy)"

definition SEC1_P192_22_dIUT :: nat where 
  "SEC1_P192_22_dIUT = 0xc5a6491d78844d6617ef33be6b8bd54da221450885d5950f"

definition SEC1_P192_22_QIUTx :: nat where 
  "SEC1_P192_22_QIUTx = 0xdb1b24f7466bc154e9d7d2c3ca52dcfe0bfc9563c5fdb6f3"

definition SEC1_P192_22_QIUTy :: nat where 
  "SEC1_P192_22_QIUTy = 0x1c74fbbf5bd99921f1a9a744f8e1cf770bd6a76a772b3003"

definition SEC1_P192_22_QIUT :: "int point" where
  "SEC1_P192_22_QIUT = Point (int SEC1_P192_22_QIUTx) (int SEC1_P192_22_QIUTy)"

definition SEC1_P192_22_ZIUT :: nat where 
  "SEC1_P192_22_ZIUT = 0xec5580eabca9f3389d2b427ddf6e49e26d629afd03fa766e"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_22_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_22_dIUT"
  by eval

lemma SEC1_P192_22_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_22_dIUT = SEC1_P192_22_QIUT" 
  by eval

lemma SEC1_P192_22_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_22_QCAVS" 
  by eval

lemma SEC1_P192_22_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_22_QCAVS" 
  using SEC1_P192_22_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_22_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_22_dIUT SEC1_P192_22_QCAVS = Some SEC1_P192_22_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_23›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_23_QCAVSx :: nat where 
  "SEC1_P192_23_QCAVSx = 0x3cec21b28668a12a2cf78e1a8e55d0efe065152fffc34718"

definition SEC1_P192_23_QCAVSy :: nat where 
  "SEC1_P192_23_QCAVSy = 0x1029557beba4ff1992bd21c23cb4825f6dae70e3318fd1ca"

definition SEC1_P192_23_QCAVS :: "int point" where
  "SEC1_P192_23_QCAVS = Point (int SEC1_P192_23_QCAVSx) (int SEC1_P192_23_QCAVSy)"

definition SEC1_P192_23_dIUT :: nat where 
  "SEC1_P192_23_dIUT = 0x2ba2703c5e23f6463c5b88dc37292fabd3399b5e1fb67c05"

definition SEC1_P192_23_QIUTx :: nat where 
  "SEC1_P192_23_QIUTx = 0x7543148906cef9b37a71a7c08363cdd3bba50142d65241aa"

definition SEC1_P192_23_QIUTy :: nat where 
  "SEC1_P192_23_QIUTy = 0x8b3a6973de8dc271e27c1ead1e962fdaae3710c724daac38"

definition SEC1_P192_23_QIUT :: "int point" where
  "SEC1_P192_23_QIUT = Point (int SEC1_P192_23_QIUTx) (int SEC1_P192_23_QIUTy)"

definition SEC1_P192_23_ZIUT :: nat where 
  "SEC1_P192_23_ZIUT = 0x7f3929dd3cbf7673bc30d859d90b880307475f800660ea32"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_23_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_23_dIUT"
  by eval

lemma SEC1_P192_23_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_23_dIUT = SEC1_P192_23_QIUT" 
  by eval

lemma SEC1_P192_23_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_23_QCAVS" 
  by eval

lemma SEC1_P192_23_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_23_QCAVS" 
  using SEC1_P192_23_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_23_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_23_dIUT SEC1_P192_23_QCAVS = Some SEC1_P192_23_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P192_24›

subsubsection‹Given Test Vector Values›

definition SEC1_P192_24_QCAVSx :: nat where 
  "SEC1_P192_24_QCAVSx = 0x7082644715b8b731f8228b5118e7270d34d181f361a221fc"

definition SEC1_P192_24_QCAVSy :: nat where 
  "SEC1_P192_24_QCAVSy = 0x464649d6c88ca89614488a1cc7b8442bb42f9fb3020a3d76"

definition SEC1_P192_24_QCAVS :: "int point" where
  "SEC1_P192_24_QCAVS = Point (int SEC1_P192_24_QCAVSx) (int SEC1_P192_24_QCAVSy)"

definition SEC1_P192_24_dIUT :: nat where 
  "SEC1_P192_24_dIUT = 0x836118c6248f882e9147976f764826c1a28755a6102977d5"

definition SEC1_P192_24_QIUTx :: nat where 
  "SEC1_P192_24_QIUTx = 0xfcd345a976c720caaa97de6697226825615e1287a9eff67e"

definition SEC1_P192_24_QIUTy :: nat where 
  "SEC1_P192_24_QIUTy = 0x58ea42edbeeafca9ff44cfd7f29abd2cbde7626d79e422c9"

definition SEC1_P192_24_QIUT :: "int point" where
  "SEC1_P192_24_QIUT = Point (int SEC1_P192_24_QIUTx) (int SEC1_P192_24_QIUTy)"

definition SEC1_P192_24_ZIUT :: nat where 
  "SEC1_P192_24_ZIUT = 0x72e88f3ea67d46d46dbf83926e7e2a6b85b54536741e6d2c"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P192_24_dIUT_valid: 
  "SEC1_P192_ECprivateKeyValid SEC1_P192_24_dIUT"
  by eval

lemma SEC1_P192_24_IUT_validPair: 
  "SEC1_P192_ECkeyGen SEC1_P192_24_dIUT = SEC1_P192_24_QIUT" 
  by eval

lemma SEC1_P192_24_QCAVS_partialValid: 
  "SEC1_P192_ECpublicKeyPartialValid SEC1_P192_24_QCAVS" 
  by eval

lemma SEC1_P192_24_QCAVS_validPublicKey: 
  "SEC1_P192_ECpublicKeyValid SEC1_P192_24_QCAVS" 
  using SEC1_P192_24_QCAVS_partialValid SEC1_P192.partValidImpliesValidIFheq1 P192_h_def by blast

lemma SEC1_P192_24_ECDHprim_Check: 
  "SEC1_P192_ECDHprim SEC1_P192_24_dIUT SEC1_P192_24_QCAVS = Some SEC1_P192_24_ZIUT"
  by eval
  
section‹NIST Test Vectors for ECDH: Curve P-224›
  
subsection‹Test Vector: SEC1_P224_0›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_0_QCAVSx :: nat where 
  "SEC1_P224_0_QCAVSx = 0xaf33cd0629bc7e996320a3f40368f74de8704fa37b8fab69abaae280"

definition SEC1_P224_0_QCAVSy :: nat where 
  "SEC1_P224_0_QCAVSy = 0x882092ccbba7930f419a8a4f9bb16978bbc3838729992559a6f2e2d7"

definition SEC1_P224_0_QCAVS :: "int point" where
  "SEC1_P224_0_QCAVS = Point (int SEC1_P224_0_QCAVSx) (int SEC1_P224_0_QCAVSy)"

definition SEC1_P224_0_dIUT :: nat where 
  "SEC1_P224_0_dIUT = 0x8346a60fc6f293ca5a0d2af68ba71d1dd389e5e40837942df3e43cbd"

definition SEC1_P224_0_QIUTx :: nat where 
  "SEC1_P224_0_QIUTx = 0x8de2e26adf72c582d6568ef638c4fd59b18da171bdf501f1d929e048"

definition SEC1_P224_0_QIUTy :: nat where 
  "SEC1_P224_0_QIUTy = 0x4a68a1c2b0fb22930d120555c1ece50ea98dea8407f71be36efac0de"

definition SEC1_P224_0_QIUT :: "int point" where
  "SEC1_P224_0_QIUT = Point (int SEC1_P224_0_QIUTx) (int SEC1_P224_0_QIUTy)"

definition SEC1_P224_0_ZIUT :: nat where 
  "SEC1_P224_0_ZIUT = 0x7d96f9a3bd3c05cf5cc37feb8b9d5209d5c2597464dec3e9983743e8"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_0_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_0_dIUT"
  by eval

lemma SEC1_P224_0_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_0_dIUT = SEC1_P224_0_QIUT" 
  by eval

lemma SEC1_P224_0_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_0_QCAVS" 
  by eval

lemma SEC1_P224_0_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_0_QCAVS" 
  using SEC1_P224_0_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_0_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_0_dIUT SEC1_P224_0_QCAVS = Some SEC1_P224_0_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_1›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_1_QCAVSx :: nat where 
  "SEC1_P224_1_QCAVSx = 0x13bfcd4f8e9442393cab8fb46b9f0566c226b22b37076976f0617a46"

definition SEC1_P224_1_QCAVSy :: nat where 
  "SEC1_P224_1_QCAVSy = 0xeeb2427529b288c63c2f8963c1e473df2fca6caa90d52e2f8db56dd4"

definition SEC1_P224_1_QCAVS :: "int point" where
  "SEC1_P224_1_QCAVS = Point (int SEC1_P224_1_QCAVSx) (int SEC1_P224_1_QCAVSy)"

definition SEC1_P224_1_dIUT :: nat where 
  "SEC1_P224_1_dIUT = 0x043cb216f4b72cdf7629d63720a54aee0c99eb32d74477dac0c2f73d"

definition SEC1_P224_1_QIUTx :: nat where 
  "SEC1_P224_1_QIUTx = 0x2f90f5c8eac9c7decdbb97b6c2f715ab725e4fe40fe6d746efbf4e1b"

definition SEC1_P224_1_QIUTy :: nat where 
  "SEC1_P224_1_QIUTy = 0x66897351454f927a309b269c5a6d31338be4c19a5acfc32cf656f45c"

definition SEC1_P224_1_QIUT :: "int point" where
  "SEC1_P224_1_QIUT = Point (int SEC1_P224_1_QIUTx) (int SEC1_P224_1_QIUTy)"

definition SEC1_P224_1_ZIUT :: nat where 
  "SEC1_P224_1_ZIUT = 0xee93ce06b89ff72009e858c68eb708e7bc79ee0300f73bed69bbca09"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_1_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_1_dIUT"
  by eval

lemma SEC1_P224_1_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_1_dIUT = SEC1_P224_1_QIUT" 
  by eval

lemma SEC1_P224_1_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_1_QCAVS" 
  by eval

lemma SEC1_P224_1_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_1_QCAVS" 
  using SEC1_P224_1_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_1_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_1_dIUT SEC1_P224_1_QCAVS = Some SEC1_P224_1_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_2›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_2_QCAVSx :: nat where 
  "SEC1_P224_2_QCAVSx = 0x756dd806b9d9c34d899691ecb45b771af468ec004486a0fdd283411e"

definition SEC1_P224_2_QCAVSy :: nat where 
  "SEC1_P224_2_QCAVSy = 0x4d02c2ca617bb2c5d9613f25dd72413d229fd2901513aa29504eeefb"

definition SEC1_P224_2_QCAVS :: "int point" where
  "SEC1_P224_2_QCAVS = Point (int SEC1_P224_2_QCAVSx) (int SEC1_P224_2_QCAVSy)"

definition SEC1_P224_2_dIUT :: nat where 
  "SEC1_P224_2_dIUT = 0x5ad0dd6dbabb4f3c2ea5fe32e561b2ca55081486df2c7c15c9622b08"

definition SEC1_P224_2_QIUTx :: nat where 
  "SEC1_P224_2_QIUTx = 0x005bca45d793e7fe99a843704ed838315ab14a5f6277507e9bc37531"

definition SEC1_P224_2_QIUTy :: nat where 
  "SEC1_P224_2_QIUTy = 0x43e9d421e1486ae5893bfd23c210e5c140d7c6b1ada59d842c9a98de"

definition SEC1_P224_2_QIUT :: "int point" where
  "SEC1_P224_2_QIUT = Point (int SEC1_P224_2_QIUTx) (int SEC1_P224_2_QIUTy)"

definition SEC1_P224_2_ZIUT :: nat where 
  "SEC1_P224_2_ZIUT = 0x3fcc01e34d4449da2a974b23fc36f9566754259d39149790cfa1ebd3"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_2_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_2_dIUT"
  by eval

lemma SEC1_P224_2_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_2_dIUT = SEC1_P224_2_QIUT" 
  by eval

lemma SEC1_P224_2_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_2_QCAVS" 
  by eval

lemma SEC1_P224_2_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_2_QCAVS" 
  using SEC1_P224_2_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_2_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_2_dIUT SEC1_P224_2_QCAVS = Some SEC1_P224_2_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_3›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_3_QCAVSx :: nat where 
  "SEC1_P224_3_QCAVSx = 0x0f537bf1c1122c55656d25e8aa8417e0b44b1526ae0523144f9921c4"

definition SEC1_P224_3_QCAVSy :: nat where 
  "SEC1_P224_3_QCAVSy = 0xf79b26d30e491a773696cc2c79b4f0596bc5b9eebaf394d162fb8684"

definition SEC1_P224_3_QCAVS :: "int point" where
  "SEC1_P224_3_QCAVS = Point (int SEC1_P224_3_QCAVSx) (int SEC1_P224_3_QCAVSy)"

definition SEC1_P224_3_dIUT :: nat where 
  "SEC1_P224_3_dIUT = 0x0aa6ff55a5d820efcb4e7d10b845ea3c9f9bc5dff86106db85318e22"

definition SEC1_P224_3_QIUTx :: nat where 
  "SEC1_P224_3_QIUTx = 0x2f96754131e0968198aa78fbe8c201dc5f3581c792de487340d32448"

definition SEC1_P224_3_QIUTy :: nat where 
  "SEC1_P224_3_QIUTy = 0x61e8a5cd79615203b6d89e9496f9e236fe3b6be8731e743d615519c6"

definition SEC1_P224_3_QIUT :: "int point" where
  "SEC1_P224_3_QIUT = Point (int SEC1_P224_3_QIUTx) (int SEC1_P224_3_QIUTy)"

definition SEC1_P224_3_ZIUT :: nat where 
  "SEC1_P224_3_ZIUT = 0x49129628b23afcef48139a3f6f59ff5e9811aa746aa4ff33c24bb940"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_3_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_3_dIUT"
  by eval

lemma SEC1_P224_3_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_3_dIUT = SEC1_P224_3_QIUT" 
  by eval

lemma SEC1_P224_3_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_3_QCAVS" 
  by eval

lemma SEC1_P224_3_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_3_QCAVS" 
  using SEC1_P224_3_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_3_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_3_dIUT SEC1_P224_3_QCAVS = Some SEC1_P224_3_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_4›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_4_QCAVSx :: nat where 
  "SEC1_P224_4_QCAVSx = 0x2b3631d2b06179b3174a100f7f57131eeea8947be0786c3dc64b2239"

definition SEC1_P224_4_QCAVSy :: nat where 
  "SEC1_P224_4_QCAVSy = 0x83de29ae3dad31adc0236c6de7f14561ca2ea083c5270c78a2e6cbc0"

definition SEC1_P224_4_QCAVS :: "int point" where
  "SEC1_P224_4_QCAVS = Point (int SEC1_P224_4_QCAVSx) (int SEC1_P224_4_QCAVSy)"

definition SEC1_P224_4_dIUT :: nat where 
  "SEC1_P224_4_dIUT = 0xefe6e6e25affaf54c98d002abbc6328da159405a1b752e32dc23950a"

definition SEC1_P224_4_QIUTx :: nat where 
  "SEC1_P224_4_QIUTx = 0x355e962920bde043695f6bffb4b355c63da6f5de665ed46f2ec817e2"

definition SEC1_P224_4_QIUTy :: nat where 
  "SEC1_P224_4_QIUTy = 0x748e095368f62e1d364edd461719793b404adbdaacbcadd88922ff37"

definition SEC1_P224_4_QIUT :: "int point" where
  "SEC1_P224_4_QIUT = Point (int SEC1_P224_4_QIUTx) (int SEC1_P224_4_QIUTy)"

definition SEC1_P224_4_ZIUT :: nat where 
  "SEC1_P224_4_ZIUT = 0xfcdc69a40501d308a6839653a8f04309ec00233949522902ffa5eac6"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_4_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_4_dIUT"
  by eval

lemma SEC1_P224_4_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_4_dIUT = SEC1_P224_4_QIUT" 
  by eval

lemma SEC1_P224_4_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_4_QCAVS" 
  by eval

lemma SEC1_P224_4_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_4_QCAVS" 
  using SEC1_P224_4_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_4_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_4_dIUT SEC1_P224_4_QCAVS = Some SEC1_P224_4_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_5›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_5_QCAVSx :: nat where 
  "SEC1_P224_5_QCAVSx = 0x4511403de29059f69a475c5a6a5f6cabed5d9f014436a8cb70a02338"

definition SEC1_P224_5_QCAVSy :: nat where 
  "SEC1_P224_5_QCAVSy = 0x7d2d1b62aa046df9340f9c37a087a06b32cf7f08a223f992812a828b"

definition SEC1_P224_5_QCAVS :: "int point" where
  "SEC1_P224_5_QCAVS = Point (int SEC1_P224_5_QCAVSx) (int SEC1_P224_5_QCAVSy)"

definition SEC1_P224_5_dIUT :: nat where 
  "SEC1_P224_5_dIUT = 0x61cb2932524001e5e9eeed6df7d9c8935ee3322029edd7aa8acbfd51"

definition SEC1_P224_5_QIUTx :: nat where 
  "SEC1_P224_5_QIUTx = 0xd50e4adabfd989d7dbc7cf4052546cc7c447a97630436997ad4b9536"

definition SEC1_P224_5_QIUTy :: nat where 
  "SEC1_P224_5_QIUTy = 0x5bea503473c5eaef9552d42c40b1f2f7ca292733b255b9bbe1b12337"

definition SEC1_P224_5_QIUT :: "int point" where
  "SEC1_P224_5_QIUT = Point (int SEC1_P224_5_QIUTx) (int SEC1_P224_5_QIUTy)"

definition SEC1_P224_5_ZIUT :: nat where 
  "SEC1_P224_5_ZIUT = 0x827e9025cb62e0e837c596063f3b9b5a0f7afd8d8783200086d61ec1"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_5_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_5_dIUT"
  by eval

lemma SEC1_P224_5_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_5_dIUT = SEC1_P224_5_QIUT" 
  by eval

lemma SEC1_P224_5_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_5_QCAVS" 
  by eval

lemma SEC1_P224_5_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_5_QCAVS" 
  using SEC1_P224_5_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_5_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_5_dIUT SEC1_P224_5_QCAVS = Some SEC1_P224_5_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_6›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_6_QCAVSx :: nat where 
  "SEC1_P224_6_QCAVSx = 0x314a0b26dd31c248845d7cc17b61cad4608259bed85a58d1f1ffd378"

definition SEC1_P224_6_QCAVSy :: nat where 
  "SEC1_P224_6_QCAVSy = 0x66e4b350352e119eecada382907f3619fd748ea73ae4899dfd496302"

definition SEC1_P224_6_QCAVS :: "int point" where
  "SEC1_P224_6_QCAVS = Point (int SEC1_P224_6_QCAVSx) (int SEC1_P224_6_QCAVSy)"

definition SEC1_P224_6_dIUT :: nat where 
  "SEC1_P224_6_dIUT = 0x8c7ace347171f92def98d845475fc82e1d1496da81ee58f505b985fa"

definition SEC1_P224_6_QIUTx :: nat where 
  "SEC1_P224_6_QIUTx = 0xb1a8dcac89aca2799320b451df1c7ff4d97567abb68141c0d95fc2aa"

definition SEC1_P224_6_QIUTy :: nat where 
  "SEC1_P224_6_QIUTy = 0x3524950902b1510bdc987d860afc27ad871ceaea66935abd3c0a99a8"

definition SEC1_P224_6_QIUT :: "int point" where
  "SEC1_P224_6_QIUT = Point (int SEC1_P224_6_QIUTx) (int SEC1_P224_6_QIUTy)"

definition SEC1_P224_6_ZIUT :: nat where 
  "SEC1_P224_6_ZIUT = 0x335ba51228d94acbed851ca7821c801d5cb1c7975d7aa90a7159f8fa"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_6_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_6_dIUT"
  by eval

lemma SEC1_P224_6_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_6_dIUT = SEC1_P224_6_QIUT" 
  by eval

lemma SEC1_P224_6_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_6_QCAVS" 
  by eval

lemma SEC1_P224_6_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_6_QCAVS" 
  using SEC1_P224_6_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_6_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_6_dIUT SEC1_P224_6_QCAVS = Some SEC1_P224_6_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_7›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_7_QCAVSx :: nat where 
  "SEC1_P224_7_QCAVSx = 0xabe6843beec2fd9e5fb64730d0be4d165438ce922ed75dd80b4603e5"

definition SEC1_P224_7_QCAVSy :: nat where 
  "SEC1_P224_7_QCAVSy = 0x6afe8673a96c4ba9900ad85995e631e436c6cc88a2c2b47b7c4886b8"

definition SEC1_P224_7_QCAVS :: "int point" where
  "SEC1_P224_7_QCAVS = Point (int SEC1_P224_7_QCAVSx) (int SEC1_P224_7_QCAVSy)"

definition SEC1_P224_7_dIUT :: nat where 
  "SEC1_P224_7_dIUT = 0x382feb9b9ba10f189d99e71a89cdfe44cb554cec13a212840977fb68"

definition SEC1_P224_7_QIUTx :: nat where 
  "SEC1_P224_7_QIUTx = 0xabb6f1e3773ff8fc73aea2a0b107809ce70adcefed6e41fc5cb43045"

definition SEC1_P224_7_QIUTy :: nat where 
  "SEC1_P224_7_QIUTy = 0xa963897ae906c10a055eeadb97ffdd6f748d3e5621e5fff304e48ba7"

definition SEC1_P224_7_QIUT :: "int point" where
  "SEC1_P224_7_QIUT = Point (int SEC1_P224_7_QIUTx) (int SEC1_P224_7_QIUTy)"

definition SEC1_P224_7_ZIUT :: nat where 
  "SEC1_P224_7_ZIUT = 0x8c2e627594206b34f7356d3426eb3d79f518ef843fbe94014cceace3"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_7_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_7_dIUT"
  by eval

lemma SEC1_P224_7_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_7_dIUT = SEC1_P224_7_QIUT" 
  by eval

lemma SEC1_P224_7_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_7_QCAVS" 
  by eval

lemma SEC1_P224_7_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_7_QCAVS" 
  using SEC1_P224_7_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_7_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_7_dIUT SEC1_P224_7_QCAVS = Some SEC1_P224_7_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_8›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_8_QCAVSx :: nat where 
  "SEC1_P224_8_QCAVSx = 0x13cf9d6d2c9aae8274c27d446afd0c888ffdd52ae299a35984d4f527"

definition SEC1_P224_8_QCAVSy :: nat where 
  "SEC1_P224_8_QCAVSy = 0xdcbee75b515751f8ee2ae355e8afd5de21c62a939a6507b538cbc4af"

definition SEC1_P224_8_QCAVS :: "int point" where
  "SEC1_P224_8_QCAVS = Point (int SEC1_P224_8_QCAVSx) (int SEC1_P224_8_QCAVSy)"

definition SEC1_P224_8_dIUT :: nat where 
  "SEC1_P224_8_dIUT = 0xe0d62035101ef487c485c60fb4500eebe6a32ec64dbe97dbe0232c46"

definition SEC1_P224_8_QIUTx :: nat where 
  "SEC1_P224_8_QIUTx = 0x88537735e9b23e3e0e076f135a82d33f9bffb465f3abce8322a62a62"

definition SEC1_P224_8_QIUTy :: nat where 
  "SEC1_P224_8_QIUTy = 0xb4c8c123673197875c0bd14ed097606d330fba2b9200ef65a44764d3"

definition SEC1_P224_8_QIUT :: "int point" where
  "SEC1_P224_8_QIUT = Point (int SEC1_P224_8_QIUTx) (int SEC1_P224_8_QIUTy)"

definition SEC1_P224_8_ZIUT :: nat where 
  "SEC1_P224_8_ZIUT = 0x632abb662728dbc994508873d5c527ca5ef923c0d31fa6c47ef4c825"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_8_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_8_dIUT"
  by eval

lemma SEC1_P224_8_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_8_dIUT = SEC1_P224_8_QIUT" 
  by eval

lemma SEC1_P224_8_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_8_QCAVS" 
  by eval

lemma SEC1_P224_8_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_8_QCAVS" 
  using SEC1_P224_8_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_8_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_8_dIUT SEC1_P224_8_QCAVS = Some SEC1_P224_8_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_9›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_9_QCAVSx :: nat where 
  "SEC1_P224_9_QCAVSx = 0x965b637c0dfbc0cf954035686d70f7ec30929e664e521dbaa2280659"

definition SEC1_P224_9_QCAVSy :: nat where 
  "SEC1_P224_9_QCAVSy = 0x82a58ff61bc90019bbcbb5875d3863db0bc2a1fa34b0ad4de1a83f99"

definition SEC1_P224_9_QCAVS :: "int point" where
  "SEC1_P224_9_QCAVS = Point (int SEC1_P224_9_QCAVSx) (int SEC1_P224_9_QCAVSy)"

definition SEC1_P224_9_dIUT :: nat where 
  "SEC1_P224_9_dIUT = 0xb96ade5b73ba72aa8b6e4d74d7bf9c58e962ff78eb542287c7b44ba2"

definition SEC1_P224_9_QIUTx :: nat where 
  "SEC1_P224_9_QIUTx = 0x37682926a54f70a4c1748f54d50d5b00138a055f924f2c65e5b0bbe4"

definition SEC1_P224_9_QIUTy :: nat where 
  "SEC1_P224_9_QIUTy = 0x596afefcdd640d29635015b89bdddd1f8c2723686d332e7a06ca8799"

definition SEC1_P224_9_QIUT :: "int point" where
  "SEC1_P224_9_QIUT = Point (int SEC1_P224_9_QIUTx) (int SEC1_P224_9_QIUTy)"

definition SEC1_P224_9_ZIUT :: nat where 
  "SEC1_P224_9_ZIUT = 0x34641141aab05ef58bd376d609345901fb8f63477c6be9097f037f1f"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_9_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_9_dIUT"
  by eval

lemma SEC1_P224_9_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_9_dIUT = SEC1_P224_9_QIUT" 
  by eval

lemma SEC1_P224_9_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_9_QCAVS" 
  by eval

lemma SEC1_P224_9_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_9_QCAVS" 
  using SEC1_P224_9_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_9_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_9_dIUT SEC1_P224_9_QCAVS = Some SEC1_P224_9_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_10›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_10_QCAVSx :: nat where 
  "SEC1_P224_10_QCAVSx = 0x73cc645372ca2e71637cda943d8148f3382ab6dd0f2e1a49da94e134"

definition SEC1_P224_10_QCAVSy :: nat where 
  "SEC1_P224_10_QCAVSy = 0xdf5c355c23e6e232ebc3bee2ab1873ee0d83e3382f8e6fe613f6343c"

definition SEC1_P224_10_QCAVS :: "int point" where
  "SEC1_P224_10_QCAVS = Point (int SEC1_P224_10_QCAVSx) (int SEC1_P224_10_QCAVSy)"

definition SEC1_P224_10_dIUT :: nat where 
  "SEC1_P224_10_dIUT = 0xa40d7e12049c71e6522c7ff2384224061c3a457058b310557655b854"

definition SEC1_P224_10_QIUTx :: nat where 
  "SEC1_P224_10_QIUTx = 0x399801243bfe0c2da9b0a53c8ca57f2eee87aaa94a8e4d5e029f42ca"

definition SEC1_P224_10_QIUTy :: nat where 
  "SEC1_P224_10_QIUTy = 0xaa49e6d4b47cee7a5c4ab71d5a67da84e0b9b425ce3e70da68c889e7"

definition SEC1_P224_10_QIUT :: "int point" where
  "SEC1_P224_10_QIUT = Point (int SEC1_P224_10_QIUTx) (int SEC1_P224_10_QIUTy)"

definition SEC1_P224_10_ZIUT :: nat where 
  "SEC1_P224_10_ZIUT = 0x4f74ac8507501a32bfc5a78d8271c200e835966e187e8d00011a8c75"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_10_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_10_dIUT"
  by eval

lemma SEC1_P224_10_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_10_dIUT = SEC1_P224_10_QIUT" 
  by eval

lemma SEC1_P224_10_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_10_QCAVS" 
  by eval

lemma SEC1_P224_10_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_10_QCAVS" 
  using SEC1_P224_10_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_10_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_10_dIUT SEC1_P224_10_QCAVS = Some SEC1_P224_10_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_11›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_11_QCAVSx :: nat where 
  "SEC1_P224_11_QCAVSx = 0x546578216250354e449e21546dd11cd1c5174236739acad9ce0f4512"

definition SEC1_P224_11_QCAVSy :: nat where 
  "SEC1_P224_11_QCAVSy = 0xd2a22fcd66d1abedc767668327c5cb9c599043276239cf3c8516af24"

definition SEC1_P224_11_QCAVS :: "int point" where
  "SEC1_P224_11_QCAVS = Point (int SEC1_P224_11_QCAVSx) (int SEC1_P224_11_QCAVSy)"

definition SEC1_P224_11_dIUT :: nat where 
  "SEC1_P224_11_dIUT = 0xad2519bc724d484e02a69f05149bb047714bf0f5986fac2e222cd946"

definition SEC1_P224_11_QIUTx :: nat where 
  "SEC1_P224_11_QIUTx = 0xdf9c1e0ef15e53b9f626e2be1cbe893639c06f3e0439ee95d7d4b1e3"

definition SEC1_P224_11_QIUTy :: nat where 
  "SEC1_P224_11_QIUTy = 0x7a52a7386adda243efdf8941085c84e31239cab92b8017336748965e"

definition SEC1_P224_11_QIUT :: "int point" where
  "SEC1_P224_11_QIUT = Point (int SEC1_P224_11_QIUTx) (int SEC1_P224_11_QIUTy)"

definition SEC1_P224_11_ZIUT :: nat where 
  "SEC1_P224_11_ZIUT = 0xad09c9ae4d2324ea81bb555b200d3c003e22a6870ee03b52df49e4de"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_11_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_11_dIUT"
  by eval

lemma SEC1_P224_11_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_11_dIUT = SEC1_P224_11_QIUT" 
  by eval

lemma SEC1_P224_11_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_11_QCAVS" 
  by eval

lemma SEC1_P224_11_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_11_QCAVS" 
  using SEC1_P224_11_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_11_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_11_dIUT SEC1_P224_11_QCAVS = Some SEC1_P224_11_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_12›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_12_QCAVSx :: nat where 
  "SEC1_P224_12_QCAVSx = 0x1d46b1dc3a28123cb51346e67baec56404868678faf7d0e8b2afa22a"

definition SEC1_P224_12_QCAVSy :: nat where 
  "SEC1_P224_12_QCAVSy = 0x0ec9e65ec97e218373e7fc115c2274d5b829a60d93f71e01d58136c3"

definition SEC1_P224_12_QCAVS :: "int point" where
  "SEC1_P224_12_QCAVS = Point (int SEC1_P224_12_QCAVSx) (int SEC1_P224_12_QCAVSy)"

definition SEC1_P224_12_dIUT :: nat where 
  "SEC1_P224_12_dIUT = 0x3d312a9b9d8ed09140900bbac1e095527ebc9e3c6493bcf3666e3a29"

definition SEC1_P224_12_QIUTx :: nat where 
  "SEC1_P224_12_QIUTx = 0xb4a0198dc8810e884425b750928b0c960c31f7a99663400b01a179df"

definition SEC1_P224_12_QIUTy :: nat where 
  "SEC1_P224_12_QIUTy = 0x812b601bfc0738242c6f86f830f27acd632ca618a0b5280c9d5769f7"

definition SEC1_P224_12_QIUT :: "int point" where
  "SEC1_P224_12_QIUT = Point (int SEC1_P224_12_QIUTx) (int SEC1_P224_12_QIUTy)"

definition SEC1_P224_12_ZIUT :: nat where 
  "SEC1_P224_12_ZIUT = 0xef029c28c68064b8abd2965a38c404fb5e944ace57e8638daba9d3cd"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_12_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_12_dIUT"
  by eval

lemma SEC1_P224_12_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_12_dIUT = SEC1_P224_12_QIUT" 
  by eval

lemma SEC1_P224_12_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_12_QCAVS" 
  by eval

lemma SEC1_P224_12_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_12_QCAVS" 
  using SEC1_P224_12_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_12_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_12_dIUT SEC1_P224_12_QCAVS = Some SEC1_P224_12_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_13›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_13_QCAVSx :: nat where 
  "SEC1_P224_13_QCAVSx = 0x266d038cc7a4fe21f6c976318e827b82bb5b8f7443a55298136506e0"

definition SEC1_P224_13_QCAVSy :: nat where 
  "SEC1_P224_13_QCAVSy = 0xdf123d98a7a20bbdf3943df2e3563422f8c0cf74d53aaabdd7c973ba"

definition SEC1_P224_13_QCAVS :: "int point" where
  "SEC1_P224_13_QCAVS = Point (int SEC1_P224_13_QCAVSx) (int SEC1_P224_13_QCAVSy)"

definition SEC1_P224_13_dIUT :: nat where 
  "SEC1_P224_13_dIUT = 0x8ce0822dc24c153995755ac350737ef506641c7d752b4f9300c612ed"

definition SEC1_P224_13_QIUTx :: nat where 
  "SEC1_P224_13_QIUTx = 0x00dfc7ec137690cd6d12fdb2fd0b8c5314582108769c2b722ffb3958"

definition SEC1_P224_13_QIUTy :: nat where 
  "SEC1_P224_13_QIUTy = 0x5eef3da4ba458127346bb64023868bddb7558a2ecfc813645f4ce9fe"

definition SEC1_P224_13_QIUT :: "int point" where
  "SEC1_P224_13_QIUT = Point (int SEC1_P224_13_QIUTx) (int SEC1_P224_13_QIUTy)"

definition SEC1_P224_13_ZIUT :: nat where 
  "SEC1_P224_13_ZIUT = 0xf83c16661dfcbad021cc3b5a5af51d9a18db4653866b3ff90787ce3e"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_13_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_13_dIUT"
  by eval

lemma SEC1_P224_13_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_13_dIUT = SEC1_P224_13_QIUT" 
  by eval

lemma SEC1_P224_13_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_13_QCAVS" 
  by eval

lemma SEC1_P224_13_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_13_QCAVS" 
  using SEC1_P224_13_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_13_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_13_dIUT SEC1_P224_13_QCAVS = Some SEC1_P224_13_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_14›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_14_QCAVSx :: nat where 
  "SEC1_P224_14_QCAVSx = 0xeb0a09f7a1c236a61f595809ec5670efd92e4598d5e613e092cdfdca"

definition SEC1_P224_14_QCAVSy :: nat where 
  "SEC1_P224_14_QCAVSy = 0x50787ae2f2f15b88bc10f7b5f0aee1418373f16153aebd1fba54288d"

definition SEC1_P224_14_QCAVS :: "int point" where
  "SEC1_P224_14_QCAVS = Point (int SEC1_P224_14_QCAVSx) (int SEC1_P224_14_QCAVSy)"

definition SEC1_P224_14_dIUT :: nat where 
  "SEC1_P224_14_dIUT = 0x0ff9b485325ab77f29e7bc379fed74bfac859482da0dee7528c19db2"

definition SEC1_P224_14_QIUTx :: nat where 
  "SEC1_P224_14_QIUTx = 0x7e603e6976db83c36011508fa695d1b515249e2e54b48fcbcfb90247"

definition SEC1_P224_14_QIUTy :: nat where 
  "SEC1_P224_14_QIUTy = 0x0179a600ce86adfca9b1b931fa5173d618da09e841803d19b0264286"

definition SEC1_P224_14_QIUT :: "int point" where
  "SEC1_P224_14_QIUT = Point (int SEC1_P224_14_QIUTx) (int SEC1_P224_14_QIUTy)"

definition SEC1_P224_14_ZIUT :: nat where 
  "SEC1_P224_14_ZIUT = 0xf51258c63f232e55a66aa25ebd597b2018d1052c02eeb63866758005"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_14_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_14_dIUT"
  by eval

lemma SEC1_P224_14_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_14_dIUT = SEC1_P224_14_QIUT" 
  by eval

lemma SEC1_P224_14_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_14_QCAVS" 
  by eval

lemma SEC1_P224_14_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_14_QCAVS" 
  using SEC1_P224_14_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_14_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_14_dIUT SEC1_P224_14_QCAVS = Some SEC1_P224_14_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_15›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_15_QCAVSx :: nat where 
  "SEC1_P224_15_QCAVSx = 0x6b2f6b18a587f562ffc61bd9b0047322286986a78f1fd139b84f7c24"

definition SEC1_P224_15_QCAVSy :: nat where 
  "SEC1_P224_15_QCAVSy = 0x7096908e4615266be59a53cd655515056ff92370a6271a5d3823d704"

definition SEC1_P224_15_QCAVS :: "int point" where
  "SEC1_P224_15_QCAVS = Point (int SEC1_P224_15_QCAVSx) (int SEC1_P224_15_QCAVSy)"

definition SEC1_P224_15_dIUT :: nat where 
  "SEC1_P224_15_dIUT = 0x19cf5ff6306467f28b9fe0675a43c0582552c8c12e59ce7c38f292b1"

definition SEC1_P224_15_QIUTx :: nat where 
  "SEC1_P224_15_QIUTx = 0xfc20e906e609c112cfc2e0fea6303882c5db94e87e022373ab2c082a"

definition SEC1_P224_15_QIUTy :: nat where 
  "SEC1_P224_15_QIUTy = 0xaecdf1daa71782bc5a26bbbd8d7e8a76490e26abc17dffc774bd7341"

definition SEC1_P224_15_QIUT :: "int point" where
  "SEC1_P224_15_QIUT = Point (int SEC1_P224_15_QIUTx) (int SEC1_P224_15_QIUTy)"

definition SEC1_P224_15_ZIUT :: nat where 
  "SEC1_P224_15_ZIUT = 0x7fdc969a186ff18429f2a276dac43beea21182d82ce2e5a0876552b1"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_15_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_15_dIUT"
  by eval

lemma SEC1_P224_15_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_15_dIUT = SEC1_P224_15_QIUT" 
  by eval

lemma SEC1_P224_15_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_15_QCAVS" 
  by eval

lemma SEC1_P224_15_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_15_QCAVS" 
  using SEC1_P224_15_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_15_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_15_dIUT SEC1_P224_15_QCAVS = Some SEC1_P224_15_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_16›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_16_QCAVSx :: nat where 
  "SEC1_P224_16_QCAVSx = 0x328101ba826acd75ff9f34d5574ce0dbc92f709bad8d7a33c47940c1"

definition SEC1_P224_16_QCAVSy :: nat where 
  "SEC1_P224_16_QCAVSy = 0xdf39f1ea88488c55d5538160878b9ced18a887ea261dd712d14024ff"

definition SEC1_P224_16_QCAVS :: "int point" where
  "SEC1_P224_16_QCAVS = Point (int SEC1_P224_16_QCAVSx) (int SEC1_P224_16_QCAVSy)"

definition SEC1_P224_16_dIUT :: nat where 
  "SEC1_P224_16_dIUT = 0x90a15368e3532c0b1e51e55d139447c2c89bc160719d697291ea7c14"

definition SEC1_P224_16_QIUTx :: nat where 
  "SEC1_P224_16_QIUTx = 0xc6837d506e976da7db3ad1267c359dff2ea6fb0b7f7f8e77024c59e9"

definition SEC1_P224_16_QIUTy :: nat where 
  "SEC1_P224_16_QIUTy = 0x67eb491d2fc8a530c46525d2a8b2d7c1df5fba1ae740a4649c683ee6"

definition SEC1_P224_16_QIUT :: "int point" where
  "SEC1_P224_16_QIUT = Point (int SEC1_P224_16_QIUTx) (int SEC1_P224_16_QIUTy)"

definition SEC1_P224_16_ZIUT :: nat where 
  "SEC1_P224_16_ZIUT = 0x3d60ab6db2b3ffe2d29ccff46d056e54230cf34982e241556ed2920c"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_16_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_16_dIUT"
  by eval

lemma SEC1_P224_16_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_16_dIUT = SEC1_P224_16_QIUT" 
  by eval

lemma SEC1_P224_16_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_16_QCAVS" 
  by eval

lemma SEC1_P224_16_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_16_QCAVS" 
  using SEC1_P224_16_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_16_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_16_dIUT SEC1_P224_16_QCAVS = Some SEC1_P224_16_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_17›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_17_QCAVSx :: nat where 
  "SEC1_P224_17_QCAVSx = 0x0081e34270871e2ebbd94183f617b4ae15f0416dd634fe6e934cf3c0"

definition SEC1_P224_17_QCAVSy :: nat where 
  "SEC1_P224_17_QCAVSy = 0x3a1e9f38a7b90b7317d26b9f6311063ab58b268cf489b2e50386d5d6"

definition SEC1_P224_17_QCAVS :: "int point" where
  "SEC1_P224_17_QCAVS = Point (int SEC1_P224_17_QCAVSx) (int SEC1_P224_17_QCAVSy)"

definition SEC1_P224_17_dIUT :: nat where 
  "SEC1_P224_17_dIUT = 0x8e0838e05e1721491067e1cabc2e8051b290e2616eec427b7121897d"

definition SEC1_P224_17_QIUTx :: nat where 
  "SEC1_P224_17_QIUTx = 0xe9150f770075626019e18f95473b71e6828041791d3f08d3faeeaa2b"

definition SEC1_P224_17_QIUTy :: nat where 
  "SEC1_P224_17_QIUTy = 0x475f70735eaae52308a3b763dc88efe18ab590ebafa035f6e08b001c"

definition SEC1_P224_17_QIUT :: "int point" where
  "SEC1_P224_17_QIUT = Point (int SEC1_P224_17_QIUTx) (int SEC1_P224_17_QIUTy)"

definition SEC1_P224_17_ZIUT :: nat where 
  "SEC1_P224_17_ZIUT = 0x9116d72786f4db5df7a8b43078c6ab9160d423513d35ea5e2559306d"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_17_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_17_dIUT"
  by eval

lemma SEC1_P224_17_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_17_dIUT = SEC1_P224_17_QIUT" 
  by eval

lemma SEC1_P224_17_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_17_QCAVS" 
  by eval

lemma SEC1_P224_17_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_17_QCAVS" 
  using SEC1_P224_17_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_17_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_17_dIUT SEC1_P224_17_QCAVS = Some SEC1_P224_17_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_18›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_18_QCAVSx :: nat where 
  "SEC1_P224_18_QCAVSx = 0x2623632fdf0bd856805a69aa186d4133ef5904e1f655a972d66cce07"

definition SEC1_P224_18_QCAVSy :: nat where 
  "SEC1_P224_18_QCAVSy = 0x2cef9728dd06fb8b50150f529b695076d4507983912585c89bd0682e"

definition SEC1_P224_18_QCAVS :: "int point" where
  "SEC1_P224_18_QCAVS = Point (int SEC1_P224_18_QCAVSx) (int SEC1_P224_18_QCAVSy)"

definition SEC1_P224_18_dIUT :: nat where 
  "SEC1_P224_18_dIUT = 0x38106e93f16a381adb1d72cee3da66ae462ad4bbfea9ecdf35d0814e"

definition SEC1_P224_18_QIUTx :: nat where 
  "SEC1_P224_18_QIUTx = 0x7be6c4c917829ab657dd79e8637d7aefd2f81f0de7654d957e97658d"

definition SEC1_P224_18_QIUTy :: nat where 
  "SEC1_P224_18_QIUTy = 0x430d22d9e8438310f61e0d43f25fa3e34585f432baad27db3021bf0d"

definition SEC1_P224_18_QIUT :: "int point" where
  "SEC1_P224_18_QIUT = Point (int SEC1_P224_18_QIUTx) (int SEC1_P224_18_QIUTy)"

definition SEC1_P224_18_ZIUT :: nat where 
  "SEC1_P224_18_ZIUT = 0x207c53dcefac789aaa0276d9200b3a940ce5f2296f4cb2e81a185d3d"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_18_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_18_dIUT"
  by eval

lemma SEC1_P224_18_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_18_dIUT = SEC1_P224_18_QIUT" 
  by eval

lemma SEC1_P224_18_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_18_QCAVS" 
  by eval

lemma SEC1_P224_18_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_18_QCAVS" 
  using SEC1_P224_18_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_18_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_18_dIUT SEC1_P224_18_QCAVS = Some SEC1_P224_18_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_19›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_19_QCAVSx :: nat where 
  "SEC1_P224_19_QCAVSx = 0x8ee4d1dcc31dee4bf6fe21ca8a587721d910acfb122c16c2a77a8152"

definition SEC1_P224_19_QCAVSy :: nat where 
  "SEC1_P224_19_QCAVSy = 0x4ebf323fff04eb477069a0ac68b345f6b1ae134efc31940e513cb99f"

definition SEC1_P224_19_QCAVS :: "int point" where
  "SEC1_P224_19_QCAVS = Point (int SEC1_P224_19_QCAVSx) (int SEC1_P224_19_QCAVSy)"

definition SEC1_P224_19_dIUT :: nat where 
  "SEC1_P224_19_dIUT = 0xe5d1718431cf50f6cbd1bc8019fa16762dfa12c989e5999977fb4ea2"

definition SEC1_P224_19_QIUTx :: nat where 
  "SEC1_P224_19_QIUTx = 0x2ea4966e7f92ed7f5cc61fde792045f63b731d6e7d0de2577f2d8ece"

definition SEC1_P224_19_QIUTy :: nat where 
  "SEC1_P224_19_QIUTy = 0x1c4a7b1ede6f839162292df424be78e8176fb6f942a3c02391700f31"

definition SEC1_P224_19_QIUT :: "int point" where
  "SEC1_P224_19_QIUT = Point (int SEC1_P224_19_QIUTx) (int SEC1_P224_19_QIUTy)"

definition SEC1_P224_19_ZIUT :: nat where 
  "SEC1_P224_19_ZIUT = 0x10e467da34f48ad7072005bccd6da1b2ba3f71eafa1c393842f91d74"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_19_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_19_dIUT"
  by eval

lemma SEC1_P224_19_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_19_dIUT = SEC1_P224_19_QIUT" 
  by eval

lemma SEC1_P224_19_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_19_QCAVS" 
  by eval

lemma SEC1_P224_19_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_19_QCAVS" 
  using SEC1_P224_19_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_19_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_19_dIUT SEC1_P224_19_QCAVS = Some SEC1_P224_19_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_20›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_20_QCAVSx :: nat where 
  "SEC1_P224_20_QCAVSx = 0x97dcbe6d28335882a6d193cc54a1063dd0775dc328565300bb99e691"

definition SEC1_P224_20_QCAVSy :: nat where 
  "SEC1_P224_20_QCAVSy = 0xdad11dd5ece8cfd9f97c9a526e4a1506e6355969ee87826fc38bcd24"

definition SEC1_P224_20_QCAVS :: "int point" where
  "SEC1_P224_20_QCAVS = Point (int SEC1_P224_20_QCAVSx) (int SEC1_P224_20_QCAVSy)"

definition SEC1_P224_20_dIUT :: nat where 
  "SEC1_P224_20_dIUT = 0x3d635691b62a9a927c633951c9369c8862bd2119d30970c2644727d6"

definition SEC1_P224_20_QIUTx :: nat where 
  "SEC1_P224_20_QIUTx = 0x438bbb980517afb20be1d674e3ac2b31cef07a9b23fb8f6e38e0d6c0"

definition SEC1_P224_20_QIUTy :: nat where 
  "SEC1_P224_20_QIUTy = 0x0be5f1c47d58d21b6ed28423b32f5a94750da47edcef33ea79942afd"

definition SEC1_P224_20_QIUT :: "int point" where
  "SEC1_P224_20_QIUT = Point (int SEC1_P224_20_QIUTx) (int SEC1_P224_20_QIUTy)"

definition SEC1_P224_20_ZIUT :: nat where 
  "SEC1_P224_20_ZIUT = 0x82fd2f9c60c4f999ac00bbe64bfc11da8ff8cda2e499fced65230bb1"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_20_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_20_dIUT"
  by eval

lemma SEC1_P224_20_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_20_dIUT = SEC1_P224_20_QIUT" 
  by eval

lemma SEC1_P224_20_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_20_QCAVS" 
  by eval

lemma SEC1_P224_20_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_20_QCAVS" 
  using SEC1_P224_20_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_20_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_20_dIUT SEC1_P224_20_QCAVS = Some SEC1_P224_20_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_21›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_21_QCAVSx :: nat where 
  "SEC1_P224_21_QCAVSx = 0xce9126dd53972dea1de1d11efef900de34b661859c4648c5c0e534f7"

definition SEC1_P224_21_QCAVSy :: nat where 
  "SEC1_P224_21_QCAVSy = 0xe113b6f2c1659d07f2716e64a83c18bbce344dd2121fe85168eae085"

definition SEC1_P224_21_QCAVS :: "int point" where
  "SEC1_P224_21_QCAVS = Point (int SEC1_P224_21_QCAVSx) (int SEC1_P224_21_QCAVSy)"

definition SEC1_P224_21_dIUT :: nat where 
  "SEC1_P224_21_dIUT = 0xacf3c85bbdc379f02f5ea36e7f0f53095a9e7046a28685a8659bf798"

definition SEC1_P224_21_QIUTx :: nat where 
  "SEC1_P224_21_QIUTx = 0xff7511215c71d796bd646e8474be4416b91684ce0d269ef6f422013b"

definition SEC1_P224_21_QIUTy :: nat where 
  "SEC1_P224_21_QIUTy = 0xb7bf5e79b5a9393bb9ea42c0bdb2d3c2dc806e1a7306aa58e4fdbea5"

definition SEC1_P224_21_QIUT :: "int point" where
  "SEC1_P224_21_QIUT = Point (int SEC1_P224_21_QIUTx) (int SEC1_P224_21_QIUTy)"

definition SEC1_P224_21_ZIUT :: nat where 
  "SEC1_P224_21_ZIUT = 0x530f7e7fc932613b29c981f261cb036cba3f1df3864e0e1cba2685a2"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_21_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_21_dIUT"
  by eval

lemma SEC1_P224_21_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_21_dIUT = SEC1_P224_21_QIUT" 
  by eval

lemma SEC1_P224_21_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_21_QCAVS" 
  by eval

lemma SEC1_P224_21_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_21_QCAVS" 
  using SEC1_P224_21_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_21_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_21_dIUT SEC1_P224_21_QCAVS = Some SEC1_P224_21_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_22›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_22_QCAVSx :: nat where 
  "SEC1_P224_22_QCAVSx = 0x84419967d6cfad41e75a02b6da605a97949a183a97c306c4b46e66a5"

definition SEC1_P224_22_QCAVSy :: nat where 
  "SEC1_P224_22_QCAVSy = 0x5cc9b259718b1bc8b144fde633a894616ffd59a3a6d5d8e942c7cbb7"

definition SEC1_P224_22_QCAVS :: "int point" where
  "SEC1_P224_22_QCAVS = Point (int SEC1_P224_22_QCAVSx) (int SEC1_P224_22_QCAVSy)"

definition SEC1_P224_22_dIUT :: nat where 
  "SEC1_P224_22_dIUT = 0xcffd62cb00a0e3163fbf2c397fadc9618210f86b4f54a675287305f0"

definition SEC1_P224_22_QIUTx :: nat where 
  "SEC1_P224_22_QIUTx = 0x04bf4d948f4430d18b4ed6c96dbaf981fa11a403ed16887f06754981"

definition SEC1_P224_22_QIUTy :: nat where 
  "SEC1_P224_22_QIUTy = 0x7c1326a9cef51f79d4e78303d6064b459f612584ac2fdf593d7d5d84"

definition SEC1_P224_22_QIUT :: "int point" where
  "SEC1_P224_22_QIUT = Point (int SEC1_P224_22_QIUTx) (int SEC1_P224_22_QIUTy)"

definition SEC1_P224_22_ZIUT :: nat where 
  "SEC1_P224_22_ZIUT = 0x49f6fd0139248ef4df2db05d1319bd5b1489e249827a45a8a5f12427"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_22_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_22_dIUT"
  by eval

lemma SEC1_P224_22_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_22_dIUT = SEC1_P224_22_QIUT" 
  by eval

lemma SEC1_P224_22_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_22_QCAVS" 
  by eval

lemma SEC1_P224_22_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_22_QCAVS" 
  using SEC1_P224_22_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_22_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_22_dIUT SEC1_P224_22_QCAVS = Some SEC1_P224_22_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_23›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_23_QCAVSx :: nat where 
  "SEC1_P224_23_QCAVSx = 0x7c9cac35768063c2827f60a7f51388f2a8f4b7f8cd736bd6bc337477"

definition SEC1_P224_23_QCAVSy :: nat where 
  "SEC1_P224_23_QCAVSy = 0x29ee6b849c6025d577dbcc55fbd17018f4edbc2ef105b004d6257bcd"

definition SEC1_P224_23_QCAVS :: "int point" where
  "SEC1_P224_23_QCAVS = Point (int SEC1_P224_23_QCAVSx) (int SEC1_P224_23_QCAVSy)"

definition SEC1_P224_23_dIUT :: nat where 
  "SEC1_P224_23_dIUT = 0x85f903e43943d13c68932e710e80de52cbc0b8f1a1418ea4da079299"

definition SEC1_P224_23_QIUTx :: nat where 
  "SEC1_P224_23_QIUTx = 0x970a4a7e01d4188497ceb46955eb1b842d9085819a9b925c84529d3d"

definition SEC1_P224_23_QIUTy :: nat where 
  "SEC1_P224_23_QIUTy = 0xdfa2526480f833ea0edbd204e4e365fef3472888fe7d9691c3ebc09f"

definition SEC1_P224_23_QIUT :: "int point" where
  "SEC1_P224_23_QIUT = Point (int SEC1_P224_23_QIUTx) (int SEC1_P224_23_QIUTy)"

definition SEC1_P224_23_ZIUT :: nat where 
  "SEC1_P224_23_ZIUT = 0x8f7e34e597ae8093b98270a74a8dfcdbed457f42f43df487c5487161"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_23_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_23_dIUT"
  by eval

lemma SEC1_P224_23_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_23_dIUT = SEC1_P224_23_QIUT" 
  by eval

lemma SEC1_P224_23_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_23_QCAVS" 
  by eval

lemma SEC1_P224_23_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_23_QCAVS" 
  using SEC1_P224_23_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_23_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_23_dIUT SEC1_P224_23_QCAVS = Some SEC1_P224_23_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P224_24›

subsubsection‹Given Test Vector Values›

definition SEC1_P224_24_QCAVSx :: nat where 
  "SEC1_P224_24_QCAVSx = 0x085a7642ad8e59b1a3e8726a7547afbecffdac1dab7e57230c6a9df4"

definition SEC1_P224_24_QCAVSy :: nat where 
  "SEC1_P224_24_QCAVSy = 0xf91c36d881fe9b8047a3530713554a1af4c25c5a8e654dcdcf689f2e"

definition SEC1_P224_24_QCAVS :: "int point" where
  "SEC1_P224_24_QCAVS = Point (int SEC1_P224_24_QCAVSx) (int SEC1_P224_24_QCAVSy)"

definition SEC1_P224_24_dIUT :: nat where 
  "SEC1_P224_24_dIUT = 0xcce64891a3d0129fee0d4a96cfbe7ac470b85e967529057cfa31a1d9"

definition SEC1_P224_24_QIUTx :: nat where 
  "SEC1_P224_24_QIUTx = 0xa6b29632db94da2125dc1cf80e03702687b2acc1122022fa2174765a"

definition SEC1_P224_24_QIUTy :: nat where 
  "SEC1_P224_24_QIUTy = 0x61723edd73e10daed73775278f1958ba56f1fc9d085ebc2b64c84fe5"

definition SEC1_P224_24_QIUT :: "int point" where
  "SEC1_P224_24_QIUT = Point (int SEC1_P224_24_QIUTx) (int SEC1_P224_24_QIUTy)"

definition SEC1_P224_24_ZIUT :: nat where 
  "SEC1_P224_24_ZIUT = 0x71954e2261e8510be1a060733671d2e9d0a2d012eb4e09556d697d2a"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P224_24_dIUT_valid: 
  "SEC1_P224_ECprivateKeyValid SEC1_P224_24_dIUT"
  by eval

lemma SEC1_P224_24_IUT_validPair: 
  "SEC1_P224_ECkeyGen SEC1_P224_24_dIUT = SEC1_P224_24_QIUT" 
  by eval

lemma SEC1_P224_24_QCAVS_partialValid: 
  "SEC1_P224_ECpublicKeyPartialValid SEC1_P224_24_QCAVS" 
  by eval

lemma SEC1_P224_24_QCAVS_validPublicKey: 
  "SEC1_P224_ECpublicKeyValid SEC1_P224_24_QCAVS" 
  using SEC1_P224_24_QCAVS_partialValid SEC1_P224.partValidImpliesValidIFheq1 P224_h_def by blast

lemma SEC1_P224_24_ECDHprim_Check: 
  "SEC1_P224_ECDHprim SEC1_P224_24_dIUT SEC1_P224_24_QCAVS = Some SEC1_P224_24_ZIUT"
  by eval
  
section‹NIST Test Vectors for ECDH: Curve P-256›
  
subsection‹Test Vector: SEC1_P256_0›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_0_QCAVSx :: nat where 
  "SEC1_P256_0_QCAVSx = 0x700c48f77f56584c5cc632ca65640db91b6bacce3a4df6b42ce7cc838833d287"

definition SEC1_P256_0_QCAVSy :: nat where 
  "SEC1_P256_0_QCAVSy = 0xdb71e509e3fd9b060ddb20ba5c51dcc5948d46fbf640dfe0441782cab85fa4ac"

definition SEC1_P256_0_QCAVS :: "int point" where
  "SEC1_P256_0_QCAVS = Point (int SEC1_P256_0_QCAVSx) (int SEC1_P256_0_QCAVSy)"

definition SEC1_P256_0_dIUT :: nat where 
  "SEC1_P256_0_dIUT = 0x7d7dc5f71eb29ddaf80d6214632eeae03d9058af1fb6d22ed80badb62bc1a534"

definition SEC1_P256_0_QIUTx :: nat where 
  "SEC1_P256_0_QIUTx = 0xead218590119e8876b29146ff89ca61770c4edbbf97d38ce385ed281d8a6b230"

definition SEC1_P256_0_QIUTy :: nat where 
  "SEC1_P256_0_QIUTy = 0x28af61281fd35e2fa7002523acc85a429cb06ee6648325389f59edfce1405141"

definition SEC1_P256_0_QIUT :: "int point" where
  "SEC1_P256_0_QIUT = Point (int SEC1_P256_0_QIUTx) (int SEC1_P256_0_QIUTy)"

definition SEC1_P256_0_ZIUT :: nat where 
  "SEC1_P256_0_ZIUT = 0x46fc62106420ff012e54a434fbdd2d25ccc5852060561e68040dd7778997bd7b"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_0_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_0_dIUT"
  by eval

lemma SEC1_P256_0_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_0_dIUT = SEC1_P256_0_QIUT" 
  by eval

lemma SEC1_P256_0_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_0_QCAVS" 
  by eval

lemma SEC1_P256_0_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_0_QCAVS" 
  using SEC1_P256_0_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_0_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_0_dIUT SEC1_P256_0_QCAVS = Some SEC1_P256_0_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_1›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_1_QCAVSx :: nat where 
  "SEC1_P256_1_QCAVSx = 0x809f04289c64348c01515eb03d5ce7ac1a8cb9498f5caa50197e58d43a86a7ae"

definition SEC1_P256_1_QCAVSy :: nat where 
  "SEC1_P256_1_QCAVSy = 0xb29d84e811197f25eba8f5194092cb6ff440e26d4421011372461f579271cda3"

definition SEC1_P256_1_QCAVS :: "int point" where
  "SEC1_P256_1_QCAVS = Point (int SEC1_P256_1_QCAVSx) (int SEC1_P256_1_QCAVSy)"

definition SEC1_P256_1_dIUT :: nat where 
  "SEC1_P256_1_dIUT = 0x38f65d6dce47676044d58ce5139582d568f64bb16098d179dbab07741dd5caf5"

definition SEC1_P256_1_QIUTx :: nat where 
  "SEC1_P256_1_QIUTx = 0x119f2f047902782ab0c9e27a54aff5eb9b964829ca99c06b02ddba95b0a3f6d0"

definition SEC1_P256_1_QIUTy :: nat where 
  "SEC1_P256_1_QIUTy = 0x8f52b726664cac366fc98ac7a012b2682cbd962e5acb544671d41b9445704d1d"

definition SEC1_P256_1_QIUT :: "int point" where
  "SEC1_P256_1_QIUT = Point (int SEC1_P256_1_QIUTx) (int SEC1_P256_1_QIUTy)"

definition SEC1_P256_1_ZIUT :: nat where 
  "SEC1_P256_1_ZIUT = 0x057d636096cb80b67a8c038c890e887d1adfa4195e9b3ce241c8a778c59cda67"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_1_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_1_dIUT"
  by eval

lemma SEC1_P256_1_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_1_dIUT = SEC1_P256_1_QIUT" 
  by eval

lemma SEC1_P256_1_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_1_QCAVS" 
  by eval

lemma SEC1_P256_1_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_1_QCAVS" 
  using SEC1_P256_1_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_1_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_1_dIUT SEC1_P256_1_QCAVS = Some SEC1_P256_1_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_2›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_2_QCAVSx :: nat where 
  "SEC1_P256_2_QCAVSx = 0xa2339c12d4a03c33546de533268b4ad667debf458b464d77443636440ee7fec3"

definition SEC1_P256_2_QCAVSy :: nat where 
  "SEC1_P256_2_QCAVSy = 0xef48a3ab26e20220bcda2c1851076839dae88eae962869a497bf73cb66faf536"

definition SEC1_P256_2_QCAVS :: "int point" where
  "SEC1_P256_2_QCAVS = Point (int SEC1_P256_2_QCAVSx) (int SEC1_P256_2_QCAVSy)"

definition SEC1_P256_2_dIUT :: nat where 
  "SEC1_P256_2_dIUT = 0x1accfaf1b97712b85a6f54b148985a1bdc4c9bec0bd258cad4b3d603f49f32c8"

definition SEC1_P256_2_QIUTx :: nat where 
  "SEC1_P256_2_QIUTx = 0xd9f2b79c172845bfdb560bbb01447ca5ecc0470a09513b6126902c6b4f8d1051"

definition SEC1_P256_2_QIUTy :: nat where 
  "SEC1_P256_2_QIUTy = 0xf815ef5ec32128d3487834764678702e64e164ff7315185e23aff5facd96d7bc"

definition SEC1_P256_2_QIUT :: "int point" where
  "SEC1_P256_2_QIUT = Point (int SEC1_P256_2_QIUTx) (int SEC1_P256_2_QIUTy)"

definition SEC1_P256_2_ZIUT :: nat where 
  "SEC1_P256_2_ZIUT = 0x2d457b78b4614132477618a5b077965ec90730a8c81a1c75d6d4ec68005d67ec"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_2_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_2_dIUT"
  by eval

lemma SEC1_P256_2_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_2_dIUT = SEC1_P256_2_QIUT" 
  by eval

lemma SEC1_P256_2_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_2_QCAVS" 
  by eval

lemma SEC1_P256_2_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_2_QCAVS" 
  using SEC1_P256_2_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_2_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_2_dIUT SEC1_P256_2_QCAVS = Some SEC1_P256_2_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_3›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_3_QCAVSx :: nat where 
  "SEC1_P256_3_QCAVSx = 0xdf3989b9fa55495719b3cf46dccd28b5153f7808191dd518eff0c3cff2b705ed"

definition SEC1_P256_3_QCAVSy :: nat where 
  "SEC1_P256_3_QCAVSy = 0x422294ff46003429d739a33206c8752552c8ba54a270defc06e221e0feaf6ac4"

definition SEC1_P256_3_QCAVS :: "int point" where
  "SEC1_P256_3_QCAVS = Point (int SEC1_P256_3_QCAVSx) (int SEC1_P256_3_QCAVSy)"

definition SEC1_P256_3_dIUT :: nat where 
  "SEC1_P256_3_dIUT = 0x207c43a79bfee03db6f4b944f53d2fb76cc49ef1c9c4d34d51b6c65c4db6932d"

definition SEC1_P256_3_QIUTx :: nat where 
  "SEC1_P256_3_QIUTx = 0x24277c33f450462dcb3d4801d57b9ced05188f16c28eda873258048cd1607e0d"

definition SEC1_P256_3_QIUTy :: nat where 
  "SEC1_P256_3_QIUTy = 0xc4789753e2b1f63b32ff014ec42cd6a69fac81dfe6d0d6fd4af372ae27c46f88"

definition SEC1_P256_3_QIUT :: "int point" where
  "SEC1_P256_3_QIUT = Point (int SEC1_P256_3_QIUTx) (int SEC1_P256_3_QIUTy)"

definition SEC1_P256_3_ZIUT :: nat where 
  "SEC1_P256_3_ZIUT = 0x96441259534b80f6aee3d287a6bb17b5094dd4277d9e294f8fe73e48bf2a0024"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_3_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_3_dIUT"
  by eval

lemma SEC1_P256_3_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_3_dIUT = SEC1_P256_3_QIUT" 
  by eval

lemma SEC1_P256_3_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_3_QCAVS" 
  by eval

lemma SEC1_P256_3_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_3_QCAVS" 
  using SEC1_P256_3_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_3_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_3_dIUT SEC1_P256_3_QCAVS = Some SEC1_P256_3_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_4›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_4_QCAVSx :: nat where 
  "SEC1_P256_4_QCAVSx = 0x41192d2813e79561e6a1d6f53c8bc1a433a199c835e141b05a74a97b0faeb922"

definition SEC1_P256_4_QCAVSy :: nat where 
  "SEC1_P256_4_QCAVSy = 0x1af98cc45e98a7e041b01cf35f462b7562281351c8ebf3ffa02e33a0722a1328"

definition SEC1_P256_4_QCAVS :: "int point" where
  "SEC1_P256_4_QCAVS = Point (int SEC1_P256_4_QCAVSx) (int SEC1_P256_4_QCAVSy)"

definition SEC1_P256_4_dIUT :: nat where 
  "SEC1_P256_4_dIUT = 0x59137e38152350b195c9718d39673d519838055ad908dd4757152fd8255c09bf"

definition SEC1_P256_4_QIUTx :: nat where 
  "SEC1_P256_4_QIUTx = 0xa8c5fdce8b62c5ada598f141adb3b26cf254c280b2857a63d2ad783a73115f6b"

definition SEC1_P256_4_QIUTy :: nat where 
  "SEC1_P256_4_QIUTy = 0x806e1aafec4af80a0d786b3de45375b517a7e5b51ffb2c356537c9e6ef227d4a"

definition SEC1_P256_4_QIUT :: "int point" where
  "SEC1_P256_4_QIUT = Point (int SEC1_P256_4_QIUTx) (int SEC1_P256_4_QIUTy)"

definition SEC1_P256_4_ZIUT :: nat where 
  "SEC1_P256_4_ZIUT = 0x19d44c8d63e8e8dd12c22a87b8cd4ece27acdde04dbf47f7f27537a6999a8e62"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_4_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_4_dIUT"
  by eval

lemma SEC1_P256_4_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_4_dIUT = SEC1_P256_4_QIUT" 
  by eval

lemma SEC1_P256_4_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_4_QCAVS" 
  by eval

lemma SEC1_P256_4_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_4_QCAVS" 
  using SEC1_P256_4_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_4_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_4_dIUT SEC1_P256_4_QCAVS = Some SEC1_P256_4_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_5›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_5_QCAVSx :: nat where 
  "SEC1_P256_5_QCAVSx = 0x33e82092a0f1fb38f5649d5867fba28b503172b7035574bf8e5b7100a3052792"

definition SEC1_P256_5_QCAVSy :: nat where 
  "SEC1_P256_5_QCAVSy = 0xf2cf6b601e0a05945e335550bf648d782f46186c772c0f20d3cd0d6b8ca14b2f"

definition SEC1_P256_5_QCAVS :: "int point" where
  "SEC1_P256_5_QCAVS = Point (int SEC1_P256_5_QCAVSx) (int SEC1_P256_5_QCAVSy)"

definition SEC1_P256_5_dIUT :: nat where 
  "SEC1_P256_5_dIUT = 0xf5f8e0174610a661277979b58ce5c90fee6c9b3bb346a90a7196255e40b132ef"

definition SEC1_P256_5_QIUTx :: nat where 
  "SEC1_P256_5_QIUTx = 0x7b861dcd2844a5a8363f6b8ef8d493640f55879217189d80326aad9480dfc149"

definition SEC1_P256_5_QIUTy :: nat where 
  "SEC1_P256_5_QIUTy = 0xc4675b45eeb306405f6c33c38bc69eb2bdec9b75ad5af4706aab84543b9cc63a"

definition SEC1_P256_5_QIUT :: "int point" where
  "SEC1_P256_5_QIUT = Point (int SEC1_P256_5_QIUTx) (int SEC1_P256_5_QIUTy)"

definition SEC1_P256_5_ZIUT :: nat where 
  "SEC1_P256_5_ZIUT = 0x664e45d5bba4ac931cd65d52017e4be9b19a515f669bea4703542a2c525cd3d3"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_5_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_5_dIUT"
  by eval

lemma SEC1_P256_5_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_5_dIUT = SEC1_P256_5_QIUT" 
  by eval

lemma SEC1_P256_5_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_5_QCAVS" 
  by eval

lemma SEC1_P256_5_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_5_QCAVS" 
  using SEC1_P256_5_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_5_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_5_dIUT SEC1_P256_5_QCAVS = Some SEC1_P256_5_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_6›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_6_QCAVSx :: nat where 
  "SEC1_P256_6_QCAVSx = 0x6a9e0c3f916e4e315c91147be571686d90464e8bf981d34a90b6353bca6eeba7"

definition SEC1_P256_6_QCAVSy :: nat where 
  "SEC1_P256_6_QCAVSy = 0x40f9bead39c2f2bcc2602f75b8a73ec7bdffcbcead159d0174c6c4d3c5357f05"

definition SEC1_P256_6_QCAVS :: "int point" where
  "SEC1_P256_6_QCAVS = Point (int SEC1_P256_6_QCAVSx) (int SEC1_P256_6_QCAVSy)"

definition SEC1_P256_6_dIUT :: nat where 
  "SEC1_P256_6_dIUT = 0x3b589af7db03459c23068b64f63f28d3c3c6bc25b5bf76ac05f35482888b5190"

definition SEC1_P256_6_QIUTx :: nat where 
  "SEC1_P256_6_QIUTx = 0x9fb38e2d58ea1baf7622e96720101cae3cde4ba6c1e9fa26d9b1de0899102863"

definition SEC1_P256_6_QIUTy :: nat where 
  "SEC1_P256_6_QIUTy = 0xd5561b900406edf50802dd7d73e89395f8aed72fba0e1d1b61fe1d22302260f0"

definition SEC1_P256_6_QIUT :: "int point" where
  "SEC1_P256_6_QIUT = Point (int SEC1_P256_6_QIUTx) (int SEC1_P256_6_QIUTy)"

definition SEC1_P256_6_ZIUT :: nat where 
  "SEC1_P256_6_ZIUT = 0xca342daa50dc09d61be7c196c85e60a80c5cb04931746820be548cdde055679d"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_6_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_6_dIUT"
  by eval

lemma SEC1_P256_6_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_6_dIUT = SEC1_P256_6_QIUT" 
  by eval

lemma SEC1_P256_6_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_6_QCAVS" 
  by eval

lemma SEC1_P256_6_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_6_QCAVS" 
  using SEC1_P256_6_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_6_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_6_dIUT SEC1_P256_6_QCAVS = Some SEC1_P256_6_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_7›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_7_QCAVSx :: nat where 
  "SEC1_P256_7_QCAVSx = 0xa9c0acade55c2a73ead1a86fb0a9713223c82475791cd0e210b046412ce224bb"

definition SEC1_P256_7_QCAVSy :: nat where 
  "SEC1_P256_7_QCAVSy = 0xf6de0afa20e93e078467c053d241903edad734c6b403ba758c2b5ff04c9d4229"

definition SEC1_P256_7_QCAVS :: "int point" where
  "SEC1_P256_7_QCAVS = Point (int SEC1_P256_7_QCAVSx) (int SEC1_P256_7_QCAVSy)"

definition SEC1_P256_7_dIUT :: nat where 
  "SEC1_P256_7_dIUT = 0xd8bf929a20ea7436b2461b541a11c80e61d826c0a4c9d322b31dd54e7f58b9c8"

definition SEC1_P256_7_QIUTx :: nat where 
  "SEC1_P256_7_QIUTx = 0x20f07631e4a6512a89ad487c4e9d63039e579cb0d7a556cb9e661cd59c1e7fa4"

definition SEC1_P256_7_QIUTy :: nat where 
  "SEC1_P256_7_QIUTy = 0x6de91846b3eee8a5ec09c2ab1f41e21bd83620ccdd1bdce3ab7ea6e02dd274f5"

definition SEC1_P256_7_QIUT :: "int point" where
  "SEC1_P256_7_QIUT = Point (int SEC1_P256_7_QIUTx) (int SEC1_P256_7_QIUTy)"

definition SEC1_P256_7_ZIUT :: nat where 
  "SEC1_P256_7_ZIUT = 0x35aa9b52536a461bfde4e85fc756be928c7de97923f0416c7a3ac8f88b3d4489"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_7_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_7_dIUT"
  by eval

lemma SEC1_P256_7_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_7_dIUT = SEC1_P256_7_QIUT" 
  by eval

lemma SEC1_P256_7_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_7_QCAVS" 
  by eval

lemma SEC1_P256_7_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_7_QCAVS" 
  using SEC1_P256_7_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_7_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_7_dIUT SEC1_P256_7_QCAVS = Some SEC1_P256_7_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_8›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_8_QCAVSx :: nat where 
  "SEC1_P256_8_QCAVSx = 0x94e94f16a98255fff2b9ac0c9598aac35487b3232d3231bd93b7db7df36f9eb9"

definition SEC1_P256_8_QCAVSy :: nat where 
  "SEC1_P256_8_QCAVSy = 0xd8049a43579cfa90b8093a94416cbefbf93386f15b3f6e190b6e3455fedfe69a"

definition SEC1_P256_8_QCAVS :: "int point" where
  "SEC1_P256_8_QCAVS = Point (int SEC1_P256_8_QCAVSx) (int SEC1_P256_8_QCAVSy)"

definition SEC1_P256_8_dIUT :: nat where 
  "SEC1_P256_8_dIUT = 0x0f9883ba0ef32ee75ded0d8bda39a5146a29f1f2507b3bd458dbea0b2bb05b4d"

definition SEC1_P256_8_QIUTx :: nat where 
  "SEC1_P256_8_QIUTx = 0xabb61b423be5d6c26e21c605832c9142dc1dfe5a5fff28726737936e6fbf516d"

definition SEC1_P256_8_QIUTy :: nat where 
  "SEC1_P256_8_QIUTy = 0x733d2513ef58beab202090586fac91bf0fee31e80ab33473ab23a2d89e58fad6"

definition SEC1_P256_8_QIUT :: "int point" where
  "SEC1_P256_8_QIUT = Point (int SEC1_P256_8_QIUTx) (int SEC1_P256_8_QIUTy)"

definition SEC1_P256_8_ZIUT :: nat where 
  "SEC1_P256_8_ZIUT = 0x605c16178a9bc875dcbff54d63fe00df699c03e8a888e9e94dfbab90b25f39b4"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_8_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_8_dIUT"
  by eval

lemma SEC1_P256_8_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_8_dIUT = SEC1_P256_8_QIUT" 
  by eval

lemma SEC1_P256_8_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_8_QCAVS" 
  by eval

lemma SEC1_P256_8_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_8_QCAVS" 
  using SEC1_P256_8_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_8_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_8_dIUT SEC1_P256_8_QCAVS = Some SEC1_P256_8_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_9›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_9_QCAVSx :: nat where 
  "SEC1_P256_9_QCAVSx = 0xe099bf2a4d557460b5544430bbf6da11004d127cb5d67f64ab07c94fcdf5274f"

definition SEC1_P256_9_QCAVSy :: nat where 
  "SEC1_P256_9_QCAVSy = 0xd9c50dbe70d714edb5e221f4e020610eeb6270517e688ca64fb0e98c7ef8c1c5"

definition SEC1_P256_9_QCAVS :: "int point" where
  "SEC1_P256_9_QCAVS = Point (int SEC1_P256_9_QCAVSx) (int SEC1_P256_9_QCAVSy)"

definition SEC1_P256_9_dIUT :: nat where 
  "SEC1_P256_9_dIUT = 0x2beedb04b05c6988f6a67500bb813faf2cae0d580c9253b6339e4a3337bb6c08"

definition SEC1_P256_9_QIUTx :: nat where 
  "SEC1_P256_9_QIUTx = 0x3d63e429cb5fa895a9247129bf4e48e89f35d7b11de8158efeb3e106a2a87395"

definition SEC1_P256_9_QIUTy :: nat where 
  "SEC1_P256_9_QIUTy = 0x0cae9e477ef41e7c8c1064379bb7b554ddcbcae79f9814281f1e50f0403c61f3"

definition SEC1_P256_9_QIUT :: "int point" where
  "SEC1_P256_9_QIUT = Point (int SEC1_P256_9_QIUTx) (int SEC1_P256_9_QIUTy)"

definition SEC1_P256_9_ZIUT :: nat where 
  "SEC1_P256_9_ZIUT = 0xf96e40a1b72840854bb62bc13c40cc2795e373d4e715980b261476835a092e0b"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_9_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_9_dIUT"
  by eval

lemma SEC1_P256_9_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_9_dIUT = SEC1_P256_9_QIUT" 
  by eval

lemma SEC1_P256_9_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_9_QCAVS" 
  by eval

lemma SEC1_P256_9_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_9_QCAVS" 
  using SEC1_P256_9_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_9_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_9_dIUT SEC1_P256_9_QCAVS = Some SEC1_P256_9_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_10›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_10_QCAVSx :: nat where 
  "SEC1_P256_10_QCAVSx = 0xf75a5fe56bda34f3c1396296626ef012dc07e4825838778a645c8248cff01658"

definition SEC1_P256_10_QCAVSy :: nat where 
  "SEC1_P256_10_QCAVSy = 0x33bbdf1b1772d8059df568b061f3f1122f28a8d819167c97be448e3dc3fb0c3c"

definition SEC1_P256_10_QCAVS :: "int point" where
  "SEC1_P256_10_QCAVS = Point (int SEC1_P256_10_QCAVSx) (int SEC1_P256_10_QCAVSy)"

definition SEC1_P256_10_dIUT :: nat where 
  "SEC1_P256_10_dIUT = 0x77c15dcf44610e41696bab758943eff1409333e4d5a11bbe72c8f6c395e9f848"

definition SEC1_P256_10_QIUTx :: nat where 
  "SEC1_P256_10_QIUTx = 0xad5d13c3db508ddcd38457e5991434a251bed49cf5ddcb59cdee73865f138c9f"

definition SEC1_P256_10_QIUTy :: nat where 
  "SEC1_P256_10_QIUTy = 0x62cec1e70588aa4fdfc7b9a09daa678081c04e1208b9d662b8a2214bf8e81a21"

definition SEC1_P256_10_QIUT :: "int point" where
  "SEC1_P256_10_QIUT = Point (int SEC1_P256_10_QIUTx) (int SEC1_P256_10_QIUTy)"

definition SEC1_P256_10_ZIUT :: nat where 
  "SEC1_P256_10_ZIUT = 0x8388fa79c4babdca02a8e8a34f9e43554976e420a4ad273c81b26e4228e9d3a3"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_10_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_10_dIUT"
  by eval

lemma SEC1_P256_10_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_10_dIUT = SEC1_P256_10_QIUT" 
  by eval

lemma SEC1_P256_10_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_10_QCAVS" 
  by eval

lemma SEC1_P256_10_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_10_QCAVS" 
  using SEC1_P256_10_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_10_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_10_dIUT SEC1_P256_10_QCAVS = Some SEC1_P256_10_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_11›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_11_QCAVSx :: nat where 
  "SEC1_P256_11_QCAVSx = 0x2db4540d50230756158abf61d9835712b6486c74312183ccefcaef2797b7674d"

definition SEC1_P256_11_QCAVSy :: nat where 
  "SEC1_P256_11_QCAVSy = 0x62f57f314e3f3495dc4e099012f5e0ba71770f9660a1eada54104cdfde77243e"

definition SEC1_P256_11_QCAVS :: "int point" where
  "SEC1_P256_11_QCAVS = Point (int SEC1_P256_11_QCAVSx) (int SEC1_P256_11_QCAVSy)"

definition SEC1_P256_11_dIUT :: nat where 
  "SEC1_P256_11_dIUT = 0x42a83b985011d12303db1a800f2610f74aa71cdf19c67d54ce6c9ed951e9093e"

definition SEC1_P256_11_QIUTx :: nat where 
  "SEC1_P256_11_QIUTx = 0xab48caa61ea35f13f8ed07ffa6a13e8db224dfecfae1a7df8b1bb6ebaf0cb97d"

definition SEC1_P256_11_QIUTy :: nat where 
  "SEC1_P256_11_QIUTy = 0x1274530ca2c385a3218bddfbcbf0b4024c9badd5243bff834ebff24a8618dccb"

definition SEC1_P256_11_QIUT :: "int point" where
  "SEC1_P256_11_QIUT = Point (int SEC1_P256_11_QIUTx) (int SEC1_P256_11_QIUTy)"

definition SEC1_P256_11_ZIUT :: nat where 
  "SEC1_P256_11_ZIUT = 0x72877cea33ccc4715038d4bcbdfe0e43f42a9e2c0c3b017fc2370f4b9acbda4a"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_11_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_11_dIUT"
  by eval

lemma SEC1_P256_11_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_11_dIUT = SEC1_P256_11_QIUT" 
  by eval

lemma SEC1_P256_11_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_11_QCAVS" 
  by eval

lemma SEC1_P256_11_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_11_QCAVS" 
  using SEC1_P256_11_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_11_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_11_dIUT SEC1_P256_11_QCAVS = Some SEC1_P256_11_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_12›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_12_QCAVSx :: nat where 
  "SEC1_P256_12_QCAVSx = 0xcd94fc9497e8990750309e9a8534fd114b0a6e54da89c4796101897041d14ecb"

definition SEC1_P256_12_QCAVSy :: nat where 
  "SEC1_P256_12_QCAVSy = 0xc3def4b5fe04faee0a11932229fff563637bfdee0e79c6deeaf449f85401c5c4"

definition SEC1_P256_12_QCAVS :: "int point" where
  "SEC1_P256_12_QCAVS = Point (int SEC1_P256_12_QCAVSx) (int SEC1_P256_12_QCAVSy)"

definition SEC1_P256_12_dIUT :: nat where 
  "SEC1_P256_12_dIUT = 0xceed35507b5c93ead5989119b9ba342cfe38e6e638ba6eea343a55475de2800b"

definition SEC1_P256_12_QIUTx :: nat where 
  "SEC1_P256_12_QIUTx = 0x9a8cd9bd72e71752df91440f77c547509a84df98114e7de4f26cdb39234a625d"

definition SEC1_P256_12_QIUTy :: nat where 
  "SEC1_P256_12_QIUTy = 0xd07cfc84c8e144fab2839f5189bb1d7c88631d579bbc58012ed9a2327da52f62"

definition SEC1_P256_12_QIUT :: "int point" where
  "SEC1_P256_12_QIUT = Point (int SEC1_P256_12_QIUTx) (int SEC1_P256_12_QIUTy)"

definition SEC1_P256_12_ZIUT :: nat where 
  "SEC1_P256_12_ZIUT = 0xe4e7408d85ff0e0e9c838003f28cdbd5247cdce31f32f62494b70e5f1bc36307"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_12_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_12_dIUT"
  by eval

lemma SEC1_P256_12_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_12_dIUT = SEC1_P256_12_QIUT" 
  by eval

lemma SEC1_P256_12_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_12_QCAVS" 
  by eval

lemma SEC1_P256_12_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_12_QCAVS" 
  using SEC1_P256_12_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_12_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_12_dIUT SEC1_P256_12_QCAVS = Some SEC1_P256_12_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_13›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_13_QCAVSx :: nat where 
  "SEC1_P256_13_QCAVSx = 0x15b9e467af4d290c417402e040426fe4cf236bae72baa392ed89780dfccdb471"

definition SEC1_P256_13_QCAVSy :: nat where 
  "SEC1_P256_13_QCAVSy = 0xcdf4e9170fb904302b8fd93a820ba8cc7ed4efd3a6f2d6b05b80b2ff2aee4e77"

definition SEC1_P256_13_QCAVS :: "int point" where
  "SEC1_P256_13_QCAVS = Point (int SEC1_P256_13_QCAVSx) (int SEC1_P256_13_QCAVSy)"

definition SEC1_P256_13_dIUT :: nat where 
  "SEC1_P256_13_dIUT = 0x43e0e9d95af4dc36483cdd1968d2b7eeb8611fcce77f3a4e7d059ae43e509604"

definition SEC1_P256_13_QIUTx :: nat where 
  "SEC1_P256_13_QIUTx = 0xf989cf8ee956a82e7ebd9881cdbfb2fd946189b08db53559bc8cfdd48071eb14"

definition SEC1_P256_13_QIUTy :: nat where 
  "SEC1_P256_13_QIUTy = 0x5eff28f1a18a616b04b7d337868679f6dd84f9a7b3d7b6f8af276c19611a541d"

definition SEC1_P256_13_QIUT :: "int point" where
  "SEC1_P256_13_QIUT = Point (int SEC1_P256_13_QIUTx) (int SEC1_P256_13_QIUTy)"

definition SEC1_P256_13_ZIUT :: nat where 
  "SEC1_P256_13_ZIUT = 0xed56bcf695b734142c24ecb1fc1bb64d08f175eb243a31f37b3d9bb4407f3b96"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_13_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_13_dIUT"
  by eval

lemma SEC1_P256_13_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_13_dIUT = SEC1_P256_13_QIUT" 
  by eval

lemma SEC1_P256_13_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_13_QCAVS" 
  by eval

lemma SEC1_P256_13_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_13_QCAVS" 
  using SEC1_P256_13_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_13_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_13_dIUT SEC1_P256_13_QCAVS = Some SEC1_P256_13_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_14›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_14_QCAVSx :: nat where 
  "SEC1_P256_14_QCAVSx = 0x49c503ba6c4fa605182e186b5e81113f075bc11dcfd51c932fb21e951eee2fa1"

definition SEC1_P256_14_QCAVSy :: nat where 
  "SEC1_P256_14_QCAVSy = 0x8af706ff0922d87b3f0c5e4e31d8b259aeb260a9269643ed520a13bb25da5924"

definition SEC1_P256_14_QCAVS :: "int point" where
  "SEC1_P256_14_QCAVS = Point (int SEC1_P256_14_QCAVSx) (int SEC1_P256_14_QCAVSy)"

definition SEC1_P256_14_dIUT :: nat where 
  "SEC1_P256_14_dIUT = 0xb2f3600df3368ef8a0bb85ab22f41fc0e5f4fdd54be8167a5c3cd4b08db04903"

definition SEC1_P256_14_QIUTx :: nat where 
  "SEC1_P256_14_QIUTx = 0x69c627625b36a429c398b45c38677cb35d8beb1cf78a571e40e99fe4eac1cd4e"

definition SEC1_P256_14_QIUTy :: nat where 
  "SEC1_P256_14_QIUTy = 0x81690112b0a88f20f7136b28d7d47e5fbc2ada3c8edd87589bc19ec9590637bd"

definition SEC1_P256_14_QIUT :: "int point" where
  "SEC1_P256_14_QIUT = Point (int SEC1_P256_14_QIUTx) (int SEC1_P256_14_QIUTy)"

definition SEC1_P256_14_ZIUT :: nat where 
  "SEC1_P256_14_ZIUT = 0xbc5c7055089fc9d6c89f83c1ea1ada879d9934b2ea28fcf4e4a7e984b28ad2cf"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_14_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_14_dIUT"
  by eval

lemma SEC1_P256_14_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_14_dIUT = SEC1_P256_14_QIUT" 
  by eval

lemma SEC1_P256_14_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_14_QCAVS" 
  by eval

lemma SEC1_P256_14_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_14_QCAVS" 
  using SEC1_P256_14_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_14_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_14_dIUT SEC1_P256_14_QCAVS = Some SEC1_P256_14_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_15›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_15_QCAVSx :: nat where 
  "SEC1_P256_15_QCAVSx = 0x19b38de39fdd2f70f7091631a4f75d1993740ba9429162c2a45312401636b29c"

definition SEC1_P256_15_QCAVSy :: nat where 
  "SEC1_P256_15_QCAVSy = 0x09aed7232b28e060941741b6828bcdfa2bc49cc844f3773611504f82a390a5ae"

definition SEC1_P256_15_QCAVS :: "int point" where
  "SEC1_P256_15_QCAVS = Point (int SEC1_P256_15_QCAVSx) (int SEC1_P256_15_QCAVSy)"

definition SEC1_P256_15_dIUT :: nat where 
  "SEC1_P256_15_dIUT = 0x4002534307f8b62a9bf67ff641ddc60fef593b17c3341239e95bdb3e579bfdc8"

definition SEC1_P256_15_QIUTx :: nat where 
  "SEC1_P256_15_QIUTx = 0x5fe964671315a18aa68a2a6e3dd1fde7e23b8ce7181471cfac43c99e1ae80262"

definition SEC1_P256_15_QIUTy :: nat where 
  "SEC1_P256_15_QIUTy = 0xd5827be282e62c84de531b963884ba832db5d6b2c3a256f0e604fe7e6b8a7f72"

definition SEC1_P256_15_QIUT :: "int point" where
  "SEC1_P256_15_QIUT = Point (int SEC1_P256_15_QIUTx) (int SEC1_P256_15_QIUTy)"

definition SEC1_P256_15_ZIUT :: nat where 
  "SEC1_P256_15_ZIUT = 0x9a4e8e657f6b0e097f47954a63c75d74fcba71a30d83651e3e5a91aa7ccd8343"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_15_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_15_dIUT"
  by eval

lemma SEC1_P256_15_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_15_dIUT = SEC1_P256_15_QIUT" 
  by eval

lemma SEC1_P256_15_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_15_QCAVS" 
  by eval

lemma SEC1_P256_15_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_15_QCAVS" 
  using SEC1_P256_15_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_15_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_15_dIUT SEC1_P256_15_QCAVS = Some SEC1_P256_15_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_16›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_16_QCAVSx :: nat where 
  "SEC1_P256_16_QCAVSx = 0x2c91c61f33adfe9311c942fdbff6ba47020feff416b7bb63cec13faf9b099954"

definition SEC1_P256_16_QCAVSy :: nat where 
  "SEC1_P256_16_QCAVSy = 0x6cab31b06419e5221fca014fb84ec870622a1b12bab5ae43682aa7ea73ea08d0"

definition SEC1_P256_16_QCAVS :: "int point" where
  "SEC1_P256_16_QCAVS = Point (int SEC1_P256_16_QCAVSx) (int SEC1_P256_16_QCAVSy)"

definition SEC1_P256_16_dIUT :: nat where 
  "SEC1_P256_16_dIUT = 0x4dfa12defc60319021b681b3ff84a10a511958c850939ed45635934ba4979147"

definition SEC1_P256_16_QIUTx :: nat where 
  "SEC1_P256_16_QIUTx = 0xc9b2b8496f1440bd4a2d1e52752fd372835b364885e154a7dac49295f281ec7c"

definition SEC1_P256_16_QIUTy :: nat where 
  "SEC1_P256_16_QIUTy = 0xfbe6b926a8a4de26ccc83b802b1212400754be25d9f3eeaf008b09870ae76321"

definition SEC1_P256_16_QIUT :: "int point" where
  "SEC1_P256_16_QIUT = Point (int SEC1_P256_16_QIUTx) (int SEC1_P256_16_QIUTy)"

definition SEC1_P256_16_ZIUT :: nat where 
  "SEC1_P256_16_ZIUT = 0x3ca1fc7ad858fb1a6aba232542f3e2a749ffc7203a2374a3f3d3267f1fc97b78"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_16_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_16_dIUT"
  by eval

lemma SEC1_P256_16_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_16_dIUT = SEC1_P256_16_QIUT" 
  by eval

lemma SEC1_P256_16_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_16_QCAVS" 
  by eval

lemma SEC1_P256_16_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_16_QCAVS" 
  using SEC1_P256_16_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_16_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_16_dIUT SEC1_P256_16_QCAVS = Some SEC1_P256_16_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_17›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_17_QCAVSx :: nat where 
  "SEC1_P256_17_QCAVSx = 0xa28a2edf58025668f724aaf83a50956b7ac1cfbbff79b08c3bf87dfd2828d767"

definition SEC1_P256_17_QCAVSy :: nat where 
  "SEC1_P256_17_QCAVSy = 0xdfa7bfffd4c766b86abeaf5c99b6e50cb9ccc9d9d00b7ffc7804b0491b67bc03"

definition SEC1_P256_17_QCAVS :: "int point" where
  "SEC1_P256_17_QCAVS = Point (int SEC1_P256_17_QCAVSx) (int SEC1_P256_17_QCAVSy)"

definition SEC1_P256_17_dIUT :: nat where 
  "SEC1_P256_17_dIUT = 0x1331f6d874a4ed3bc4a2c6e9c74331d3039796314beee3b7152fcdba5556304e"

definition SEC1_P256_17_QIUTx :: nat where 
  "SEC1_P256_17_QIUTx = 0x59e1e101521046ad9cf1d082e9d2ec7dd22530cce064991f1e55c5bcf5fcb591"

definition SEC1_P256_17_QIUTy :: nat where 
  "SEC1_P256_17_QIUTy = 0x482f4f673176c8fdaa0bb6e59b15a3e47454e3a04297d3863c9338d98add1f37"

definition SEC1_P256_17_QIUT :: "int point" where
  "SEC1_P256_17_QIUT = Point (int SEC1_P256_17_QIUTx) (int SEC1_P256_17_QIUTy)"

definition SEC1_P256_17_ZIUT :: nat where 
  "SEC1_P256_17_ZIUT = 0x1aaabe7ee6e4a6fa732291202433a237df1b49bc53866bfbe00db96a0f58224f"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_17_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_17_dIUT"
  by eval

lemma SEC1_P256_17_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_17_dIUT = SEC1_P256_17_QIUT" 
  by eval

lemma SEC1_P256_17_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_17_QCAVS" 
  by eval

lemma SEC1_P256_17_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_17_QCAVS" 
  using SEC1_P256_17_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_17_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_17_dIUT SEC1_P256_17_QCAVS = Some SEC1_P256_17_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_18›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_18_QCAVSx :: nat where 
  "SEC1_P256_18_QCAVSx = 0xa2ef857a081f9d6eb206a81c4cf78a802bdf598ae380c8886ecd85fdc1ed7644"

definition SEC1_P256_18_QCAVSy :: nat where 
  "SEC1_P256_18_QCAVSy = 0x563c4c20419f07bc17d0539fade1855e34839515b892c0f5d26561f97fa04d1a"

definition SEC1_P256_18_QCAVS :: "int point" where
  "SEC1_P256_18_QCAVS = Point (int SEC1_P256_18_QCAVSx) (int SEC1_P256_18_QCAVSy)"

definition SEC1_P256_18_dIUT :: nat where 
  "SEC1_P256_18_dIUT = 0xdd5e9f70ae740073ca0204df60763fb6036c45709bf4a7bb4e671412fad65da3"

definition SEC1_P256_18_QIUTx :: nat where 
  "SEC1_P256_18_QIUTx = 0x30b9db2e2e977bcdc98cb87dd736cbd8e78552121925cf16e1933657c2fb2314"

definition SEC1_P256_18_QIUTy :: nat where 
  "SEC1_P256_18_QIUTy = 0x6a45028800b81291bce5c2e1fed7ded650620ebbe6050c6f3a7f0dfb4673ab5c"

definition SEC1_P256_18_QIUT :: "int point" where
  "SEC1_P256_18_QIUT = Point (int SEC1_P256_18_QIUTx) (int SEC1_P256_18_QIUTy)"

definition SEC1_P256_18_ZIUT :: nat where 
  "SEC1_P256_18_ZIUT = 0x430e6a4fba4449d700d2733e557f66a3bf3d50517c1271b1ddae1161b7ac798c"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_18_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_18_dIUT"
  by eval

lemma SEC1_P256_18_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_18_dIUT = SEC1_P256_18_QIUT" 
  by eval

lemma SEC1_P256_18_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_18_QCAVS" 
  by eval

lemma SEC1_P256_18_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_18_QCAVS" 
  using SEC1_P256_18_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_18_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_18_dIUT SEC1_P256_18_QCAVS = Some SEC1_P256_18_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_19›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_19_QCAVSx :: nat where 
  "SEC1_P256_19_QCAVSx = 0xccd8a2d86bc92f2e01bce4d6922cf7fe1626aed044685e95e2eebd464505f01f"

definition SEC1_P256_19_QCAVSy :: nat where 
  "SEC1_P256_19_QCAVSy = 0xe9ddd583a9635a667777d5b8a8f31b0f79eba12c75023410b54b8567dddc0f38"

definition SEC1_P256_19_QCAVS :: "int point" where
  "SEC1_P256_19_QCAVS = Point (int SEC1_P256_19_QCAVSx) (int SEC1_P256_19_QCAVSy)"

definition SEC1_P256_19_dIUT :: nat where 
  "SEC1_P256_19_dIUT = 0x5ae026cfc060d55600717e55b8a12e116d1d0df34af831979057607c2d9c2f76"

definition SEC1_P256_19_QIUTx :: nat where 
  "SEC1_P256_19_QIUTx = 0x46c9ebd1a4a3c8c0b6d572b5dcfba12467603208a9cb5d2acfbb733c40cf6391"

definition SEC1_P256_19_QIUTy :: nat where 
  "SEC1_P256_19_QIUTy = 0x46c913a27d044185d38b467ace011e04d4d9bbbb8cb9ae25fa92aaf15a595e86"

definition SEC1_P256_19_QIUT :: "int point" where
  "SEC1_P256_19_QIUT = Point (int SEC1_P256_19_QIUTx) (int SEC1_P256_19_QIUTy)"

definition SEC1_P256_19_ZIUT :: nat where 
  "SEC1_P256_19_ZIUT = 0x1ce9e6740529499f98d1f1d71329147a33df1d05e4765b539b11cf615d6974d3"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_19_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_19_dIUT"
  by eval

lemma SEC1_P256_19_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_19_dIUT = SEC1_P256_19_QIUT" 
  by eval

lemma SEC1_P256_19_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_19_QCAVS" 
  by eval

lemma SEC1_P256_19_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_19_QCAVS" 
  using SEC1_P256_19_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_19_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_19_dIUT SEC1_P256_19_QCAVS = Some SEC1_P256_19_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_20›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_20_QCAVSx :: nat where 
  "SEC1_P256_20_QCAVSx = 0xc188ffc8947f7301fb7b53e36746097c2134bf9cc981ba74b4e9c4361f595e4e"

definition SEC1_P256_20_QCAVSy :: nat where 
  "SEC1_P256_20_QCAVSy = 0xbf7d2f2056e72421ef393f0c0f2b0e00130e3cac4abbcc00286168e85ec55051"

definition SEC1_P256_20_QCAVS :: "int point" where
  "SEC1_P256_20_QCAVS = Point (int SEC1_P256_20_QCAVSx) (int SEC1_P256_20_QCAVSy)"

definition SEC1_P256_20_dIUT :: nat where 
  "SEC1_P256_20_dIUT = 0xb601ac425d5dbf9e1735c5e2d5bdb79ca98b3d5be4a2cfd6f2273f150e064d9d"

definition SEC1_P256_20_QIUTx :: nat where 
  "SEC1_P256_20_QIUTx = 0x7c9e950841d26c8dde8994398b8f5d475a022bc63de7773fcf8d552e01f1ba0a"

definition SEC1_P256_20_QIUTy :: nat where 
  "SEC1_P256_20_QIUTy = 0xcc42b9885c9b3bee0f8d8c57d3a8f6355016c019c4062fa22cff2f209b5cc2e1"

definition SEC1_P256_20_QIUT :: "int point" where
  "SEC1_P256_20_QIUT = Point (int SEC1_P256_20_QIUTx) (int SEC1_P256_20_QIUTy)"

definition SEC1_P256_20_ZIUT :: nat where 
  "SEC1_P256_20_ZIUT = 0x4690e3743c07d643f1bc183636ab2a9cb936a60a802113c49bb1b3f2d0661660"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_20_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_20_dIUT"
  by eval

lemma SEC1_P256_20_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_20_dIUT = SEC1_P256_20_QIUT" 
  by eval

lemma SEC1_P256_20_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_20_QCAVS" 
  by eval

lemma SEC1_P256_20_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_20_QCAVS" 
  using SEC1_P256_20_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_20_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_20_dIUT SEC1_P256_20_QCAVS = Some SEC1_P256_20_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_21›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_21_QCAVSx :: nat where 
  "SEC1_P256_21_QCAVSx = 0x317e1020ff53fccef18bf47bb7f2dd7707fb7b7a7578e04f35b3beed222a0eb6"

definition SEC1_P256_21_QCAVSy :: nat where 
  "SEC1_P256_21_QCAVSy = 0x09420ce5a19d77c6fe1ee587e6a49fbaf8f280e8df033d75403302e5a27db2ae"

definition SEC1_P256_21_QCAVS :: "int point" where
  "SEC1_P256_21_QCAVS = Point (int SEC1_P256_21_QCAVSx) (int SEC1_P256_21_QCAVSy)"

definition SEC1_P256_21_dIUT :: nat where 
  "SEC1_P256_21_dIUT = 0xfefb1dda1845312b5fce6b81b2be205af2f3a274f5a212f66c0d9fc33d7ae535"

definition SEC1_P256_21_QIUTx :: nat where 
  "SEC1_P256_21_QIUTx = 0x38b54db85500cb20c61056edd3d88b6a9dc26780a047f213a6e1b900f76596eb"

definition SEC1_P256_21_QIUTy :: nat where 
  "SEC1_P256_21_QIUTy = 0x6387e4e5781571e4eb8ae62991a33b5dc33301c5bc7e125d53794a39160d8fd0"

definition SEC1_P256_21_QIUT :: "int point" where
  "SEC1_P256_21_QIUT = Point (int SEC1_P256_21_QIUTx) (int SEC1_P256_21_QIUTy)"

definition SEC1_P256_21_ZIUT :: nat where 
  "SEC1_P256_21_ZIUT = 0x30c2261bd0004e61feda2c16aa5e21ffa8d7e7f7dbf6ec379a43b48e4b36aeb0"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_21_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_21_dIUT"
  by eval

lemma SEC1_P256_21_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_21_dIUT = SEC1_P256_21_QIUT" 
  by eval

lemma SEC1_P256_21_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_21_QCAVS" 
  by eval

lemma SEC1_P256_21_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_21_QCAVS" 
  using SEC1_P256_21_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_21_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_21_dIUT SEC1_P256_21_QCAVS = Some SEC1_P256_21_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_22›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_22_QCAVSx :: nat where 
  "SEC1_P256_22_QCAVSx = 0x45fb02b2ceb9d7c79d9c2fa93e9c7967c2fa4df5789f9640b24264b1e524fcb1"

definition SEC1_P256_22_QCAVSy :: nat where 
  "SEC1_P256_22_QCAVSy = 0x5c6e8ecf1f7d3023893b7b1ca1e4d178972ee2a230757ddc564ffe37f5c5a321"

definition SEC1_P256_22_QCAVS :: "int point" where
  "SEC1_P256_22_QCAVS = Point (int SEC1_P256_22_QCAVSx) (int SEC1_P256_22_QCAVSy)"

definition SEC1_P256_22_dIUT :: nat where 
  "SEC1_P256_22_dIUT = 0x334ae0c4693d23935a7e8e043ebbde21e168a7cba3fa507c9be41d7681e049ce"

definition SEC1_P256_22_QIUTx :: nat where 
  "SEC1_P256_22_QIUTx = 0x3f2bf1589abf3047bf3e54ac9a95379bff95f8f55405f64eca36a7eebe8ffca7"

definition SEC1_P256_22_QIUTy :: nat where 
  "SEC1_P256_22_QIUTy = 0x5212a94e66c5ae9a8991872f66a72723d80ec5b2e925745c456f5371943b3a06"

definition SEC1_P256_22_QIUT :: "int point" where
  "SEC1_P256_22_QIUT = Point (int SEC1_P256_22_QIUTx) (int SEC1_P256_22_QIUTy)"

definition SEC1_P256_22_ZIUT :: nat where 
  "SEC1_P256_22_ZIUT = 0x2adae4a138a239dcd93c243a3803c3e4cf96e37fe14e6a9b717be9599959b11c"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_22_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_22_dIUT"
  by eval

lemma SEC1_P256_22_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_22_dIUT = SEC1_P256_22_QIUT" 
  by eval

lemma SEC1_P256_22_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_22_QCAVS" 
  by eval

lemma SEC1_P256_22_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_22_QCAVS" 
  using SEC1_P256_22_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_22_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_22_dIUT SEC1_P256_22_QCAVS = Some SEC1_P256_22_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_23›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_23_QCAVSx :: nat where 
  "SEC1_P256_23_QCAVSx = 0xa19ef7bff98ada781842fbfc51a47aff39b5935a1c7d9625c8d323d511c92de6"

definition SEC1_P256_23_QCAVSy :: nat where 
  "SEC1_P256_23_QCAVSy = 0xe9c184df75c955e02e02e400ffe45f78f339e1afe6d056fb3245f4700ce606ef"

definition SEC1_P256_23_QCAVS :: "int point" where
  "SEC1_P256_23_QCAVS = Point (int SEC1_P256_23_QCAVSx) (int SEC1_P256_23_QCAVSy)"

definition SEC1_P256_23_dIUT :: nat where 
  "SEC1_P256_23_dIUT = 0x2c4bde40214fcc3bfc47d4cf434b629acbe9157f8fd0282540331de7942cf09d"

definition SEC1_P256_23_QIUTx :: nat where 
  "SEC1_P256_23_QIUTx = 0x29c0807f10cbc42fb45c9989da50681eead716daa7b9e91fd32e062f5eb92ca0"

definition SEC1_P256_23_QIUTy :: nat where 
  "SEC1_P256_23_QIUTy = 0xff1d6d1955d7376b2da24fe1163a271659136341bc2eb1195fc706dc62e7f34d"

definition SEC1_P256_23_QIUT :: "int point" where
  "SEC1_P256_23_QIUT = Point (int SEC1_P256_23_QIUTx) (int SEC1_P256_23_QIUTy)"

definition SEC1_P256_23_ZIUT :: nat where 
  "SEC1_P256_23_ZIUT = 0x2e277ec30f5ea07d6ce513149b9479b96e07f4b6913b1b5c11305c1444a1bc0b"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_23_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_23_dIUT"
  by eval

lemma SEC1_P256_23_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_23_dIUT = SEC1_P256_23_QIUT" 
  by eval

lemma SEC1_P256_23_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_23_QCAVS" 
  by eval

lemma SEC1_P256_23_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_23_QCAVS" 
  using SEC1_P256_23_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_23_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_23_dIUT SEC1_P256_23_QCAVS = Some SEC1_P256_23_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P256_24›

subsubsection‹Given Test Vector Values›

definition SEC1_P256_24_QCAVSx :: nat where 
  "SEC1_P256_24_QCAVSx = 0x356c5a444c049a52fee0adeb7e5d82ae5aa83030bfff31bbf8ce2096cf161c4b"

definition SEC1_P256_24_QCAVSy :: nat where 
  "SEC1_P256_24_QCAVSy = 0x57d128de8b2a57a094d1a001e572173f96e8866ae352bf29cddaf92fc85b2f92"

definition SEC1_P256_24_QCAVS :: "int point" where
  "SEC1_P256_24_QCAVS = Point (int SEC1_P256_24_QCAVSx) (int SEC1_P256_24_QCAVSy)"

definition SEC1_P256_24_dIUT :: nat where 
  "SEC1_P256_24_dIUT = 0x85a268f9d7772f990c36b42b0a331adc92b5941de0b862d5d89a347cbf8faab0"

definition SEC1_P256_24_QIUTx :: nat where 
  "SEC1_P256_24_QIUTx = 0x9cf4b98581ca1779453cc816ff28b4100af56cf1bf2e5bc312d83b6b1b21d333"

definition SEC1_P256_24_QIUTy :: nat where 
  "SEC1_P256_24_QIUTy = 0x7a5504fcac5231a0d12d658218284868229c844a04a3450d6c7381abe080bf3b"

definition SEC1_P256_24_QIUT :: "int point" where
  "SEC1_P256_24_QIUT = Point (int SEC1_P256_24_QIUTx) (int SEC1_P256_24_QIUTy)"

definition SEC1_P256_24_ZIUT :: nat where 
  "SEC1_P256_24_ZIUT = 0x1e51373bd2c6044c129c436e742a55be2a668a85ae08441b6756445df5493857"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P256_24_dIUT_valid: 
  "SEC1_P256_ECprivateKeyValid SEC1_P256_24_dIUT"
  by eval

lemma SEC1_P256_24_IUT_validPair: 
  "SEC1_P256_ECkeyGen SEC1_P256_24_dIUT = SEC1_P256_24_QIUT" 
  by eval

lemma SEC1_P256_24_QCAVS_partialValid: 
  "SEC1_P256_ECpublicKeyPartialValid SEC1_P256_24_QCAVS" 
  by eval

lemma SEC1_P256_24_QCAVS_validPublicKey: 
  "SEC1_P256_ECpublicKeyValid SEC1_P256_24_QCAVS" 
  using SEC1_P256_24_QCAVS_partialValid SEC1_P256.partValidImpliesValidIFheq1 P256_h_def by blast

lemma SEC1_P256_24_ECDHprim_Check: 
  "SEC1_P256_ECDHprim SEC1_P256_24_dIUT SEC1_P256_24_QCAVS = Some SEC1_P256_24_ZIUT"
  by eval
  
section‹NIST Test Vectors for ECDH: Curve P-384›
  
subsection‹Test Vector: SEC1_P384_0›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_0_QCAVSx :: nat where 
  "SEC1_P384_0_QCAVSx = 0xa7c76b970c3b5fe8b05d2838ae04ab47697b9eaf52e764592efda27fe7513272734466b400091adbf2d68c58e0c50066"

definition SEC1_P384_0_QCAVSy :: nat where 
  "SEC1_P384_0_QCAVSy = 0xac68f19f2e1cb879aed43a9969b91a0839c4c38a49749b661efedf243451915ed0905a32b060992b468c64766fc8437a"

definition SEC1_P384_0_QCAVS :: "int point" where
  "SEC1_P384_0_QCAVS = Point (int SEC1_P384_0_QCAVSx) (int SEC1_P384_0_QCAVSy)"

definition SEC1_P384_0_dIUT :: nat where 
  "SEC1_P384_0_dIUT = 0x3cc3122a68f0d95027ad38c067916ba0eb8c38894d22e1b15618b6818a661774ad463b205da88cf699ab4d43c9cf98a1"

definition SEC1_P384_0_QIUTx :: nat where 
  "SEC1_P384_0_QIUTx = 0x9803807f2f6d2fd966cdd0290bd410c0190352fbec7ff6247de1302df86f25d34fe4a97bef60cff548355c015dbb3e5f"

definition SEC1_P384_0_QIUTy :: nat where 
  "SEC1_P384_0_QIUTy = 0xba26ca69ec2f5b5d9dad20cc9da711383a9dbe34ea3fa5a2af75b46502629ad54dd8b7d73a8abb06a3a3be47d650cc99"

definition SEC1_P384_0_QIUT :: "int point" where
  "SEC1_P384_0_QIUT = Point (int SEC1_P384_0_QIUTx) (int SEC1_P384_0_QIUTy)"

definition SEC1_P384_0_ZIUT :: nat where 
  "SEC1_P384_0_ZIUT = 0x5f9d29dc5e31a163060356213669c8ce132e22f57c9a04f40ba7fcead493b457e5621e766c40a2e3d4d6a04b25e533f1"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_0_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_0_dIUT"
  by eval

lemma SEC1_P384_0_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_0_dIUT = SEC1_P384_0_QIUT" 
  by eval

lemma SEC1_P384_0_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_0_QCAVS" 
  by eval

lemma SEC1_P384_0_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_0_QCAVS" 
  using SEC1_P384_0_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_0_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_0_dIUT SEC1_P384_0_QCAVS = Some SEC1_P384_0_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_1›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_1_QCAVSx :: nat where 
  "SEC1_P384_1_QCAVSx = 0x30f43fcf2b6b00de53f624f1543090681839717d53c7c955d1d69efaf0349b7363acb447240101cbb3af6641ce4b88e0"

definition SEC1_P384_1_QCAVSy :: nat where 
  "SEC1_P384_1_QCAVSy = 0x25e46c0c54f0162a77efcc27b6ea792002ae2ba82714299c860857a68153ab62e525ec0530d81b5aa15897981e858757"

definition SEC1_P384_1_QCAVS :: "int point" where
  "SEC1_P384_1_QCAVS = Point (int SEC1_P384_1_QCAVSx) (int SEC1_P384_1_QCAVSy)"

definition SEC1_P384_1_dIUT :: nat where 
  "SEC1_P384_1_dIUT = 0x92860c21bde06165f8e900c687f8ef0a05d14f290b3f07d8b3a8cc6404366e5d5119cd6d03fb12dc58e89f13df9cd783"

definition SEC1_P384_1_QIUTx :: nat where 
  "SEC1_P384_1_QIUTx = 0xea4018f5a307c379180bf6a62fd2ceceebeeb7d4df063a66fb838aa35243419791f7e2c9d4803c9319aa0eb03c416b66"

definition SEC1_P384_1_QIUTy :: nat where 
  "SEC1_P384_1_QIUTy = 0x68835a91484f05ef028284df6436fb88ffebabcdd69ab0133e6735a1bcfb37203d10d340a8328a7b68770ca75878a1a6"

definition SEC1_P384_1_QIUT :: "int point" where
  "SEC1_P384_1_QIUT = Point (int SEC1_P384_1_QIUTx) (int SEC1_P384_1_QIUTy)"

definition SEC1_P384_1_ZIUT :: nat where 
  "SEC1_P384_1_ZIUT = 0xa23742a2c267d7425fda94b93f93bbcc24791ac51cd8fd501a238d40812f4cbfc59aac9520d758cf789c76300c69d2ff"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_1_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_1_dIUT"
  by eval

lemma SEC1_P384_1_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_1_dIUT = SEC1_P384_1_QIUT" 
  by eval

lemma SEC1_P384_1_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_1_QCAVS" 
  by eval

lemma SEC1_P384_1_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_1_QCAVS" 
  using SEC1_P384_1_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_1_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_1_dIUT SEC1_P384_1_QCAVS = Some SEC1_P384_1_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_2›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_2_QCAVSx :: nat where 
  "SEC1_P384_2_QCAVSx = 0x1aefbfa2c6c8c855a1a216774550b79a24cda37607bb1f7cc906650ee4b3816d68f6a9c75da6e4242cebfb6652f65180"

definition SEC1_P384_2_QCAVSy :: nat where 
  "SEC1_P384_2_QCAVSy = 0x419d28b723ebadb7658fcebb9ad9b7adea674f1da3dc6b6397b55da0f61a3eddacb4acdb14441cb214b04a0844c02fa3"

definition SEC1_P384_2_QCAVS :: "int point" where
  "SEC1_P384_2_QCAVS = Point (int SEC1_P384_2_QCAVSx) (int SEC1_P384_2_QCAVSy)"

definition SEC1_P384_2_dIUT :: nat where 
  "SEC1_P384_2_dIUT = 0x12cf6a223a72352543830f3f18530d5cb37f26880a0b294482c8a8ef8afad09aa78b7dc2f2789a78c66af5d1cc553853"

definition SEC1_P384_2_QIUTx :: nat where 
  "SEC1_P384_2_QIUTx = 0xfcfcea085e8cf74d0dced1620ba8423694f903a219bbf901b0b59d6ac81baad316a242ba32bde85cb248119b852fab66"

definition SEC1_P384_2_QIUTy :: nat where 
  "SEC1_P384_2_QIUTy = 0x972e3c68c7ab402c5836f2a16ed451a33120a7750a6039f3ff15388ee622b7065f7122bf6d51aefbc29b37b03404581b"

definition SEC1_P384_2_QIUT :: "int point" where
  "SEC1_P384_2_QIUT = Point (int SEC1_P384_2_QIUTx) (int SEC1_P384_2_QIUTy)"

definition SEC1_P384_2_ZIUT :: nat where 
  "SEC1_P384_2_ZIUT = 0x3d2e640f350805eed1ff43b40a72b2abed0a518bcebe8f2d15b111b6773223da3c3489121db173d414b5bd5ad7153435"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_2_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_2_dIUT"
  by eval

lemma SEC1_P384_2_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_2_dIUT = SEC1_P384_2_QIUT" 
  by eval

lemma SEC1_P384_2_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_2_QCAVS" 
  by eval

lemma SEC1_P384_2_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_2_QCAVS" 
  using SEC1_P384_2_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_2_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_2_dIUT SEC1_P384_2_QCAVS = Some SEC1_P384_2_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_3›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_3_QCAVSx :: nat where 
  "SEC1_P384_3_QCAVSx = 0x8bc089326ec55b9cf59b34f0eb754d93596ca290fcb3444c83d4de3a5607037ec397683f8cef07eab2fe357eae36c449"

definition SEC1_P384_3_QCAVSy :: nat where 
  "SEC1_P384_3_QCAVSy = 0xd9d16ce8ac85b3f1e94568521aae534e67139e310ec72693526aa2e927b5b322c95a1a033c229cb6770c957cd3148dd7"

definition SEC1_P384_3_QCAVS :: "int point" where
  "SEC1_P384_3_QCAVS = Point (int SEC1_P384_3_QCAVSx) (int SEC1_P384_3_QCAVSy)"

definition SEC1_P384_3_dIUT :: nat where 
  "SEC1_P384_3_dIUT = 0x8dd48063a3a058c334b5cc7a4ce07d02e5ee6d8f1f3c51a1600962cbab462690ae3cd974fb39e40b0e843daa0fd32de1"

definition SEC1_P384_3_QIUTx :: nat where 
  "SEC1_P384_3_QIUTx = 0xe38c9846248123c3421861ea4d32669a7b5c3c08376ad28104399494c84ff5efa3894adb2c6cbe8c3c913ef2eec5bd3c"

definition SEC1_P384_3_QIUTy :: nat where 
  "SEC1_P384_3_QIUTy = 0x9fa84024a1028796df84021f7b6c9d02f0f4bd1a612a03cbf75a0beea43fef8ae84b48c60172aadf09c1ad016d0bf3ce"

definition SEC1_P384_3_QIUT :: "int point" where
  "SEC1_P384_3_QIUT = Point (int SEC1_P384_3_QIUTx) (int SEC1_P384_3_QIUTy)"

definition SEC1_P384_3_ZIUT :: nat where 
  "SEC1_P384_3_ZIUT = 0x6a42cfc392aba0bfd3d17b7ccf062b91fc09bbf3417612d02a90bdde62ae40c54bb2e56e167d6b70db670097eb8db854"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_3_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_3_dIUT"
  by eval

lemma SEC1_P384_3_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_3_dIUT = SEC1_P384_3_QIUT" 
  by eval

lemma SEC1_P384_3_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_3_QCAVS" 
  by eval

lemma SEC1_P384_3_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_3_QCAVS" 
  using SEC1_P384_3_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_3_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_3_dIUT SEC1_P384_3_QCAVS = Some SEC1_P384_3_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_4›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_4_QCAVSx :: nat where 
  "SEC1_P384_4_QCAVSx = 0xeb952e2d9ac0c20c6cc48fb225c2ad154f53c8750b003fd3b4ed8ed1dc0defac61bcdde02a2bcfee7067d75d342ed2b0"

definition SEC1_P384_4_QCAVSy :: nat where 
  "SEC1_P384_4_QCAVSy = 0xf1828205baece82d1b267d0d7ff2f9c9e15b69a72df47058a97f3891005d1fb38858f5603de840e591dfa4f6e7d489e1"

definition SEC1_P384_4_QCAVS :: "int point" where
  "SEC1_P384_4_QCAVS = Point (int SEC1_P384_4_QCAVSx) (int SEC1_P384_4_QCAVSy)"

definition SEC1_P384_4_dIUT :: nat where 
  "SEC1_P384_4_dIUT = 0x84ece6cc3429309bd5b23e959793ed2b111ec5cb43b6c18085fcaea9efa0685d98a6262ee0d330ee250bc8a67d0e733f"

definition SEC1_P384_4_QIUTx :: nat where 
  "SEC1_P384_4_QIUTx = 0x3222063a2997b302ee60ee1961108ff4c7acf1c0ef1d5fb0d164b84bce71c431705cb9aea9a45f5d73806655a058bee3"

definition SEC1_P384_4_QIUTy :: nat where 
  "SEC1_P384_4_QIUTy = 0xe61fa9e7fbe7cd43abf99596a3d3a039e99fa9dc93b0bdd9cad81966d17eeaf557068afa7c78466bb5b22032d1100fa6"

definition SEC1_P384_4_QIUT :: "int point" where
  "SEC1_P384_4_QIUT = Point (int SEC1_P384_4_QIUTx) (int SEC1_P384_4_QIUTy)"

definition SEC1_P384_4_ZIUT :: nat where 
  "SEC1_P384_4_ZIUT = 0xce7ba454d4412729a32bb833a2d1fd2ae612d4667c3a900e069214818613447df8c611de66da200db7c375cf913e4405"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_4_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_4_dIUT"
  by eval

lemma SEC1_P384_4_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_4_dIUT = SEC1_P384_4_QIUT" 
  by eval

lemma SEC1_P384_4_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_4_QCAVS" 
  by eval

lemma SEC1_P384_4_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_4_QCAVS" 
  using SEC1_P384_4_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_4_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_4_dIUT SEC1_P384_4_QCAVS = Some SEC1_P384_4_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_5›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_5_QCAVSx :: nat where 
  "SEC1_P384_5_QCAVSx = 0x441d029e244eb7168d647d4df50db5f4e4974ab3fdaf022aff058b3695d0b8c814cc88da6285dc6df1ac55c553885003"

definition SEC1_P384_5_QCAVSy :: nat where 
  "SEC1_P384_5_QCAVSy = 0xe8025ac23a41d4b1ea2aa46c50c6e479946b59b6d76497cd9249977e0bfe4a6262622f13d42a3c43d66bdbb30403c345"

definition SEC1_P384_5_QCAVS :: "int point" where
  "SEC1_P384_5_QCAVS = Point (int SEC1_P384_5_QCAVSx) (int SEC1_P384_5_QCAVSy)"

definition SEC1_P384_5_dIUT :: nat where 
  "SEC1_P384_5_dIUT = 0x68fce2121dc3a1e37b10f1dde309f9e2e18fac47cd1770951451c3484cdb77cb136d00e731260597cc2859601c01a25b"

definition SEC1_P384_5_QIUTx :: nat where 
  "SEC1_P384_5_QIUTx = 0x868be0e694841830e424d913d8e7d86b84ee1021d82b0ecf523f09fe89a76c0c95c49f2dfbcf829c1e39709d55efbb3b"

definition SEC1_P384_5_QIUTy :: nat where 
  "SEC1_P384_5_QIUTy = 0x9195eb183675b40fd92f51f37713317e4a9b4f715c8ab22e0773b1bc71d3a219f05b8116074658ee86b52e36f3897116"

definition SEC1_P384_5_QIUT :: "int point" where
  "SEC1_P384_5_QIUT = Point (int SEC1_P384_5_QIUTx) (int SEC1_P384_5_QIUTy)"

definition SEC1_P384_5_ZIUT :: nat where 
  "SEC1_P384_5_ZIUT = 0xba69f0acdf3e1ca95caaac4ecaf475bbe51b54777efce01ca381f45370e486fe87f9f419b150c61e329a286d1aa265ec"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_5_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_5_dIUT"
  by eval

lemma SEC1_P384_5_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_5_dIUT = SEC1_P384_5_QIUT" 
  by eval

lemma SEC1_P384_5_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_5_QCAVS" 
  by eval

lemma SEC1_P384_5_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_5_QCAVS" 
  using SEC1_P384_5_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_5_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_5_dIUT SEC1_P384_5_QCAVS = Some SEC1_P384_5_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_6›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_6_QCAVSx :: nat where 
  "SEC1_P384_6_QCAVSx = 0x3d4e6bf08a73404accc1629873468e4269e82d90d832e58ad72142639b5a056ad8d35c66c60e8149fac0c797bceb7c2f"

definition SEC1_P384_6_QCAVSy :: nat where 
  "SEC1_P384_6_QCAVSy = 0x9b0308dc7f0e6d29f8c277acbc65a21e5adb83d11e6873bc0a07fda0997f482504602f59e10bc5cb476b83d0a4f75e71"

definition SEC1_P384_6_QCAVS :: "int point" where
  "SEC1_P384_6_QCAVS = Point (int SEC1_P384_6_QCAVSx) (int SEC1_P384_6_QCAVSy)"

definition SEC1_P384_6_dIUT :: nat where 
  "SEC1_P384_6_dIUT = 0xb1764c54897e7aae6de9e7751f2f37de849291f88f0f91093155b858d1cc32a3a87980f706b86cc83f927bdfdbeae0bd"

definition SEC1_P384_6_QIUTx :: nat where 
  "SEC1_P384_6_QIUTx = 0xc371222feaa6770c6f3ea3e0dac9740def4fcf821378b7f91ff937c21e0470f70f3a31d5c6b2912195f10926942b48ae"

definition SEC1_P384_6_QIUTy :: nat where 
  "SEC1_P384_6_QIUTy = 0x047d6b4d765123563f81116bc665b7b8cc6207830d805fd84da7cb805a65baa7c12fd592d1b5b5e3e65d9672a9ef7662"

definition SEC1_P384_6_QIUT :: "int point" where
  "SEC1_P384_6_QIUT = Point (int SEC1_P384_6_QIUTx) (int SEC1_P384_6_QIUTy)"

definition SEC1_P384_6_ZIUT :: nat where 
  "SEC1_P384_6_ZIUT = 0x1a6688ee1d6e59865d8e3ada37781d36bb0c2717eef92e61964d3927cb765c2965ea80f7f63e58c322ba0397faeaf62b"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_6_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_6_dIUT"
  by eval

lemma SEC1_P384_6_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_6_dIUT = SEC1_P384_6_QIUT" 
  by eval

lemma SEC1_P384_6_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_6_QCAVS" 
  by eval

lemma SEC1_P384_6_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_6_QCAVS" 
  using SEC1_P384_6_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_6_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_6_dIUT SEC1_P384_6_QCAVS = Some SEC1_P384_6_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_7›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_7_QCAVSx :: nat where 
  "SEC1_P384_7_QCAVSx = 0xf5f6bef1d110da03be0017eac760cc34b24d092f736f237bc7054b3865312a813bcb62d297fb10a4f7abf54708fe2d3d"

definition SEC1_P384_7_QCAVSy :: nat where 
  "SEC1_P384_7_QCAVSy = 0x06fdf8d7dc032f4e10010bf19cbf6159321252ff415fb91920d438f24e67e60c2eb0463204679fa356af44cea9c9ebf5"

definition SEC1_P384_7_QCAVS :: "int point" where
  "SEC1_P384_7_QCAVS = Point (int SEC1_P384_7_QCAVSx) (int SEC1_P384_7_QCAVSy)"

definition SEC1_P384_7_dIUT :: nat where 
  "SEC1_P384_7_dIUT = 0xf0f7a96e70d98fd5a30ad6406cf56eb5b72a510e9f192f50e1f84524dbf3d2439f7287bb36f5aa912a79deaab4adea82"

definition SEC1_P384_7_QIUTx :: nat where 
  "SEC1_P384_7_QIUTx = 0x99c8c41cb1ab5e0854a346e4b08a537c1706a61553387c8d94943ab15196d40dbaa55b8210a77a5d00915f2c4ea69eab"

definition SEC1_P384_7_QIUTy :: nat where 
  "SEC1_P384_7_QIUTy = 0x5531065bdcf17bfb3cb55a02e41a57c7f694c383ad289f900fbd656c2233a93c92e933e7a26f54cbb56f0ad875c51bb0"

definition SEC1_P384_7_QIUT :: "int point" where
  "SEC1_P384_7_QIUT = Point (int SEC1_P384_7_QIUTx) (int SEC1_P384_7_QIUTy)"

definition SEC1_P384_7_ZIUT :: nat where 
  "SEC1_P384_7_ZIUT = 0xd06a568bf2336b90cbac325161be7695eacb2295f599500d787f072612aca313ee5d874f807ddef6c1f023fe2b6e7cd0"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_7_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_7_dIUT"
  by eval

lemma SEC1_P384_7_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_7_dIUT = SEC1_P384_7_QIUT" 
  by eval

lemma SEC1_P384_7_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_7_QCAVS" 
  by eval

lemma SEC1_P384_7_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_7_QCAVS" 
  using SEC1_P384_7_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_7_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_7_dIUT SEC1_P384_7_QCAVS = Some SEC1_P384_7_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_8›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_8_QCAVSx :: nat where 
  "SEC1_P384_8_QCAVSx = 0x7cdec77e0737ea37c67b89b7137fe38818010f4464438ee4d1d35a0c488cad3fde2f37d00885d36d3b795b9f93d23a67"

definition SEC1_P384_8_QCAVSy :: nat where 
  "SEC1_P384_8_QCAVSy = 0x28c42ee8d6027c56cf979ba4c229fdb01d234944f8ac433650112c3cf0f02844e888a3569dfef7828a8a884589aa055e"

definition SEC1_P384_8_QCAVS :: "int point" where
  "SEC1_P384_8_QCAVS = Point (int SEC1_P384_8_QCAVSx) (int SEC1_P384_8_QCAVSy)"

definition SEC1_P384_8_dIUT :: nat where 
  "SEC1_P384_8_dIUT = 0x9efb87ddc61d43c482ba66e1b143aef678fbd0d1bebc2000941fabe677fe5b706bf78fce36d100b17cc787ead74bbca2"

definition SEC1_P384_8_QIUTx :: nat where 
  "SEC1_P384_8_QIUTx = 0x4c34efee8f0c95565d2065d1bbac2a2dd25ae964320eb6bccedc5f3a9b42a881a1afca1bb6b880584fa27b01c193cd92"

definition SEC1_P384_8_QIUTy :: nat where 
  "SEC1_P384_8_QIUTy = 0xd8fb01dbf7cd0a3868c26b951f393c3c56c2858cee901f7793ff5d271925d13a41f8e52409f4eba1990f33acb0bac669"

definition SEC1_P384_8_QIUT :: "int point" where
  "SEC1_P384_8_QIUT = Point (int SEC1_P384_8_QIUTx) (int SEC1_P384_8_QIUTy)"

definition SEC1_P384_8_ZIUT :: nat where 
  "SEC1_P384_8_ZIUT = 0xbb3b1eda9c6560d82ff5bee403339f1e80342338a991344853b56b24f109a4d94b92f654f0425edd4c205903d7586104"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_8_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_8_dIUT"
  by eval

lemma SEC1_P384_8_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_8_dIUT = SEC1_P384_8_QIUT" 
  by eval

lemma SEC1_P384_8_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_8_QCAVS" 
  by eval

lemma SEC1_P384_8_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_8_QCAVS" 
  using SEC1_P384_8_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_8_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_8_dIUT SEC1_P384_8_QCAVS = Some SEC1_P384_8_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_9›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_9_QCAVSx :: nat where 
  "SEC1_P384_9_QCAVSx = 0x8eeea3a319c8df99fbc29cb55f243a720d95509515ee5cc587a5c5ae22fbbd009e626db3e911def0b99a4f7ae304b1ba"

definition SEC1_P384_9_QCAVSy :: nat where 
  "SEC1_P384_9_QCAVSy = 0x73877dc94db9adddc0d9a4b24e8976c22d73c844370e1ee857f8d1b129a3bd5f63f40caf3bd0533e38a5f5777074ff9e"

definition SEC1_P384_9_QCAVS :: "int point" where
  "SEC1_P384_9_QCAVS = Point (int SEC1_P384_9_QCAVSx) (int SEC1_P384_9_QCAVSy)"

definition SEC1_P384_9_dIUT :: nat where 
  "SEC1_P384_9_dIUT = 0xd787a57fde22ec656a0a525cf3c738b30d73af61e743ea90893ecb2d7b622add2f94ee25c2171467afb093f3f84d0018"

definition SEC1_P384_9_QIUTx :: nat where 
  "SEC1_P384_9_QIUTx = 0x171546923b87b2cbbad664f01ce932bf09d6a6118168678446bfa9f0938608cb4667a98f4ec8ac1462285c2508f74862"

definition SEC1_P384_9_QIUTy :: nat where 
  "SEC1_P384_9_QIUTy = 0xfa41cb4db68ae71f1f8a3e8939dc52c2dec61a83c983beb2a02baf29ec49278088882ed0cf56c74b5c173b552ccf63cf"

definition SEC1_P384_9_QIUT :: "int point" where
  "SEC1_P384_9_QIUT = Point (int SEC1_P384_9_QIUTx) (int SEC1_P384_9_QIUTy)"

definition SEC1_P384_9_ZIUT :: nat where 
  "SEC1_P384_9_ZIUT = 0x1e97b60add7cb35c7403dd884c0a75795b7683fff8b49f9d8672a8206bfdcf0a106b8768f983258c74167422e44e4d14"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_9_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_9_dIUT"
  by eval

lemma SEC1_P384_9_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_9_dIUT = SEC1_P384_9_QIUT" 
  by eval

lemma SEC1_P384_9_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_9_QCAVS" 
  by eval

lemma SEC1_P384_9_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_9_QCAVS" 
  using SEC1_P384_9_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_9_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_9_dIUT SEC1_P384_9_QCAVS = Some SEC1_P384_9_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_10›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_10_QCAVSx :: nat where 
  "SEC1_P384_10_QCAVSx = 0xa721f6a2d4527411834b13d4d3a33c29beb83ab7682465c6cbaf6624aca6ea58c30eb0f29dd842886695400d7254f20f"

definition SEC1_P384_10_QCAVSy :: nat where 
  "SEC1_P384_10_QCAVSy = 0x14ba6e26355109ad35129366d5e3a640ae798505a7fa55a96a36b5dad33de00474f6670f522214dd7952140ab0a7eb68"

definition SEC1_P384_10_QCAVS :: "int point" where
  "SEC1_P384_10_QCAVS = Point (int SEC1_P384_10_QCAVSx) (int SEC1_P384_10_QCAVSy)"

definition SEC1_P384_10_dIUT :: nat where 
  "SEC1_P384_10_dIUT = 0x83d70f7b164d9f4c227c767046b20eb34dfc778f5387e32e834b1e6daec20edb8ca5bb4192093f543b68e6aeb7ce788b"

definition SEC1_P384_10_QIUTx :: nat where 
  "SEC1_P384_10_QIUTx = 0x57cd770f3bbcbe0c78c770eab0b169bc45e139f86378ffae1c2b16966727c2f2eb724572b8f3eb228d130db4ff862c63"

definition SEC1_P384_10_QIUTy :: nat where 
  "SEC1_P384_10_QIUTy = 0x7ec5c8813b685558d83e924f14bc719f6eb7ae0cbb2c474227c5bda88637a4f26c64817929af999592da6f787490332f"

definition SEC1_P384_10_QIUT :: "int point" where
  "SEC1_P384_10_QIUT = Point (int SEC1_P384_10_QIUTx) (int SEC1_P384_10_QIUTy)"

definition SEC1_P384_10_ZIUT :: nat where 
  "SEC1_P384_10_ZIUT = 0x1023478840e54775bfc69293a3cf97f5bc914726455c66538eb5623e218feef7df4befa23e09d77145ad577db32b41f9"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_10_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_10_dIUT"
  by eval

lemma SEC1_P384_10_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_10_dIUT = SEC1_P384_10_QIUT" 
  by eval

lemma SEC1_P384_10_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_10_QCAVS" 
  by eval

lemma SEC1_P384_10_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_10_QCAVS" 
  using SEC1_P384_10_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_10_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_10_dIUT SEC1_P384_10_QCAVS = Some SEC1_P384_10_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_11›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_11_QCAVSx :: nat where 
  "SEC1_P384_11_QCAVSx = 0xd882a8505c2d5cb9b8851fc676677bb0087681ad53faceba1738286b45827561e7da37b880276c656cfc38b32ade847e"

definition SEC1_P384_11_QCAVSy :: nat where 
  "SEC1_P384_11_QCAVSy = 0x34b314bdc134575654573cffaf40445da2e6aaf987f7e913cd4c3091523058984a25d8f21da8326192456c6a0fa5f60c"

definition SEC1_P384_11_QCAVS :: "int point" where
  "SEC1_P384_11_QCAVS = Point (int SEC1_P384_11_QCAVSx) (int SEC1_P384_11_QCAVSy)"

definition SEC1_P384_11_dIUT :: nat where 
  "SEC1_P384_11_dIUT = 0x8f558e05818b88ed383d5fca962e53413db1a0e4637eda194f761944cbea114ab9d5da175a7d57882550b0e432f395a9"

definition SEC1_P384_11_QIUTx :: nat where 
  "SEC1_P384_11_QIUTx = 0x9a2f57f4867ce753d72b0d95195df6f96c1fae934f602efd7b6a54582f556cfa539d89005ca2edac08ad9b72dd1f60ba"

definition SEC1_P384_11_QIUTy :: nat where 
  "SEC1_P384_11_QIUTy = 0xd9b94ee82da9cc601f346044998ba387aee56404dc6ecc8ab2b590443319d0b2b6176f9d0eac2d44678ed561607d09a9"

definition SEC1_P384_11_QIUT :: "int point" where
  "SEC1_P384_11_QIUT = Point (int SEC1_P384_11_QIUTx) (int SEC1_P384_11_QIUTy)"

definition SEC1_P384_11_ZIUT :: nat where 
  "SEC1_P384_11_ZIUT = 0x6ad6b9dc8a6cf0d3691c501cbb967867f6e4bbb764b60dbff8fcff3ed42dbba39d63cf325b4b4078858495ddee75f954"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_11_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_11_dIUT"
  by eval

lemma SEC1_P384_11_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_11_dIUT = SEC1_P384_11_QIUT" 
  by eval

lemma SEC1_P384_11_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_11_QCAVS" 
  by eval

lemma SEC1_P384_11_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_11_QCAVS" 
  using SEC1_P384_11_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_11_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_11_dIUT SEC1_P384_11_QCAVS = Some SEC1_P384_11_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_12›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_12_QCAVSx :: nat where 
  "SEC1_P384_12_QCAVSx = 0x815c9d773dbf5fb6a1b86799966247f4006a23c92e68c55e9eaa998b17d8832dd4d84d927d831d4f68dac67c6488219f"

definition SEC1_P384_12_QCAVSy :: nat where 
  "SEC1_P384_12_QCAVSy = 0xe79269948b2611484560fd490feec887cb55ef99a4b524880fa7499d6a07283aae2afa33feab97deca40bc606c4d8764"

definition SEC1_P384_12_QCAVS :: "int point" where
  "SEC1_P384_12_QCAVS = Point (int SEC1_P384_12_QCAVSx) (int SEC1_P384_12_QCAVSy)"

definition SEC1_P384_12_dIUT :: nat where 
  "SEC1_P384_12_dIUT = 0x0f5dee0affa7bbf239d5dff32987ebb7cf84fcceed643e1d3c62d0b3352aec23b6e5ac7fa4105c8cb26126ad2d1892cb"

definition SEC1_P384_12_QIUTx :: nat where 
  "SEC1_P384_12_QIUTx = 0x23346bdfbc9d7c7c736e02bdf607671ff6082fdd27334a8bc75f3b23681ebe614d0597dd614fae58677c835a9f0b273b"

definition SEC1_P384_12_QIUTy :: nat where 
  "SEC1_P384_12_QIUTy = 0x82ba36290d2f94db41479eb45ab4eaf67928a2315138d59eecc9b5285dfddd6714f77557216ea44cc6fc119d8243efaf"

definition SEC1_P384_12_QIUT :: "int point" where
  "SEC1_P384_12_QIUT = Point (int SEC1_P384_12_QIUTx) (int SEC1_P384_12_QIUTy)"

definition SEC1_P384_12_ZIUT :: nat where 
  "SEC1_P384_12_ZIUT = 0xcc9e063566d46b357b3fcae21827377331e5e290a36e60cd7c39102b828ae0b918dc5a02216b07fe6f1958d834e42437"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_12_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_12_dIUT"
  by eval

lemma SEC1_P384_12_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_12_dIUT = SEC1_P384_12_QIUT" 
  by eval

lemma SEC1_P384_12_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_12_QCAVS" 
  by eval

lemma SEC1_P384_12_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_12_QCAVS" 
  using SEC1_P384_12_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_12_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_12_dIUT SEC1_P384_12_QCAVS = Some SEC1_P384_12_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_13›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_13_QCAVSx :: nat where 
  "SEC1_P384_13_QCAVSx = 0x1c0eeda7a2be000c5bdcda0478aed4db733d2a9e341224379123ad847030f29e3b168fa18e89a3c0fba2a6ce1c28fc3b"

definition SEC1_P384_13_QCAVSy :: nat where 
  "SEC1_P384_13_QCAVSy = 0xec8c1c83c118c4dbea94271869f2d868eb65e8b44e21e6f14b0f4d9b38c068daefa27114255b9a41d084cc4a1ad85456"

definition SEC1_P384_13_QCAVS :: "int point" where
  "SEC1_P384_13_QCAVS = Point (int SEC1_P384_13_QCAVSx) (int SEC1_P384_13_QCAVSy)"

definition SEC1_P384_13_dIUT :: nat where 
  "SEC1_P384_13_dIUT = 0x037b633b5b8ba857c0fc85656868232e2febf59578718391b81da8541a00bfe53c30ae04151847f27499f8d7abad8cf4"

definition SEC1_P384_13_QIUTx :: nat where 
  "SEC1_P384_13_QIUTx = 0x8878ac8a947f7d5cb2b47aad24fbb8210d86126585399a2871f84aa9c5fde3074ae540c6bf82275ca822d0feb862bc74"

definition SEC1_P384_13_QIUTy :: nat where 
  "SEC1_P384_13_QIUTy = 0x632f5cd2f900c2711c32f8930728eb647d31edd8d650f9654e7d33e5ed1b475489d08daa30d8cbcba6bfc3b60d9b5a37"

definition SEC1_P384_13_QIUT :: "int point" where
  "SEC1_P384_13_QIUT = Point (int SEC1_P384_13_QIUTx) (int SEC1_P384_13_QIUTy)"

definition SEC1_P384_13_ZIUT :: nat where 
  "SEC1_P384_13_ZIUT = 0xdeff7f03bd09865baf945e73edff6d5122c03fb561db87dec8662e09bed4340b28a9efe118337bb7d3d4f7f568635ff9"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_13_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_13_dIUT"
  by eval

lemma SEC1_P384_13_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_13_dIUT = SEC1_P384_13_QIUT" 
  by eval

lemma SEC1_P384_13_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_13_QCAVS" 
  by eval

lemma SEC1_P384_13_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_13_QCAVS" 
  using SEC1_P384_13_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_13_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_13_dIUT SEC1_P384_13_QCAVS = Some SEC1_P384_13_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_14›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_14_QCAVSx :: nat where 
  "SEC1_P384_14_QCAVSx = 0xc95c185e256bf997f30b311548ae7f768a38dee43eeeef43083f3077be70e2bf39ac1d4daf360c514c8c6be623443d1a"

definition SEC1_P384_14_QCAVSy :: nat where 
  "SEC1_P384_14_QCAVSy = 0x3e63a663eaf75d8a765ab2b9a35513d7933fa5e26420a5244550ec6c3b6f033b96db2aca3d6ac6aab052ce929595aea5"

definition SEC1_P384_14_QCAVS :: "int point" where
  "SEC1_P384_14_QCAVS = Point (int SEC1_P384_14_QCAVSx) (int SEC1_P384_14_QCAVSy)"

definition SEC1_P384_14_dIUT :: nat where 
  "SEC1_P384_14_dIUT = 0xe3d07106bedcc096e7d91630ffd3094df2c7859db8d7edbb2e37b4ac47f429a637d06a67d2fba33838764ef203464991"

definition SEC1_P384_14_QIUTx :: nat where 
  "SEC1_P384_14_QIUTx = 0xe74a1a2b85f1cbf8dbbdf050cf1aff8acb02fda2fb6591f9d3cfe4e79d0ae938a9c1483e7b75f8db24505d65065cdb18"

definition SEC1_P384_14_QIUTy :: nat where 
  "SEC1_P384_14_QIUTy = 0x1773ee591822f7abaa856a1a60bc0a5203548dbd1cb5025466eff8481bd07614eaa04a16c3db76905913e972a5b6b59d"

definition SEC1_P384_14_QIUT :: "int point" where
  "SEC1_P384_14_QIUT = Point (int SEC1_P384_14_QIUTx) (int SEC1_P384_14_QIUTy)"

definition SEC1_P384_14_ZIUT :: nat where 
  "SEC1_P384_14_ZIUT = 0xc8b1038f735ad3bb3e4637c3e47eab487637911a6b7950a4e461948329d3923b969e5db663675623611a457fcda35a71"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_14_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_14_dIUT"
  by eval

lemma SEC1_P384_14_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_14_dIUT = SEC1_P384_14_QIUT" 
  by eval

lemma SEC1_P384_14_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_14_QCAVS" 
  by eval

lemma SEC1_P384_14_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_14_QCAVS" 
  using SEC1_P384_14_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_14_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_14_dIUT SEC1_P384_14_QCAVS = Some SEC1_P384_14_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_15›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_15_QCAVSx :: nat where 
  "SEC1_P384_15_QCAVSx = 0x3497238a7e6ad166df2dac039aa4dac8d17aa925e7c7631eb3b56e3aaa1c545fcd54d2e5985807910fb202b1fc191d2a"

definition SEC1_P384_15_QCAVSy :: nat where 
  "SEC1_P384_15_QCAVSy = 0xa49e5c487dcc7aa40a8f234c979446040d9174e3ad357d404d7765183195aed3f913641b90c81a306ebf0d8913861316"

definition SEC1_P384_15_QCAVS :: "int point" where
  "SEC1_P384_15_QCAVS = Point (int SEC1_P384_15_QCAVSx) (int SEC1_P384_15_QCAVSy)"

definition SEC1_P384_15_dIUT :: nat where 
  "SEC1_P384_15_dIUT = 0xf3f9b0c65a49a506632c8a45b10f66b5316f9eeb06fae218f2da62333f99905117b141c760e8974efc4af10570635791"

definition SEC1_P384_15_QIUTx :: nat where 
  "SEC1_P384_15_QIUTx = 0xa4ad77aa7d86e5361118a6b921710c820721210712f4c347985fdee58aa4effa1e28be80a17b120b139f96300f89b49b"

definition SEC1_P384_15_QIUTy :: nat where 
  "SEC1_P384_15_QIUTy = 0x1ddf22e07e03f1560d8f45a480094560dba9fae7f9531130c1b57ebb95982496524f31d3797793396fa823f22bdb4328"

definition SEC1_P384_15_QIUT :: "int point" where
  "SEC1_P384_15_QIUT = Point (int SEC1_P384_15_QIUTx) (int SEC1_P384_15_QIUTy)"

definition SEC1_P384_15_ZIUT :: nat where 
  "SEC1_P384_15_ZIUT = 0xd337eaa32b9f716b8747b005b97a553c59dab0c51df41a2d49039cdae705aa75c7b9e7bc0b6a0e8c578c902bc4fff23e"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_15_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_15_dIUT"
  by eval

lemma SEC1_P384_15_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_15_dIUT = SEC1_P384_15_QIUT" 
  by eval

lemma SEC1_P384_15_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_15_QCAVS" 
  by eval

lemma SEC1_P384_15_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_15_QCAVS" 
  using SEC1_P384_15_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_15_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_15_dIUT SEC1_P384_15_QCAVS = Some SEC1_P384_15_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_16›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_16_QCAVSx :: nat where 
  "SEC1_P384_16_QCAVSx = 0x90a34737d45b1aa65f74e0bd0659bc118f8e4b774b761944ffa6573c6df4f41dec0d11b697abd934d390871d4b453240"

definition SEC1_P384_16_QCAVSy :: nat where 
  "SEC1_P384_16_QCAVSy = 0x9b590719bb3307c149a7817be355d684893a307764b512eeffe07cb699edb5a6ffbf8d6032e6c79d5e93e94212c2aa4e"

definition SEC1_P384_16_QCAVS :: "int point" where
  "SEC1_P384_16_QCAVS = Point (int SEC1_P384_16_QCAVSx) (int SEC1_P384_16_QCAVSy)"

definition SEC1_P384_16_dIUT :: nat where 
  "SEC1_P384_16_dIUT = 0x59fce7fad7de28bac0230690c95710c720e528f9a4e54d3a6a8cd5fc5c5f21637031ce1c5b4e3d39647d8dcb9b794664"

definition SEC1_P384_16_QIUTx :: nat where 
  "SEC1_P384_16_QIUTx = 0x9c43bf971edf09402876ee742095381f78b1bd3aa39b5132af75dbfe7e98bd78bde10fe2e903c2b6379e1deee175a1b0"

definition SEC1_P384_16_QIUTy :: nat where 
  "SEC1_P384_16_QIUTy = 0xa6c58ecea5a477bb01bd543b339f1cc49f1371a2cda4d46eb4e53e250597942351a99665a122ffea9bde0636c375daf2"

definition SEC1_P384_16_QIUT :: "int point" where
  "SEC1_P384_16_QIUT = Point (int SEC1_P384_16_QIUTx) (int SEC1_P384_16_QIUTy)"

definition SEC1_P384_16_ZIUT :: nat where 
  "SEC1_P384_16_ZIUT = 0x32d292b695a4488e42a7b7922e1ae537d76a3d21a0b2e36875f60e9f6d3e8779c2afb3a413b9dd79ae18e70b47d337c1"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_16_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_16_dIUT"
  by eval

lemma SEC1_P384_16_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_16_dIUT = SEC1_P384_16_QIUT" 
  by eval

lemma SEC1_P384_16_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_16_QCAVS" 
  by eval

lemma SEC1_P384_16_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_16_QCAVS" 
  using SEC1_P384_16_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_16_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_16_dIUT SEC1_P384_16_QCAVS = Some SEC1_P384_16_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_17›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_17_QCAVSx :: nat where 
  "SEC1_P384_17_QCAVSx = 0xdda546acfc8f903d11e2e3920669636d44b2068aeb66ff07aa266f0030e1535b0ed0203cb8a460ac990f1394faf22f1d"

definition SEC1_P384_17_QCAVSy :: nat where 
  "SEC1_P384_17_QCAVSy = 0x15bbb2597913035faadf413476f4c70f7279769a40c986f470c427b4ee4962abdf8173bbad81874772925fd32f0b159f"

definition SEC1_P384_17_QCAVS :: "int point" where
  "SEC1_P384_17_QCAVS = Point (int SEC1_P384_17_QCAVSx) (int SEC1_P384_17_QCAVSy)"

definition SEC1_P384_17_dIUT :: nat where 
  "SEC1_P384_17_dIUT = 0x3e49fbf950a424c5d80228dc4bc35e9f6c6c0c1d04440998da0a609a877575dbe437d6a5cedaa2ddd2a1a17fd112aded"

definition SEC1_P384_17_QIUTx :: nat where 
  "SEC1_P384_17_QIUTx = 0x5a949594228b1a3d6f599eb3db0d06070fbc551c657b58234ba164ce3fe415fa5f3eb823c08dc29b8c341219c77b6b3d"

definition SEC1_P384_17_QIUTy :: nat where 
  "SEC1_P384_17_QIUTy = 0x2baad447c8c290cfed25edd9031c41d0b76921457327f42db31122b81f337bbf0b1039ec830ce9061a3761953c75e4a8"

definition SEC1_P384_17_QIUT :: "int point" where
  "SEC1_P384_17_QIUT = Point (int SEC1_P384_17_QIUTx) (int SEC1_P384_17_QIUTy)"

definition SEC1_P384_17_ZIUT :: nat where 
  "SEC1_P384_17_ZIUT = 0x1220e7e6cad7b25df98e5bbdcc6c0b65ca6c2a50c5ff6c41dca71e475646fd489615979ca92fb4389aeadefde79a24f1"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_17_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_17_dIUT"
  by eval

lemma SEC1_P384_17_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_17_dIUT = SEC1_P384_17_QIUT" 
  by eval

lemma SEC1_P384_17_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_17_QCAVS" 
  by eval

lemma SEC1_P384_17_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_17_QCAVS" 
  using SEC1_P384_17_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_17_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_17_dIUT SEC1_P384_17_QCAVS = Some SEC1_P384_17_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_18›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_18_QCAVSx :: nat where 
  "SEC1_P384_18_QCAVSx = 0x788be2336c52f4454d63ee944b1e49bfb619a08371048e6da92e584eae70bde1f171c4df378bd1f3c0ab03048a237802"

definition SEC1_P384_18_QCAVSy :: nat where 
  "SEC1_P384_18_QCAVSy = 0x4673ebd8db604eaf41711748bab2968a23ca4476ce144e728247f08af752929157b5830f1e26067466bdfa8b65145a33"

definition SEC1_P384_18_QCAVS :: "int point" where
  "SEC1_P384_18_QCAVS = Point (int SEC1_P384_18_QCAVSx) (int SEC1_P384_18_QCAVSy)"

definition SEC1_P384_18_dIUT :: nat where 
  "SEC1_P384_18_dIUT = 0x50ccc1f7076e92f4638e85f2db98e0b483e6e2204c92bdd440a6deea04e37a07c6e72791c190ad4e4e86e01efba84269"

definition SEC1_P384_18_QIUTx :: nat where 
  "SEC1_P384_18_QIUTx = 0x756c07df0ce32c839dac9fb4733c9c28b70113a676a7057c38d223f22a3a9095a8d564653af528e04c7e1824be4a6512"

definition SEC1_P384_18_QIUTy :: nat where 
  "SEC1_P384_18_QIUTy = 0x17c2ce6962cbd2a2e066297b39d57dd9bb4680f0191d390f70b4e461419b2972ce68ad46127fdda6c39195774ea86df3"

definition SEC1_P384_18_QIUT :: "int point" where
  "SEC1_P384_18_QIUT = Point (int SEC1_P384_18_QIUTx) (int SEC1_P384_18_QIUTy)"

definition SEC1_P384_18_ZIUT :: nat where 
  "SEC1_P384_18_ZIUT = 0x793bb9cd22a93cf468faf804a38d12b78cb12189ec679ddd2e9aa21fa9a5a0b049ab16a23574fe04c1c3c02343b91beb"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_18_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_18_dIUT"
  by eval

lemma SEC1_P384_18_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_18_dIUT = SEC1_P384_18_QIUT" 
  by eval

lemma SEC1_P384_18_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_18_QCAVS" 
  by eval

lemma SEC1_P384_18_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_18_QCAVS" 
  using SEC1_P384_18_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_18_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_18_dIUT SEC1_P384_18_QCAVS = Some SEC1_P384_18_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_19›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_19_QCAVSx :: nat where 
  "SEC1_P384_19_QCAVSx = 0xd09bb822eb99e38060954747c82bb3278cf96bbf36fece3400f4c873838a40c135eb3babb9293bd1001bf3ecdee7bf26"

definition SEC1_P384_19_QCAVSy :: nat where 
  "SEC1_P384_19_QCAVSy = 0xd416db6e1b87bbb7427788a3b6c7a7ab2c165b1e366f9608df512037584f213a648d47f16ac326e19aae972f63fd76c9"

definition SEC1_P384_19_QCAVS :: "int point" where
  "SEC1_P384_19_QCAVS = Point (int SEC1_P384_19_QCAVSx) (int SEC1_P384_19_QCAVSy)"

definition SEC1_P384_19_dIUT :: nat where 
  "SEC1_P384_19_dIUT = 0x06f132b71f74d87bf99857e1e4350a594e5fe35533b888552ceccbc0d8923c902e36141d7691e28631b8bc9bafe5e064"

definition SEC1_P384_19_QIUTx :: nat where 
  "SEC1_P384_19_QIUTx = 0x2a3cc6b8ff5cde926e7e3a189a1bd029c9b586351af8838f4f201cb8f4b70ef3b0da06d352c80fc26baf8f42b784459e"

definition SEC1_P384_19_QIUTy :: nat where 
  "SEC1_P384_19_QIUTy = 0xbf9985960176da6d23c7452a2954ffcbbcb24249b43019a2a023e0b3dabd461f19ad3e775c364f3f11ad49f3099400d3"

definition SEC1_P384_19_QIUT :: "int point" where
  "SEC1_P384_19_QIUT = Point (int SEC1_P384_19_QIUTx) (int SEC1_P384_19_QIUTy)"

definition SEC1_P384_19_ZIUT :: nat where 
  "SEC1_P384_19_ZIUT = 0x012d191cf7404a523678c6fc075de8285b243720a903047708bb33e501e0dbee5bcc40d7c3ef6c6da39ea24d830da1e8"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_19_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_19_dIUT"
  by eval

lemma SEC1_P384_19_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_19_dIUT = SEC1_P384_19_QIUT" 
  by eval

lemma SEC1_P384_19_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_19_QCAVS" 
  by eval

lemma SEC1_P384_19_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_19_QCAVS" 
  using SEC1_P384_19_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_19_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_19_dIUT SEC1_P384_19_QCAVS = Some SEC1_P384_19_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_20›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_20_QCAVSx :: nat where 
  "SEC1_P384_20_QCAVSx = 0x13741262ede5861dad71063dfd204b91ea1d3b7c631df68eb949969527d79a1dc59295ef7d2bca6743e8cd77b04d1b58"

definition SEC1_P384_20_QCAVSy :: nat where 
  "SEC1_P384_20_QCAVSy = 0x0baaeadc7e19d74a8a04451a135f1be1b02fe299f9dc00bfdf201e83d995c6950bcc1cb89d6f7b30bf54656b9a4da586"

definition SEC1_P384_20_QCAVS :: "int point" where
  "SEC1_P384_20_QCAVS = Point (int SEC1_P384_20_QCAVSx) (int SEC1_P384_20_QCAVSy)"

definition SEC1_P384_20_dIUT :: nat where 
  "SEC1_P384_20_dIUT = 0x12048ebb4331ec19a1e23f1a2c773b664ccfe90a28bfb846fc12f81dff44b7443c77647164bf1e9e67fd2c07a6766241"

definition SEC1_P384_20_QIUTx :: nat where 
  "SEC1_P384_20_QIUTx = 0xbc18836bc7a9fdf54b5352f37d7528ab8fa8ec544a8c6180511cbfdd49cce377c39e34c031b5240dc9980503ed2f262c"

definition SEC1_P384_20_QIUTy :: nat where 
  "SEC1_P384_20_QIUTy = 0x8086cbe338191080f0b7a16c7afc4c7b0326f9ac66f58552ef4bb9d24de3429ed5d3277ed58fcf48f2b5f61326bec6c6"

definition SEC1_P384_20_QIUT :: "int point" where
  "SEC1_P384_20_QIUT = Point (int SEC1_P384_20_QIUTx) (int SEC1_P384_20_QIUTy)"

definition SEC1_P384_20_ZIUT :: nat where 
  "SEC1_P384_20_ZIUT = 0xad0fd3ddffe8884b9263f3c15fe1f07f2a5a22ffdc7e967085eea45f0cd959f20f18f522763e28bcc925e496a52dda98"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_20_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_20_dIUT"
  by eval

lemma SEC1_P384_20_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_20_dIUT = SEC1_P384_20_QIUT" 
  by eval

lemma SEC1_P384_20_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_20_QCAVS" 
  by eval

lemma SEC1_P384_20_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_20_QCAVS" 
  using SEC1_P384_20_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_20_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_20_dIUT SEC1_P384_20_QCAVS = Some SEC1_P384_20_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_21›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_21_QCAVSx :: nat where 
  "SEC1_P384_21_QCAVSx = 0x9e22cbc18657f516a864b37b783348b66f1aa9626cd631f4fa1bd32ad88cf11db52057c660860d39d11fbf024fabd444"

definition SEC1_P384_21_QCAVSy :: nat where 
  "SEC1_P384_21_QCAVSy = 0x6b0d53c79681c28116df71e9cee74fd56c8b7f04b39f1198cc72284e98be9562e35926fb4f48a9fbecafe729309e8b6f"

definition SEC1_P384_21_QCAVS :: "int point" where
  "SEC1_P384_21_QCAVS = Point (int SEC1_P384_21_QCAVSx) (int SEC1_P384_21_QCAVSy)"

definition SEC1_P384_21_dIUT :: nat where 
  "SEC1_P384_21_dIUT = 0x34d61a699ca576169fcdc0cc7e44e4e1221db0fe63d16850c8104029f7d48449714b9884328cae189978754ab460b486"

definition SEC1_P384_21_QIUTx :: nat where 
  "SEC1_P384_21_QIUTx = 0x867f81104ccd6b163a7902b670ef406042cb0cce7dcdc63d1dfc91b2c40e3cdf7595834bf9eceb79849f1636fc8462fc"

definition SEC1_P384_21_QIUTy :: nat where 
  "SEC1_P384_21_QIUTy = 0x9d4bde8e875ec49697d258d1d59465f8431c6f5531e1c59e9f9ebe3cf164a8d9ce10a12f1979283a959bad244dd83863"

definition SEC1_P384_21_QIUT :: "int point" where
  "SEC1_P384_21_QIUT = Point (int SEC1_P384_21_QIUTx) (int SEC1_P384_21_QIUTy)"

definition SEC1_P384_21_ZIUT :: nat where 
  "SEC1_P384_21_ZIUT = 0xdc4ca392dc15e20185f2c6a8ea5ec31dfc96f56153a47394b3072b13d0015f5d4ae13beb3bed54d65848f9b8383e6c95"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_21_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_21_dIUT"
  by eval

lemma SEC1_P384_21_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_21_dIUT = SEC1_P384_21_QIUT" 
  by eval

lemma SEC1_P384_21_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_21_QCAVS" 
  by eval

lemma SEC1_P384_21_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_21_QCAVS" 
  using SEC1_P384_21_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_21_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_21_dIUT SEC1_P384_21_QCAVS = Some SEC1_P384_21_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_22›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_22_QCAVSx :: nat where 
  "SEC1_P384_22_QCAVSx = 0x2db5da5f940eaa884f4db5ec2139b0469f38e4e6fbbcc52df15c0f7cf7fcb1808c749764b6be85d2fdc5b16f58ad5dc0"

definition SEC1_P384_22_QCAVSy :: nat where 
  "SEC1_P384_22_QCAVSy = 0x22e8b02dcf33e1b5a083849545f84ad5e43f77cb71546dbbac0d11bdb2ee202e9d3872e8d028c08990746c5e1dde9989"

definition SEC1_P384_22_QCAVS :: "int point" where
  "SEC1_P384_22_QCAVS = Point (int SEC1_P384_22_QCAVSx) (int SEC1_P384_22_QCAVSy)"

definition SEC1_P384_22_dIUT :: nat where 
  "SEC1_P384_22_dIUT = 0xdc60fa8736d702135ff16aab992bb88eac397f5972456c72ec447374d0d8ce61153831bfc86ad5a6eb5b60bfb96a862c"

definition SEC1_P384_22_QIUTx :: nat where 
  "SEC1_P384_22_QIUTx = 0xb69beede85d0f829fec1b893ccb9c3e052ff692e13b974537bc5b0f9feaf7b22e84f03231629b24866bdb4b8cf908914"

definition SEC1_P384_22_QIUTy :: nat where 
  "SEC1_P384_22_QIUTy = 0x66f85e2bfcaba2843285b0e14ebc07ef7dafff8b424416fee647b59897b619f20eed95a632e6a4206bf7da429c04c560"

definition SEC1_P384_22_QIUT :: "int point" where
  "SEC1_P384_22_QIUT = Point (int SEC1_P384_22_QIUTx) (int SEC1_P384_22_QIUTy)"

definition SEC1_P384_22_ZIUT :: nat where 
  "SEC1_P384_22_ZIUT = 0xd765b208112d2b9ed5ad10c4046e2e3b0dbf57c469329519e239ac28b25c7d852bf757d5de0ee271cadd021d86cfd347"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_22_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_22_dIUT"
  by eval

lemma SEC1_P384_22_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_22_dIUT = SEC1_P384_22_QIUT" 
  by eval

lemma SEC1_P384_22_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_22_QCAVS" 
  by eval

lemma SEC1_P384_22_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_22_QCAVS" 
  using SEC1_P384_22_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_22_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_22_dIUT SEC1_P384_22_QCAVS = Some SEC1_P384_22_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_23›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_23_QCAVSx :: nat where 
  "SEC1_P384_23_QCAVSx = 0x329647baa354224eb4414829c5368c82d7893b39804e08cbb2180f459befc4b347a389a70c91a23bd9d30c83be5295d3"

definition SEC1_P384_23_QCAVSy :: nat where 
  "SEC1_P384_23_QCAVSy = 0xcc8f61923fad2aa8e505d6cfa126b9fabd5af9dce290b75660ef06d1caa73681d06089c33bc4246b3aa30dbcd2435b12"

definition SEC1_P384_23_QCAVS :: "int point" where
  "SEC1_P384_23_QCAVS = Point (int SEC1_P384_23_QCAVSx) (int SEC1_P384_23_QCAVSy)"

definition SEC1_P384_23_dIUT :: nat where 
  "SEC1_P384_23_dIUT = 0x6fa6a1c704730987aa634b0516a826aba8c6d6411d3a4c89772d7a62610256a2e2f289f5c3440b0ec1e70fa339e251ce"

definition SEC1_P384_23_QIUTx :: nat where 
  "SEC1_P384_23_QIUTx = 0x53de1fc1328e8de14aecab29ad8a40d6b13768f86f7d298433d20fec791f86f8bc73f358098b256a298bb488de257bf4"

definition SEC1_P384_23_QIUTy :: nat where 
  "SEC1_P384_23_QIUTy = 0xac28944fd27f17b82946c04c66c41f0053d3692f275da55cd8739a95bd8cd3af2f96e4de959ea8344d8945375905858b"

definition SEC1_P384_23_QIUT :: "int point" where
  "SEC1_P384_23_QIUT = Point (int SEC1_P384_23_QIUTx) (int SEC1_P384_23_QIUTy)"

definition SEC1_P384_23_ZIUT :: nat where 
  "SEC1_P384_23_ZIUT = 0xd3778850aeb58804fbe9dfe6f38b9fa8e20c2ca4e0dec335aafceca0333e3f2490b53c0c1a14a831ba37c4b9d74be0f2"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_23_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_23_dIUT"
  by eval

lemma SEC1_P384_23_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_23_dIUT = SEC1_P384_23_QIUT" 
  by eval

lemma SEC1_P384_23_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_23_QCAVS" 
  by eval

lemma SEC1_P384_23_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_23_QCAVS" 
  using SEC1_P384_23_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_23_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_23_dIUT SEC1_P384_23_QCAVS = Some SEC1_P384_23_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P384_24›

subsubsection‹Given Test Vector Values›

definition SEC1_P384_24_QCAVSx :: nat where 
  "SEC1_P384_24_QCAVSx = 0x29d8a36d22200a75b7aea1bb47cdfcb1b7fd66de967041434728ab5d533a060df732130600fe6f75852a871fb2938e39"

definition SEC1_P384_24_QCAVSy :: nat where 
  "SEC1_P384_24_QCAVSy = 0xe19b53db528395de897a45108967715eb8cb55c3fcbf23379372c0873a058d57544b102ecce722b2ccabb1a603774fd5"

definition SEC1_P384_24_QCAVS :: "int point" where
  "SEC1_P384_24_QCAVS = Point (int SEC1_P384_24_QCAVSx) (int SEC1_P384_24_QCAVSy)"

definition SEC1_P384_24_dIUT :: nat where 
  "SEC1_P384_24_dIUT = 0x74ad8386c1cb2ca0fcdeb31e0869bb3f48c036afe2ef110ca302bc8b910f621c9fcc54cec32bb89ec7caa84c7b8e54a8"

definition SEC1_P384_24_QIUTx :: nat where 
  "SEC1_P384_24_QIUTx = 0x27a3e83cfb9d5122e73129d801615857da7cc089cccc9c54ab3032a19e0a0a9f677346e37f08a0b3ed8da6e5dd691063"

definition SEC1_P384_24_QIUTy :: nat where 
  "SEC1_P384_24_QIUTy = 0x8d60e44aa5e0fd30c918456796af37f0e41957901645e5c596c6d989f5859b03a0bd7d1f4e77936fff3c74d204e5388e"

definition SEC1_P384_24_QIUT :: "int point" where
  "SEC1_P384_24_QIUT = Point (int SEC1_P384_24_QIUTx) (int SEC1_P384_24_QIUTy)"

definition SEC1_P384_24_ZIUT :: nat where 
  "SEC1_P384_24_ZIUT = 0x81e1e71575bb4505498de097350186430a6242fa6c57b85a5f984a23371123d2d1424eefbf804258392bc723e4ef1e35"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P384_24_dIUT_valid: 
  "SEC1_P384_ECprivateKeyValid SEC1_P384_24_dIUT"
  by eval

lemma SEC1_P384_24_IUT_validPair: 
  "SEC1_P384_ECkeyGen SEC1_P384_24_dIUT = SEC1_P384_24_QIUT" 
  by eval

lemma SEC1_P384_24_QCAVS_partialValid: 
  "SEC1_P384_ECpublicKeyPartialValid SEC1_P384_24_QCAVS" 
  by eval

lemma SEC1_P384_24_QCAVS_validPublicKey: 
  "SEC1_P384_ECpublicKeyValid SEC1_P384_24_QCAVS" 
  using SEC1_P384_24_QCAVS_partialValid SEC1_P384.partValidImpliesValidIFheq1 P384_h_def by blast

lemma SEC1_P384_24_ECDHprim_Check: 
  "SEC1_P384_ECDHprim SEC1_P384_24_dIUT SEC1_P384_24_QCAVS = Some SEC1_P384_24_ZIUT"
  by eval
  
section‹NIST Test Vectors for ECDH: Curve P-521›
  
subsection‹Test Vector: SEC1_P521_0›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_0_QCAVSx :: nat where 
  "SEC1_P521_0_QCAVSx = 0x000000685a48e86c79f0f0875f7bc18d25eb5fc8c0b07e5da4f4370f3a9490340854334b1e1b87fa395464c60626124a4e70d0f785601d37c09870ebf176666877a2046d"

definition SEC1_P521_0_QCAVSy :: nat where 
  "SEC1_P521_0_QCAVSy = 0x000001ba52c56fc8776d9e8f5db4f0cc27636d0b741bbe05400697942e80b739884a83bde99e0f6716939e632bc8986fa18dccd443a348b6c3e522497955a4f3c302f676"

definition SEC1_P521_0_QCAVS :: "int point" where
  "SEC1_P521_0_QCAVS = Point (int SEC1_P521_0_QCAVSx) (int SEC1_P521_0_QCAVSy)"

definition SEC1_P521_0_dIUT :: nat where 
  "SEC1_P521_0_dIUT = 0x0000017eecc07ab4b329068fba65e56a1f8890aa935e57134ae0ffcce802735151f4eac6564f6ee9974c5e6887a1fefee5743ae2241bfeb95d5ce31ddcb6f9edb4d6fc47"

definition SEC1_P521_0_QIUTx :: nat where 
  "SEC1_P521_0_QIUTx = 0x000000602f9d0cf9e526b29e22381c203c48a886c2b0673033366314f1ffbcba240ba42f4ef38a76174635f91e6b4ed34275eb01c8467d05ca80315bf1a7bbd945f550a5"

definition SEC1_P521_0_QIUTy :: nat where 
  "SEC1_P521_0_QIUTy = 0x000001b7c85f26f5d4b2d7355cf6b02117659943762b6d1db5ab4f1dbc44ce7b2946eb6c7de342962893fd387d1b73d7a8672d1f236961170b7eb3579953ee5cdc88cd2d"

definition SEC1_P521_0_QIUT :: "int point" where
  "SEC1_P521_0_QIUT = Point (int SEC1_P521_0_QIUTx) (int SEC1_P521_0_QIUTy)"

definition SEC1_P521_0_ZIUT :: nat where 
  "SEC1_P521_0_ZIUT = 0x005fc70477c3e63bc3954bd0df3ea0d1f41ee21746ed95fc5e1fdf90930d5e136672d72cc770742d1711c3c3a4c334a0ad9759436a4d3c5bf6e74b9578fac148c831"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_0_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_0_dIUT"
  by eval

lemma SEC1_P521_0_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_0_dIUT = SEC1_P521_0_QIUT" 
  by eval

lemma SEC1_P521_0_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_0_QCAVS" 
  by eval

lemma SEC1_P521_0_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_0_QCAVS" 
  using SEC1_P521_0_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_0_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_0_dIUT SEC1_P521_0_QCAVS = Some SEC1_P521_0_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_1›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_1_QCAVSx :: nat where 
  "SEC1_P521_1_QCAVSx = 0x000001df277c152108349bc34d539ee0cf06b24f5d3500677b4445453ccc21409453aafb8a72a0be9ebe54d12270aa51b3ab7f316aa5e74a951c5e53f74cd95fc29aee7a"

definition SEC1_P521_1_QCAVSy :: nat where 
  "SEC1_P521_1_QCAVSy = 0x0000013d52f33a9f3c14384d1587fa8abe7aed74bc33749ad9c570b471776422c7d4505d9b0a96b3bfac041e4c6a6990ae7f700e5b4a6640229112deafa0cd8bb0d089b0"

definition SEC1_P521_1_QCAVS :: "int point" where
  "SEC1_P521_1_QCAVS = Point (int SEC1_P521_1_QCAVSx) (int SEC1_P521_1_QCAVSy)"

definition SEC1_P521_1_dIUT :: nat where 
  "SEC1_P521_1_dIUT = 0x000000816f19c1fb10ef94d4a1d81c156ec3d1de08b66761f03f06ee4bb9dcebbbfe1eaa1ed49a6a990838d8ed318c14d74cc872f95d05d07ad50f621ceb620cd905cfb8"

definition SEC1_P521_1_QIUTx :: nat where 
  "SEC1_P521_1_QIUTx = 0x000000d45615ed5d37fde699610a62cd43ba76bedd8f85ed31005fe00d6450fbbd101291abd96d4945a8b57bc73b3fe9f4671105309ec9b6879d0551d930dac8ba45d255"

definition SEC1_P521_1_QIUTy :: nat where 
  "SEC1_P521_1_QIUTy = 0x000001425332844e592b440c0027972ad1526431c06732df19cd46a242172d4dd67c2c8c99dfc22e49949a56cf90c6473635ce82f25b33682fb19bc33bd910ed8ce3a7fa"

definition SEC1_P521_1_QIUT :: "int point" where
  "SEC1_P521_1_QIUT = Point (int SEC1_P521_1_QIUTx) (int SEC1_P521_1_QIUTy)"

definition SEC1_P521_1_ZIUT :: nat where 
  "SEC1_P521_1_ZIUT = 0x000b3920ac830ade812c8f96805da2236e002acbbf13596a9ab254d44d0e91b6255ebf1229f366fb5a05c5884ef46032c26d42189273ca4efa4c3db6bd12a6853759"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_1_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_1_dIUT"
  by eval

lemma SEC1_P521_1_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_1_dIUT = SEC1_P521_1_QIUT" 
  by eval

lemma SEC1_P521_1_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_1_QCAVS" 
  by eval

lemma SEC1_P521_1_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_1_QCAVS" 
  using SEC1_P521_1_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_1_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_1_dIUT SEC1_P521_1_QCAVS = Some SEC1_P521_1_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_2›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_2_QCAVSx :: nat where 
  "SEC1_P521_2_QCAVSx = 0x00000092db3142564d27a5f0006f819908fba1b85038a5bc2509906a497daac67fd7aee0fc2daba4e4334eeaef0e0019204b471cd88024f82115d8149cc0cf4f7ce1a4d5"

definition SEC1_P521_2_QCAVSy :: nat where 
  "SEC1_P521_2_QCAVSy = 0x0000016bad0623f517b158d9881841d2571efbad63f85cbe2e581960c5d670601a6760272675a548996217e4ab2b8ebce31d71fca63fcc3c08e91c1d8edd91cf6fe845f8"

definition SEC1_P521_2_QCAVS :: "int point" where
  "SEC1_P521_2_QCAVS = Point (int SEC1_P521_2_QCAVSx) (int SEC1_P521_2_QCAVSy)"

definition SEC1_P521_2_dIUT :: nat where 
  "SEC1_P521_2_dIUT = 0x0000012f2e0c6d9e9d117ceb9723bced02eb3d4eebf5feeaf8ee0113ccd8057b13ddd416e0b74280c2d0ba8ed291c443bc1b141caf8afb3a71f97f57c225c03e1e4d42b0"

definition SEC1_P521_2_QIUTx :: nat where 
  "SEC1_P521_2_QIUTx = 0x000000717fcb3d4a40d103871ede044dc803db508aaa4ae74b70b9fb8d8dfd84bfecfad17871879698c292d2fd5e17b4f9343636c531a4fac68a35a93665546b9a878679"

definition SEC1_P521_2_QIUTy :: nat where 
  "SEC1_P521_2_QIUTy = 0x000000f3d96a8637036993ab5d244500fff9d2772112826f6436603d3eb234a44d5c4e5c577234679c4f9df725ee5b9118f23d8a58d0cc01096daf70e8dfec0128bdc2e8"

definition SEC1_P521_2_QIUT :: "int point" where
  "SEC1_P521_2_QIUT = Point (int SEC1_P521_2_QIUTx) (int SEC1_P521_2_QIUTy)"

definition SEC1_P521_2_ZIUT :: nat where 
  "SEC1_P521_2_ZIUT = 0x006b380a6e95679277cfee4e8353bf96ef2a1ebdd060749f2f046fe571053740bbcc9a0b55790bc9ab56c3208aa05ddf746a10a3ad694daae00d980d944aabc6a08f"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_2_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_2_dIUT"
  by eval

lemma SEC1_P521_2_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_2_dIUT = SEC1_P521_2_QIUT" 
  by eval

lemma SEC1_P521_2_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_2_QCAVS" 
  by eval

lemma SEC1_P521_2_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_2_QCAVS" 
  using SEC1_P521_2_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_2_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_2_dIUT SEC1_P521_2_QCAVS = Some SEC1_P521_2_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_3›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_3_QCAVSx :: nat where 
  "SEC1_P521_3_QCAVSx = 0x000000fdd40d9e9d974027cb3bae682162eac1328ad61bc4353c45bf5afe76bf607d2894c8cce23695d920f2464fda4773d4693be4b3773584691bdb0329b7f4c86cc299"

definition SEC1_P521_3_QCAVSy :: nat where 
  "SEC1_P521_3_QCAVSy = 0x00000034ceac6a3fef1c3e1c494bfe8d872b183832219a7e14da414d4e3474573671ec19b033be831b915435905925b44947c592959945b4eb7c951c3b9c8cf52530ba23"

definition SEC1_P521_3_QCAVS :: "int point" where
  "SEC1_P521_3_QCAVS = Point (int SEC1_P521_3_QCAVSx) (int SEC1_P521_3_QCAVSy)"

definition SEC1_P521_3_dIUT :: nat where 
  "SEC1_P521_3_dIUT = 0x000000e548a79d8b05f923b9825d11b656f222e8cb98b0f89de1d317184dc5a698f7c71161ee7dc11cd31f4f4f8ae3a981e1a3e78bdebb97d7c204b9261b4ef92e0918e0"

definition SEC1_P521_3_QIUTx :: nat where 
  "SEC1_P521_3_QIUTx = 0x0000000ce800217ed243dd10a79ad73df578aa8a3f9194af528cd1094bbfee27a3b5481ad5862c8876c0c3f91294c0ab3aa806d9020cbaa2ed72b7fecdc5a09a6dad6f32"

definition SEC1_P521_3_QIUTy :: nat where 
  "SEC1_P521_3_QIUTy = 0x000001543c9ab45b12469232918e21d5a351f9a4b9cbf9efb2afcc402fa9b31650bec2d641a05c440d35331c0893d11fb13151335988b303341301a73dc5f61d574e67d9"

definition SEC1_P521_3_QIUT :: "int point" where
  "SEC1_P521_3_QIUT = Point (int SEC1_P521_3_QIUTx) (int SEC1_P521_3_QIUTy)"

definition SEC1_P521_3_ZIUT :: nat where 
  "SEC1_P521_3_ZIUT = 0x00fbbcd0b8d05331fef6086f22a6cce4d35724ab7a2f49dd8458d0bfd57a0b8b70f246c17c4468c076874b0dff7a0336823b19e98bf1cec05e4beffb0591f97713c6"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_3_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_3_dIUT"
  by eval

lemma SEC1_P521_3_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_3_dIUT = SEC1_P521_3_QIUT" 
  by eval

lemma SEC1_P521_3_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_3_QCAVS" 
  by eval

lemma SEC1_P521_3_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_3_QCAVS" 
  using SEC1_P521_3_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_3_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_3_dIUT SEC1_P521_3_QCAVS = Some SEC1_P521_3_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_4›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_4_QCAVSx :: nat where 
  "SEC1_P521_4_QCAVSx = 0x00000098d99dee0816550e84dbfced7e88137fddcf581a725a455021115fe49f8dc3cf233cd9ea0e6f039dc7919da973cdceaca205da39e0bd98c8062536c47f258f44b5"

definition SEC1_P521_4_QCAVSy :: nat where 
  "SEC1_P521_4_QCAVSy = 0x000000cd225c8797371be0c4297d2b457740100c774141d8f214c23b61aa2b6cd4806b9b70722aa4965fb622f42b7391e27e5ec21c5679c5b06b59127372997d421adc1e"

definition SEC1_P521_4_QCAVS :: "int point" where
  "SEC1_P521_4_QCAVS = Point (int SEC1_P521_4_QCAVSx) (int SEC1_P521_4_QCAVSy)"

definition SEC1_P521_4_dIUT :: nat where 
  "SEC1_P521_4_dIUT = 0x000001c8aae94bb10b8ca4f7be577b4fb32bb2381032c4942c24fc2d753e7cc5e47b483389d9f3b956d20ee9001b1eef9f23545f72c5602140046839e963313c3decc864"

definition SEC1_P521_4_QIUTx :: nat where 
  "SEC1_P521_4_QIUTx = 0x00000106a14e2ee8ff970aa8ab0c79b97a33bba2958e070b75b94736b77bbe3f777324fa52872771aa88a63a9e8490c3378df4dc760cd14d62be700779dd1a4377943656"

definition SEC1_P521_4_QIUTy :: nat where 
  "SEC1_P521_4_QIUTy = 0x0000002366ce3941e0b284b1aa81215d0d3b9778fce23c8cd1e4ed6fa0abf62156c91d4b3eb55999c3471bed275e9e60e5aa9d690d310bfb15c9c5bbd6f5e9eb39682b74"

definition SEC1_P521_4_QIUT :: "int point" where
  "SEC1_P521_4_QIUT = Point (int SEC1_P521_4_QIUTx) (int SEC1_P521_4_QIUTy)"

definition SEC1_P521_4_ZIUT :: nat where 
  "SEC1_P521_4_ZIUT = 0x0145cfa38f25943516c96a5fd4bfebb2f645d10520117aa51971eff442808a23b4e23c187e639ff928c3725fbd1c0c2ad0d4aeb207bc1a6fb6cb6d467888dc044b3c"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_4_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_4_dIUT"
  by eval

lemma SEC1_P521_4_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_4_dIUT = SEC1_P521_4_QIUT" 
  by eval

lemma SEC1_P521_4_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_4_QCAVS" 
  by eval

lemma SEC1_P521_4_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_4_QCAVS" 
  using SEC1_P521_4_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_4_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_4_dIUT SEC1_P521_4_QCAVS = Some SEC1_P521_4_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_5›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_5_QCAVSx :: nat where 
  "SEC1_P521_5_QCAVSx = 0x0000007ae115adaaf041691ab6b7fb8c921f99d8ed32d283d67084e80b9ad9c40c56cd98389fb0a849d9ecf7268c297b6f93406119f40e32b5773ed25a28a9a85c4a7588"

definition SEC1_P521_5_QCAVSy :: nat where 
  "SEC1_P521_5_QCAVSy = 0x000001a28e004e37eeaefe1f4dbb71f1878696141af3a10a9691c4ed93487214643b761fa4b0fbeeb247cf6d3fba7a60697536ad03f49b80a9d1cb079673654977c5fa94"

definition SEC1_P521_5_QCAVS :: "int point" where
  "SEC1_P521_5_QCAVS = Point (int SEC1_P521_5_QCAVSx) (int SEC1_P521_5_QCAVSy)"

definition SEC1_P521_5_dIUT :: nat where 
  "SEC1_P521_5_dIUT = 0x0000009b0af137c9696c75b7e6df7b73156bb2d45f482e5a4217324f478b10ceb76af09724cf86afa316e7f89918d31d54824a5c33107a483c15c15b96edc661340b1c0e"

definition SEC1_P521_5_QIUTx :: nat where 
  "SEC1_P521_5_QIUTx = 0x000000748cdbb875d35f4bccb62abe20e82d32e4c14dc2feb5b87da2d0ccb11c9b6d4b7737b6c46f0dfb4d896e2db92fcf53cdbbae2a404c0babd564ad7adeac6273efa3"

definition SEC1_P521_5_QIUTy :: nat where 
  "SEC1_P521_5_QIUTy = 0x000001984acab8d8f173323de0bb60274b228871609373bb22a17287e9dec7495873abc09a8915b54c8455c8e02f654f602e23a2bbd7a9ebb74f3009bd65ecc650814cc0"

definition SEC1_P521_5_QIUT :: "int point" where
  "SEC1_P521_5_QIUT = Point (int SEC1_P521_5_QIUTx) (int SEC1_P521_5_QIUTy)"

definition SEC1_P521_5_ZIUT :: nat where 
  "SEC1_P521_5_ZIUT = 0x005c5721e96c273319fd60ecc46b5962f698e974b429f28fe6962f4ac656be2eb8674c4aafc037eab48ece612953b1e8d861016b6ad0c79805784c67f73ada96f351"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_5_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_5_dIUT"
  by eval

lemma SEC1_P521_5_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_5_dIUT = SEC1_P521_5_QIUT" 
  by eval

lemma SEC1_P521_5_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_5_QCAVS" 
  by eval

lemma SEC1_P521_5_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_5_QCAVS" 
  using SEC1_P521_5_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_5_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_5_dIUT SEC1_P521_5_QCAVS = Some SEC1_P521_5_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_6›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_6_QCAVSx :: nat where 
  "SEC1_P521_6_QCAVSx = 0x0000012588115e6f7f7bdcfdf57f03b169b479758baafdaf569d04135987b2ce6164c02a57685eb5276b5dae6295d3fe90620f38b5535c6d2260c173e61eb888ca920203"

definition SEC1_P521_6_QCAVSy :: nat where 
  "SEC1_P521_6_QCAVSy = 0x000001542c169cf97c2596fe2ddd848a222e367c5f7e6267ebc1bcd9ab5dcf49158f1a48e4af29a897b7e6a82091c2db874d8e7abf0f58064691344154f396dbaed188b6"

definition SEC1_P521_6_QCAVS :: "int point" where
  "SEC1_P521_6_QCAVS = Point (int SEC1_P521_6_QCAVSx) (int SEC1_P521_6_QCAVSy)"

definition SEC1_P521_6_dIUT :: nat where 
  "SEC1_P521_6_dIUT = 0x000001e48faacee6dec83ffcde944cf6bdf4ce4bae72747888ebafee455b1e91584971efb49127976a52f4142952f7c207ec0265f2b718cf3ead96ea4f62c752e4f7acd3"

definition SEC1_P521_6_QIUTx :: nat where 
  "SEC1_P521_6_QIUTx = 0x0000010eb1b4d9172bcc23f4f20cc9560fc54928c3f34ea61c00391dc766c76ed9fa608449377d1e4fadd1236025417330b4b91086704ace3e4e6484c606e2a943478c86"

definition SEC1_P521_6_QIUTy :: nat where 
  "SEC1_P521_6_QIUTy = 0x00000149413864069825ee1d0828da9f4a97713005e9bd1adbc3b38c5b946900721a960fe96ad2c1b3a44fe3de9156136d44cb17cbc2415729bb782e16bfe2deb3069e43"

definition SEC1_P521_6_QIUT :: "int point" where
  "SEC1_P521_6_QIUT = Point (int SEC1_P521_6_QIUTx) (int SEC1_P521_6_QIUTy)"

definition SEC1_P521_6_ZIUT :: nat where 
  "SEC1_P521_6_ZIUT = 0x01736d9717429b4f412e903febe2f9e0fffd81355d6ce2c06ff3f66a3be15ceec6e65e308347593f00d7f33591da4043c30763d72749f72cdceebe825e4b34ecd570"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_6_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_6_dIUT"
  by eval

lemma SEC1_P521_6_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_6_dIUT = SEC1_P521_6_QIUT" 
  by eval

lemma SEC1_P521_6_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_6_QCAVS" 
  by eval

lemma SEC1_P521_6_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_6_QCAVS" 
  using SEC1_P521_6_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_6_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_6_dIUT SEC1_P521_6_QCAVS = Some SEC1_P521_6_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_7›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_7_QCAVSx :: nat where 
  "SEC1_P521_7_QCAVSx = 0x00000169491d55bd09049fdf4c2a53a660480fee4c03a0538675d1cd09b5bba78dac48543ef118a1173b3fbf8b20e39ce0e6b890a163c50f9645b3d21d1cbb3b60a6fff4"

definition SEC1_P521_7_QCAVSy :: nat where 
  "SEC1_P521_7_QCAVSy = 0x00000083494b2eba76910fed33c761804515011fab50e3b377abd8a8a045d886d2238d2c268ac1b6ec88bd71b7ba78e2c33c152e4bf7da5d565e4acbecf5e92c7ad662bb"

definition SEC1_P521_7_QCAVS :: "int point" where
  "SEC1_P521_7_QCAVS = Point (int SEC1_P521_7_QCAVSx) (int SEC1_P521_7_QCAVSy)"

definition SEC1_P521_7_dIUT :: nat where 
  "SEC1_P521_7_dIUT = 0x000000c29aa223ea8d64b4a1eda27f39d3bc98ea0148dd98c1cbe595f8fd2bfbde119c9e017a50f5d1fc121c08c1cef31b758859556eb3e0e042d8dd6aaac57a05ca61e3"

definition SEC1_P521_7_QIUTx :: nat where 
  "SEC1_P521_7_QIUTx = 0x0000001511c848ef60d5419a98d10204db0fe58224124370061bcfa4e9249d50618c56bf3722471b259f38263bb7b280d23caf2a1ee8737f9371cdb2732cdc958369930c"

definition SEC1_P521_7_QIUTy :: nat where 
  "SEC1_P521_7_QIUTy = 0x000001d461681ae6d8c49b4c5f4d6016143fb1bd7491573e3ed0e6c48b82e821644f87f82f0e5f08fd16f1f98fa17586200ab02ed8c627b35c3f27617ec5fd92f456203f"

definition SEC1_P521_7_QIUT :: "int point" where
  "SEC1_P521_7_QIUT = Point (int SEC1_P521_7_QIUTx) (int SEC1_P521_7_QIUTy)"

definition SEC1_P521_7_ZIUT :: nat where 
  "SEC1_P521_7_ZIUT = 0x018f2ae9476c771726a77780208dedfefa205488996b18fecc50bfd4c132753f5766b2cd744afa9918606de2e016effc63622e9029e76dc6e3f0c69f7aeced565c2c"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_7_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_7_dIUT"
  by eval

lemma SEC1_P521_7_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_7_dIUT = SEC1_P521_7_QIUT" 
  by eval

lemma SEC1_P521_7_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_7_QCAVS" 
  by eval

lemma SEC1_P521_7_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_7_QCAVS" 
  using SEC1_P521_7_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_7_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_7_dIUT SEC1_P521_7_QCAVS = Some SEC1_P521_7_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_8›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_8_QCAVSx :: nat where 
  "SEC1_P521_8_QCAVSx = 0x0000008415f5bbd0eee387d6c09d0ef8acaf29c66db45d6ba101860ae45d3c60e1e0e3f7247a4626a60fdd404965c3566c79f6449e856ce0bf94619f97da8da24bd2cfb6"

definition SEC1_P521_8_QCAVSy :: nat where 
  "SEC1_P521_8_QCAVSy = 0x000000fdd7c59c58c361bc50a7a5d0d36f723b17c4f2ad2b03c24d42dc50f74a8c465a0afc4683f10fab84652dfe9e928c2626b5456453e1573ff60be1507467d431fbb2"

definition SEC1_P521_8_QCAVS :: "int point" where
  "SEC1_P521_8_QCAVS = Point (int SEC1_P521_8_QCAVSx) (int SEC1_P521_8_QCAVSy)"

definition SEC1_P521_8_dIUT :: nat where 
  "SEC1_P521_8_dIUT = 0x00000028692be2bf5c4b48939846fb3d5bce74654bb2646e15f8389e23708a1afadf561511ea0d9957d0b53453819d60fba8f65a18f7b29df021b1bb01cd163293acc3cc"

definition SEC1_P521_8_QIUTx :: nat where 
  "SEC1_P521_8_QIUTx = 0x000001cfdc10c799f5c79cb6930a65fba351748e07567993e5e410ef4cacc4cd8a25784991eb4674e41050f930c7190ac812b9245f48a7973b658daf408822fe5b85f668"

definition SEC1_P521_8_QIUTy :: nat where 
  "SEC1_P521_8_QIUTy = 0x00000180d9ddfc9af77b9c4a6f02a834db15e535e0b3845b2cce30388301b51cecbe3276307ef439b5c9e6a72dc2d94d879bc395052dbb4a5787d06efb280210fb8be037"

definition SEC1_P521_8_QIUT :: "int point" where
  "SEC1_P521_8_QIUT = Point (int SEC1_P521_8_QIUTx) (int SEC1_P521_8_QIUTy)"

definition SEC1_P521_8_ZIUT :: nat where 
  "SEC1_P521_8_ZIUT = 0x0105a346988b92ed8c7a25ce4d79d21bc86cfcc7f99c6cd19dbb4a39f48ab943b79e4f0647348da0b80bd864b85c6b8d92536d6aa544dc7537a00c858f8b66319e25"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_8_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_8_dIUT"
  by eval

lemma SEC1_P521_8_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_8_dIUT = SEC1_P521_8_QIUT" 
  by eval

lemma SEC1_P521_8_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_8_QCAVS" 
  by eval

lemma SEC1_P521_8_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_8_QCAVS" 
  using SEC1_P521_8_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_8_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_8_dIUT SEC1_P521_8_QCAVS = Some SEC1_P521_8_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_9›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_9_QCAVSx :: nat where 
  "SEC1_P521_9_QCAVSx = 0x000001c721eea805a5cba29f34ba5758775be0cf6160e6c08723f5ab17bf96a1ff2bd9427961a4f34b07fc0b14ca4b2bf6845debd5a869f124ebfa7aa72fe565050b7f18"

definition SEC1_P521_9_QCAVSy :: nat where 
  "SEC1_P521_9_QCAVSy = 0x000000b6e89eb0e1dcf181236f7c548fd1a8c16b258b52c1a9bfd3fe8f22841b26763265f074c4ccf2d634ae97b701956f67a11006c52d97197d92f585f5748bc2672eeb"

definition SEC1_P521_9_QCAVS :: "int point" where
  "SEC1_P521_9_QCAVS = Point (int SEC1_P521_9_QCAVSx) (int SEC1_P521_9_QCAVSy)"

definition SEC1_P521_9_dIUT :: nat where 
  "SEC1_P521_9_dIUT = 0x000001194d1ee613f5366cbc44b504d21a0cf6715e209cd358f2dd5f3e71cc0d67d0e964168c42a084ebda746f9863a86bacffc819f1edf1b8c727ccfb3047240a57c435"

definition SEC1_P521_9_QIUTx :: nat where 
  "SEC1_P521_9_QIUTx = 0x0000016bd15c8a58d366f7f2b2f298cc87b7485e9ee70d11d12448b8377c0a82c7626f67aff7f97be7a3546bf417eeeddf75a93c130191c84108042ea2fca17fd3f80d14"

definition SEC1_P521_9_QIUTy :: nat where 
  "SEC1_P521_9_QIUTy = 0x000001560502d04b74fce1743aab477a9d1eac93e5226981fdb97a7478ce4ce566ff7243931284fad850b0c2bcae0ddd2d97790160c1a2e77c3ed6c95ecc44b89e2637fc"

definition SEC1_P521_9_QIUT :: "int point" where
  "SEC1_P521_9_QIUT = Point (int SEC1_P521_9_QIUTx) (int SEC1_P521_9_QIUTy)"

definition SEC1_P521_9_ZIUT :: nat where 
  "SEC1_P521_9_ZIUT = 0x004531b3d2c6cd12f21604c8610e6723dbf4daf80b5a459d6ba5814397d1c1f7a21d7c114be964e27376aaebe3a7bc3d6af7a7f8c7befb611afe487ff032921f750f"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_9_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_9_dIUT"
  by eval

lemma SEC1_P521_9_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_9_dIUT = SEC1_P521_9_QIUT" 
  by eval

lemma SEC1_P521_9_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_9_QCAVS" 
  by eval

lemma SEC1_P521_9_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_9_QCAVS" 
  using SEC1_P521_9_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_9_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_9_dIUT SEC1_P521_9_QCAVS = Some SEC1_P521_9_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_10›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_10_QCAVSx :: nat where 
  "SEC1_P521_10_QCAVSx = 0x000001c35823e440a9363ab98d9fc7a7bc0c0532dc7977a79165599bf1a9cc64c00fb387b42cca365286e8430360bfad3643bc31354eda50dc936c329ecdb60905c40fcb"

definition SEC1_P521_10_QCAVSy :: nat where 
  "SEC1_P521_10_QCAVSy = 0x000000d9e7f433531e44df4f6d514201cbaabb06badd6783e01111726d815531d233c5cdb722893ffbb2027259d594de77438809738120c6f783934f926c3fb69b40c409"

definition SEC1_P521_10_QCAVS :: "int point" where
  "SEC1_P521_10_QCAVS = Point (int SEC1_P521_10_QCAVSx) (int SEC1_P521_10_QCAVSy)"

definition SEC1_P521_10_dIUT :: nat where 
  "SEC1_P521_10_dIUT = 0x000001fd90e3e416e98aa3f2b6afa7f3bf368e451ad9ca5bd54b5b14aee2ed6723dde5181f5085b68169b09fbec721372ccf6b284713f9a6356b8d560a8ff78ca3737c88"

definition SEC1_P521_10_QIUTx :: nat where 
  "SEC1_P521_10_QIUTx = 0x000001ebea1b10d3e3b971b7efb69fc878de11c7f472e4e4d384c31b8d6288d8071517acade9b39796c7af5163bcf71aeda777533f382c6cf0a4d9bbb938c85f44b78037"

definition SEC1_P521_10_QIUTy :: nat where 
  "SEC1_P521_10_QIUTy = 0x0000016b0e3e19c2996b2cbd1ff64730e7ca90edca1984f9b2951333535e5748baa34a99f61ff4d5f812079e0f01e87789f34efdad8098015ee74a4f846dd190d16dc6e1"

definition SEC1_P521_10_QIUT :: "int point" where
  "SEC1_P521_10_QIUT = Point (int SEC1_P521_10_QIUTx) (int SEC1_P521_10_QIUTy)"

definition SEC1_P521_10_ZIUT :: nat where 
  "SEC1_P521_10_ZIUT = 0x0100c8935969077bae0ba89ef0df8161d975ec5870ac811ae7e65ca5394efba4f0633d41bf79ea5e5b9496bbd7aae000b0594baa82ef8f244e6984ae87ae1ed124b7"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_10_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_10_dIUT"
  by eval

lemma SEC1_P521_10_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_10_dIUT = SEC1_P521_10_QIUT" 
  by eval

lemma SEC1_P521_10_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_10_QCAVS" 
  by eval

lemma SEC1_P521_10_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_10_QCAVS" 
  using SEC1_P521_10_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_10_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_10_dIUT SEC1_P521_10_QCAVS = Some SEC1_P521_10_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_11›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_11_QCAVSx :: nat where 
  "SEC1_P521_11_QCAVSx = 0x000000093057fb862f2ad2e82e581baeb3324e7b32946f2ba845a9beeed87d6995f54918ec6619b9931955d5a89d4d74adf1046bb362192f2ef6bd3e3d2d04dd1f87054a"

definition SEC1_P521_11_QCAVSy :: nat where 
  "SEC1_P521_11_QCAVSy = 0x000000aa3fb2448335f694e3cda4ae0cc71b1b2f2a206fa802d7262f19983c44674fe15327acaac1fa40424c395a6556cb8167312527fae5865ecffc14bbdc17da78cdcf"

definition SEC1_P521_11_QCAVS :: "int point" where
  "SEC1_P521_11_QCAVS = Point (int SEC1_P521_11_QCAVSx) (int SEC1_P521_11_QCAVSy)"

definition SEC1_P521_11_dIUT :: nat where 
  "SEC1_P521_11_dIUT = 0x0000009012ecfdadc85ced630afea534cdc8e9d1ab8be5f3753dcf5f2b09b40eda66fc6858549bc36e6f8df55998cfa9a0703aecf6c42799c245011064f530c09db98369"

definition SEC1_P521_11_QIUTx :: nat where 
  "SEC1_P521_11_QIUTx = 0x000000234e32be0a907131d2d128a6477e0caceb86f02479745e0fe245cb332de631c078871160482eeef584e274df7fa412cea3e1e91f71ecba8781d9205d48386341ad"

definition SEC1_P521_11_QIUTy :: nat where 
  "SEC1_P521_11_QIUTy = 0x000001cf86455b09b1c005cffba8d76289a3759628c874beea462f51f30bd581e3803134307dedbb771b3334ee15be2e242cd79c3407d2f58935456c6941dd9b6d155a46"

definition SEC1_P521_11_QIUT :: "int point" where
  "SEC1_P521_11_QIUT = Point (int SEC1_P521_11_QIUTx) (int SEC1_P521_11_QIUTy)"

definition SEC1_P521_11_ZIUT :: nat where 
  "SEC1_P521_11_ZIUT = 0x017f36af19303841d13a389d95ec0b801c7f9a679a823146c75c17bc44256e9ad422a4f8b31f14647b2c7d317b933f7c2946c4b8abd1d56d620fab1b5ff1a3adc71f"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_11_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_11_dIUT"
  by eval

lemma SEC1_P521_11_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_11_dIUT = SEC1_P521_11_QIUT" 
  by eval

lemma SEC1_P521_11_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_11_QCAVS" 
  by eval

lemma SEC1_P521_11_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_11_QCAVS" 
  using SEC1_P521_11_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_11_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_11_dIUT SEC1_P521_11_QCAVS = Some SEC1_P521_11_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_12›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_12_QCAVSx :: nat where 
  "SEC1_P521_12_QCAVSx = 0x00000083192ed0b1cb31f75817794937f66ad91cf74552cd510cedb9fd641310422af5d09f221cad249ee814d16dd7ac84ded9eacdc28340fcfc9c0c06abe30a2fc28cd8"

definition SEC1_P521_12_QCAVSy :: nat where 
  "SEC1_P521_12_QCAVSy = 0x0000002212ed868c9ba0fb2c91e2c39ba93996a3e4ebf45f2852d0928c48930e875cc7b428d0e7f3f4d503e5d60c68cb49b13c2480cd486bed9200caddaddfe4ff8e3562"

definition SEC1_P521_12_QCAVS :: "int point" where
  "SEC1_P521_12_QCAVS = Point (int SEC1_P521_12_QCAVSx) (int SEC1_P521_12_QCAVSy)"

definition SEC1_P521_12_dIUT :: nat where 
  "SEC1_P521_12_dIUT = 0x000001b5ff847f8eff20b88cfad42c06e58c3742f2f8f1fdfd64b539ba48c25926926bd5e332b45649c0b184f77255e9d58fe8afa1a6d968e2cb1d4637777120c765c128"

definition SEC1_P521_12_QIUTx :: nat where 
  "SEC1_P521_12_QIUTx = 0x000001de3dc9263bc8c4969dc684be0eec54befd9a9f3dba194d8658a789341bf0d78d84da6735227cafaf09351951691197573c8c360a11e5285712b8bbdf5ac91b977c"

definition SEC1_P521_12_QIUTy :: nat where 
  "SEC1_P521_12_QIUTy = 0x000000812de58cd095ec2e5a9b247eb3ed41d8bef6aeace194a7a05b65aa5d289fbc9b1770ec84bb6be0c2c64cc37c1d54a7f5d71377a9adbe20f26f6f2b544a821ea831"

definition SEC1_P521_12_QIUT :: "int point" where
  "SEC1_P521_12_QIUT = Point (int SEC1_P521_12_QIUTx) (int SEC1_P521_12_QIUTy)"

definition SEC1_P521_12_ZIUT :: nat where 
  "SEC1_P521_12_ZIUT = 0x00062f9fc29ae1a68b2ee0dcf956cbd38c88ae5f645eaa546b00ebe87a7260bf724be20d34b9d02076655c933d056b21e304c24ddb1dedf1dd76de611fc4a2340336"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_12_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_12_dIUT"
  by eval

lemma SEC1_P521_12_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_12_dIUT = SEC1_P521_12_QIUT" 
  by eval

lemma SEC1_P521_12_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_12_QCAVS" 
  by eval

lemma SEC1_P521_12_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_12_QCAVS" 
  using SEC1_P521_12_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_12_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_12_dIUT SEC1_P521_12_QCAVS = Some SEC1_P521_12_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_13›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_13_QCAVSx :: nat where 
  "SEC1_P521_13_QCAVSx = 0x000001a89b636a93e5d2ba6c2292bf23033a84f06a3ac1220ea71e806afbe097a804cc67e9baa514cfb6c12c9194be30212bf7aae7fdf6d376c212f0554e656463ffab7e"

definition SEC1_P521_13_QCAVSy :: nat where 
  "SEC1_P521_13_QCAVSy = 0x00000182efcaf70fc412d336602e014da47256a0b606f2addcce8053bf817ac8656bb4e42f14c8cbf2a68f488ab35dcdf64056271dee1f606a440ba4bd4e5a11b8b8e54f"

definition SEC1_P521_13_QCAVS :: "int point" where
  "SEC1_P521_13_QCAVS = Point (int SEC1_P521_13_QCAVSx) (int SEC1_P521_13_QCAVSy)"

definition SEC1_P521_13_dIUT :: nat where 
  "SEC1_P521_13_dIUT = 0x0000011a6347d4e801c91923488354cc533e7e35fddf81ff0fb7f56bb0726e0c29ee5dcdc5f394ba54cf57269048aab6e055895c8da24b8b0639a742314390cc04190ed6"

definition SEC1_P521_13_QIUTx :: nat where 
  "SEC1_P521_13_QIUTx = 0x000000fe30267f33ba5cdefc25cbb3c9320dad9ccb1d7d376644620ca4fadee5626a3cede25ad254624def727a7048f7145f76162aa98042f9b123b2076f8e8cf59b3fdf"

definition SEC1_P521_13_QIUTy :: nat where 
  "SEC1_P521_13_QIUTy = 0x0000001145dc6631953b6e2945e94301d6cbb098fe4b04f7ee9b09411df104dc82d7d79ec46a01ed0f2d3e7db6eb680694bdeb107c1078aec6cabd9ebee3d342fe7e54df"

definition SEC1_P521_13_QIUT :: "int point" where
  "SEC1_P521_13_QIUT = Point (int SEC1_P521_13_QIUTx) (int SEC1_P521_13_QIUTy)"

definition SEC1_P521_13_ZIUT :: nat where 
  "SEC1_P521_13_ZIUT = 0x0128ab09bfec5406799e610f772ba17e892249fa8e0e7b18a04b9197034b250b48294f1867fb9641518f92766066a07a8b917b0e76879e1011e51ccbd9f540c54d4f"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_13_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_13_dIUT"
  by eval

lemma SEC1_P521_13_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_13_dIUT = SEC1_P521_13_QIUT" 
  by eval

lemma SEC1_P521_13_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_13_QCAVS" 
  by eval

lemma SEC1_P521_13_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_13_QCAVS" 
  using SEC1_P521_13_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_13_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_13_dIUT SEC1_P521_13_QCAVS = Some SEC1_P521_13_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_14›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_14_QCAVSx :: nat where 
  "SEC1_P521_14_QCAVSx = 0x0000017200b3f16a68cbaed2bf78ba8cddfb6cffac262bba00fbc25f9dc72a07ce59372904899f364c44cb264c097b647d4412bee3e519892d534d9129f8a28f7500fee7"

definition SEC1_P521_14_QCAVSy :: nat where 
  "SEC1_P521_14_QCAVSy = 0x000000baba8d672a4f4a3b63de48b96f56e18df5d68f7d70d5109833f43770d6732e06b39ad60d93e5b43db8789f1ec0aba47286a39ea584235acea757dbf13d53b58364"

definition SEC1_P521_14_QCAVS :: "int point" where
  "SEC1_P521_14_QCAVS = Point (int SEC1_P521_14_QCAVSx) (int SEC1_P521_14_QCAVSy)"

definition SEC1_P521_14_dIUT :: nat where 
  "SEC1_P521_14_dIUT = 0x00000022b6d2a22d71dfaa811d2d9f9f31fbed27f2e1f3d239538ddf3e4cc8c39a330266db25b7bc0a9704f17bde7f3592bf5f1f2d4b56013aacc3d8d1bc02f00d3146cc"

definition SEC1_P521_14_QIUTx :: nat where 
  "SEC1_P521_14_QIUTx = 0x000000ba38cfbf9fd2518a3f61d43549e7a6a6d28b2be57ffd3e0faceb636b34ed17e044a9f249dae8fc132e937e2d9349cd2ed77bb1049ceb692a2ec5b17ad61502a64c"

definition SEC1_P521_14_QIUTy :: nat where 
  "SEC1_P521_14_QIUTy = 0x0000001ec91d3058573fa6c0564a02a1a010160c313bc7c73510dc983e5461682b5be00dbce7e2c682ad73f29ca822cdc111f68fabe33a7b384a648342c3cdb9f050bcdb"

definition SEC1_P521_14_QIUT :: "int point" where
  "SEC1_P521_14_QIUT = Point (int SEC1_P521_14_QIUTx) (int SEC1_P521_14_QIUTy)"

definition SEC1_P521_14_ZIUT :: nat where 
  "SEC1_P521_14_ZIUT = 0x0101e462e9d9159968f6440e956f11dcf2227ae4aea81667122b6af9239a291eb5d6cf5a4087f358525fcacfa46bb2db01a75af1ba519b2d31da33eda87a9d565748"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_14_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_14_dIUT"
  by eval

lemma SEC1_P521_14_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_14_dIUT = SEC1_P521_14_QIUT" 
  by eval

lemma SEC1_P521_14_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_14_QCAVS" 
  by eval

lemma SEC1_P521_14_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_14_QCAVS" 
  using SEC1_P521_14_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_14_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_14_dIUT SEC1_P521_14_QCAVS = Some SEC1_P521_14_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_15›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_15_QCAVSx :: nat where 
  "SEC1_P521_15_QCAVSx = 0x0000004efd5dbd2f979e3831ce98f82355d6ca14a5757842875882990ab85ab9b7352dd6b9b2f4ea9a1e95c3880d65d1f3602f9ca653dc346fac858658d75626f4d4fb08"

definition SEC1_P521_15_QCAVSy :: nat where 
  "SEC1_P521_15_QCAVSy = 0x00000061cf15dbdaa7f31589c98400373da284506d70c89f074ed262a9e28140796b7236c2eef99016085e71552ff488c72b7339fefb7915c38459cb20ab85aec4e45052"

definition SEC1_P521_15_QCAVS :: "int point" where
  "SEC1_P521_15_QCAVS = Point (int SEC1_P521_15_QCAVSx) (int SEC1_P521_15_QCAVSy)"

definition SEC1_P521_15_dIUT :: nat where 
  "SEC1_P521_15_dIUT = 0x0000005bacfff268acf6553c3c583b464ea36a1d35e2b257a5d49eb3419d5a095087c2fb4d15cf5bf5af816d0f3ff7586490ccd3ddc1a98b39ce63749c6288ce0dbdac7d"

definition SEC1_P521_15_QIUTx :: nat where 
  "SEC1_P521_15_QIUTx = 0x00000036e488da7581472a9d8e628c58d6ad727311b7e6a3f6ae33a8544f34b09280249020be7196916fafd90e2ec54b66b5468d2361b99b56fa00d7ac37abb8c6f16653"

definition SEC1_P521_15_QIUTy :: nat where 
  "SEC1_P521_15_QIUTy = 0x0000011edb9fb8adb6a43f4f5f5fdc1421c9fe04fc8ba46c9b66334e3af927c8befb4307104f299acec4e30f812d9345c9720d19869dbfffd4ca3e7d2713eb5fc3f42615"

definition SEC1_P521_15_QIUT :: "int point" where
  "SEC1_P521_15_QIUT = Point (int SEC1_P521_15_QIUTx) (int SEC1_P521_15_QIUTy)"

definition SEC1_P521_15_ZIUT :: nat where 
  "SEC1_P521_15_ZIUT = 0x0141d6a4b719ab67eaf04a92c0a41e2dda78f4354fb90bdc35202cc7699b9b04d49616f82255debf7bbec045ae58f982a66905fcfae69d689785e38c868eb4a27e7b"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_15_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_15_dIUT"
  by eval

lemma SEC1_P521_15_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_15_dIUT = SEC1_P521_15_QIUT" 
  by eval

lemma SEC1_P521_15_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_15_QCAVS" 
  by eval

lemma SEC1_P521_15_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_15_QCAVS" 
  using SEC1_P521_15_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_15_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_15_dIUT SEC1_P521_15_QCAVS = Some SEC1_P521_15_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_16›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_16_QCAVSx :: nat where 
  "SEC1_P521_16_QCAVSx = 0x00000129891de0cf3cf82e8c2cf1bf90bb296fe00ab08ca45bb7892e0e227a504fdd05d2381a4448b68adff9c4153c87eacb78330d8bd52515f9f9a0b58e85f446bb4e10"

definition SEC1_P521_16_QCAVSy :: nat where 
  "SEC1_P521_16_QCAVSy = 0x0000009edd679696d3d1d0ef327f200383253f6413683d9e4fcc87bb35f112c2f110098d15e5701d7ceee416291ff5fed85e687f727388b9afe26a4f6feed560b218e6bb"

definition SEC1_P521_16_QCAVS :: "int point" where
  "SEC1_P521_16_QCAVS = Point (int SEC1_P521_16_QCAVSx) (int SEC1_P521_16_QCAVSy)"

definition SEC1_P521_16_dIUT :: nat where 
  "SEC1_P521_16_dIUT = 0x0000008e2c93c5423876223a637cad367c8589da69a2d0fc68612f31923ae50219df2452e7cc92615b67f17b57ffd2f52b19154bb40d7715336420fde2e89fee244f59dc"

definition SEC1_P521_16_QIUTx :: nat where 
  "SEC1_P521_16_QIUTx = 0x000000fa3b35118d6c422570f724a26f90b2833b19239174cea081c53133f64db60d6940ea1261299c04c1f4587cdb0c4c39616479c1bb0c146799a118032dcf98f899c0"

definition SEC1_P521_16_QIUTy :: nat where 
  "SEC1_P521_16_QIUTy = 0x00000069f040229006151fa32b51f679c8816f7c17506b403809dc77cd58a2aec430d94d13b6c916de99f355aa45fcfbc6853d686c71be496a067d24bfaea4818fc51f75"

definition SEC1_P521_16_QIUT :: "int point" where
  "SEC1_P521_16_QIUT = Point (int SEC1_P521_16_QIUTx) (int SEC1_P521_16_QIUTy)"

definition SEC1_P521_16_ZIUT :: nat where 
  "SEC1_P521_16_ZIUT = 0x00345e26e0abb1aac12b75f3a9cf41efe1c336396dffa4a067a4c2cfeb878c68b2b045faa4e5b4e6fa4678f5b603c351903b14bf9a6a70c439257199a640890b61d1"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_16_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_16_dIUT"
  by eval

lemma SEC1_P521_16_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_16_dIUT = SEC1_P521_16_QIUT" 
  by eval

lemma SEC1_P521_16_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_16_QCAVS" 
  by eval

lemma SEC1_P521_16_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_16_QCAVS" 
  using SEC1_P521_16_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_16_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_16_dIUT SEC1_P521_16_QCAVS = Some SEC1_P521_16_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_17›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_17_QCAVSx :: nat where 
  "SEC1_P521_17_QCAVSx = 0x000001a3c20240e59f5b7a3e17c275d2314ba1741210ad58b71036f8c83cc1f6b0f409dfdd9113e94b67ec39c3291426c23ffcc447054670d2908ff8fe67dc2306034c5c"

definition SEC1_P521_17_QCAVSy :: nat where 
  "SEC1_P521_17_QCAVSy = 0x000001d2825bfd3af8b1e13205780c137fe938f84fde40188e61ea02cead81badfdb425c29f7d7fb0324debadc10bbb93de68f62c35069268283f5265865db57a79f7bf7"

definition SEC1_P521_17_QCAVS :: "int point" where
  "SEC1_P521_17_QCAVS = Point (int SEC1_P521_17_QCAVSx) (int SEC1_P521_17_QCAVSy)"

definition SEC1_P521_17_dIUT :: nat where 
  "SEC1_P521_17_dIUT = 0x00000004d49d39d40d8111bf16d28c5936554326b197353eebbcf47545393bc8d3aaf98f14f5be7074bfb38e6cc97b989754074daddb3045f4e4ce745669fdb3ec0d5fa8"

definition SEC1_P521_17_QIUTx :: nat where 
  "SEC1_P521_17_QIUTx = 0x0000012ec226d050ce07c79b3df4d0f0891f9f7adf462e8c98dbc1a2a14f5e53a3f5ad894433587cc429a8be9ea1d84fa33b1803690dae04da7218d30026157fc995cf52"

definition SEC1_P521_17_QIUTy :: nat where 
  "SEC1_P521_17_QIUTy = 0x0000004837dfbf3426f57b5c793269130abb9a38f618532211931154db4eeb9aede88e57290f842ea0f2ea9a5f74c6203a3920fe4e305f6118f676b154e1d75b9cb5eb88"

definition SEC1_P521_17_QIUT :: "int point" where
  "SEC1_P521_17_QIUT = Point (int SEC1_P521_17_QIUTx) (int SEC1_P521_17_QIUTy)"

definition SEC1_P521_17_ZIUT :: nat where 
  "SEC1_P521_17_ZIUT = 0x006fe9de6fb8e672e7fd150fdc5e617fabb0d43906354ccfd224757c7276f7a1010091b17ed072074f8d10a5ec971eb35a5cb7076603b7bc38d432cbc059f80f9488"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_17_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_17_dIUT"
  by eval

lemma SEC1_P521_17_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_17_dIUT = SEC1_P521_17_QIUT" 
  by eval

lemma SEC1_P521_17_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_17_QCAVS" 
  by eval

lemma SEC1_P521_17_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_17_QCAVS" 
  using SEC1_P521_17_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_17_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_17_dIUT SEC1_P521_17_QCAVS = Some SEC1_P521_17_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_18›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_18_QCAVSx :: nat where 
  "SEC1_P521_18_QCAVSx = 0x0000007e2d138f2832e345ae8ff65957e40e5ec7163f016bdf6d24a2243daa631d878a4a16783990c722382130f9e51f0c1bd6ff5ac96780e48b68f5dec95f42e6144bb5"

definition SEC1_P521_18_QCAVSy :: nat where 
  "SEC1_P521_18_QCAVSy = 0x000000b0de5c896791f52886b0f09913e26e78dd0b69798fc4df6d95e3ca708ecbcbcce1c1895f5561bbabaae372e9e67e6e1a3be60e19b470cdf673ec1fc393d3426e20"

definition SEC1_P521_18_QCAVS :: "int point" where
  "SEC1_P521_18_QCAVS = Point (int SEC1_P521_18_QCAVSx) (int SEC1_P521_18_QCAVSy)"

definition SEC1_P521_18_dIUT :: nat where 
  "SEC1_P521_18_dIUT = 0x0000011a5d1cc79cd2bf73ea106f0e60a5ace220813b53e27b739864334a07c03367efda7a4619fa6eef3a9746492283b3c445610a023a9cc49bf4591140384fca5c8bb5"

definition SEC1_P521_18_QIUTx :: nat where 
  "SEC1_P521_18_QIUTx = 0x000000eb07c7332eedb7d3036059d35f7d2288d4377d5f42337ad3964079fb120ccd4c8bd384b585621055217023acd9a94fcb3b965bfb394675e788ade41a1de73e620c"

definition SEC1_P521_18_QIUTy :: nat where 
  "SEC1_P521_18_QIUTy = 0x000000491a835de2e6e7deb7e090f4a11f2c460c0b1f3d5e94ee8d751014dc720784fd3b54500c86ebaef18429f09e8e876d5d1538968a030d7715dde99f0d8f06e29d59"

definition SEC1_P521_18_QIUT :: "int point" where
  "SEC1_P521_18_QIUT = Point (int SEC1_P521_18_QIUTx) (int SEC1_P521_18_QIUTy)"

definition SEC1_P521_18_ZIUT :: nat where 
  "SEC1_P521_18_ZIUT = 0x01e4e759ecedce1013baf73e6fcc0b92451d03bdd50489b78871c333114990c9ba6a9b2fc7b1a2d9a1794c1b60d9279af6f146f0bbfb0683140403bfa4ccdb524a29"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_18_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_18_dIUT"
  by eval

lemma SEC1_P521_18_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_18_dIUT = SEC1_P521_18_QIUT" 
  by eval

lemma SEC1_P521_18_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_18_QCAVS" 
  by eval

lemma SEC1_P521_18_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_18_QCAVS" 
  using SEC1_P521_18_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_18_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_18_dIUT SEC1_P521_18_QCAVS = Some SEC1_P521_18_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_19›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_19_QCAVSx :: nat where 
  "SEC1_P521_19_QCAVSx = 0x000000118c36022209b1af8ebad1a12b566fc48744576e1199fe80de1cdf851cdf03e5b9091a8f7e079e83b7f827259b691d0c22ee29d6bdf73ec7bbfd746f2cd97a357d"

definition SEC1_P521_19_QCAVSy :: nat where 
  "SEC1_P521_19_QCAVSy = 0x000000da5ff4904548a342e2e7ba6a1f4ee5f840411a96cf63e6fe622f22c13e614e0a847c11a1ab3f1d12cc850c32e095614ca8f7e2721477b486e9ff40372977c3f65c"

definition SEC1_P521_19_QCAVS :: "int point" where
  "SEC1_P521_19_QCAVS = Point (int SEC1_P521_19_QCAVSx) (int SEC1_P521_19_QCAVSy)"

definition SEC1_P521_19_dIUT :: nat where 
  "SEC1_P521_19_dIUT = 0x0000010c908caf1be74c616b625fc8c1f514446a6aec83b5937141d6afbb0a8c7666a7746fa1f7a6664a2123e8cdf6cd8bf836c56d3c0ebdcc980e43a186f938f3a78ae7"

definition SEC1_P521_19_QIUTx :: nat where 
  "SEC1_P521_19_QIUTx = 0x00000031890f4c7abec3f723362285d77d2636f876817db3bbc88b01e773597b969ff6f013ea470c854ab4a7739004eb8cbea69b82ddf36acadd406871798ecb2ac3aa7f"

definition SEC1_P521_19_QIUTy :: nat where 
  "SEC1_P521_19_QIUTy = 0x000000d8b429ae3250266b9643c0c765a60dc10155bc2531cf8627296f4978b6640a9e600e19d0037d58503fa80799546a814d7478a550aa90e5ebeb052527faaeae5d08"

definition SEC1_P521_19_QIUT :: "int point" where
  "SEC1_P521_19_QIUT = Point (int SEC1_P521_19_QIUTx) (int SEC1_P521_19_QIUTy)"

definition SEC1_P521_19_ZIUT :: nat where 
  "SEC1_P521_19_ZIUT = 0x0163c9191d651039a5fe985a0eea1eba018a40ab1937fcd2b61220820ee8f2302e9799f6edfc3f5174f369d672d377ea8954a8d0c8b851e81a56fda95212a6578f0e"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_19_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_19_dIUT"
  by eval

lemma SEC1_P521_19_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_19_dIUT = SEC1_P521_19_QIUT" 
  by eval

lemma SEC1_P521_19_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_19_QCAVS" 
  by eval

lemma SEC1_P521_19_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_19_QCAVS" 
  using SEC1_P521_19_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_19_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_19_dIUT SEC1_P521_19_QCAVS = Some SEC1_P521_19_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_20›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_20_QCAVSx :: nat where 
  "SEC1_P521_20_QCAVSx = 0x000001780edff1ca1c03cfbe593edc6c049bcb2860294a92c355489d9afb2e702075ade1c953895a456230a0cde905de4a3f38573dbfcccd67ad6e7e93f0b5581e926a5d"

definition SEC1_P521_20_QCAVSy :: nat where 
  "SEC1_P521_20_QCAVSy = 0x000000a5481962c9162962e7f0ebdec936935d0eaa813e8226d40d7f6119bfd940602380c86721e61db1830f51e139f210000bcec0d8edd39e54d73a9a129f95cd5fa979"

definition SEC1_P521_20_QCAVS :: "int point" where
  "SEC1_P521_20_QCAVS = Point (int SEC1_P521_20_QCAVSx) (int SEC1_P521_20_QCAVSy)"

definition SEC1_P521_20_dIUT :: nat where 
  "SEC1_P521_20_dIUT = 0x000001b37d6b7288de671360425d3e5ac1ccb21815079d8d73431e9b74a6f0e7ae004a357575b11ad66642ce8b775593eba9d98bf25c75ef0b4d3a2098bbc641f59a2b77"

definition SEC1_P521_20_QIUTx :: nat where 
  "SEC1_P521_20_QIUTx = 0x000000189a5ee34de7e35aefeaeef9220c18071b4c29a4c3bd9d954458bd3e82a7a34da34cff5579b8101c065b1f2f527cf4581501e28ef5671873e65267733d003520af"

definition SEC1_P521_20_QIUTy :: nat where 
  "SEC1_P521_20_QIUTy = 0x000001eb4bc50a7b4d4599d7e3fa773ddb9eb252c9b3422872e544bdf75c7bf60f5166ddc11eb08fa7c30822dabaee373ab468eb2d922e484e2a527fff2ebb804b7d9a37"

definition SEC1_P521_20_QIUT :: "int point" where
  "SEC1_P521_20_QIUT = Point (int SEC1_P521_20_QIUTx) (int SEC1_P521_20_QIUTy)"

definition SEC1_P521_20_ZIUT :: nat where 
  "SEC1_P521_20_ZIUT = 0x015d613e267a36342e0d125cdad643d80d97ed0600afb9e6b9545c9e64a98cc6da7c5aaa3a8da0bdd9dd3b97e9788218a80abafc106ef065c8f1c4e1119ef58d298b"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_20_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_20_dIUT"
  by eval

lemma SEC1_P521_20_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_20_dIUT = SEC1_P521_20_QIUT" 
  by eval

lemma SEC1_P521_20_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_20_QCAVS" 
  by eval

lemma SEC1_P521_20_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_20_QCAVS" 
  using SEC1_P521_20_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_20_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_20_dIUT SEC1_P521_20_QCAVS = Some SEC1_P521_20_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_21›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_21_QCAVSx :: nat where 
  "SEC1_P521_21_QCAVSx = 0x0000016dacffa183e5303083a334f765de724ec5ec9402026d4797884a9828a0d321a8cfac74ab737fe20a7d6befcfc73b6a35c1c7b01d373e31abc192d48a4241a35803"

definition SEC1_P521_21_QCAVSy :: nat where 
  "SEC1_P521_21_QCAVSy = 0x0000011e5327cac22d305e7156e559176e19bee7e4f2f59e86f1a9d0b6603b6a7df1069bde6387feb71587b8ffce5b266e1bae86de29378a34e5c74b6724c4d40a719923"

definition SEC1_P521_21_QCAVS :: "int point" where
  "SEC1_P521_21_QCAVS = Point (int SEC1_P521_21_QCAVSx) (int SEC1_P521_21_QCAVSy)"

definition SEC1_P521_21_dIUT :: nat where 
  "SEC1_P521_21_dIUT = 0x000000f2661ac762f60c5fff23be5d969ccd4ec6f98e4e72618d12bdcdb9b4102162333788c0bae59f91cdfc172c7a1681ee44d96ab2135a6e5f3415ebbcd55165b1afb0"

definition SEC1_P521_21_QIUTx :: nat where 
  "SEC1_P521_21_QIUTx = 0x000000a8e25a6902d687b4787cdc94c364ac7cecc5c495483ed363dc0aa95ee2bd739c4c4d46b17006c728b076350d7d7e54c6822f52f47162a25109aaaba690cab696ec"

definition SEC1_P521_21_QIUTy :: nat where 
  "SEC1_P521_21_QIUTy = 0x00000168d2f08fe19e4dc9ee7a195b03c9f7fe6676f9f520b6270557504e72ca4394a2c6918625e15ac0c51b8f95cd560123653fb8e8ee6db961e2c4c62cc54e92e2a2a9"

definition SEC1_P521_21_QIUT :: "int point" where
  "SEC1_P521_21_QIUT = Point (int SEC1_P521_21_QIUTx) (int SEC1_P521_21_QIUTy)"

definition SEC1_P521_21_ZIUT :: nat where 
  "SEC1_P521_21_ZIUT = 0x014d6082a3b5ced1ab8ca265a8106f302146c4acb8c30bb14a4c991e3c82a9731288bdb91e0e85bda313912d06384fc44f2153fb13506fa9cf43c9aab5750988c943"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_21_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_21_dIUT"
  by eval

lemma SEC1_P521_21_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_21_dIUT = SEC1_P521_21_QIUT" 
  by eval

lemma SEC1_P521_21_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_21_QCAVS" 
  by eval

lemma SEC1_P521_21_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_21_QCAVS" 
  using SEC1_P521_21_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_21_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_21_dIUT SEC1_P521_21_QCAVS = Some SEC1_P521_21_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_22›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_22_QCAVSx :: nat where 
  "SEC1_P521_22_QCAVSx = 0x000000a091421d3703e3b341e9f1e7d58f8cf7bdbd1798d001967b801d1cec27e605c580b2387c1cb464f55ce7ac80334102ab03cfb86d88af76c9f4129c01bedd3bbfc4"

definition SEC1_P521_22_QCAVSy :: nat where 
  "SEC1_P521_22_QCAVSy = 0x0000008c9c577a8e6fc446815e9d40baa66025f15dae285f19eb668ee60ae9c98e7ecdbf2b2a68e22928059f67db188007161d3ecf397e0883f0c4eb7eaf7827a62205cc"

definition SEC1_P521_22_QCAVS :: "int point" where
  "SEC1_P521_22_QCAVS = Point (int SEC1_P521_22_QCAVSx) (int SEC1_P521_22_QCAVSy)"

definition SEC1_P521_22_dIUT :: nat where 
  "SEC1_P521_22_dIUT = 0x000000f430ca1261f09681a9282e9e970a9234227b1d5e58d558c3cc6eff44d1bdf53de16ad5ee2b18b92d62fc79586116b0efc15f79340fb7eaf5ce6c44341dcf8dde27"

definition SEC1_P521_22_QIUTx :: nat where 
  "SEC1_P521_22_QIUTx = 0x0000006c1d9b5eca87de1fb871a0a32f807c725adccde9b3967453a71347d608f0c030cd09e338cdecbf4a02015bc8a6e8d3e2595fe773ffc2fc4e4a55d0b1a2cc00323b"

definition SEC1_P521_22_QIUTy :: nat where 
  "SEC1_P521_22_QIUTy = 0x000001141b2109e7f4981c952aa818a2b9f6f5c41feccdb7a7a45b9b4b672937771b008cae5f934dfe3fed10d383ab1f38769c92ce88d9be5414817ecb073a31ab368ccb"

definition SEC1_P521_22_QIUT :: "int point" where
  "SEC1_P521_22_QIUT = Point (int SEC1_P521_22_QIUTx) (int SEC1_P521_22_QIUTy)"

definition SEC1_P521_22_ZIUT :: nat where 
  "SEC1_P521_22_ZIUT = 0x0020c00747cb8d492fd497e0fec54644bf027d418ab686381f109712a99cabe328b9743d2225836f9ad66e5d7fed1de247e0da92f60d5b31f9e47672e57f710598f4"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_22_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_22_dIUT"
  by eval

lemma SEC1_P521_22_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_22_dIUT = SEC1_P521_22_QIUT" 
  by eval

lemma SEC1_P521_22_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_22_QCAVS" 
  by eval

lemma SEC1_P521_22_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_22_QCAVS" 
  using SEC1_P521_22_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_22_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_22_dIUT SEC1_P521_22_QCAVS = Some SEC1_P521_22_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_23›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_23_QCAVSx :: nat where 
  "SEC1_P521_23_QCAVSx = 0x0000004f38816681771289ce0cb83a5e29a1ab06fc91f786994b23708ff08a08a0f675b809ae99e9f9967eb1a49f196057d69e50d6dedb4dd2d9a81c02bdcc8f7f518460"

definition SEC1_P521_23_QCAVSy :: nat where 
  "SEC1_P521_23_QCAVSy = 0x0000009efb244c8b91087de1eed766500f0e81530752d469256ef79f6b965d8a2232a0c2dbc4e8e1d09214bab38485be6e357c4200d073b52f04e4a16fc6f5247187aecb"

definition SEC1_P521_23_QCAVS :: "int point" where
  "SEC1_P521_23_QCAVS = Point (int SEC1_P521_23_QCAVSx) (int SEC1_P521_23_QCAVSy)"

definition SEC1_P521_23_dIUT :: nat where 
  "SEC1_P521_23_dIUT = 0x0000005dc33aeda03c2eb233014ee468dff753b72f73b00991043ea353828ae69d4cd0fadeda7bb278b535d7c57406ff2e6e473a5a4ff98e90f90d6dadd25100e8d85666"

definition SEC1_P521_23_QIUTx :: nat where 
  "SEC1_P521_23_QIUTx = 0x000000c825ba307373cec8dd2498eef82e21fd9862168dbfeb83593980ca9f82875333899fe94f137daf1c4189eb502937c3a367ea7951ed8b0f3377fcdf2922021d46a5"

definition SEC1_P521_23_QIUTy :: nat where 
  "SEC1_P521_23_QIUTy = 0x0000016b8a2540d5e65493888bc337249e67c0a68774f3e8d81e3b4574a0125165f0bd58b8af9de74b35832539f95c3cd9f1b759408560aa6851ae3ac7555347b0d3b13b"

definition SEC1_P521_23_QIUT :: "int point" where
  "SEC1_P521_23_QIUT = Point (int SEC1_P521_23_QIUTx) (int SEC1_P521_23_QIUTy)"

definition SEC1_P521_23_ZIUT :: nat where 
  "SEC1_P521_23_ZIUT = 0x00c2bfafcd7fbd3e2fd1c750fdea61e70bd4787a7e68468c574ee99ebc47eedef064e8944a73bcb7913dbab5d93dca660d216c553622362794f7a2acc71022bdb16f"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_23_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_23_dIUT"
  by eval

lemma SEC1_P521_23_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_23_dIUT = SEC1_P521_23_QIUT" 
  by eval

lemma SEC1_P521_23_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_23_QCAVS" 
  by eval

lemma SEC1_P521_23_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_23_QCAVS" 
  using SEC1_P521_23_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_23_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_23_dIUT SEC1_P521_23_QCAVS = Some SEC1_P521_23_ZIUT"
  by eval
  
subsection‹Test Vector: SEC1_P521_24›

subsubsection‹Given Test Vector Values›

definition SEC1_P521_24_QCAVSx :: nat where 
  "SEC1_P521_24_QCAVSx = 0x000001a32099b02c0bd85371f60b0dd20890e6c7af048c8179890fda308b359dbbc2b7a832bb8c6526c4af99a7ea3f0b3cb96ae1eb7684132795c478ad6f962e4a6f446d"

definition SEC1_P521_24_QCAVSy :: nat where 
  "SEC1_P521_24_QCAVSy = 0x0000017627357b39e9d7632a1370b3e93c1afb5c851b910eb4ead0c9d387df67cde85003e0e427552f1cd09059aad0262e235cce5fba8cedc4fdc1463da76dcd4b6d1a46"

definition SEC1_P521_24_QCAVS :: "int point" where
  "SEC1_P521_24_QCAVS = Point (int SEC1_P521_24_QCAVSx) (int SEC1_P521_24_QCAVSy)"

definition SEC1_P521_24_dIUT :: nat where 
  "SEC1_P521_24_dIUT = 0x000000df14b1f1432a7b0fb053965fd8643afee26b2451ecb6a8a53a655d5fbe16e4c64ce8647225eb11e7fdcb23627471dffc5c2523bd2ae89957cba3a57a23933e5a78"

definition SEC1_P521_24_QIUTx :: nat where 
  "SEC1_P521_24_QIUTx = 0x0000004e8583bbbb2ecd93f0714c332dff5ab3bc6396e62f3c560229664329baa5138c3bb1c36428abd4e23d17fcb7a2cfcc224b2e734c8941f6f121722d7b6b94154576"

definition SEC1_P521_24_QIUTy :: nat where 
  "SEC1_P521_24_QIUTy = 0x000001cf0874f204b0363f020864672fadbf87c8811eb147758b254b74b14fae742159f0f671a018212bbf25b8519e126d4cad778cfff50d288fd39ceb0cac635b175ec0"

definition SEC1_P521_24_QIUT :: "int point" where
  "SEC1_P521_24_QIUT = Point (int SEC1_P521_24_QIUTx) (int SEC1_P521_24_QIUTy)"

definition SEC1_P521_24_ZIUT :: nat where 
  "SEC1_P521_24_ZIUT = 0x01aaf24e5d47e4080c18c55ea35581cd8da30f1a079565045d2008d51b12d0abb4411cda7a0785b15d149ed301a3697062f42da237aa7f07e0af3fd00eb1800d9c41"

subsubsection‹Checks›
text‹ Check that dIUT is a valid private key for the given curve.  Check that QIUT is the 
public key gotten from appling the EC key generation routine to dIUT.  Check that QCAVS is a partially
valid public key (meaning, it is a point on the curve but not the point at infinity).  Because the 
cofactor of the curve is h=1, then we know that QCAVS is a (totally) valid public key.  Finally show 
that the output of the Diffie-Hellman primitive applied to dIUT and QCAVS is the shared secret ZIUT.›

lemma SEC1_P521_24_dIUT_valid: 
  "SEC1_P521_ECprivateKeyValid SEC1_P521_24_dIUT"
  by eval

lemma SEC1_P521_24_IUT_validPair: 
  "SEC1_P521_ECkeyGen SEC1_P521_24_dIUT = SEC1_P521_24_QIUT" 
  by eval

lemma SEC1_P521_24_QCAVS_partialValid: 
  "SEC1_P521_ECpublicKeyPartialValid SEC1_P521_24_QCAVS" 
  by eval

lemma SEC1_P521_24_QCAVS_validPublicKey: 
  "SEC1_P521_ECpublicKeyValid SEC1_P521_24_QCAVS" 
  using SEC1_P521_24_QCAVS_partialValid SEC1_P521.partValidImpliesValidIFheq1 P521_h_def by blast

lemma SEC1_P521_24_ECDHprim_Check: 
  "SEC1_P521_ECDHprim SEC1_P521_24_dIUT SEC1_P521_24_QCAVS = Some SEC1_P521_24_ZIUT"
  by eval
  
section‹NIST Test Vectors for Public Key Validation (PKV): Curve P-192›
  
subsection‹Test Vector: PKV_P192_01›

definition PKV_P192_01_Qx :: nat where
  "PKV_P192_01_Qx = 0x472a620598e6715eff9cc022805d8cc8e8219f0e32042538"

definition PKV_P192_01_Qy :: nat where
  "PKV_P192_01_Qy = 0x1971ca86edb3471b2a16b9aae9de90f366f371b26385027e6"

definition PKV_P192_01_Q :: "int point" where
  "PKV_P192_01_Q = Point (int PKV_P192_01_Qx) (int PKV_P192_01_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P192_01_OutOfRange: 
  "P192_p  PKV_P192_01_Qx  P192_p  PKV_P192_01_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P192_02›

definition PKV_P192_02_Qx :: nat where
  "PKV_P192_02_Qx = 0x192a2b854bf4e70d5a8fecc98f43b4a744b26808f8cf4c60d"

definition PKV_P192_02_Qy :: nat where
  "PKV_P192_02_Qy = 0x2c0b29190588eabf08dfe160ac8d3ab6f5d5cc73678ebae8"

definition PKV_P192_02_Q :: "int point" where
  "PKV_P192_02_Q = Point (int PKV_P192_02_Qx) (int PKV_P192_02_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P192_02_OutOfRange: 
  "P192_p  PKV_P192_02_Qx  P192_p  PKV_P192_02_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P192_03›

definition PKV_P192_03_Qx :: nat where
  "PKV_P192_03_Qx = 0xc07ce28e4c846d7327f0554119ddb7e865fa1dd448ba2b40"

definition PKV_P192_03_Qy :: nat where
  "PKV_P192_03_Qy = 0x33aefa3177b99901d9ab6c12eb0749197420296ccb9d4e4a"

definition PKV_P192_03_Q :: "int point" where
  "PKV_P192_03_Q = Point (int PKV_P192_03_Qx) (int PKV_P192_03_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P192_03_NotOnCurve:
  "¬ SEC1_P192_on_curve' PKV_P192_03_Q" 
  by eval

subsection‹Test Vector: PKV_P192_04›

definition PKV_P192_04_Qx :: nat where
  "PKV_P192_04_Qx = 0xf77c2e5946d99932b2a01c1c73a296ecde568978103d8e2b"

definition PKV_P192_04_Qy :: nat where
  "PKV_P192_04_Qy = 0xde46b2d5c94dc11b53578eafaa23f96de9747b086979416c"

definition PKV_P192_04_Q :: "int point" where
  "PKV_P192_04_Q = Point (int PKV_P192_04_Qx) (int PKV_P192_04_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P192_04_NotOnCurve:
  "¬ SEC1_P192_on_curve' PKV_P192_04_Q" 
  by eval

subsection‹Test Vector: PKV_P192_05›

definition PKV_P192_05_Qx :: nat where
  "PKV_P192_05_Qx = 0xa9ee43654ce0b400c40b8d4ad9a4a53cac9662c56f3c4bde"

definition PKV_P192_05_Qy :: nat where
  "PKV_P192_05_Qy = 0xf376f44a79c50ac3fdfd25ab684af439b5938bee11a63f21"

definition PKV_P192_05_Q :: "int point" where
  "PKV_P192_05_Q = Point (int PKV_P192_05_Qx) (int PKV_P192_05_Qy)"

text‹Result: P (0 )›

lemma PKV_P192_05_OnCurve:
  "SEC1_P192_on_curve' PKV_P192_05_Q" 
  by eval

subsection‹Test Vector: PKV_P192_06›

definition PKV_P192_06_Qx :: nat where
  "PKV_P192_06_Qx = 0x1016451af3e7a7fa2ce3cf2acbe07fa7fa19a5f14455bf2ec"

definition PKV_P192_06_Qy :: nat where
  "PKV_P192_06_Qy = 0xc074630aea063e00bb41e6fbf752dd4f8e5bc742bf3363eb"

definition PKV_P192_06_Q :: "int point" where
  "PKV_P192_06_Q = Point (int PKV_P192_06_Qx) (int PKV_P192_06_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P192_06_OutOfRange: 
  "P192_p  PKV_P192_06_Qx  P192_p  PKV_P192_06_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P192_07›

definition PKV_P192_07_Qx :: nat where
  "PKV_P192_07_Qx = 0x18eea61787fbcd90f73f947346cdf13f05b4170e3e7456165"

definition PKV_P192_07_Qy :: nat where
  "PKV_P192_07_Qy = 0x5514c7b6e0eecc4e9c1ad99710f009a550bf3f952bb16593"

definition PKV_P192_07_Q :: "int point" where
  "PKV_P192_07_Q = Point (int PKV_P192_07_Qx) (int PKV_P192_07_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P192_07_OutOfRange: 
  "P192_p  PKV_P192_07_Qx  P192_p  PKV_P192_07_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P192_08›

definition PKV_P192_08_Qx :: nat where
  "PKV_P192_08_Qx = 0x39ed11c88869f6c4705125d9d5fc7c6b1e3d22b2fa7a6b57"

definition PKV_P192_08_Qy :: nat where
  "PKV_P192_08_Qy = 0xd0cf50208f6b1a61ba346a3f3f8f58128c8199e5405a6f11"

definition PKV_P192_08_Q :: "int point" where
  "PKV_P192_08_Q = Point (int PKV_P192_08_Qx) (int PKV_P192_08_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P192_08_NotOnCurve:
  "¬ SEC1_P192_on_curve' PKV_P192_08_Q" 
  by eval

subsection‹Test Vector: PKV_P192_09›

definition PKV_P192_09_Qx :: nat where
  "PKV_P192_09_Qx = 0x2addae388ff40a6176fe562b161da7b4957efc897d2a3c90"

definition PKV_P192_09_Qy :: nat where
  "PKV_P192_09_Qy = 0xa3f0772b1cd64f64dc0c84e59a21ac2dedb0e952e73d772a"

definition PKV_P192_09_Q :: "int point" where
  "PKV_P192_09_Q = Point (int PKV_P192_09_Qx) (int PKV_P192_09_Qy)"

text‹Result: P (0 )›

lemma PKV_P192_09_OnCurve:
  "SEC1_P192_on_curve' PKV_P192_09_Q" 
  by eval

subsection‹Test Vector: PKV_P192_10›

definition PKV_P192_10_Qx :: nat where
  "PKV_P192_10_Qx = 0xb4e02a84d98e3f0be8893538f23dc647fdfa7440169198fc"

definition PKV_P192_10_Qy :: nat where
  "PKV_P192_10_Qy = 0x7b26155e6d89a42e4bb3eb4d2dee6b39f42166b365774eb9"

definition PKV_P192_10_Q :: "int point" where
  "PKV_P192_10_Q = Point (int PKV_P192_10_Qx) (int PKV_P192_10_Qy)"

text‹Result: P (0 )›

lemma PKV_P192_10_OnCurve:
  "SEC1_P192_on_curve' PKV_P192_10_Q" 
  by eval

subsection‹Test Vector: PKV_P192_11›

definition PKV_P192_11_Qx :: nat where
  "PKV_P192_11_Qx = 0x87d67f9b7cced918d827ffc086cfd6a181fc61b2f56e000b"

definition PKV_P192_11_Qy :: nat where
  "PKV_P192_11_Qy = 0xc6c8d686c61a816d25c085db665f018e31ad6f71ee24d895"

definition PKV_P192_11_Q :: "int point" where
  "PKV_P192_11_Q = Point (int PKV_P192_11_Qx) (int PKV_P192_11_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P192_11_NotOnCurve:
  "¬ SEC1_P192_on_curve' PKV_P192_11_Q" 
  by eval

subsection‹Test Vector: PKV_P192_12›

definition PKV_P192_12_Qx :: nat where
  "PKV_P192_12_Qx = 0xb88fe1ee132f0e2c4e3fe10d580461120cf32f64ec4204eb"

definition PKV_P192_12_Qy :: nat where
  "PKV_P192_12_Qy = 0x92e65bc720f0360755d23dd8ca42f9705310f4d432850d84"

definition PKV_P192_12_Q :: "int point" where
  "PKV_P192_12_Q = Point (int PKV_P192_12_Qx) (int PKV_P192_12_Qy)"

text‹Result: P (0 )›

lemma PKV_P192_12_OnCurve:
  "SEC1_P192_on_curve' PKV_P192_12_Q" 
  by eval

section‹NIST Test Vectors for Public Key Validation (PKV): Curve P-224›
  
subsection‹Test Vector: PKV_P224_01›

definition PKV_P224_01_Qx :: nat where
  "PKV_P224_01_Qx = 0x7fb5f1881188022d9b0a8b808834dfdc0c7388459350983e4aee1412"

definition PKV_P224_01_Qy :: nat where
  "PKV_P224_01_Qy = 0x661c3e17be24b038b1532d90675747482c238911093fb7035cc9b2cd"

definition PKV_P224_01_Q :: "int point" where
  "PKV_P224_01_Q = Point (int PKV_P224_01_Qx) (int PKV_P224_01_Qy)"

text‹Result: P (0 )›

lemma PKV_P224_01_OnCurve:
  "SEC1_P224_on_curve' PKV_P224_01_Q" 
  by eval

subsection‹Test Vector: PKV_P224_02›

definition PKV_P224_02_Qx :: nat where
  "PKV_P224_02_Qx = 0x7a9369e2173bbf29589bf47e3ae0ccf47df6d2268c2292f906cc9261"

definition PKV_P224_02_Qy :: nat where
  "PKV_P224_02_Qy = 0x11afc53c7c1b085029f53b41fcd5a336bafb35b89d302f2bd04df44e6"

definition PKV_P224_02_Q :: "int point" where
  "PKV_P224_02_Q = Point (int PKV_P224_02_Qx) (int PKV_P224_02_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P224_02_OutOfRange: 
  "P224_p  PKV_P224_02_Qx  P224_p  PKV_P224_02_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P224_03›

definition PKV_P224_03_Qx :: nat where
  "PKV_P224_03_Qx = 0x4946a9935f9e416fdaf8bebdbc2ac454db06c6bba64b9d18b2b4f758"

definition PKV_P224_03_Qy :: nat where
  "PKV_P224_03_Qy = 0x2a69cc394e21b913244e01a4c45ff00f1b6d0a63ec2a738955cd1714"

definition PKV_P224_03_Q :: "int point" where
  "PKV_P224_03_Q = Point (int PKV_P224_03_Qx) (int PKV_P224_03_Qy)"

text‹Result: P (0 )›

lemma PKV_P224_03_OnCurve:
  "SEC1_P224_on_curve' PKV_P224_03_Q" 
  by eval

subsection‹Test Vector: PKV_P224_04›

definition PKV_P224_04_Qx :: nat where
  "PKV_P224_04_Qx = 0x1803faeef9b40957f59ab97d543f86690afd7471dfb8b04b84ea31085"

definition PKV_P224_04_Qy :: nat where
  "PKV_P224_04_Qy = 0x738cc29474ca048930b7f1a29db3773d11839ed83a6993e3f23692d7"

definition PKV_P224_04_Q :: "int point" where
  "PKV_P224_04_Q = Point (int PKV_P224_04_Qx) (int PKV_P224_04_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P224_04_OutOfRange: 
  "P224_p  PKV_P224_04_Qx  P224_p  PKV_P224_04_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P224_05›

definition PKV_P224_05_Qx :: nat where
  "PKV_P224_05_Qx = 0x9ec284178ce9605003e67662caaf19049f784fbe20d0bdbfd38e762d"

definition PKV_P224_05_Qy :: nat where
  "PKV_P224_05_Qy = 0x7f59ec5820f0ac70148ec71000b806a928704ea270254e529e05828b"

definition PKV_P224_05_Q :: "int point" where
  "PKV_P224_05_Q = Point (int PKV_P224_05_Qx) (int PKV_P224_05_Qy)"

text‹Result: P (0 )›

lemma PKV_P224_05_OnCurve:
  "SEC1_P224_on_curve' PKV_P224_05_Q" 
  by eval

subsection‹Test Vector: PKV_P224_06›

definition PKV_P224_06_Qx :: nat where
  "PKV_P224_06_Qx = 0xc01795a001b6b8a5b3db9acbdb55c2f97f4a50aa0a0cfed1d50a4c28"

definition PKV_P224_06_Qy :: nat where
  "PKV_P224_06_Qy = 0xb79dbe52a47a4640100cc939b435377f0bcb8db4ec52ecaadac5d919"

definition PKV_P224_06_Q :: "int point" where
  "PKV_P224_06_Q = Point (int PKV_P224_06_Qx) (int PKV_P224_06_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P224_06_NotOnCurve:
  "¬ SEC1_P224_on_curve' PKV_P224_06_Q" 
  by eval

subsection‹Test Vector: PKV_P224_07›

definition PKV_P224_07_Qx :: nat where
  "PKV_P224_07_Qx = 0xfbe3bff58dc58ca1ef9dc942fd43cdadbd060d70e0b1e6b9583a2228"

definition PKV_P224_07_Qy :: nat where
  "PKV_P224_07_Qy = 0xca844b43c237d497c34b986c681bf3cc54f968c0db74b2e1d9fe9d94"

definition PKV_P224_07_Q :: "int point" where
  "PKV_P224_07_Q = Point (int PKV_P224_07_Qx) (int PKV_P224_07_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P224_07_NotOnCurve:
  "¬ SEC1_P224_on_curve' PKV_P224_07_Q" 
  by eval

subsection‹Test Vector: PKV_P224_08›

definition PKV_P224_08_Qx :: nat where
  "PKV_P224_08_Qx = 0xcbe83c33848dd5a89ea8c45d23b99f23254e2077bd9ab26f6b5bed9f"

definition PKV_P224_08_Qy :: nat where
  "PKV_P224_08_Qy = 0xc0d09533d78a96e39028162534d74b097364095e2dc60776938af83b"

definition PKV_P224_08_Q :: "int point" where
  "PKV_P224_08_Q = Point (int PKV_P224_08_Qx) (int PKV_P224_08_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P224_08_NotOnCurve:
  "¬ SEC1_P224_on_curve' PKV_P224_08_Q" 
  by eval

subsection‹Test Vector: PKV_P224_09›

definition PKV_P224_09_Qx :: nat where
  "PKV_P224_09_Qx = 0x491e8d6c73708104c9530878f866e585cba008ef70baa46a809a2c03"

definition PKV_P224_09_Qy :: nat where
  "PKV_P224_09_Qy = 0x924a28ace8db9a88f7f874a1f24ac7f0bf56484f2130d5be5a8a1721"

definition PKV_P224_09_Q :: "int point" where
  "PKV_P224_09_Q = Point (int PKV_P224_09_Qx) (int PKV_P224_09_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P224_09_NotOnCurve:
  "¬ SEC1_P224_on_curve' PKV_P224_09_Q" 
  by eval

subsection‹Test Vector: PKV_P224_10›

definition PKV_P224_10_Qx :: nat where
  "PKV_P224_10_Qx = 0x1a89dc6a91002c9d25a3c4621fb5606b52531fd8e48a44119f442f749"

definition PKV_P224_10_Qy :: nat where
  "PKV_P224_10_Qy = 0x62f556641faa83059425026ca18ecbd219fe6d5df3b7713ce8b168cd"

definition PKV_P224_10_Q :: "int point" where
  "PKV_P224_10_Q = Point (int PKV_P224_10_Qx) (int PKV_P224_10_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P224_10_OutOfRange: 
  "P224_p  PKV_P224_10_Qx  P224_p  PKV_P224_10_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P224_11›

definition PKV_P224_11_Qx :: nat where
  "PKV_P224_11_Qx = 0xfcaef937ce0075a8adbff9ceb504357313ca150f6c402625832f22f0"

definition PKV_P224_11_Qy :: nat where
  "PKV_P224_11_Qy = 0x55b249ced1fa80dae295a532b8e54880c9d5b11921f1ab2f64f8da13"

definition PKV_P224_11_Q :: "int point" where
  "PKV_P224_11_Q = Point (int PKV_P224_11_Qx) (int PKV_P224_11_Qy)"

text‹Result: P (0 )›

lemma PKV_P224_11_OnCurve:
  "SEC1_P224_on_curve' PKV_P224_11_Q" 
  by eval

subsection‹Test Vector: PKV_P224_12›

definition PKV_P224_12_Qx :: nat where
  "PKV_P224_12_Qx = 0x182a4cee32c06292556f4e29950f5b2db9ad627a56e92680358d6cac4"

definition PKV_P224_12_Qy :: nat where
  "PKV_P224_12_Qy = 0xfa2a87aa3757ae9fa00d11db57089632c4f9e33fb214b9324cf75bd9"

definition PKV_P224_12_Q :: "int point" where
  "PKV_P224_12_Q = Point (int PKV_P224_12_Qx) (int PKV_P224_12_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P224_12_OutOfRange: 
  "P224_p  PKV_P224_12_Qx  P224_p  PKV_P224_12_Qy" 
  by eval
      
section‹NIST Test Vectors for Public Key Validation (PKV): Curve P-256›
  
subsection‹Test Vector: PKV_P256_01›

definition PKV_P256_01_Qx :: nat where
  "PKV_P256_01_Qx = 0xe0f7449c5588f24492c338f2bc8f7865f755b958d48edb0f2d0056e50c3fd5b7"

definition PKV_P256_01_Qy :: nat where
  "PKV_P256_01_Qy = 0x86d7e9255d0f4b6f44fa2cd6f8ba3c0aa828321d6d8cc430ca6284ce1d5b43a0"

definition PKV_P256_01_Q :: "int point" where
  "PKV_P256_01_Q = Point (int PKV_P256_01_Qx) (int PKV_P256_01_Qy)"

text‹Result: P (0 )›

lemma PKV_P256_01_OnCurve:
  "SEC1_P256_on_curve' PKV_P256_01_Q" 
  by eval

subsection‹Test Vector: PKV_P256_02›

definition PKV_P256_02_Qx :: nat where
  "PKV_P256_02_Qx = 0xd17c446237d9df87266ba3a91ff27f45abfdcb77bfd83536e92903efb861a9a9"

definition PKV_P256_02_Qy :: nat where
  "PKV_P256_02_Qy = 0x1eabb6a349ce2cd447d777b6739c5fc066add2002d2029052c408d0701066231c"

definition PKV_P256_02_Q :: "int point" where
  "PKV_P256_02_Q = Point (int PKV_P256_02_Qx) (int PKV_P256_02_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P256_02_OutOfRange: 
  "P256_p  PKV_P256_02_Qx  P256_p  PKV_P256_02_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P256_03›

definition PKV_P256_03_Qx :: nat where
  "PKV_P256_03_Qx = 0x17875397ae87369365656d490e8ce956911bd97607f2aff41b56f6f3a61989826"

definition PKV_P256_03_Qy :: nat where
  "PKV_P256_03_Qy = 0x980a3c4f61b9692633fbba5ef04c9cb546dd05cdec9fa8428b8849670e2fba92"

definition PKV_P256_03_Q :: "int point" where
  "PKV_P256_03_Q = Point (int PKV_P256_03_Qx) (int PKV_P256_03_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P256_03_OutOfRange: 
  "P256_p  PKV_P256_03_Qx  P256_p  PKV_P256_03_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P256_04›

definition PKV_P256_04_Qx :: nat where
  "PKV_P256_04_Qx = 0xf2d1c0dc0852c3d8a2a2500a23a44813ccce1ac4e58444175b440469ffc12273"

definition PKV_P256_04_Qy :: nat where
  "PKV_P256_04_Qy = 0x32bfe992831b305d8c37b9672df5d29fcb5c29b4a40534683e3ace23d24647dd"

definition PKV_P256_04_Q :: "int point" where
  "PKV_P256_04_Q = Point (int PKV_P256_04_Qx) (int PKV_P256_04_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P256_04_NotOnCurve:
  "¬ SEC1_P256_on_curve' PKV_P256_04_Q" 
  by eval

subsection‹Test Vector: PKV_P256_05›

definition PKV_P256_05_Qx :: nat where
  "PKV_P256_05_Qx = 0x10b0ca230fff7c04768f4b3d5c75fa9f6c539bea644dffbec5dc796a213061b58"

definition PKV_P256_05_Qy :: nat where
  "PKV_P256_05_Qy = 0xf5edf37c11052b75f771b7f9fa050e353e464221fec916684ed45b6fead38205"

definition PKV_P256_05_Q :: "int point" where
  "PKV_P256_05_Q = Point (int PKV_P256_05_Qx) (int PKV_P256_05_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P256_05_OutOfRange: 
  "P256_p  PKV_P256_05_Qx  P256_p  PKV_P256_05_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P256_06›

definition PKV_P256_06_Qx :: nat where
  "PKV_P256_06_Qx = 0x2c1052f25360a15062d204a056274e93cbe8fc4c4e9b9561134ad5c15ce525da"

definition PKV_P256_06_Qy :: nat where
  "PKV_P256_06_Qy = 0xced9783713a8a2a09eff366987639c625753295d9a85d0f5325e32dedbcada0b"

definition PKV_P256_06_Q :: "int point" where
  "PKV_P256_06_Q = Point (int PKV_P256_06_Qx) (int PKV_P256_06_Qy)"

text‹Result: P (0 )›

lemma PKV_P256_06_OnCurve:
  "SEC1_P256_on_curve' PKV_P256_06_Q" 
  by eval

subsection‹Test Vector: PKV_P256_07›

definition PKV_P256_07_Qx :: nat where
  "PKV_P256_07_Qx = 0xa40d077a87dae157d93dcccf3fe3aca9c6479a75aa2669509d2ef05c7de6782f"

definition PKV_P256_07_Qy :: nat where
  "PKV_P256_07_Qy = 0x503d86b87d743ba20804fd7e7884aa017414a7b5b5963e0d46e3a9611419ddf3"

definition PKV_P256_07_Q :: "int point" where
  "PKV_P256_07_Q = Point (int PKV_P256_07_Qx) (int PKV_P256_07_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P256_07_NotOnCurve:
  "¬ SEC1_P256_on_curve' PKV_P256_07_Q" 
  by eval

subsection‹Test Vector: PKV_P256_08›

definition PKV_P256_08_Qx :: nat where
  "PKV_P256_08_Qx = 0x2633d398a3807b1895548adbb0ea2495ef4b930f91054891030817df87d4ac0a"

definition PKV_P256_08_Qy :: nat where
  "PKV_P256_08_Qy = 0xd6b2f738e3873cc8364a2d364038ce7d0798bb092e3dd77cbdae7c263ba618d2"

definition PKV_P256_08_Q :: "int point" where
  "PKV_P256_08_Q = Point (int PKV_P256_08_Qx) (int PKV_P256_08_Qy)"

text‹Result: P (0 )›

lemma PKV_P256_08_OnCurve:
  "SEC1_P256_on_curve' PKV_P256_08_Q" 
  by eval

subsection‹Test Vector: PKV_P256_09›

definition PKV_P256_09_Qx :: nat where
  "PKV_P256_09_Qx = 0x14bf57f76c260b51ec6bbc72dbd49f02a56eaed070b774dc4bad75a54653c3d56"

definition PKV_P256_09_Qy :: nat where
  "PKV_P256_09_Qy = 0x7a231a23bf8b3aa31d9600d888a0678677a30e573decd3dc56b33f365cc11236"

definition PKV_P256_09_Q :: "int point" where
  "PKV_P256_09_Q = Point (int PKV_P256_09_Qx) (int PKV_P256_09_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P256_09_OutOfRange: 
  "P256_p  PKV_P256_09_Qx  P256_p  PKV_P256_09_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P256_10›

definition PKV_P256_10_Qx :: nat where
  "PKV_P256_10_Qx = 0x2fa74931ae816b426f484180e517f5050c92decfc8daf756cd91f54d51b302f1"

definition PKV_P256_10_Qy :: nat where
  "PKV_P256_10_Qy = 0x5b994346137988c58c14ae2152ac2f6ad96d97decb33099bd8a0210114cd1141"

definition PKV_P256_10_Q :: "int point" where
  "PKV_P256_10_Q = Point (int PKV_P256_10_Qx) (int PKV_P256_10_Qy)"

text‹Result: P (0 )›

lemma PKV_P256_10_OnCurve:
  "SEC1_P256_on_curve' PKV_P256_10_Q" 
  by eval

subsection‹Test Vector: PKV_P256_11›

definition PKV_P256_11_Qx :: nat where
  "PKV_P256_11_Qx = 0xf8c6dd3181a76aa0e36c2790bba47041acbe7b1e473ff71eee39a824dc595ff0"

definition PKV_P256_11_Qy :: nat where
  "PKV_P256_11_Qy = 0x9c965f227f281b3072b95b8daf29e88b35284f3574462e268e529bbdc50e9e52"

definition PKV_P256_11_Q :: "int point" where
  "PKV_P256_11_Q = Point (int PKV_P256_11_Qx) (int PKV_P256_11_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P256_11_NotOnCurve:
  "¬ SEC1_P256_on_curve' PKV_P256_11_Q" 
  by eval

subsection‹Test Vector: PKV_P256_12›

definition PKV_P256_12_Qx :: nat where
  "PKV_P256_12_Qx = 0x7a81a7e0b015252928d8b36e4ca37e92fdc328eb25c774b4f872693028c4be38"

definition PKV_P256_12_Qy :: nat where
  "PKV_P256_12_Qy = 0x08862f7335147261e7b1c3d055f9a316e4cab7daf99cc09d1c647f5dd6e7d5bb"

definition PKV_P256_12_Q :: "int point" where
  "PKV_P256_12_Q = Point (int PKV_P256_12_Qx) (int PKV_P256_12_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P256_12_NotOnCurve:
  "¬ SEC1_P256_on_curve' PKV_P256_12_Q" 
  by eval

section‹NIST Test Vectors for Public Key Validation (PKV): Curve P-384›
  
subsection‹Test Vector: PKV_P384_01›

definition PKV_P384_01_Qx :: nat where
  "PKV_P384_01_Qx = 0xe87cc868cdf196471d3fc78c324be2c4a0de8dbde182afea88baa51666f3cc9993eae5f1d60d4aec58894f0357273c48"

definition PKV_P384_01_Qy :: nat where
  "PKV_P384_01_Qy = 0x187219b0adc398c835791798053cc6a0bcc6e43228ac23101ee93dfce0e508be988a55fa495eb93b832064dc035e7720"

definition PKV_P384_01_Q :: "int point" where
  "PKV_P384_01_Q = Point (int PKV_P384_01_Qx) (int PKV_P384_01_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P384_01_NotOnCurve:
  "¬ SEC1_P384_on_curve' PKV_P384_01_Q" 
  by eval

subsection‹Test Vector: PKV_P384_02›

definition PKV_P384_02_Qx :: nat where
  "PKV_P384_02_Qx = 0x6e9c7e92ee23713fabb05d0b50e088eb534fd1e2b257c03304cfa33598f88a07c7e31a13e24707a7057ca2919323058e"

definition PKV_P384_02_Qy :: nat where
  "PKV_P384_02_Qy = 0xa218a485e22eae08c3618cfd73befcfcd13c3f196c08df99d7f79ebffe9f127b896aa0cb36cfdf2fc4818b8cd766f185"

definition PKV_P384_02_Q :: "int point" where
  "PKV_P384_02_Q = Point (int PKV_P384_02_Qx) (int PKV_P384_02_Qy)"

text‹Result: P (0 )›

lemma PKV_P384_02_OnCurve:
  "SEC1_P384_on_curve' PKV_P384_02_Q" 
  by eval

subsection‹Test Vector: PKV_P384_03›

definition PKV_P384_03_Qx :: nat where
  "PKV_P384_03_Qx = 0x452eb75736ac00974f953a0ce6060c19911a3463b045cb15ad6c0fa5045d66b04252a9001e8c4a9a6a0293f127bd20d9"

definition PKV_P384_03_Qy :: nat where
  "PKV_P384_03_Qy = 0xa1da4fbf8f0726fb9e04cf3ed0404af6cafb028b924c1951165f0ffe7caf04c05444cc7defb8cb62381727b6c1589f13"

definition PKV_P384_03_Q :: "int point" where
  "PKV_P384_03_Q = Point (int PKV_P384_03_Qx) (int PKV_P384_03_Qy)"

text‹Result: P (0 )›

lemma PKV_P384_03_OnCurve:
  "SEC1_P384_on_curve' PKV_P384_03_Q" 
  by eval

subsection‹Test Vector: PKV_P384_04›

definition PKV_P384_04_Qx :: nat where
  "PKV_P384_04_Qx = 0x25e5509a54f5fa62f94551dff3dfe210db1bb2bbc8fd4e672fbd5a211f9fd2f7eadc2b83fcd4198b7f857d9a2dc39c11"

definition PKV_P384_04_Qy :: nat where
  "PKV_P384_04_Qy = 0x98a4a13bc2f2d04bebd6d4e04412a9d306e57b90364583a6ec25bf6f0175bb5b397b8cfea83fd5d1e0ad052852b4aba7"

definition PKV_P384_04_Q :: "int point" where
  "PKV_P384_04_Q = Point (int PKV_P384_04_Qx) (int PKV_P384_04_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P384_04_NotOnCurve:
  "¬ SEC1_P384_on_curve' PKV_P384_04_Q" 
  by eval

subsection‹Test Vector: PKV_P384_05›

definition PKV_P384_05_Qx :: nat where
  "PKV_P384_05_Qx = 0x11a14be72dd023667047c260dd1960dd16555289d9570001d53ea3e494c1c107800dc5b24dd4de8490a071658702a0962"

definition PKV_P384_05_Qy :: nat where
  "PKV_P384_05_Qy = 0x78d65f6975d10df838b96a16cba873b59c28f2c7d05654b8c8b78bd193694ae45d6c6e046a20b984c3467c72d49395fe"

definition PKV_P384_05_Q :: "int point" where
  "PKV_P384_05_Q = Point (int PKV_P384_05_Qx) (int PKV_P384_05_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P384_05_OutOfRange: 
  "P384_p  PKV_P384_05_Qx  P384_p  PKV_P384_05_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P384_06›

definition PKV_P384_06_Qx :: nat where
  "PKV_P384_06_Qx = 0xa953eafd9dae3862d1049dd99cf628745bfb8f1024aaa567c51e9da01eb9bda996a7b1c906b3bb44a94649df2bcef304"

definition PKV_P384_06_Qy :: nat where
  "PKV_P384_06_Qy = 0x2f66dda137d3a408e2498d532f652e668f09b86bc056ff699efcc71ed1f22967ca7a99c8bf64f246b93c1982f856ed27"

definition PKV_P384_06_Q :: "int point" where
  "PKV_P384_06_Q = Point (int PKV_P384_06_Qx) (int PKV_P384_06_Qy)"

text‹Result: P (0 )›

lemma PKV_P384_06_OnCurve:
  "SEC1_P384_on_curve' PKV_P384_06_Q" 
  by eval

subsection‹Test Vector: PKV_P384_07›

definition PKV_P384_07_Qx :: nat where
  "PKV_P384_07_Qx = 0x1bf2238026a2489fb6ac1a8d6b82fdb33b05e8d01f1e2671eb22e61734031cc63efbf7e14d23e81fd432fc9935c627cdd"

definition PKV_P384_07_Qy :: nat where
  "PKV_P384_07_Qy = 0x6b377c8b187d568b782a28b38a7861b69e3d016f9f9ebb7eff2e7732a5132785b5a32e069dcef12875a995908a8b72f1"

definition PKV_P384_07_Q :: "int point" where
  "PKV_P384_07_Q = Point (int PKV_P384_07_Qx) (int PKV_P384_07_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P384_07_OutOfRange: 
  "P384_p  PKV_P384_07_Qx  P384_p  PKV_P384_07_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P384_08›

definition PKV_P384_08_Qx :: nat where
  "PKV_P384_08_Qx = 0xa999b80932ea62b4689769225b3ff34b0709c4e32342a824799ca63dcce1f3ed8819e080fc7fa130c1881c8131f4bcb5"

definition PKV_P384_08_Qy :: nat where
  "PKV_P384_08_Qy = 0xb8c77d0868c2c159e1be6bcd60ec488ab31531c21e1cb8fe2493ed26ac848fde7d27823a9a4912650511a3d460e25ef2"

definition PKV_P384_08_Q :: "int point" where
  "PKV_P384_08_Q = Point (int PKV_P384_08_Qx) (int PKV_P384_08_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P384_08_NotOnCurve:
  "¬ SEC1_P384_on_curve' PKV_P384_08_Q" 
  by eval

subsection‹Test Vector: PKV_P384_09›

definition PKV_P384_09_Qx :: nat where
  "PKV_P384_09_Qx = 0x5cbaa8088b0804fe14c2a6fa54b1adee1690fd17a682ea9ec5202ba7575e217c37e98565b7e95e7c882bb6eef76a3df1"

definition PKV_P384_09_Qy :: nat where
  "PKV_P384_09_Qy = 0x79d8c7e96ae7a7668496317c596b24ebe56e6ea5bc64b74c38867eb2c419d8277d20b9c27a2d5c75d1c7a47885d38d0e"

definition PKV_P384_09_Q :: "int point" where
  "PKV_P384_09_Q = Point (int PKV_P384_09_Qx) (int PKV_P384_09_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P384_09_NotOnCurve:
  "¬ SEC1_P384_on_curve' PKV_P384_09_Q" 
  by eval

subsection‹Test Vector: PKV_P384_10›

definition PKV_P384_10_Qx :: nat where
  "PKV_P384_10_Qx = 0xcfb4dbdcb1a8c6e8c6b4a9dd091eed015476ebd20837de1f6261a27999a08cff345f0d4627eb7778fc3495916a6d017b"

definition PKV_P384_10_Qy :: nat where
  "PKV_P384_10_Qy = 0x1c08f7a421bc0731321374f9b31ecf5ca820c006180da4c496f29f0d0e4947f368808fd3052ee4f1afb8c2005fd0c0ee8"

definition PKV_P384_10_Q :: "int point" where
  "PKV_P384_10_Q = Point (int PKV_P384_10_Qx) (int PKV_P384_10_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P384_10_OutOfRange: 
  "P384_p  PKV_P384_10_Qx  P384_p  PKV_P384_10_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P384_11›

definition PKV_P384_11_Qx :: nat where
  "PKV_P384_11_Qx = 0x1adaff25f37c8dfd33ecf216691a2107e522c21c99e29a76d8c1757ef84cc37c73ec5c2aa3be2fb0d5f1d372e08fbf9e"

definition PKV_P384_11_Qy :: nat where
  "PKV_P384_11_Qy = 0x1f39c8f86a20c130c34f767e085217232599541516e2d79d8e526fa03082bed2a5dc5fde6fd410c30245212e7816dd014"

definition PKV_P384_11_Q :: "int point" where
  "PKV_P384_11_Q = Point (int PKV_P384_11_Qx) (int PKV_P384_11_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P384_11_OutOfRange: 
  "P384_p  PKV_P384_11_Qx  P384_p  PKV_P384_11_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P384_12›

definition PKV_P384_12_Qx :: nat where
  "PKV_P384_12_Qx = 0x31951643c18400593f2d7cb32a3acf6071b4d95b8ab80a0535aa5edc9e01145f6dcc91a9977eb450eb077112edf887b2"

definition PKV_P384_12_Qy :: nat where
  "PKV_P384_12_Qy = 0x098a9e569684ca517bfdd5bc4b57876b210c3d7598e4f989e8f88f9f103b5d90d6baaa1a6617d524001c44a677bd13d0"

definition PKV_P384_12_Q :: "int point" where
  "PKV_P384_12_Q = Point (int PKV_P384_12_Qx) (int PKV_P384_12_Qy)"

text‹Result: P (0 )›

lemma PKV_P384_12_OnCurve:
  "SEC1_P384_on_curve' PKV_P384_12_Q" 
  by eval

section‹NIST Test Vectors for Public Key Validation (PKV): Curve P-521›
  
subsection‹Test Vector: PKV_P521_01›

definition PKV_P521_01_Qx :: nat where
  "PKV_P521_01_Qx = 0x165252970b786685babd0463f7314275c44ac1b558ab5a8e4bde60a441623b204982dcba2d3c0e7d379d5b637fd3edc0b0d2e0b7a33f7b36c03bb8bf3c6c5469ebe"

definition PKV_P521_01_Qy :: nat where
  "PKV_P521_01_Qy = 0x1300db0f0bc9b21ecff15eff4ed3bbe3dc1ac403dc96c89344d0030304da7ce57f1dc757af6816279464c61a0ab33645c3cd6583842cff0928081660b30775f594f"

definition PKV_P521_01_Q :: "int point" where
  "PKV_P521_01_Q = Point (int PKV_P521_01_Qx) (int PKV_P521_01_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P521_01_NotOnCurve:
  "¬ SEC1_P521_on_curve' PKV_P521_01_Q" 
  by eval

subsection‹Test Vector: PKV_P521_02›

definition PKV_P521_02_Qx :: nat where
  "PKV_P521_02_Qx = 0x089a576c7b31c5848afb9eaeb92bb28ab2fc3bdd58338438051680e4cf0fede6caa63135fb7b6f5d7f2f0743f3d141fd1d67ef55f200d9cfb5fa7ef004929d1b938"

definition PKV_P521_02_Qy :: nat where
  "PKV_P521_02_Qy = 0x00868b8e12a141d97eb3b8c500eb211d6e70b661665edcadb7a7f989d174fd4ed5d148f6588769122b8ce8a784b3027c5777520a20a368983b01743c27e42c49039"

definition PKV_P521_02_Q :: "int point" where
  "PKV_P521_02_Q = Point (int PKV_P521_02_Qx) (int PKV_P521_02_Qy)"

text‹Result: P (0 )›

lemma PKV_P521_02_OnCurve:
  "SEC1_P521_on_curve' PKV_P521_02_Q" 
  by eval

subsection‹Test Vector: PKV_P521_03›

definition PKV_P521_03_Qx :: nat where
  "PKV_P521_03_Qx = 0x1a39c4c5d5f6af8285931694b030f6b8bbc0a012ab73c3947c72a6210643cc63673947f5847f2503bb81ae1c8b6a0d7cb0ee5675f9027ca75445aee2b6d7beb78ea"

definition PKV_P521_03_Qy :: nat where
  "PKV_P521_03_Qy = 0x148beebbe6e298779e59d8fc88cfc28f4aa784d927e5127813894b6d593760608539d2eb9db9cc8b39813a5e5e03a7d39be67c9c8a566fa8d65ff25b5bee83b0a9b"

definition PKV_P521_03_Q :: "int point" where
  "PKV_P521_03_Q = Point (int PKV_P521_03_Qx) (int PKV_P521_03_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P521_03_NotOnCurve:
  "¬ SEC1_P521_on_curve' PKV_P521_03_Q" 
  by eval

subsection‹Test Vector: PKV_P521_04›

definition PKV_P521_04_Qx :: nat where
  "PKV_P521_04_Qx = 0x0f516be607274ad187f6e064eef542f28e93010598575174bed6741c9602a16f05e2a871de673f369c35b01749557c3211c21b77c95d0d2b3451683c36546ae8386"

definition PKV_P521_04_Qy :: nat where
  "PKV_P521_04_Qy = 0x0277086ba2919d478b0d147543ab823e5dce17b8faa2d9c035ec4db8423f844891e28c8bde0c585b511b3e2a98684deed119c34657d934e9d8400e4d3ddab6f8139"

definition PKV_P521_04_Q :: "int point" where
  "PKV_P521_04_Q = Point (int PKV_P521_04_Qx) (int PKV_P521_04_Qy)"

text‹Result: P (0 )›

lemma PKV_P521_04_OnCurve:
  "SEC1_P521_on_curve' PKV_P521_04_Q" 
  by eval

subsection‹Test Vector: PKV_P521_05›

definition PKV_P521_05_Qx :: nat where
  "PKV_P521_05_Qx = 0x0602d4e6955c52cacb2451b8a465d9345703a0a2a723e953156c07524d701f3f5f696c5be70c092210bb163e0d5df75151cf48f7ee0edc360f61cc8a94be560683d"

definition PKV_P521_05_Qy :: nat where
  "PKV_P521_05_Qy = 0x0eb13acba7b5bc7d32626d6499c5906ac73de240dd7766cef84d53525cb98c4dd852ed8dab8c1b440bacacccb8f2d4024c5e6a3de80840803a0bc11e5750b53c878"

definition PKV_P521_05_Q :: "int point" where
  "PKV_P521_05_Q = Point (int PKV_P521_05_Qx) (int PKV_P521_05_Qy)"

text‹Result: P (0 )›

lemma PKV_P521_05_OnCurve:
  "SEC1_P521_on_curve' PKV_P521_05_Q" 
  by eval

subsection‹Test Vector: PKV_P521_06›

definition PKV_P521_06_Qx :: nat where
  "PKV_P521_06_Qx = 0x0bec1326722dd943f750fda964cb485f31469095e5b4b5c4fa02041e2f443693ac36c09f371aea7d501881fc4e369f68e071bb14c5003b31dce504bd14338eb8104"

definition PKV_P521_06_Qy :: nat where
  "PKV_P521_06_Qy = 0x36cd502b0a4e91e1080a5b94886a819a823687693ce878408b59183730c69e5ab2d6e05ea5f3d32855cf8c7c20da59a28913025a1fa6835a3751ec6da69502f0547"

definition PKV_P521_06_Q :: "int point" where
  "PKV_P521_06_Q = Point (int PKV_P521_06_Qx) (int PKV_P521_06_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P521_06_OutOfRange: 
  "P521_p  PKV_P521_06_Qx  P521_p  PKV_P521_06_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P521_07›

definition PKV_P521_07_Qx :: nat where
  "PKV_P521_07_Qx = 0x3e064e446ce29891240b02948288bedc37a4e4163a78702f942728e2d530cfecdc0362cf2209a706a9d4db24c1dd6aba7ad81d6ddecdf6e12073a1c31e2dacd185d"

definition PKV_P521_07_Qy :: nat where
  "PKV_P521_07_Qy = 0x12d0363dbdc4d157afd517beaecf2e6c93896a288c7cec5f9ba9394524fb6d4f647a9937fe440fda73f2e31410517ed5a814eed038356699085f9983f2ea5faccd0"

definition PKV_P521_07_Q :: "int point" where
  "PKV_P521_07_Q = Point (int PKV_P521_07_Qx) (int PKV_P521_07_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P521_07_OutOfRange: 
  "P521_p  PKV_P521_07_Qx  P521_p  PKV_P521_07_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P521_08›

definition PKV_P521_08_Qx :: nat where
  "PKV_P521_08_Qx = 0x164ce2e2fa873f5648c22ed37f26c13d3da3180a0f6c3aa4b68d0a13293784a5f1356fc2495217065de4f3b504ee2248747ef96180e102879363fa5393fe6fc5fbe"

definition PKV_P521_08_Qy :: nat where
  "PKV_P521_08_Qy = 0x23126d6903cbd7735291d77599cfe7f5e45056250c37deba2642dc0b7163ce0cf763d0d353bb9974cf15195c4bc4421bdae274492cfca739a8b8341235cc2268bc0"

definition PKV_P521_08_Q :: "int point" where
  "PKV_P521_08_Q = Point (int PKV_P521_08_Qx) (int PKV_P521_08_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P521_08_OutOfRange: 
  "P521_p  PKV_P521_08_Qx  P521_p  PKV_P521_08_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P521_09›

definition PKV_P521_09_Qx :: nat where
  "PKV_P521_09_Qx = 0x0dc2c4a23433293a771300ec79a3cd0f2e627110a97da85a82f4f85e7be9c280213048a3ad01b3e72bf54555a1b5da9945adcfed94ed8f6ed405c77506b5e00f45a"

definition PKV_P521_09_Qy :: nat where
  "PKV_P521_09_Qy = 0x18f746aacd6ed4eaaf9b038789927a30125691bc525b29592abb13cf98f64c03cb36a477dc53971563ee74f3a7614677ab6817f6e5f22ceb02c90826a33fe7c94cd"

definition PKV_P521_09_Q :: "int point" where
  "PKV_P521_09_Q = Point (int PKV_P521_09_Qx) (int PKV_P521_09_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P521_09_NotOnCurve:
  "¬ SEC1_P521_on_curve' PKV_P521_09_Q" 
  by eval

subsection‹Test Vector: PKV_P521_10›

definition PKV_P521_10_Qx :: nat where
  "PKV_P521_10_Qx = 0x16e0383adc2986d01c18d7bde3b89eb5f732b56a6424c9394ec556a4660c3b88ddbc8654345ba6cff94bb002d16bc92e5907035f933785f633698e711738160d842"

definition PKV_P521_10_Qy :: nat where
  "PKV_P521_10_Qy = 0x1cf24be44e919e1576ecf51abdea113f8bb7121d670b86d8ee93ce1e6f79b17a6394987d74e6787facef5ca655196603468afd76e5cdf54ebb1331ce183cfe28c9e"

definition PKV_P521_10_Q :: "int point" where
  "PKV_P521_10_Q = Point (int PKV_P521_10_Qx) (int PKV_P521_10_Qy)"

text‹Result: F (2 - Point not on curve)›

lemma PKV_P521_10_NotOnCurve:
  "¬ SEC1_P521_on_curve' PKV_P521_10_Q" 
  by eval

subsection‹Test Vector: PKV_P521_11›

definition PKV_P521_11_Qx :: nat where
  "PKV_P521_11_Qx = 0x3d68ed9ce2bcb68f12ac37385ccdb6ee445f7b0a8f257593735abdf8bc0b98bc5ab5c5750e2e111fec2ecde6be321522ddc90d2b54634d30d28f43024f76e4653a7"

definition PKV_P521_11_Qy :: nat where
  "PKV_P521_11_Qy = 0x03f6f5f224d6aee43d781e3ad723062a61729a6ed959cd18c75d4982961ba8033767ed1168674a545b0a693d3587fbeaebc9b116143dbe1155ead48de89d980d617"

definition PKV_P521_11_Q :: "int point" where
  "PKV_P521_11_Q = Point (int PKV_P521_11_Qx) (int PKV_P521_11_Qy)"

text‹Result: F (1 - Q_x or Q_y out of range)›

lemma PKV_P521_11_OutOfRange: 
  "P521_p  PKV_P521_11_Qx  P521_p  PKV_P521_11_Qy" 
  by eval
      
subsection‹Test Vector: PKV_P521_12›

definition PKV_P521_12_Qx :: nat where
  "PKV_P521_12_Qx = 0x0fa01f87597bb9710346da572a4804ec01c56260c2ce23bc397b3822e5b6f0b75709d58a74e2f7bc8d1021b9c5fccecc3abf2314360bfcf6643593167e6d3641852"

definition PKV_P521_12_Qy :: nat where
  "PKV_P521_12_Qy = 0x1ead5db7ddd86f2bea66d9b7dfa941ed1409fcdbf7fdd976792a69ccda08c5fdc8e7d392f5921891de5fe6336fde535b468109bba424dba3db926e4d7b1b9cf4cfe"

definition PKV_P521_12_Q :: "int point" where
  "PKV_P521_12_Q = Point (int PKV_P521_12_Qx) (int PKV_P521_12_Qy)"

text‹Result: P (0 )›

lemma PKV_P521_12_OnCurve:
  "SEC1_P521_on_curve' PKV_P521_12_Q" 
  by eval


section‹NIST Test Vectors for EC Key Pair Generation: Curve P-192›
  
subsection‹Test Vector: KeyPair_P192_01›

definition KeyPair_P192_01_d :: nat where
  "KeyPair_P192_01_d = 0xe5ce89a34adddf25ff3bf1ffe6803f57d0220de3118798ea"

definition KeyPair_P192_01_Qx :: nat where
  "KeyPair_P192_01_Qx = 0x8abf7b3ceb2b02438af19543d3e5b1d573fa9ac60085840f"

definition KeyPair_P192_01_Qy :: nat where
  "KeyPair_P192_01_Qy = 0xa87f80182dcd56a6a061f81f7da393e7cffd5e0738c6b245"

definition KeyPair_P192_01_Q :: "int point" where
  "KeyPair_P192_01_Q = Point (int KeyPair_P192_01_Qx) (int KeyPair_P192_01_Qy)"

lemma KeyPair_P192_01_check: 
  "SEC1_P192_ECkeyGen KeyPair_P192_01_d = KeyPair_P192_01_Q" 
  by eval

subsection‹Test Vector: KeyPair_P192_02›

definition KeyPair_P192_02_d :: nat where
  "KeyPair_P192_02_d = 0x7d14435714ad13ff23341cb567cc91198ff8617cc39751b2"

definition KeyPair_P192_02_Qx :: nat where
  "KeyPair_P192_02_Qx = 0x39dc723b19527daa1e80425209c56463481b9b47c51f8cbd"

definition KeyPair_P192_02_Qy :: nat where
  "KeyPair_P192_02_Qy = 0x432a3e84f2a16418834fabaf6b7d2341669512951f1672ad"

definition KeyPair_P192_02_Q :: "int point" where
  "KeyPair_P192_02_Q = Point (int KeyPair_P192_02_Qx) (int KeyPair_P192_02_Qy)"

lemma KeyPair_P192_02_check: 
  "SEC1_P192_ECkeyGen KeyPair_P192_02_d = KeyPair_P192_02_Q" 
  by eval

subsection‹Test Vector: KeyPair_P192_03›

definition KeyPair_P192_03_d :: nat where
  "KeyPair_P192_03_d = 0x12039a122de1725d8d0e369b2fb536f7a38414a67cf69a83"

definition KeyPair_P192_03_Qx :: nat where
  "KeyPair_P192_03_Qx = 0xeb7e6011e825eb9f0c19d7d8695cf8da5805bf1f86019cef"

definition KeyPair_P192_03_Qy :: nat where
  "KeyPair_P192_03_Qy = 0x59694469ad604f6b407a9cba3d696b33ff410a9357b34baa"

definition KeyPair_P192_03_Q :: "int point" where
  "KeyPair_P192_03_Q = Point (int KeyPair_P192_03_Qx) (int KeyPair_P192_03_Qy)"

lemma KeyPair_P192_03_check: 
  "SEC1_P192_ECkeyGen KeyPair_P192_03_d = KeyPair_P192_03_Q" 
  by eval

subsection‹Test Vector: KeyPair_P192_04›

definition KeyPair_P192_04_d :: nat where
  "KeyPair_P192_04_d = 0xc9000be980277861ba12aef988c4fcee9fcc7976cdb52c24"

definition KeyPair_P192_04_Qx :: nat where
  "KeyPair_P192_04_Qx = 0x5c066c35a41fa529e0c2f222b554dba8248ec5f8a5d97db9"

definition KeyPair_P192_04_Qy :: nat where
  "KeyPair_P192_04_Qy = 0x87d8232e5ee59c028b3d92879748096f6ae174acf08f86ca"

definition KeyPair_P192_04_Q :: "int point" where
  "KeyPair_P192_04_Q = Point (int KeyPair_P192_04_Qx) (int KeyPair_P192_04_Qy)"

lemma KeyPair_P192_04_check: 
  "SEC1_P192_ECkeyGen KeyPair_P192_04_d = KeyPair_P192_04_Q" 
  by eval

subsection‹Test Vector: KeyPair_P192_05›

definition KeyPair_P192_05_d :: nat where
  "KeyPair_P192_05_d = 0x33d3e07b943e37455588cff5e45ca817ae800a1302bb01e3"

definition KeyPair_P192_05_Qx :: nat where
  "KeyPair_P192_05_Qx = 0x7adb43a9da0f8cd12fab9b87f2e03c956810b4d239fb14c7"

definition KeyPair_P192_05_Qy :: nat where
  "KeyPair_P192_05_Qy = 0x651ec5a42a7787ab26fce3af2ed10255de0d229842334f09"

definition KeyPair_P192_05_Q :: "int point" where
  "KeyPair_P192_05_Q = Point (int KeyPair_P192_05_Qx) (int KeyPair_P192_05_Qy)"

lemma KeyPair_P192_05_check: 
  "SEC1_P192_ECkeyGen KeyPair_P192_05_d = KeyPair_P192_05_Q" 
  by eval

subsection‹Test Vector: KeyPair_P192_06›

definition KeyPair_P192_06_d :: nat where
  "KeyPair_P192_06_d = 0xe23b51e2a07e73e23ff3399978f537dfb2532af873d1ceae"

definition KeyPair_P192_06_Qx :: nat where
  "KeyPair_P192_06_Qx = 0xbd6b11dbf12b786e0a972eadd87ec45c29e1d92f4f1e5f2b"

definition KeyPair_P192_06_Qy :: nat where
  "KeyPair_P192_06_Qy = 0x32e821b746f3ebdb5919458b08601ca259bae36c25cf1529"

definition KeyPair_P192_06_Q :: "int point" where
  "KeyPair_P192_06_Q = Point (int KeyPair_P192_06_Qx) (int KeyPair_P192_06_Qy)"

lemma KeyPair_P192_06_check: 
  "SEC1_P192_ECkeyGen KeyPair_P192_06_d = KeyPair_P192_06_Q" 
  by eval

subsection‹Test Vector: KeyPair_P192_07›

definition KeyPair_P192_07_d :: nat where
  "KeyPair_P192_07_d = 0x59c9a7db3e58ee05cca57a26faa4e459605ca606bda62a9b"

definition KeyPair_P192_07_Qx :: nat where
  "KeyPair_P192_07_Qx = 0xe8f3af0e8d2d25f1b32f5f131f00c88ef807a33886106106"

definition KeyPair_P192_07_Qy :: nat where
  "KeyPair_P192_07_Qy = 0x3661a1803542e39c00634b8e3186a019dc39b151bd6bb6ae"

definition KeyPair_P192_07_Q :: "int point" where
  "KeyPair_P192_07_Q = Point (int KeyPair_P192_07_Qx) (int KeyPair_P192_07_Qy)"

lemma KeyPair_P192_07_check: 
  "SEC1_P192_ECkeyGen KeyPair_P192_07_d = KeyPair_P192_07_Q" 
  by eval

subsection‹Test Vector: KeyPair_P192_08›

definition KeyPair_P192_08_d :: nat where
  "KeyPair_P192_08_d = 0x36bf9605bfec53fcf22cb1e0cce77b40e41b092b3ae6d009"

definition KeyPair_P192_08_Qx :: nat where
  "KeyPair_P192_08_Qx = 0x8475e3b373cae91c8229b958a0238984aef5a9f4b6cd30a6"

definition KeyPair_P192_08_Qy :: nat where
  "KeyPair_P192_08_Qy = 0xded68658891512f35e6c003beaf4475b628376c87d1083ba"

definition KeyPair_P192_08_Q :: "int point" where
  "KeyPair_P192_08_Q = Point (int KeyPair_P192_08_Qx) (int KeyPair_P192_08_Qy)"

lemma KeyPair_P192_08_check: 
  "SEC1_P192_ECkeyGen KeyPair_P192_08_d = KeyPair_P192_08_Q" 
  by eval

subsection‹Test Vector: KeyPair_P192_09›

definition KeyPair_P192_09_d :: nat where
  "KeyPair_P192_09_d = 0x6774fde78e05c49a819e8a15f375a5944e289abd615a9d59"

definition KeyPair_P192_09_Qx :: nat where
  "KeyPair_P192_09_Qx = 0x7fc1af2b7080521311a84e8894a09c954105ece28cd508aa"

definition KeyPair_P192_09_Qy :: nat where
  "KeyPair_P192_09_Qy = 0x5d09f61268d10bb1d6d27f1d7fb3d87fb79d032ad35ac691"

definition KeyPair_P192_09_Q :: "int point" where
  "KeyPair_P192_09_Q = Point (int KeyPair_P192_09_Qx) (int KeyPair_P192_09_Qy)"

lemma KeyPair_P192_09_check: 
  "SEC1_P192_ECkeyGen KeyPair_P192_09_d = KeyPair_P192_09_Q" 
  by eval

subsection‹Test Vector: KeyPair_P192_10›

definition KeyPair_P192_10_d :: nat where
  "KeyPair_P192_10_d = 0xec08d03c8b42b1c79bfb3e8eb38b1553db63599a511dd5b9"

definition KeyPair_P192_10_Qx :: nat where
  "KeyPair_P192_10_Qx = 0x65c0b405cdcee79cdea70ad70cf303b4614f2406fd48ad3c"

definition KeyPair_P192_10_Qy :: nat where
  "KeyPair_P192_10_Qy = 0xf622c2bbc0443b9249e2f0a1d93b8364203458a2fa7d4264"

definition KeyPair_P192_10_Q :: "int point" where
  "KeyPair_P192_10_Q = Point (int KeyPair_P192_10_Qx) (int KeyPair_P192_10_Qy)"

lemma KeyPair_P192_10_check: 
  "SEC1_P192_ECkeyGen KeyPair_P192_10_d = KeyPair_P192_10_Q" 
  by eval

section‹NIST Test Vectors for EC Key Pair Generation: Curve P-224›
  
subsection‹Test Vector: KeyPair_P224_01›

definition KeyPair_P224_01_d :: nat where
  "KeyPair_P224_01_d = 0xe7c92383846a4e6887a10498d8eaca2bd0487d985bd7d3f92ce3ab30"

definition KeyPair_P224_01_Qx :: nat where
  "KeyPair_P224_01_Qx = 0x0a3682d2aaa4dd931bee042d32e95755507ab164b12f84843f4b7b96"

definition KeyPair_P224_01_Qy :: nat where
  "KeyPair_P224_01_Qy = 0xa6313a938eff7a293222e0e3c7b4c6132489b33255a61c3fc1ce2256"

definition KeyPair_P224_01_Q :: "int point" where
  "KeyPair_P224_01_Q = Point (int KeyPair_P224_01_Qx) (int KeyPair_P224_01_Qy)"

lemma KeyPair_P224_01_check: 
  "SEC1_P224_ECkeyGen KeyPair_P224_01_d = KeyPair_P224_01_Q" 
  by eval

subsection‹Test Vector: KeyPair_P224_02›

definition KeyPair_P224_02_d :: nat where
  "KeyPair_P224_02_d = 0x7f29534466bcb399777a0c7d3d4eff787d96db26ac3561f9d43cccd9"

definition KeyPair_P224_02_Qx :: nat where
  "KeyPair_P224_02_Qx = 0xd64ad34b097c4ff4ebee68f11f04e195ef0c7d123eb21c11f13a91ad"

definition KeyPair_P224_02_Qy :: nat where
  "KeyPair_P224_02_Qy = 0xd155b803686c3e2aefde4a626dba1e09722af3617aaf67ad59844b09"

definition KeyPair_P224_02_Q :: "int point" where
  "KeyPair_P224_02_Q = Point (int KeyPair_P224_02_Qx) (int KeyPair_P224_02_Qy)"

lemma KeyPair_P224_02_check: 
  "SEC1_P224_ECkeyGen KeyPair_P224_02_d = KeyPair_P224_02_Q" 
  by eval

subsection‹Test Vector: KeyPair_P224_03›

definition KeyPair_P224_03_d :: nat where
  "KeyPair_P224_03_d = 0x2ce71aafdad95b69fdda27f441b2f28da06db5d17adb468af3e351a0"

definition KeyPair_P224_03_Qx :: nat where
  "KeyPair_P224_03_Qx = 0x86e8c2c2078692b9308e6ad34e1dfc51a6676d3c8a4a5e7b3295876a"

definition KeyPair_P224_03_Qy :: nat where
  "KeyPair_P224_03_Qy = 0x7c83d5d358916a0d990be417da7879f7e05a1576a7942c1bc075c0d3"

definition KeyPair_P224_03_Q :: "int point" where
  "KeyPair_P224_03_Q = Point (int KeyPair_P224_03_Qx) (int KeyPair_P224_03_Qy)"

lemma KeyPair_P224_03_check: 
  "SEC1_P224_ECkeyGen KeyPair_P224_03_d = KeyPair_P224_03_Q" 
  by eval

subsection‹Test Vector: KeyPair_P224_04›

definition KeyPair_P224_04_d :: nat where
  "KeyPair_P224_04_d = 0x68736e34687ee8408917200baa0a30a87c3d1b2d04fe617c0ba212e9"

definition KeyPair_P224_04_Qx :: nat where
  "KeyPair_P224_04_Qx = 0x50990cc0d778df3c2b7fbc610e1b562a92632f98dfb799118f108e22"

definition KeyPair_P224_04_Qy :: nat where
  "KeyPair_P224_04_Qy = 0xb9bacfd53f09d5de324bb13b866c1062bb616b0aa000d304509d0a54"

definition KeyPair_P224_04_Q :: "int point" where
  "KeyPair_P224_04_Q = Point (int KeyPair_P224_04_Qx) (int KeyPair_P224_04_Qy)"

lemma KeyPair_P224_04_check: 
  "SEC1_P224_ECkeyGen KeyPair_P224_04_d = KeyPair_P224_04_Q" 
  by eval

subsection‹Test Vector: KeyPair_P224_05›

definition KeyPair_P224_05_d :: nat where
  "KeyPair_P224_05_d = 0xd6901df3d4020dae1132b6f4e028dbead231d1f3d53c8eb22d85a6d7"

definition KeyPair_P224_05_Qx :: nat where
  "KeyPair_P224_05_Qx = 0xfe90c3e6d3e8a67759fbac03ef656bdabaee8c32724d9ad59cd0e379"

definition KeyPair_P224_05_Qy :: nat where
  "KeyPair_P224_05_Qy = 0xf80833779516d2ff7b5f1dc2e46d5c635d3a361356b220ac19dc0511"

definition KeyPair_P224_05_Q :: "int point" where
  "KeyPair_P224_05_Q = Point (int KeyPair_P224_05_Qx) (int KeyPair_P224_05_Qy)"

lemma KeyPair_P224_05_check: 
  "SEC1_P224_ECkeyGen KeyPair_P224_05_d = KeyPair_P224_05_Q" 
  by eval

subsection‹Test Vector: KeyPair_P224_06›

definition KeyPair_P224_06_d :: nat where
  "KeyPair_P224_06_d = 0x7b9bc29a4e737d97af8b148b2a56f0a0dd22b8e69db10172ddcb8e59"

definition KeyPair_P224_06_Qx :: nat where
  "KeyPair_P224_06_Qx = 0x893e9480585b457b17b6b252e2f0c319c1fcf203558419e7c2edbfb5"

definition KeyPair_P224_06_Qy :: nat where
  "KeyPair_P224_06_Qy = 0xd11a2968026b5884b0fec07969b7524ee87e02ff3dff30359fa56cb1"

definition KeyPair_P224_06_Q :: "int point" where
  "KeyPair_P224_06_Q = Point (int KeyPair_P224_06_Qx) (int KeyPair_P224_06_Qy)"

lemma KeyPair_P224_06_check: 
  "SEC1_P224_ECkeyGen KeyPair_P224_06_d = KeyPair_P224_06_Q" 
  by eval

subsection‹Test Vector: KeyPair_P224_07›

definition KeyPair_P224_07_d :: nat where
  "KeyPair_P224_07_d = 0xc078ad3141b567848dc389c3eb1d6d6733929c73d6bcb94333805376"

definition KeyPair_P224_07_Qx :: nat where
  "KeyPair_P224_07_Qx = 0xcd2008a5f860d45eb249258b4708de54643d50c761885c21f204feb0"

definition KeyPair_P224_07_Qy :: nat where
  "KeyPair_P224_07_Qy = 0x72b480f1e2b4f57bd0ad56be3317edd7b8ecced650fa82576ea896b8"

definition KeyPair_P224_07_Q :: "int point" where
  "KeyPair_P224_07_Q = Point (int KeyPair_P224_07_Qx) (int KeyPair_P224_07_Qy)"

lemma KeyPair_P224_07_check: 
  "SEC1_P224_ECkeyGen KeyPair_P224_07_d = KeyPair_P224_07_Q" 
  by eval

subsection‹Test Vector: KeyPair_P224_08›

definition KeyPair_P224_08_d :: nat where
  "KeyPair_P224_08_d = 0x6140b0af79c0b92e938bc8a5fc2ccbd62b7e7825c1bb5c89d4434856"

definition KeyPair_P224_08_Qx :: nat where
  "KeyPair_P224_08_Qx = 0xf40ea8fb8e3c2815df388a4a0dde5390c6291c24e20eb61ebfcf9b79"

definition KeyPair_P224_08_Qy :: nat where
  "KeyPair_P224_08_Qy = 0x1fb6a9d6ccdcd1765d88b185ce1aef4fa32c3a3465d334dde2468aa3"

definition KeyPair_P224_08_Q :: "int point" where
  "KeyPair_P224_08_Q = Point (int KeyPair_P224_08_Qx) (int KeyPair_P224_08_Qy)"

lemma KeyPair_P224_08_check: 
  "SEC1_P224_ECkeyGen KeyPair_P224_08_d = KeyPair_P224_08_Q" 
  by eval

subsection‹Test Vector: KeyPair_P224_09›

definition KeyPair_P224_09_d :: nat where
  "KeyPair_P224_09_d = 0x30f4113572e0fb39209aad2698608bc47f8c0d8c1061e954f72bc805"

definition KeyPair_P224_09_Qx :: nat where
  "KeyPair_P224_09_Qx = 0xb73221c141720c43d50183e975b0b28825e2f7a792a59fe984b9ef44"

definition KeyPair_P224_09_Qy :: nat where
  "KeyPair_P224_09_Qy = 0xa694d0f15642b2dd12ca17852a284e04543f5949161085296d932e63"

definition KeyPair_P224_09_Q :: "int point" where
  "KeyPair_P224_09_Q = Point (int KeyPair_P224_09_Qx) (int KeyPair_P224_09_Qy)"

lemma KeyPair_P224_09_check: 
  "SEC1_P224_ECkeyGen KeyPair_P224_09_d = KeyPair_P224_09_Q" 
  by eval

subsection‹Test Vector: KeyPair_P224_10›

definition KeyPair_P224_10_d :: nat where
  "KeyPair_P224_10_d = 0x9c1005987043a7b8f41efcccf3ac3e413ef31938097d2ac8bee9ec8e"

definition KeyPair_P224_10_Qx :: nat where
  "KeyPair_P224_10_Qx = 0x3c47a265cdb7dd2ec7c3fb97d95fccd5ae0ee55e9efb3bef7bcdfc7a"

definition KeyPair_P224_10_Qy :: nat where
  "KeyPair_P224_10_Qy = 0x62c2e871d67133111f05540a7204f9ae0d023ed7738703238ec18d85"

definition KeyPair_P224_10_Q :: "int point" where
  "KeyPair_P224_10_Q = Point (int KeyPair_P224_10_Qx) (int KeyPair_P224_10_Qy)"

lemma KeyPair_P224_10_check: 
  "SEC1_P224_ECkeyGen KeyPair_P224_10_d = KeyPair_P224_10_Q" 
  by eval

section‹NIST Test Vectors for EC Key Pair Generation: Curve P-256›
  
subsection‹Test Vector: KeyPair_P256_01›

definition KeyPair_P256_01_d :: nat where
  "KeyPair_P256_01_d = 0xc9806898a0334916c860748880a541f093b579a9b1f32934d86c363c39800357"

definition KeyPair_P256_01_Qx :: nat where
  "KeyPair_P256_01_Qx = 0xd0720dc691aa80096ba32fed1cb97c2b620690d06de0317b8618d5ce65eb728f"

definition KeyPair_P256_01_Qy :: nat where
  "KeyPair_P256_01_Qy = 0x9681b517b1cda17d0d83d335d9c4a8a9a9b0b1b3c7106d8f3c72bc5093dc275f"

definition KeyPair_P256_01_Q :: "int point" where
  "KeyPair_P256_01_Q = Point (int KeyPair_P256_01_Qx) (int KeyPair_P256_01_Qy)"

lemma KeyPair_P256_01_check: 
  "SEC1_P256_ECkeyGen KeyPair_P256_01_d = KeyPair_P256_01_Q" 
  by eval

subsection‹Test Vector: KeyPair_P256_02›

definition KeyPair_P256_02_d :: nat where
  "KeyPair_P256_02_d = 0x710735c8388f48c684a97bd66751cc5f5a122d6b9a96a2dbe73662f78217446d"

definition KeyPair_P256_02_Qx :: nat where
  "KeyPair_P256_02_Qx = 0xf6836a8add91cb182d8d258dda6680690eb724a66dc3bb60d2322565c39e4ab9"

definition KeyPair_P256_02_Qy :: nat where
  "KeyPair_P256_02_Qy = 0x1f837aa32864870cb8e8d0ac2ff31f824e7beddc4bb7ad72c173ad974b289dc2"

definition KeyPair_P256_02_Q :: "int point" where
  "KeyPair_P256_02_Q = Point (int KeyPair_P256_02_Qx) (int KeyPair_P256_02_Qy)"

lemma KeyPair_P256_02_check: 
  "SEC1_P256_ECkeyGen KeyPair_P256_02_d = KeyPair_P256_02_Q" 
  by eval

subsection‹Test Vector: KeyPair_P256_03›

definition KeyPair_P256_03_d :: nat where
  "KeyPair_P256_03_d = 0x78d5d8b7b3e2c16b3e37e7e63becd8ceff61e2ce618757f514620ada8a11f6e4"

definition KeyPair_P256_03_Qx :: nat where
  "KeyPair_P256_03_Qx = 0x76711126cbb2af4f6a5fe5665dad4c88d27b6cb018879e03e54f779f203a854e"

definition KeyPair_P256_03_Qy :: nat where
  "KeyPair_P256_03_Qy = 0xa26df39960ab5248fd3620fd018398e788bd89a3cea509b352452b69811e6856"

definition KeyPair_P256_03_Q :: "int point" where
  "KeyPair_P256_03_Q = Point (int KeyPair_P256_03_Qx) (int KeyPair_P256_03_Qy)"

lemma KeyPair_P256_03_check: 
  "SEC1_P256_ECkeyGen KeyPair_P256_03_d = KeyPair_P256_03_Q" 
  by eval

subsection‹Test Vector: KeyPair_P256_04›

definition KeyPair_P256_04_d :: nat where
  "KeyPair_P256_04_d = 0x2a61a0703860585fe17420c244e1de5a6ac8c25146b208ef88ad51ae34c8cb8c"

definition KeyPair_P256_04_Qx :: nat where
  "KeyPair_P256_04_Qx = 0xe1aa7196ceeac088aaddeeba037abb18f67e1b55c0a5c4e71ec70ad666fcddc8"

definition KeyPair_P256_04_Qy :: nat where
  "KeyPair_P256_04_Qy = 0xd7d35bdce6dedc5de98a7ecb27a9cd066a08f586a733b59f5a2cdb54f971d5c8"

definition KeyPair_P256_04_Q :: "int point" where
  "KeyPair_P256_04_Q = Point (int KeyPair_P256_04_Qx) (int KeyPair_P256_04_Qy)"

lemma KeyPair_P256_04_check: 
  "SEC1_P256_ECkeyGen KeyPair_P256_04_d = KeyPair_P256_04_Q" 
  by eval

subsection‹Test Vector: KeyPair_P256_05›

definition KeyPair_P256_05_d :: nat where
  "KeyPair_P256_05_d = 0x01b965b45ff386f28c121c077f1d7b2710acc6b0cb58d8662d549391dcf5a883"

definition KeyPair_P256_05_Qx :: nat where
  "KeyPair_P256_05_Qx = 0x1f038c5422e88eec9e88b815e8f6b3e50852333fc423134348fc7d79ef8e8a10"

definition KeyPair_P256_05_Qy :: nat where
  "KeyPair_P256_05_Qy = 0x43a047cb20e94b4ffb361ef68952b004c0700b2962e0c0635a70269bc789b849"

definition KeyPair_P256_05_Q :: "int point" where
  "KeyPair_P256_05_Q = Point (int KeyPair_P256_05_Qx) (int KeyPair_P256_05_Qy)"

lemma KeyPair_P256_05_check: 
  "SEC1_P256_ECkeyGen KeyPair_P256_05_d = KeyPair_P256_05_Q" 
  by eval

subsection‹Test Vector: KeyPair_P256_06›

definition KeyPair_P256_06_d :: nat where
  "KeyPair_P256_06_d = 0xfac92c13d374c53a085376fe4101618e1e181b5a63816a84a0648f3bdc24e519"

definition KeyPair_P256_06_Qx :: nat where
  "KeyPair_P256_06_Qx = 0x7258f2ab96fc84ef6ccb33e308cd392d8b568ea635730ceb4ebd72fa870583b9"

definition KeyPair_P256_06_Qy :: nat where
  "KeyPair_P256_06_Qy = 0x489807ca55bdc29ca5c8fe69b94f227b0345cccdbe89975e75d385cc2f6bb1e2"

definition KeyPair_P256_06_Q :: "int point" where
  "KeyPair_P256_06_Q = Point (int KeyPair_P256_06_Qx) (int KeyPair_P256_06_Qy)"

lemma KeyPair_P256_06_check: 
  "SEC1_P256_ECkeyGen KeyPair_P256_06_d = KeyPair_P256_06_Q" 
  by eval

subsection‹Test Vector: KeyPair_P256_07›

definition KeyPair_P256_07_d :: nat where
  "KeyPair_P256_07_d = 0xf257a192dde44227b3568008ff73bcf599a5c45b32ab523b5b21ca582fef5a0a"

definition KeyPair_P256_07_Qx :: nat where
  "KeyPair_P256_07_Qx = 0xd2e01411817b5512b79bbbe14d606040a4c90deb09e827d25b9f2fc068997872"

definition KeyPair_P256_07_Qy :: nat where
  "KeyPair_P256_07_Qy = 0x503f138f8bab1df2c4507ff663a1fdf7f710e7adb8e7841eaa902703e314e793"

definition KeyPair_P256_07_Q :: "int point" where
  "KeyPair_P256_07_Q = Point (int KeyPair_P256_07_Qx) (int KeyPair_P256_07_Qy)"

lemma KeyPair_P256_07_check: 
  "SEC1_P256_ECkeyGen KeyPair_P256_07_d = KeyPair_P256_07_Q" 
  by eval

subsection‹Test Vector: KeyPair_P256_08›

definition KeyPair_P256_08_d :: nat where
  "KeyPair_P256_08_d = 0xadd67e57c42a3d28708f0235eb86885a4ea68e0d8cfd76eb46134c596522abfd"

definition KeyPair_P256_08_Qx :: nat where
  "KeyPair_P256_08_Qx = 0x55bed2d9c029b7f230bde934c7124ed52b1330856f13cbac65a746f9175f85d7"

definition KeyPair_P256_08_Qy :: nat where
  "KeyPair_P256_08_Qy = 0x32805e311d583b4e007c40668185e85323948e21912b6b0d2cda8557389ae7b0"

definition KeyPair_P256_08_Q :: "int point" where
  "KeyPair_P256_08_Q = Point (int KeyPair_P256_08_Qx) (int KeyPair_P256_08_Qy)"

lemma KeyPair_P256_08_check: 
  "SEC1_P256_ECkeyGen KeyPair_P256_08_d = KeyPair_P256_08_Q" 
  by eval

subsection‹Test Vector: KeyPair_P256_09›

definition KeyPair_P256_09_d :: nat where
  "KeyPair_P256_09_d = 0x4494860fd2c805c5c0d277e58f802cff6d731f76314eb1554142a637a9bc5538"

definition KeyPair_P256_09_Qx :: nat where
  "KeyPair_P256_09_Qx = 0x5190277a0c14d8a3d289292f8a544ce6ea9183200e51aec08440e0c1a463a4e4"

definition KeyPair_P256_09_Qy :: nat where
  "KeyPair_P256_09_Qy = 0xecd98514821bd5aaf3419ab79b71780569470e4fed3da3c1353b28fe137f36eb"

definition KeyPair_P256_09_Q :: "int point" where
  "KeyPair_P256_09_Q = Point (int KeyPair_P256_09_Qx) (int KeyPair_P256_09_Qy)"

lemma KeyPair_P256_09_check: 
  "SEC1_P256_ECkeyGen KeyPair_P256_09_d = KeyPair_P256_09_Q" 
  by eval

subsection‹Test Vector: KeyPair_P256_10›

definition KeyPair_P256_10_d :: nat where
  "KeyPair_P256_10_d = 0xd40b07b1ea7b86d4709ef9dc634c61229feb71abd63dc7fc85ef46711a87b210"

definition KeyPair_P256_10_Qx :: nat where
  "KeyPair_P256_10_Qx = 0xfbcea7c2827e0e8085d7707b23a3728823ea6f4878b24747fb4fd2842d406c73"

definition KeyPair_P256_10_Qy :: nat where
  "KeyPair_P256_10_Qy = 0x2393c85f1f710c5afc115a39ba7e18abe03f19c9d4bb3d47d19468b818efa535"

definition KeyPair_P256_10_Q :: "int point" where
  "KeyPair_P256_10_Q = Point (int KeyPair_P256_10_Qx) (int KeyPair_P256_10_Qy)"

lemma KeyPair_P256_10_check: 
  "SEC1_P256_ECkeyGen KeyPair_P256_10_d = KeyPair_P256_10_Q" 
  by eval

section‹NIST Test Vectors for EC Key Pair Generation: Curve P-384›
  
subsection‹Test Vector: KeyPair_P384_01›

definition KeyPair_P384_01_d :: nat where
  "KeyPair_P384_01_d = 0x5394f7973ea868c52bf3ff8d8ceeb4db90a683653b12485d5f627c3ce5abd8978fc9673d14a71d925747931662493c37"

definition KeyPair_P384_01_Qx :: nat where
  "KeyPair_P384_01_Qx = 0xfd3c84e5689bed270e601b3d80f90d67a9ae451cce890f53e583229ad0e2ee645611fa9936dfa45306ec18066774aa24"

definition KeyPair_P384_01_Qy :: nat where
  "KeyPair_P384_01_Qy = 0xb83ca4126cfc4c4d1d18a4b6c21c7f699d5123dd9c24f66f833846eeb58296196b42ec06425db5b70a4b81b7fcf705a0"

definition KeyPair_P384_01_Q :: "int point" where
  "KeyPair_P384_01_Q = Point (int KeyPair_P384_01_Qx) (int KeyPair_P384_01_Qy)"

lemma KeyPair_P384_01_check: 
  "SEC1_P384_ECkeyGen KeyPair_P384_01_d = KeyPair_P384_01_Q" 
  by eval

subsection‹Test Vector: KeyPair_P384_02›

definition KeyPair_P384_02_d :: nat where
  "KeyPair_P384_02_d = 0x9b90d800abc37df43536e0dc321d43e6aeb5317fcb5118a0e827c8165b1cb05051ef12794b5278a293accbc0b1beb2c2"

definition KeyPair_P384_02_Qx :: nat where
  "KeyPair_P384_02_Qx = 0x732b0f83d303475584d88ed91cc74b367e9ffbfcc2d044d1485417d2731fa4f3b70347388e2308e9e43bdbf952465393"

definition KeyPair_P384_02_Qy :: nat where
  "KeyPair_P384_02_Qy = 0xd8d232a2c995a6ff133893dcfa9b559c11376eb999abf55edd51cc5edb7935500f80f55ca1a542a1b87f6c8c643b83d6"

definition KeyPair_P384_02_Q :: "int point" where
  "KeyPair_P384_02_Q = Point (int KeyPair_P384_02_Qx) (int KeyPair_P384_02_Qy)"

lemma KeyPair_P384_02_check: 
  "SEC1_P384_ECkeyGen KeyPair_P384_02_d = KeyPair_P384_02_Q" 
  by eval

subsection‹Test Vector: KeyPair_P384_03›

definition KeyPair_P384_03_d :: nat where
  "KeyPair_P384_03_d = 0x9fdc866600017ee5be419ef11f13a537a403ed16743d16d43fba9938893fa9771c20b971faa4719744cb40af3b73bf84"

definition KeyPair_P384_03_Qx :: nat where
  "KeyPair_P384_03_Qx = 0x315ec7497da462a6bc8fd16bbe9c84c7d96e8b9a49d6ce72936ef3a55d12c73e43f375a2d49015d8990eff598f3e4621"

definition KeyPair_P384_03_Qy :: nat where
  "KeyPair_P384_03_Qy = 0x8b84927898f8c79e8e5a101329454fd15b6333475e90e14bca1668323f854b10a2e95aeae619b4f68801ba3db1840a36"

definition KeyPair_P384_03_Q :: "int point" where
  "KeyPair_P384_03_Q = Point (int KeyPair_P384_03_Qx) (int KeyPair_P384_03_Qy)"

lemma KeyPair_P384_03_check: 
  "SEC1_P384_ECkeyGen KeyPair_P384_03_d = KeyPair_P384_03_Q" 
  by eval

subsection‹Test Vector: KeyPair_P384_04›

definition KeyPair_P384_04_d :: nat where
  "KeyPair_P384_04_d = 0xdc6dc5ed9ee57eb5e39878f97aeac359f8258b9c381fb1987abb061184013220a76d667fcf3a53088eae8da3f8d9b520"

definition KeyPair_P384_04_Qx :: nat where
  "KeyPair_P384_04_Qx = 0x285d0daaf8f6315a05fde9e6ec2f1d6ade8ab158abd4a40995f34a749421a2fbb0a642031e1be371b07931caa41bc523"

definition KeyPair_P384_04_Qy :: nat where
  "KeyPair_P384_04_Qy = 0x2920820f7eaebbd05ac5920d2ebd3561c12b635a1454476669a2d09b31fca66f883dacc2eb4ac078d203970d49a039f3"

definition KeyPair_P384_04_Q :: "int point" where
  "KeyPair_P384_04_Q = Point (int KeyPair_P384_04_Qx) (int KeyPair_P384_04_Qy)"

lemma KeyPair_P384_04_check: 
  "SEC1_P384_ECkeyGen KeyPair_P384_04_d = KeyPair_P384_04_Q" 
  by eval

subsection‹Test Vector: KeyPair_P384_05›

definition KeyPair_P384_05_d :: nat where
  "KeyPair_P384_05_d = 0x0f4144cf327d6feae06091d4a8c710011ce32c987af1451fccc471bc3f482bc741368e88d04324c93924dcb3b5274b02"

definition KeyPair_P384_05_Qx :: nat where
  "KeyPair_P384_05_Qx = 0xb8880cd358dbc4d8743b6bae8cd1296235a78720bad3c8786fa5db6a632ea82320cc705386810fa9b49d1c9c8164ba19"

definition KeyPair_P384_05_Qy :: nat where
  "KeyPair_P384_05_Qy = 0x03f930af97df0862640f1b9116f3d253c39180a34e3cf18fd54178fcad835ba1c4eae2f02136748bc57cde57e4cedb99"

definition KeyPair_P384_05_Q :: "int point" where
  "KeyPair_P384_05_Q = Point (int KeyPair_P384_05_Qx) (int KeyPair_P384_05_Qy)"

lemma KeyPair_P384_05_check: 
  "SEC1_P384_ECkeyGen KeyPair_P384_05_d = KeyPair_P384_05_Q" 
  by eval

subsection‹Test Vector: KeyPair_P384_06›

definition KeyPair_P384_06_d :: nat where
  "KeyPair_P384_06_d = 0xc5099dab8c2e1914bc37331f288a8d35d55e6163a6553dbe6a4fbd60311478ff9502a0efa3f2ac5ee6705afc443bcd39"

definition KeyPair_P384_06_Qx :: nat where
  "KeyPair_P384_06_Qx = 0x1f6e499ddca491b6e5da6e34eff16847284e816443cb9699f95e1449cc361327fb244f780cd704f55c1a4112fbda1264"

definition KeyPair_P384_06_Qy :: nat where
  "KeyPair_P384_06_Qy = 0x7b60b9f805598ca6a9521e5ccfd8e2c988fc450dc70f26202dd8567ef038f51e75441ae7b97b0debe3a2ed4b0442533a"

definition KeyPair_P384_06_Q :: "int point" where
  "KeyPair_P384_06_Q = Point (int KeyPair_P384_06_Qx) (int KeyPair_P384_06_Qy)"

lemma KeyPair_P384_06_check: 
  "SEC1_P384_ECkeyGen KeyPair_P384_06_d = KeyPair_P384_06_Q" 
  by eval

subsection‹Test Vector: KeyPair_P384_07›

definition KeyPair_P384_07_d :: nat where
  "KeyPair_P384_07_d = 0x86feaa8de26a91e4684c6f06689d4cf46ddc45ba27403a980107c0d31a5b675005fcbb7a07dc28ef9ee8d536252e0624"

definition KeyPair_P384_07_Qx :: nat where
  "KeyPair_P384_07_Qx = 0xc7d4b2b050813accbccf6abd5032e549126a9d31cbc926695506ced640a1e8b7bf27649ca4b255e347e4f111b208d56b"

definition KeyPair_P384_07_Qy :: nat where
  "KeyPair_P384_07_Qy = 0xfe14fbfb45373044431fe35d3e781bcfbb274e57b886a0922a7067de34425312879bfbc9a9df7a833d0e4f8f8180edce"

definition KeyPair_P384_07_Q :: "int point" where
  "KeyPair_P384_07_Q = Point (int KeyPair_P384_07_Qx) (int KeyPair_P384_07_Qy)"

lemma KeyPair_P384_07_check: 
  "SEC1_P384_ECkeyGen KeyPair_P384_07_d = KeyPair_P384_07_Q" 
  by eval

subsection‹Test Vector: KeyPair_P384_08›

definition KeyPair_P384_08_d :: nat where
  "KeyPair_P384_08_d = 0xf0d9c1fb64733c4e91e4129e4c0ca6879e25dafd0b583bc5d1d989664e592e52390494fd65bfff76b2a46fd7db23be47"

definition KeyPair_P384_08_Qx :: nat where
  "KeyPair_P384_08_Qx = 0xca2a552a6151cff2651f62ccd3f859dba9888f9762423e780d83824777e99356c785d5c720e78dc01c3b0f01d76d1627"

definition KeyPair_P384_08_Qy :: nat where
  "KeyPair_P384_08_Qy = 0x0cc8da46a8267f0b4008ff48f1cfaffd03164d9234fe5aa85bf3db31b2cfb89595a3de2883b9ea0c7bd95d6f507b2139"

definition KeyPair_P384_08_Q :: "int point" where
  "KeyPair_P384_08_Q = Point (int KeyPair_P384_08_Qx) (int KeyPair_P384_08_Qy)"

lemma KeyPair_P384_08_check: 
  "SEC1_P384_ECkeyGen KeyPair_P384_08_d = KeyPair_P384_08_Q" 
  by eval

subsection‹Test Vector: KeyPair_P384_09›

definition KeyPair_P384_09_d :: nat where
  "KeyPair_P384_09_d = 0x3512b88c012aee202b5122353c610be0ef2a0604bfec2a6dfee7f1a80395aba6368029c388d462c53c24b5ec25055bb4"

definition KeyPair_P384_09_Qx :: nat where
  "KeyPair_P384_09_Qx = 0xfe79a325e7d410a5df04913bc91ad3a132b928679d4050226f1ad98a8caf473f7ed0b7211639f44c90deb2dd88937103"

definition KeyPair_P384_09_Qy :: nat where
  "KeyPair_P384_09_Qy = 0x69856bf3d2ce4388ca413de5625e050192552b1224234ad7ce3e71f23d164d83dacfac06b31d577a5860119c71c036eb"

definition KeyPair_P384_09_Q :: "int point" where
  "KeyPair_P384_09_Q = Point (int KeyPair_P384_09_Qx) (int KeyPair_P384_09_Qy)"

lemma KeyPair_P384_09_check: 
  "SEC1_P384_ECkeyGen KeyPair_P384_09_d = KeyPair_P384_09_Q" 
  by eval

subsection‹Test Vector: KeyPair_P384_10›

definition KeyPair_P384_10_d :: nat where
  "KeyPair_P384_10_d = 0x3a7b2a6a03c92154f4ef31dd257bd9d3397494da4c93dc033a1c7925a295ce12412797ac995191b665229ad6db881e6e"

definition KeyPair_P384_10_Qx :: nat where
  "KeyPair_P384_10_Qx = 0x9e248419857ca4a03c1f2ba99eddad34ce1904ba486c73e3903adb7fc5e63a6d397a9a1a5e06f1d62d78cb7038a518d3"

definition KeyPair_P384_10_Qy :: nat where
  "KeyPair_P384_10_Qy = 0x86a15f544ecd494b6c8cd5f561832a2d73fb83d0a8a76f10808022b47f2fff1e8730c43684ad17b6bd049c586b4d3cdf"

definition KeyPair_P384_10_Q :: "int point" where
  "KeyPair_P384_10_Q = Point (int KeyPair_P384_10_Qx) (int KeyPair_P384_10_Qy)"

lemma KeyPair_P384_10_check: 
  "SEC1_P384_ECkeyGen KeyPair_P384_10_d = KeyPair_P384_10_Q" 
  by eval

section‹NIST Test Vectors for EC Key Pair Generation: Curve P-521›
  
subsection‹Test Vector: KeyPair_P521_01›

definition KeyPair_P521_01_d :: nat where
  "KeyPair_P521_01_d = 0x0184258ea667ab99d09d4363b3f51384fc0acd2f3b66258ef31203ed30363fcda7661b6a817daaf831415a1f21cb1cda3a74cc1865f2ef40f683c14174ea72803cff"

definition KeyPair_P521_01_Qx :: nat where
  "KeyPair_P521_01_Qx = 0x019ee818048f86ada6db866b7e49a9b535750c3673cb61bbfe5585c2df263860fe4d8aa8f7486aed5ea2a4d733e346eaefa87ac515c78b9a986ee861584926ce4860"

definition KeyPair_P521_01_Qy :: nat where
  "KeyPair_P521_01_Qy = 0x01b6809c89c0aa7fb057a32acbb9ab4d7b06ba39dba8833b9b54424add2956e95fe48b7fbf60c3df5172bf386f2505f1e1bb2893da3b96d4f5ae78f2544881a238f7"

definition KeyPair_P521_01_Q :: "int point" where
  "KeyPair_P521_01_Q = Point (int KeyPair_P521_01_Qx) (int KeyPair_P521_01_Qy)"

lemma KeyPair_P521_01_check: 
  "SEC1_P521_ECkeyGen KeyPair_P521_01_d = KeyPair_P521_01_Q" 
  by eval

subsection‹Test Vector: KeyPair_P521_02›

definition KeyPair_P521_02_d :: nat where
  "KeyPair_P521_02_d = 0x014b967f6651b5e6a482fccc609ab6630b3806fe1f94f4083319b0b50575fb3436a04f508172f7fc396d6e969ca3e8d1c1e9a84d431a48b94f30566dc6808dd1d138"

definition KeyPair_P521_02_Qx :: nat where
  "KeyPair_P521_02_Qx = 0x0145f371040d3d4a24d6d3ceb2681db207b77096ab57606d92981a69ce35a0ac4628c2dc1284e4dd9715cde46f18b59e9fc98fea162ceb6e2c481ecbfad4e19d3abf"

definition KeyPair_P521_02_Qy :: nat where
  "KeyPair_P521_02_Qy = 0x0125eb751ff4fb8bb98e1fb455d2cfb35e3323de5c7280fc9e51729704f4fec51d5a6ce6c1f75dbf710e1f9d3ee9f2a77e7c12c045e729d0e9a281c37f0f07b8cf0c"

definition KeyPair_P521_02_Q :: "int point" where
  "KeyPair_P521_02_Q = Point (int KeyPair_P521_02_Qx) (int KeyPair_P521_02_Qy)"

lemma KeyPair_P521_02_check: 
  "SEC1_P521_ECkeyGen KeyPair_P521_02_d = KeyPair_P521_02_Q" 
  by eval

subsection‹Test Vector: KeyPair_P521_03›

definition KeyPair_P521_03_d :: nat where
  "KeyPair_P521_03_d = 0x7616133442038e27357db450c353bd11fba3bcac8b7b8c3ef76aadb5fe05be1dd57a22d42a5444d00dcd018d389170c54fe781cb21c36020f657d001e1cbb41dd1"

definition KeyPair_P521_03_Qx :: nat where
  "KeyPair_P521_03_Qx = 0xbbecf65446053080cc1cf955938c58eb630c84ecad2756f93b47ebfa9f9bca3fa8343539812608cab2d3a9f8079ab8311a4f269b0a3cd9e0ddd066fc4121d92f0e"

definition KeyPair_P521_03_Qy :: nat where
  "KeyPair_P521_03_Qy = 0x01dd96db411ad67997b10d42c76b8510c8a930dfa9a5927ac274b0c5021798690777b8e77e6ae2648bf513e02f586898e7dae20d71d19838a9f3175f06b057c5f2f4"

definition KeyPair_P521_03_Q :: "int point" where
  "KeyPair_P521_03_Q = Point (int KeyPair_P521_03_Qx) (int KeyPair_P521_03_Qy)"

lemma KeyPair_P521_03_check: 
  "SEC1_P521_ECkeyGen KeyPair_P521_03_d = KeyPair_P521_03_Q" 
  by eval

subsection‹Test Vector: KeyPair_P521_04›

definition KeyPair_P521_04_d :: nat where
  "KeyPair_P521_04_d = 0x013bcc0ed286861d3f5463bcfc0b68a6ec0fcf86291ba41257838b72536ada986e43e05ec4c32c0b29da632dd1ce39efc81c8278f5d18d9cf27f6e75523821a46d99"

definition KeyPair_P521_04_Qx :: nat where
  "KeyPair_P521_04_Qx = 0xa3a165c2bb535d1041d54b749e2f6e6c734a03c09df69c14a5dd2aa57790acc504548885f0bd3a44f8b66bb9c36b3ff257d7d465efb81445d4cc5a5af7f36c679c"

definition KeyPair_P521_04_Qy :: nat where
  "KeyPair_P521_04_Qy = 0x8a5d094e4f2aa18fb877d2649dfd76f9482ac2e049aefbb463f3c9061cfdfaec785df9577a090e45a17330f422fb16a16acccff9ade7b034ec544c7a8aea441c49"

definition KeyPair_P521_04_Q :: "int point" where
  "KeyPair_P521_04_Q = Point (int KeyPair_P521_04_Qx) (int KeyPair_P521_04_Qy)"

lemma KeyPair_P521_04_check: 
  "SEC1_P521_ECkeyGen KeyPair_P521_04_d = KeyPair_P521_04_Q" 
  by eval

subsection‹Test Vector: KeyPair_P521_05›

definition KeyPair_P521_05_d :: nat where
  "KeyPair_P521_05_d = 0x01f79977450ce5887ae2ef7d648ab658c056e57f0a690cf28a4e94f373f2c15eb3c0d3e0d670feca6ff02d5fd03187146eb85e09d72f8cabb1900d0c338a23080c12"

definition KeyPair_P521_05_Qx :: nat where
  "KeyPair_P521_05_Qx = 0x016d9ede24a3950098798766e57c53f2749cd0d3f56ca0904a3711c030965291edd5c6fe0903771768f42340e88e1cd2f161358972775fa53e5b87c3b660acc447e2"

definition KeyPair_P521_05_Qy :: nat where
  "KeyPair_P521_05_Qy = 0x010ca5cfe6df8e069ae1326dd9e18cca75cde7cb24b427a409025f9e12b5098a56a20bb90b1d23b75fad7a54f9e25ff892e1236d1717f1f94e18fa2289f899fe2221"

definition KeyPair_P521_05_Q :: "int point" where
  "KeyPair_P521_05_Q = Point (int KeyPair_P521_05_Qx) (int KeyPair_P521_05_Qy)"

lemma KeyPair_P521_05_check: 
  "SEC1_P521_ECkeyGen KeyPair_P521_05_d = KeyPair_P521_05_Q" 
  by eval

subsection‹Test Vector: KeyPair_P521_06›

definition KeyPair_P521_06_d :: nat where
  "KeyPair_P521_06_d = 0x9a9160d2614937c33284627826be871c26407c84d23e6d23de5f5f48b500b89b0bc07f10c4e0fb99c085d9e9d7149278f76e3fae4abaeeef2495fe3d228ef0f949"

definition KeyPair_P521_06_Qx :: nat where
  "KeyPair_P521_06_Qx = 0x019ed72e6bfc673f2a852acf9d60e2c3b19c50a56c54ac304612b26f83afe1aff4f87dca458e83b6f89ec48f8b1a20931acd3c97c71bf21b5633cf4fd68437db45c1"

definition KeyPair_P521_06_Qy :: nat where
  "KeyPair_P521_06_Qy = 0xd141dc4272ca03a528ad8fdade9ecb3070fb2d4af0bb296abdaed651b5d26573eb4443a4d0d4134ff248d8ed402c93bf6a905cb2792b9cecb4aeb69ed78f410382"

definition KeyPair_P521_06_Q :: "int point" where
  "KeyPair_P521_06_Q = Point (int KeyPair_P521_06_Qx) (int KeyPair_P521_06_Qy)"

lemma KeyPair_P521_06_check: 
  "SEC1_P521_ECkeyGen KeyPair_P521_06_d = KeyPair_P521_06_Q" 
  by eval

subsection‹Test Vector: KeyPair_P521_07›

definition KeyPair_P521_07_d :: nat where
  "KeyPair_P521_07_d = 0x2fdc02492573228ada3fa8a2db68d72e9396a2bfca9a8ebdb5c2955cc894a7493cfae001759368eb8ffc3c29b15365f6484cdd6a44e084f1d3c88dba7aa4f29c3c"

definition KeyPair_P521_07_Qx :: nat where
  "KeyPair_P521_07_Qx = 0x010ba48733fc3e8f54f601f74659bcd43fde4cf8c5a07da341ce68e792f8f70721c23dc6d9b1b401bd3254c8de546e9367f10aee947b1dd295e6d822524546ddc195"

definition KeyPair_P521_07_Qy :: nat where
  "KeyPair_P521_07_Qy = 0x01b2c0ea5c4171cdc069fc6c69e18636cfa404f487a143b3981a1f212969cdbd6601a84302867f8a4a4730fdcd0f994c226f7c02c5e664b79c34b7e5d071423ff528"

definition KeyPair_P521_07_Q :: "int point" where
  "KeyPair_P521_07_Q = Point (int KeyPair_P521_07_Qx) (int KeyPair_P521_07_Qy)"

lemma KeyPair_P521_07_check: 
  "SEC1_P521_ECkeyGen KeyPair_P521_07_d = KeyPair_P521_07_Q" 
  by eval

subsection‹Test Vector: KeyPair_P521_08›

definition KeyPair_P521_08_d :: nat where
  "KeyPair_P521_08_d = 0x01ad69406c11c66fad5fe2295f0e526622488755ecb18ba12ee51fa879ed47ff5f5b05195a821e8d36489492b5de2009f303e17b9fdf6379dae52c0178a16927ca38"

definition KeyPair_P521_08_Qx :: nat where
  "KeyPair_P521_08_Qx = 0x01f1ca24041ba73812c1124e96454545c45ab903407afce3105108362ed3cb4f7d0d5b1466074c2ef22c7fd1ebc16e74a74a163fbb2f530ef44549dad81e806f24d6"

definition KeyPair_P521_08_Qy :: nat where
  "KeyPair_P521_08_Qy = 0x6b34d6eff12bb76aee9bd7ac590e437735ae77da4a60191e8e01f1ceb8ad7c1eda4d0f84d4ed2dc72de702d351ef8f64b2cdf2a95ef185d3119f276f6ccb3c5a65"

definition KeyPair_P521_08_Q :: "int point" where
  "KeyPair_P521_08_Q = Point (int KeyPair_P521_08_Qx) (int KeyPair_P521_08_Qy)"

lemma KeyPair_P521_08_check: 
  "SEC1_P521_ECkeyGen KeyPair_P521_08_d = KeyPair_P521_08_Q" 
  by eval

subsection‹Test Vector: KeyPair_P521_09›

definition KeyPair_P521_09_d :: nat where
  "KeyPair_P521_09_d = 0x013c41b6514c608a2e4696cfc6bd2ddd36611ca5dbf6f2d2e3e32a1925c5ae4ff591dcaa75c4e8043adcb99d510cb664868bb638a2c52b81bb240a974548a68fce79"

definition KeyPair_P521_09_Qx :: nat where
  "KeyPair_P521_09_Qx = 0xc6d82f16433c71e37f2e9779be4599a3b1dda415f6c338e52df4ca70607a69637b50170f21bbb7f60b9a9c145bb63e6d4f370fcd00bfb60f7a0dc55cc44f65fc90"

definition KeyPair_P521_09_Qy :: nat where
  "KeyPair_P521_09_Qy = 0x0152344d6f2e72deb2c59ff2ae268fb067279a1942ae231734ba980c5457a6a73bbf2b13343ae44a0c8a712572851da4b91065ee0436abe811ae71883c4a2f1b797f"

definition KeyPair_P521_09_Q :: "int point" where
  "KeyPair_P521_09_Q = Point (int KeyPair_P521_09_Qx) (int KeyPair_P521_09_Qy)"

lemma KeyPair_P521_09_check: 
  "SEC1_P521_ECkeyGen KeyPair_P521_09_d = KeyPair_P521_09_Q" 
  by eval

subsection‹Test Vector: KeyPair_P521_10›

definition KeyPair_P521_10_d :: nat where
  "KeyPair_P521_10_d = 0x316e2d06fd00c9c4266ea20bf60cdf867859a6f5ba242de35054cdcf5486e5e344ab1d1bce13e2cc831137320774ec3ab0f6fb554fccec56ada267959794898028"

definition KeyPair_P521_10_Qx :: nat where
  "KeyPair_P521_10_Qx = 0xa183880e61c6e0435e591694e51f63c099fcd5b61e3ddacc4057399afc6a90321424ab0ec1699aeeb9c404616d62c23466132b52583c18d3530116b58ad41452f0"

definition KeyPair_P521_10_Qy :: nat where
  "KeyPair_P521_10_Qy = 0x191e06057e2282b4de6e0741fb37b04f0e6ae172be81267b0db3023e7a116ac5861decd54ba84e15d5fd64d6ca628461b79e120851bed1c74adebe3ddee838a170"

definition KeyPair_P521_10_Q :: "int point" where
  "KeyPair_P521_10_Q = Point (int KeyPair_P521_10_Qx) (int KeyPair_P521_10_Qy)"

lemma KeyPair_P521_10_check: 
  "SEC1_P521_ECkeyGen KeyPair_P521_10_d = KeyPair_P521_10_Q" 
  by eval


section‹NIST Test Vectors for MQV primitive: Curve P-224›
  
subsubsection‹MQV test vector: MQV_P224_TV00›
  
definition MQV_P224_TV00_d1U :: nat where
  "MQV_P224_TV00_d1U = 0x6d8c41f3d0b37c0798edd78efcf356034d44e1567b70af5e28ef727d"

definition MQV_P224_TV00_Q1U :: "int point" where 
  "MQV_P224_TV00_Q1U = Point
    0x806f58d250b67ce4184e30c51790e407c42bb8756389f21114311b9d
    0x30701e57f9e2b70ce3c9e976cd3b558258886148bf0027c0e22ad1e4"

definition MQV_P224_TV00_d2U :: nat where
 "MQV_P224_TV00_d2U = 0x2d5a929d323aa324002b599b06ead8d561b14ba8004ea5e6103991b4"

definition MQV_P224_TV00_Q2U :: "int point" where 
 "MQV_P224_TV00_Q2U = Point
    0x16fcde9b1f20df13cf6a4f431014da8186651d1f02feed59db29e1a1
    0xa61cc6ee60fec95a27698f2ca4c18bf38ca41c82ab43ae45869713fc"

definition MQV_P224_TV00_d1V :: nat where
 "MQV_P224_TV00_d1V = 0xcb041f7bbbeb61de0b3db17372d93f441ca66c9a264282e872477343"

definition MQV_P224_TV00_Q1V :: "int point" where 
 "MQV_P224_TV00_Q1V = Point
    0x5968d2e2950f42414fd75dd1d44527bafe40d026941456ab664a015d
    0xd6a2cb51a58d94e7575f8fff76adb352735f56c4ae322bc59ba98dd7"

definition MQV_P224_TV00_d2V :: nat where
 "MQV_P224_TV00_d2V = 0x525f132594d65b56e1336e46e8fdf51748c16e0678c92ddf8e16a3d2"

definition MQV_P224_TV00_Q2V :: "int point" where 
 "MQV_P224_TV00_Q2V = Point
    0x5ec307ce2d112924ca3058b24b4a2bcd1f52bd188077711ffdeb834e
    0xf6a917dc423612511f910b4d6fff603ccd3bf4b349d090a7ecb21c5c"

definition MQV_P224_TV00_Z :: nat where
  "MQV_P224_TV00_Z = 0x929d77e3b7ba428ac57539cbd0688892259307cb60b5e93886abc756"
  
lemma MQV_P224_TV00_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV00_d1U MQV_P224_TV00_Q1U 
                       MQV_P224_TV00_d2U MQV_P224_TV00_Q2U 
                       MQV_P224_TV00_Q1V MQV_P224_TV00_Q2V 
    = Some MQV_P224_TV00_Z"
  by eval

lemma MQV_P224_TV00_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV00_d1V MQV_P224_TV00_Q1V 
                       MQV_P224_TV00_d2V MQV_P224_TV00_Q2V 
                       MQV_P224_TV00_Q1U MQV_P224_TV00_Q2U 
    = Some MQV_P224_TV00_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV01›
  
definition MQV_P224_TV01_d1U :: nat where
  "MQV_P224_TV01_d1U = 0xf7c4fbe0ab04b005b8d219789923265ce2b7600b31b53142639f7878"

definition MQV_P224_TV01_Q1U :: "int point" where 
  "MQV_P224_TV01_Q1U = Point
    0x4edf993ef751be2808e9dbf2b651fe3eca2b64afda5a8c42f5e1011a
    0xfc95caec7da10aad4dd06196e9e913250f308863656811eda6cca7e7"

definition MQV_P224_TV01_d2U :: nat where
 "MQV_P224_TV01_d2U = 0x9611ef1eabe12ddf72b53ec88a2704bb41774dd0264e50e132539f32"

definition MQV_P224_TV01_Q2U :: "int point" where 
 "MQV_P224_TV01_Q2U = Point
    0x21399475d657b947a412a6bd34db141eb3d9ab7207a53afd560c2d06
    0x36062588fbc7234c8b8dcad32d5196534aae5422ea437e95845d64c6"

definition MQV_P224_TV01_d1V :: nat where
 "MQV_P224_TV01_d1V = 0xa2b84fb7f8fc767d854bdfa63cad4cdb5167846d3d37007134e36824"

definition MQV_P224_TV01_Q1V :: "int point" where 
 "MQV_P224_TV01_Q1V = Point
    0xbecfca6b2d76074fdd1e480924ab6b69afc27fce91e8fa4becef2652
    0x463f00d8de39c8ce436ea64177fe698d82c73bea1f8bbee955b62cfd"

definition MQV_P224_TV01_d2V :: nat where
 "MQV_P224_TV01_d2V = 0xb01986de4bd3c82e5c5f8900c85bc89d2396804e5045adebc481cbc5"

definition MQV_P224_TV01_Q2V :: "int point" where 
 "MQV_P224_TV01_Q2V = Point
    0x08134eec2643b0fdb4e73f7910ae058af28e9e10f0534887d47864f7
    0x2ab901a1d6745583cdfbea2bfe44647687c6f925eb286e42f15fb538"

definition MQV_P224_TV01_Z :: nat where
  "MQV_P224_TV01_Z = 0x0dcae61d73a48bf65cffaf41579c1b9fc5eaeb1ae141d2ad1e5f5fc9"
  
lemma MQV_P224_TV01_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV01_d1U MQV_P224_TV01_Q1U 
                       MQV_P224_TV01_d2U MQV_P224_TV01_Q2U 
                       MQV_P224_TV01_Q1V MQV_P224_TV01_Q2V 
    = Some MQV_P224_TV01_Z"
  by eval

lemma MQV_P224_TV01_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV01_d1V MQV_P224_TV01_Q1V 
                       MQV_P224_TV01_d2V MQV_P224_TV01_Q2V 
                       MQV_P224_TV01_Q1U MQV_P224_TV01_Q2U 
    = Some MQV_P224_TV01_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV02›
  
definition MQV_P224_TV02_d1U :: nat where
  "MQV_P224_TV02_d1U = 0x353ad62196b1e0af483366d3ea35dc1ed507a4538e4e07abde795e4d"

definition MQV_P224_TV02_Q1U :: "int point" where 
  "MQV_P224_TV02_Q1U = Point
    0xf258e262f5faf3be1c79de7d1b64fe317c4b3d14ad166cfa67e1102a
    0xd64e3a4306663bfc48b417a7ea0c72f7a64d619a1a3164c80f7cd4f8"

definition MQV_P224_TV02_d2U :: nat where
 "MQV_P224_TV02_d2U = 0x862705b6c0d0bbe4e05d9ff7ad56a22f1629d742d57b65c7cf008d0f"

definition MQV_P224_TV02_Q2U :: "int point" where 
 "MQV_P224_TV02_Q2U = Point
    0x62037b76de245213fe349facf6910172eaf2935e5f7831c972fc925a
    0xc5389c72b52eeed09cb5137c4a2860fabcf1c58312e595ba24b9f0ee"

definition MQV_P224_TV02_d1V :: nat where
 "MQV_P224_TV02_d1V = 0x84cb564445d425945caae39d85a0d6f89e7ae00267adedc90dbf3f08"

definition MQV_P224_TV02_Q1V :: "int point" where 
 "MQV_P224_TV02_Q1V = Point
    0xfc1334f4c2d57116770e4c21f4c9e633a31de526b415635407524c47
    0x03cf570dbdb8cad7be7f9e082fe857cfd1bc118fbd25d6da2184231b"

definition MQV_P224_TV02_d2V :: nat where
 "MQV_P224_TV02_d2V = 0x0d078cc4eec5e21d7bf01bf54143d47aece05236a6f1ca4701b8e906"

definition MQV_P224_TV02_Q2V :: "int point" where 
 "MQV_P224_TV02_Q2V = Point
    0xa8c9a4de9dc18c0db46dc7ae1534fdbf50f33205cde8ae4d0141fb04
    0x2d2816bc2501cc07b80a0f944f6f57014144a626a7e303bc67f52323"

definition MQV_P224_TV02_Z :: nat where
  "MQV_P224_TV02_Z = 0x028e8cd82fea565f58140f389688d799dfae7a625d858993b4b95cc8"
  
lemma MQV_P224_TV02_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV02_d1U MQV_P224_TV02_Q1U 
                       MQV_P224_TV02_d2U MQV_P224_TV02_Q2U 
                       MQV_P224_TV02_Q1V MQV_P224_TV02_Q2V 
    = Some MQV_P224_TV02_Z"
  by eval

lemma MQV_P224_TV02_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV02_d1V MQV_P224_TV02_Q1V 
                       MQV_P224_TV02_d2V MQV_P224_TV02_Q2V 
                       MQV_P224_TV02_Q1U MQV_P224_TV02_Q2U 
    = Some MQV_P224_TV02_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV03›
  
definition MQV_P224_TV03_d1U :: nat where
  "MQV_P224_TV03_d1U = 0xe2922a3a3c3dbaaa2ad35f804df99e6a7c6f2397c6de990693d7defe"

definition MQV_P224_TV03_Q1U :: "int point" where 
  "MQV_P224_TV03_Q1U = Point
    0x2416783f7b95c103c69252eb1fb2ce1524d1b4aedc9f1fb66703918b
    0xf119ff6f2915faad719bdaa042777a4cec3be0a40599b753aaa3e30e"

definition MQV_P224_TV03_d2U :: nat where
 "MQV_P224_TV03_d2U = 0xac3c9973cfa6f21537356c6805b9567d0624db6799db68fc40f05380"

definition MQV_P224_TV03_Q2U :: "int point" where 
 "MQV_P224_TV03_Q2U = Point
    0xe34f0875849b05c6db9bb01bbd22e37079fd29adabe5cee15d39e783
    0xe7b4eb9c65b1c5f06de66dd1fbe979eec71d50cbe7b6c685b3674742"

definition MQV_P224_TV03_d1V :: nat where
 "MQV_P224_TV03_d1V = 0x8f5f0145e0d81bcfbbffd79f09194852d913ea04d905f418c9011f5b"

definition MQV_P224_TV03_Q1V :: "int point" where 
 "MQV_P224_TV03_Q1V = Point
    0x6cf9633cc06a7fe7e8663493585463a11e3b792ac25931615f478748
    0xe8d094b51694dfcaeb369829d898dfc1b6ad89119858b81164743de1"

definition MQV_P224_TV03_d2V :: nat where
 "MQV_P224_TV03_d2V = 0xd288d05af9c5259e11a6eb8afad1c5790f303c0e1d47cc9e36f708ba"

definition MQV_P224_TV03_Q2V :: "int point" where 
 "MQV_P224_TV03_Q2V = Point
    0xeeb9dfcc259e4a41950c94af7834295d9d1ceb2805ce57b29900d494
    0x06d5cc6926aef07b3bcabff0cbd9b2232070fd136cbdd74147cdd3ee"

definition MQV_P224_TV03_Z :: nat where
  "MQV_P224_TV03_Z = 0xa82991a75d8099b5a986452d579e89eb093828545cda428d294972db"
  
lemma MQV_P224_TV03_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV03_d1U MQV_P224_TV03_Q1U 
                       MQV_P224_TV03_d2U MQV_P224_TV03_Q2U 
                       MQV_P224_TV03_Q1V MQV_P224_TV03_Q2V 
    = Some MQV_P224_TV03_Z"
  by eval

lemma MQV_P224_TV03_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV03_d1V MQV_P224_TV03_Q1V 
                       MQV_P224_TV03_d2V MQV_P224_TV03_Q2V 
                       MQV_P224_TV03_Q1U MQV_P224_TV03_Q2U 
    = Some MQV_P224_TV03_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV04›
  
definition MQV_P224_TV04_d1U :: nat where
  "MQV_P224_TV04_d1U = 0x226dbc1c306d6e0fa970fea19b6927c54808865b149f4299e539adb5"

definition MQV_P224_TV04_Q1U :: "int point" where 
  "MQV_P224_TV04_Q1U = Point
    0x9da3c7ba9ac68b7dcf1ff04bcd82940a18862267f845c1d2fb93659f
    0x13c8baf1dd923c6979b6e2ac9ae5296fbcf4837ba51b0837e9b168aa"

definition MQV_P224_TV04_d2U :: nat where
 "MQV_P224_TV04_d2U = 0xc9b5ae40cdd00b6533cc427731d72a182b808c293ede554143edb78a"

definition MQV_P224_TV04_Q2U :: "int point" where 
 "MQV_P224_TV04_Q2U = Point
    0x40c2338e9096159da86e33ba8ac76c61aab0a850b7832c89576f11e6
    0x41bd61d03811938baf3530d11bcb75690e4e5a5c66f0c72546819b92"

definition MQV_P224_TV04_d1V :: nat where
 "MQV_P224_TV04_d1V = 0xd941cfa2a826c56681b25d7ff60d0c06e604e0e8318bbae088603898"

definition MQV_P224_TV04_Q1V :: "int point" where 
 "MQV_P224_TV04_Q1V = Point
    0xe4b8f2a6b3b4896844e4b0968263c7bc13ccb636228be141bd28a0ac
    0x764123452f1b015a7c1af10852e5b5988e70495480044742612ec06e"

definition MQV_P224_TV04_d2V :: nat where
 "MQV_P224_TV04_d2V = 0x6b06f55264738d0232b3b0c0b116d00bce33de97aebc8493f435aed5"

definition MQV_P224_TV04_Q2V :: "int point" where 
 "MQV_P224_TV04_Q2V = Point
    0x08deef5c44b76fa5a6613b961953cf0bca5c9eb1d3c1f75ac0460770
    0x53e1ff5509cee84e6b73a922aaad1751a14c2e9b4d6a260e0cf4c58d"

definition MQV_P224_TV04_Z :: nat where
  "MQV_P224_TV04_Z = 0x0dbaf659ea87c4325b2df184206536f93e436056d66e6319643bf922"
  
lemma MQV_P224_TV04_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV04_d1U MQV_P224_TV04_Q1U 
                       MQV_P224_TV04_d2U MQV_P224_TV04_Q2U 
                       MQV_P224_TV04_Q1V MQV_P224_TV04_Q2V 
    = Some MQV_P224_TV04_Z"
  by eval

lemma MQV_P224_TV04_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV04_d1V MQV_P224_TV04_Q1V 
                       MQV_P224_TV04_d2V MQV_P224_TV04_Q2V 
                       MQV_P224_TV04_Q1U MQV_P224_TV04_Q2U 
    = Some MQV_P224_TV04_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV05›
  
definition MQV_P224_TV05_d1U :: nat where
  "MQV_P224_TV05_d1U = 0xd6af0e44ae2cfac67adb071b2dabc8abdf00c390408742625e7b4c04"

definition MQV_P224_TV05_Q1U :: "int point" where 
  "MQV_P224_TV05_Q1U = Point
    0x5708aaa920be3ba67c2cca3e5c198da6faac8efb6765ac5108918843
    0xf245becd2d03c1fe46d174fef8dc3b827e099a6c7f5362a493bc3c1b"

definition MQV_P224_TV05_d2U :: nat where
 "MQV_P224_TV05_d2U = 0x793cc76dbdb9f1fbbc78bf089e45ae6924dc31b02f441a30f77de16f"

definition MQV_P224_TV05_Q2U :: "int point" where 
 "MQV_P224_TV05_Q2U = Point
    0x42c41e7ad7fd685fdee881b0b6f6550ce3e5fb5c50f8ae981d83aa4b
    0xed54ff26fa9087b9c2402930c2fb6d452440190f1d3b3d441f721b2b"

definition MQV_P224_TV05_d1V :: nat where
 "MQV_P224_TV05_d1V = 0x82f725a7aeb06edcd96af1e0625ad0afc30a1143957f44516c63bdd9"

definition MQV_P224_TV05_Q1V :: "int point" where 
 "MQV_P224_TV05_Q1V = Point
    0xe7abd97d6a216563638413aeecb3fac2acc38603769498879d878314
    0xa0c5c3b98019c0bf8ff7aaadb3b7c6965fe5ae2419419a49bd224710"

definition MQV_P224_TV05_d2V :: nat where
 "MQV_P224_TV05_d2V = 0xc4480d254182ea2603faa8fbe7ab6520cb6fdda334cc6cde4acf9645"

definition MQV_P224_TV05_Q2V :: "int point" where 
 "MQV_P224_TV05_Q2V = Point
    0x3bcd74b8c4b6286b324c6e9039fdca2b36e888fe7c50c966482b288a
    0x997087b647ae361c3ee12854696cbf0b090174079268e8da09940102"

definition MQV_P224_TV05_Z :: nat where
  "MQV_P224_TV05_Z = 0x02fbcf5bc59e4198ee5964a7af1e5afbfe7fe77e336bf6472cda3452"
  
lemma MQV_P224_TV05_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV05_d1U MQV_P224_TV05_Q1U 
                       MQV_P224_TV05_d2U MQV_P224_TV05_Q2U 
                       MQV_P224_TV05_Q1V MQV_P224_TV05_Q2V 
    = Some MQV_P224_TV05_Z"
  by eval

lemma MQV_P224_TV05_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV05_d1V MQV_P224_TV05_Q1V 
                       MQV_P224_TV05_d2V MQV_P224_TV05_Q2V 
                       MQV_P224_TV05_Q1U MQV_P224_TV05_Q2U 
    = Some MQV_P224_TV05_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV06›
  
definition MQV_P224_TV06_d1U :: nat where
  "MQV_P224_TV06_d1U = 0x4443bddcee4270c830d86795c94e5baea973610a8556c06e0338da54"

definition MQV_P224_TV06_Q1U :: "int point" where 
  "MQV_P224_TV06_Q1U = Point
    0x778d03226e9a89f51ff6308e66a5699ce8f2823d62f4780dd9c92a80
    0x33d448309d90021baeae4e9047e1af31170d7378cc2caa99a4ab2c29"

definition MQV_P224_TV06_d2U :: nat where
 "MQV_P224_TV06_d2U = 0x51357789d9aa830b38092518c8dac85b32aa1fa65f8c35c807f8b177"

definition MQV_P224_TV06_Q2U :: "int point" where 
 "MQV_P224_TV06_Q2U = Point
    0x32f41a6160f5029d74c56e22ed8b99fbe944738f6521fac7512ae49b
    0x7ed2d659c06f4863ff0c530a7836f6c195183acc6fbe7c983777e889"

definition MQV_P224_TV06_d1V :: nat where
 "MQV_P224_TV06_d1V = 0xb3a8b3c1fa41e997bdbf7659dc0c6b676baeb691c3039761f82b912d"

definition MQV_P224_TV06_Q1V :: "int point" where 
 "MQV_P224_TV06_Q1V = Point
    0x15699f707aec413a446fc8633ba05b088b14c71d87129fdb6942aeb0
    0x55aeddfe66bc3615dda82ca1c102e057a2a29aeb5ccde9dc47d4054d"

definition MQV_P224_TV06_d2V :: nat where
 "MQV_P224_TV06_d2V = 0x807d2911472267ee14d63eeb711f6c57153f89f60010f9d3d757f191"

definition MQV_P224_TV06_Q2V :: "int point" where 
 "MQV_P224_TV06_Q2V = Point
    0x33ed6bf7d987cabe4adfd84f18194d94acab2453fe1feedd7b3cafe4
    0x0633a689099932545ccc8b9556766dc3495c56fc3c16510592f6de15"

definition MQV_P224_TV06_Z :: nat where
  "MQV_P224_TV06_Z = 0x048e1c04873f29bdf2c9aab021f9057dda993062a79755ee36b3b45b"
  
lemma MQV_P224_TV06_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV06_d1U MQV_P224_TV06_Q1U 
                       MQV_P224_TV06_d2U MQV_P224_TV06_Q2U 
                       MQV_P224_TV06_Q1V MQV_P224_TV06_Q2V 
    = Some MQV_P224_TV06_Z"
  by eval

lemma MQV_P224_TV06_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV06_d1V MQV_P224_TV06_Q1V 
                       MQV_P224_TV06_d2V MQV_P224_TV06_Q2V 
                       MQV_P224_TV06_Q1U MQV_P224_TV06_Q2U 
    = Some MQV_P224_TV06_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV07›
  
definition MQV_P224_TV07_d1U :: nat where
  "MQV_P224_TV07_d1U = 0xd2416d53247b498e33ec27ccfe87cbf5fe73481dff65094581d73474"

definition MQV_P224_TV07_Q1U :: "int point" where 
  "MQV_P224_TV07_Q1U = Point
    0xd34bbec872b3347a2adedbeab7b897929c72b2031d59f79770f476bf
    0xb6da20b1269bc57e3c1006fb555d4cd3d224125b0f18cbdbb0e35efd"

definition MQV_P224_TV07_d2U :: nat where
 "MQV_P224_TV07_d2U = 0x3b16ab616fe2fcd0ed213844b5e81fc6db8cc0cc454397ce641b6378"

definition MQV_P224_TV07_Q2U :: "int point" where 
 "MQV_P224_TV07_Q2U = Point
    0xcaf1d3a30021a43c4657cb307a06bbfc3cc6ac375e927da76ea02a42
    0x02f8a37fd6ba46a80615c544008fdbe9357c1c431aed5679ff9317eb"

definition MQV_P224_TV07_d1V :: nat where
 "MQV_P224_TV07_d1V = 0x35edf3e22749784da6430dff37bea247d062f6b7ab2d9d8006288e3f"

definition MQV_P224_TV07_Q1V :: "int point" where 
 "MQV_P224_TV07_Q1V = Point
    0xda46796bcec27d85c769a64683ba49829ff7627cdb0aa77c996516ba
    0x20422426918ffd3fe2c5cbbd07c5d444ae06a56cb0963348045115ee"

definition MQV_P224_TV07_d2V :: nat where
 "MQV_P224_TV07_d2V = 0x4370398493e96d0ea70947b240c7657ad26d9c17038726dad093ef0d"

definition MQV_P224_TV07_Q2V :: "int point" where 
 "MQV_P224_TV07_Q2V = Point
    0x6891a8f6a96ef66ce20a58e4553d373fb56ac48dd4a1b5b737920ee3
    0x56a75d18ec1868c91c393f8a568b405667738333d082cb03b5d0f428"

definition MQV_P224_TV07_Z :: nat where
  "MQV_P224_TV07_Z = 0xd6ab785ddd7e195cde74310b03b591a4391c8c7888ab46246a315807"
  
lemma MQV_P224_TV07_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV07_d1U MQV_P224_TV07_Q1U 
                       MQV_P224_TV07_d2U MQV_P224_TV07_Q2U 
                       MQV_P224_TV07_Q1V MQV_P224_TV07_Q2V 
    = Some MQV_P224_TV07_Z"
  by eval

lemma MQV_P224_TV07_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV07_d1V MQV_P224_TV07_Q1V 
                       MQV_P224_TV07_d2V MQV_P224_TV07_Q2V 
                       MQV_P224_TV07_Q1U MQV_P224_TV07_Q2U 
    = Some MQV_P224_TV07_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV08›
  
definition MQV_P224_TV08_d1U :: nat where
  "MQV_P224_TV08_d1U = 0x40e26316602cb9d8881bca47c229d72fb4679bca2663f55b2bec105c"

definition MQV_P224_TV08_Q1U :: "int point" where 
  "MQV_P224_TV08_Q1U = Point
    0xfc7e7ce2dc6e4fffd4e12988ad239755cf0b53e7af331b39a6f0963b
    0x6a3c122c6d2027aa0be361b56e134fbaa8027a1281779a4021e29125"

definition MQV_P224_TV08_d2U :: nat where
 "MQV_P224_TV08_d2U = 0xc077eedfca8db711de31148da0450cc42f9cdef365cf2347d3688bf7"

definition MQV_P224_TV08_Q2U :: "int point" where 
 "MQV_P224_TV08_Q2U = Point
    0xf0ccdfb5cd5ea928a6092beb8f4c0e1a78744f5a2d11b656975847aa
    0x25b29c1f7a6b320b69ddab6ff7760fc3367d32b60638937823aa7522"

definition MQV_P224_TV08_d1V :: nat where
 "MQV_P224_TV08_d1V = 0x1e9e5e5b18a28cdd6ca011eaef6567d08bb7222d982e523bc8c94e09"

definition MQV_P224_TV08_Q1V :: "int point" where 
 "MQV_P224_TV08_Q1V = Point
    0xb9438b01fc9484b283291dfc1c743a78e32b99378017e3d9faaaa485
    0x53e522de5b78fafa16f5abf7f077a49af6898a6c7e578e48c0ed937a"

definition MQV_P224_TV08_d2V :: nat where
 "MQV_P224_TV08_d2V = 0x1c4f57f29463578a6a73c7d3590944c72472ab8427b167d8170aa999"

definition MQV_P224_TV08_Q2V :: "int point" where 
 "MQV_P224_TV08_Q2V = Point
    0x1197e615bab9f53a9717c019be5fdaa5767dae0fbf2c7915c96933ee
    0xe5bdb744894f6f467656a26a9eff2b42a9d6c72d28669a9230941ff5"

definition MQV_P224_TV08_Z :: nat where
  "MQV_P224_TV08_Z = 0x82a84cba89bb3a171992a25b4f9f85a66d2de38259905f65d6729f86"
  
lemma MQV_P224_TV08_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV08_d1U MQV_P224_TV08_Q1U 
                       MQV_P224_TV08_d2U MQV_P224_TV08_Q2U 
                       MQV_P224_TV08_Q1V MQV_P224_TV08_Q2V 
    = Some MQV_P224_TV08_Z"
  by eval

lemma MQV_P224_TV08_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV08_d1V MQV_P224_TV08_Q1V 
                       MQV_P224_TV08_d2V MQV_P224_TV08_Q2V 
                       MQV_P224_TV08_Q1U MQV_P224_TV08_Q2U 
    = Some MQV_P224_TV08_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV09›
  
definition MQV_P224_TV09_d1U :: nat where
  "MQV_P224_TV09_d1U = 0xb3f3aee4ca8a15b9b13a2fd63762ecc224f2c19a633b168fbde532b2"

definition MQV_P224_TV09_Q1U :: "int point" where 
  "MQV_P224_TV09_Q1U = Point
    0x4a40227f164be2e91c1e2195bd5d99c3c6f5443c0efbab4efc4db681
    0x3b74c18fa4da5459e85ad3421264b3273656167a6db616995abb5608"

definition MQV_P224_TV09_d2U :: nat where
 "MQV_P224_TV09_d2U = 0xf39b73e06d8e7be32d298a5ca7d9fec68df8c9b9f7f296095912b4da"

definition MQV_P224_TV09_Q2U :: "int point" where 
 "MQV_P224_TV09_Q2U = Point
    0xaaaf894084b5e34f33f557e83f962d101e0cfec386cd46f52dacef85
    0x79939a5471d95430fb6f15f8d7d2d89b9c048167c359b43d934dabc1"

definition MQV_P224_TV09_d1V :: nat where
 "MQV_P224_TV09_d1V = 0xa49c077bd8eca4c77398666c38fd291546dee8140aeab97b374b266b"

definition MQV_P224_TV09_Q1V :: "int point" where 
 "MQV_P224_TV09_Q1V = Point
    0x1cf8cd0122c74d17d4cf98725cf88dcfbeb6c8af667b2b184e5cf847
    0x2f2e508a658bfa4202bea2add4e5041651e54996dcb7870f7f03038d"

definition MQV_P224_TV09_d2V :: nat where
 "MQV_P224_TV09_d2V = 0x020a1b140861c3a43ca3e4f34e45b3d6c4eeb5d114a35118d0214941"

definition MQV_P224_TV09_Q2V :: "int point" where 
 "MQV_P224_TV09_Q2V = Point
    0xd24ab58b79cca8f768b44d388a597db5946571d0c49db5fb6828d41c
    0x21a921c19733731cb64950a03e1a96a7f3dd90737a2f120f28437619"

definition MQV_P224_TV09_Z :: nat where
  "MQV_P224_TV09_Z = 0x14ef539913f288c4c12f6c3e855f92ba8b3c49d7ba6a1f1ac4e83b4a"
  
lemma MQV_P224_TV09_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV09_d1U MQV_P224_TV09_Q1U 
                       MQV_P224_TV09_d2U MQV_P224_TV09_Q2U 
                       MQV_P224_TV09_Q1V MQV_P224_TV09_Q2V 
    = Some MQV_P224_TV09_Z"
  by eval

lemma MQV_P224_TV09_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV09_d1V MQV_P224_TV09_Q1V 
                       MQV_P224_TV09_d2V MQV_P224_TV09_Q2V 
                       MQV_P224_TV09_Q1U MQV_P224_TV09_Q2U 
    = Some MQV_P224_TV09_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV10›
  
definition MQV_P224_TV10_d1U :: nat where
  "MQV_P224_TV10_d1U = 0x7c9b0bef261c7d1bff907f308003077c2e9bff4cc9b95dc87a114a19"

definition MQV_P224_TV10_Q1U :: "int point" where 
  "MQV_P224_TV10_Q1U = Point
    0xcf4c1cb790ed7ec6f867869711b031d1ab3309d13cb4125d3a692b10
    0xee5bcb751de2a43e0c87aa62e9b6df340659e95da67e17b546bb7572"

definition MQV_P224_TV10_d2U :: nat where
 "MQV_P224_TV10_d2U = 0x58835e0377e8067b4a8d5b30f5fe20eaf7977c7f10e618183813add9"

definition MQV_P224_TV10_Q2U :: "int point" where 
 "MQV_P224_TV10_Q2U = Point
    0xa6a6e87b9311cb82fd5ac61a00f64fa799382ca0a47dc63881af587d
    0xe3e47be7c5d0b39868e4b226f6a85ef831b452f9868c7bbe60c5dff5"

definition MQV_P224_TV10_d1V :: nat where
 "MQV_P224_TV10_d1V = 0x38753b0d6df429f36b3fd0a75c6027cb88bfce8843676a617769ecbf"

definition MQV_P224_TV10_Q1V :: "int point" where 
 "MQV_P224_TV10_Q1V = Point
    0xfd7482877034bd8e399db5568d9bcddcda7aabab85dae0a830017659
    0xc7a5216bc479bf39e40c6b5c70bade2fb375f5034a2288ebaacc1af3"

definition MQV_P224_TV10_d2V :: nat where
 "MQV_P224_TV10_d2V = 0xc646ddd76e5feb140f3dadc6d6cd45edd0091dc0f5d88ba382d27f51"

definition MQV_P224_TV10_Q2V :: "int point" where 
 "MQV_P224_TV10_Q2V = Point
    0x5f1f9a5149182842b79d11782ec23da5fec0e8ba48da4e6c2dd2290c
    0xf58886f5492b664dd4ae9c3540d1eeeb5adeb32cdb5f926f10cc85e2"

definition MQV_P224_TV10_Z :: nat where
  "MQV_P224_TV10_Z = 0xdcb931b2722b22590f9ee8fd67c5a793ae1fa1e37766ef58bcbc8bf8"
  
lemma MQV_P224_TV10_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV10_d1U MQV_P224_TV10_Q1U 
                       MQV_P224_TV10_d2U MQV_P224_TV10_Q2U 
                       MQV_P224_TV10_Q1V MQV_P224_TV10_Q2V 
    = Some MQV_P224_TV10_Z"
  by eval

lemma MQV_P224_TV10_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV10_d1V MQV_P224_TV10_Q1V 
                       MQV_P224_TV10_d2V MQV_P224_TV10_Q2V 
                       MQV_P224_TV10_Q1U MQV_P224_TV10_Q2U 
    = Some MQV_P224_TV10_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV11›
  
definition MQV_P224_TV11_d1U :: nat where
  "MQV_P224_TV11_d1U = 0x6279718e596c30be5c8b543587f05a950685e27dea7861ae932a8420"

definition MQV_P224_TV11_Q1U :: "int point" where 
  "MQV_P224_TV11_Q1U = Point
    0x25869f043a87f5cdbd37b448cedd8b99e1a487e71947badeacaccf94
    0x306e391e2106bfc55dead6ab42f059bf3e7a9d3ae139f09729a9e471"

definition MQV_P224_TV11_d2U :: nat where
 "MQV_P224_TV11_d2U = 0x2c30f863aa73f674e67a037d1e86f28865585a01e7f96b8965ca251d"

definition MQV_P224_TV11_Q2U :: "int point" where 
 "MQV_P224_TV11_Q2U = Point
    0x8d544c9dbda5575b1c9c994b5e761d68ee6168c0b2e4b54cb09b953b
    0x33bc68e29b01b2d2bc9fc911ebd5f251ea7cfb3c4961afb7df3f134e"

definition MQV_P224_TV11_d1V :: nat where
 "MQV_P224_TV11_d1V = 0x67a74346c217e2e80f8af968cf142fe96b7b622af4cca8771cf76129"

definition MQV_P224_TV11_Q1V :: "int point" where 
 "MQV_P224_TV11_Q1V = Point
    0xef06b9a41a52407769bb92fcb579eca0697e697a3141dc6285501724
    0xbb404e11c57337161bbdcd3adf4c100b1dda79788ccfe46022d75176"

definition MQV_P224_TV11_d2V :: nat where
 "MQV_P224_TV11_d2V = 0x74d1c6d9e33eaa041801719ad163796d6de037c1cd10259376579d0e"

definition MQV_P224_TV11_Q2V :: "int point" where 
 "MQV_P224_TV11_Q2V = Point
    0x6ff7c29ed5d729be5bd9478c63e00070792a60ef24cabd32dd18afc5
    0xc970c3628c8a34470cc6a0b4f4c63af24ea2d42f85a55c2a2e4f9a43"

definition MQV_P224_TV11_Z :: nat where
  "MQV_P224_TV11_Z = 0x0d99b994791d9f58a5cd7c25f2db6e0945b602023b063851b6c7677f"
  
lemma MQV_P224_TV11_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV11_d1U MQV_P224_TV11_Q1U 
                       MQV_P224_TV11_d2U MQV_P224_TV11_Q2U 
                       MQV_P224_TV11_Q1V MQV_P224_TV11_Q2V 
    = Some MQV_P224_TV11_Z"
  by eval

lemma MQV_P224_TV11_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV11_d1V MQV_P224_TV11_Q1V 
                       MQV_P224_TV11_d2V MQV_P224_TV11_Q2V 
                       MQV_P224_TV11_Q1U MQV_P224_TV11_Q2U 
    = Some MQV_P224_TV11_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV12›
  
definition MQV_P224_TV12_d1U :: nat where
  "MQV_P224_TV12_d1U = 0x9ce87a17bdcaa457212b819db2d9583e55c5839025bccd84e0166316"

definition MQV_P224_TV12_Q1U :: "int point" where 
  "MQV_P224_TV12_Q1U = Point
    0x13d9fd4a99f21387b7cd39dddc6baba9bdb40dc2cf49a8e5f86994af
    0xa935ce2561b499a33fe57b2b8b01f8f9abe48de23d591ca85c569d2b"

definition MQV_P224_TV12_d2U :: nat where
 "MQV_P224_TV12_d2U = 0xb19e99c3b02673de8f46f13e183429ea26fce49644689d00ca06c127"

definition MQV_P224_TV12_Q2U :: "int point" where 
 "MQV_P224_TV12_Q2U = Point
    0xc6ca72532f355cfd55db272430e1375ff623f2ff5b1af0d3a99c681e
    0x2172a7fdb8f58bcbaa1ad4c84e45c6c9f18411d35af6cbc5be947e89"

definition MQV_P224_TV12_d1V :: nat where
 "MQV_P224_TV12_d1V = 0x9400ab4201238bd0ce5e31791fc9a415b7f943858b8283d5b227d842"

definition MQV_P224_TV12_Q1V :: "int point" where 
 "MQV_P224_TV12_Q1V = Point
    0xa5e6009f4eb8268ae8ed73ec6acf869a839de8d73bd0edcd7f3ae3f6
    0xd541daf65c80fd6ffb62ba6a5c07fc26493a710480f250a00fd56f70"

definition MQV_P224_TV12_d2V :: nat where
 "MQV_P224_TV12_d2V = 0x0994edf688dbc7310a5d51d89ffd9232d48aa249aa94c97a3ac32185"

definition MQV_P224_TV12_Q2V :: "int point" where 
 "MQV_P224_TV12_Q2V = Point
    0xfbc7f7448175dffeea21562a7f56208bde45df51bcbdf60719ea6a73
    0xf9be7bf69ab9c5e34e9b415983306b0a660ac1362a55b1210a570ea9"

definition MQV_P224_TV12_Z :: nat where
  "MQV_P224_TV12_Z = 0x8711577bcff99d4fcc50a6b01b25b8383eba0ef71ef6986fa268a2b8"
  
lemma MQV_P224_TV12_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV12_d1U MQV_P224_TV12_Q1U 
                       MQV_P224_TV12_d2U MQV_P224_TV12_Q2U 
                       MQV_P224_TV12_Q1V MQV_P224_TV12_Q2V 
    = Some MQV_P224_TV12_Z"
  by eval

lemma MQV_P224_TV12_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV12_d1V MQV_P224_TV12_Q1V 
                       MQV_P224_TV12_d2V MQV_P224_TV12_Q2V 
                       MQV_P224_TV12_Q1U MQV_P224_TV12_Q2U 
    = Some MQV_P224_TV12_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV13›
  
definition MQV_P224_TV13_d1U :: nat where
  "MQV_P224_TV13_d1U = 0xb06ffc17a4e3671a693bd34c3fcd5be29beefdad15c6d8fcbcf6777c"

definition MQV_P224_TV13_Q1U :: "int point" where 
  "MQV_P224_TV13_Q1U = Point
    0x393edeb86f7c4ff980a23aaac80f4cb4ccde20310659dec086de7b5b
    0xc74258dabe8dba123d807c91c14fb3cf1419e284a462397f1c85bb51"

definition MQV_P224_TV13_d2U :: nat where
 "MQV_P224_TV13_d2U = 0x519fdd571f29de40c60fbf630c0a7adc4c69b5397ec87ce5e3b426ca"

definition MQV_P224_TV13_Q2U :: "int point" where 
 "MQV_P224_TV13_Q2U = Point
    0x857432cb4382dc1e073571a112bf91c2c3dd70a3be475ff3b26db9fd
    0xfe76593a064421a765f4eb461e2ed9f1cbbed709fde3b7c0e42dd2c9"

definition MQV_P224_TV13_d1V :: nat where
 "MQV_P224_TV13_d1V = 0xa5f3400cd6bbe920d9259c24779956ec51f80f3b35582a3c19d25f2c"

definition MQV_P224_TV13_Q1V :: "int point" where 
 "MQV_P224_TV13_Q1V = Point
    0x3d528dbaef30590b9587739a2a093d1ee44a5cee653d08cef0795100
    0x98f115deca11f2803a68b812e27c01520766de58a5a9667901f714a1"

definition MQV_P224_TV13_d2V :: nat where
 "MQV_P224_TV13_d2V = 0x91c8cf734e643c4dbdcd02b331c70a53825cff17d9197978f77cfc53"

definition MQV_P224_TV13_Q2V :: "int point" where 
 "MQV_P224_TV13_Q2V = Point
    0x524d4c6f210865cc8f4964c63c1f5461fa32c81be1587e82b0faee1c
    0x35adfd21560c1c452691b6c41f2ebf3525976a45161734bece2453d3"

definition MQV_P224_TV13_Z :: nat where
  "MQV_P224_TV13_Z = 0x28b23ae5157ad557a95b239faaf6b91c463c5bd8fe0f2117e20cb490"
  
lemma MQV_P224_TV13_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV13_d1U MQV_P224_TV13_Q1U 
                       MQV_P224_TV13_d2U MQV_P224_TV13_Q2U 
                       MQV_P224_TV13_Q1V MQV_P224_TV13_Q2V 
    = Some MQV_P224_TV13_Z"
  by eval

lemma MQV_P224_TV13_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV13_d1V MQV_P224_TV13_Q1V 
                       MQV_P224_TV13_d2V MQV_P224_TV13_Q2V 
                       MQV_P224_TV13_Q1U MQV_P224_TV13_Q2U 
    = Some MQV_P224_TV13_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV14›
  
definition MQV_P224_TV14_d1U :: nat where
  "MQV_P224_TV14_d1U = 0xd8f33d088144cc5026c6f6b7b948e8fe9a3e56607ee251a66acd9736"

definition MQV_P224_TV14_Q1U :: "int point" where 
  "MQV_P224_TV14_Q1U = Point
    0x7ba5224df0f968395e0266e5da94e5bdd9c981c1dca6371701e736ce
    0xbe41b1ae25099f8c42ddc99575b2f77971446cbbbf0fdaa169c873dc"

definition MQV_P224_TV14_d2U :: nat where
 "MQV_P224_TV14_d2U = 0xf784a10d843bfe40b3f3a50dfce74840701c9b690f18aa645c4432ec"

definition MQV_P224_TV14_Q2U :: "int point" where 
 "MQV_P224_TV14_Q2U = Point
    0xc7c5687d2b09d67e820a8a35072177f170d3fcc3b17c37fcd2aeba6c
    0xe2925505c1d17c41d69ab577d3bee234e40ef8aabd9c781cf26a9841"

definition MQV_P224_TV14_d1V :: nat where
 "MQV_P224_TV14_d1V = 0x3ea8d9c3176c4f8ad3d9ddfffcd874485e4262a7144d9da2b76d8d95"

definition MQV_P224_TV14_Q1V :: "int point" where 
 "MQV_P224_TV14_Q1V = Point
    0xe75eb189b63d3db177fafa9fe6b1daf67386de1b83431014ef10e12c
    0xa63681c341bc4954c5b92f0978bacf5e26ba6e9ea529d3067ee89131"

definition MQV_P224_TV14_d2V :: nat where
 "MQV_P224_TV14_d2V = 0xe380968e6691de47f0d550e72f5c4369d828e0de85f8ef80b683f468"

definition MQV_P224_TV14_Q2V :: "int point" where 
 "MQV_P224_TV14_Q2V = Point
    0xd46885385c442f9eacf65a72706af5d94a67884da358fceb13d661c6
    0xb9ee752389e3b5e725888b2aeabb9214b3ce1e9388ab29f85ab15fad"

definition MQV_P224_TV14_Z :: nat where
  "MQV_P224_TV14_Z = 0x09f65bc8c77d4e47be5b2b824fa45b8b88240840a0f785437572b68c"
  
lemma MQV_P224_TV14_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV14_d1U MQV_P224_TV14_Q1U 
                       MQV_P224_TV14_d2U MQV_P224_TV14_Q2U 
                       MQV_P224_TV14_Q1V MQV_P224_TV14_Q2V 
    = Some MQV_P224_TV14_Z"
  by eval

lemma MQV_P224_TV14_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV14_d1V MQV_P224_TV14_Q1V 
                       MQV_P224_TV14_d2V MQV_P224_TV14_Q2V 
                       MQV_P224_TV14_Q1U MQV_P224_TV14_Q2U 
    = Some MQV_P224_TV14_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV15›
  
definition MQV_P224_TV15_d1U :: nat where
  "MQV_P224_TV15_d1U = 0xa1d70408d7172c7d6fc198ddf71d845cd46f8c121ff81bc47b264deb"

definition MQV_P224_TV15_Q1U :: "int point" where 
  "MQV_P224_TV15_Q1U = Point
    0x88d8c7c465d406995b72d06013e4629594f17b4cbc2af75899d3e528
    0x5a470ecb38a6ddd83846bea4b48a99b3bf675dc4cef8233f4bb553b5"

definition MQV_P224_TV15_d2U :: nat where
 "MQV_P224_TV15_d2U = 0x743a4febe338222b7b482a817ab7766c1772b9ed7b6efbf267462423"

definition MQV_P224_TV15_Q2U :: "int point" where 
 "MQV_P224_TV15_Q2U = Point
    0x27121469e2622cef28228a9c286280d0c44859662b0b06f3ab0c4254
    0xe380ba215248439acc3d4e28487348896ba162f32beb24bc6c0114d5"

definition MQV_P224_TV15_d1V :: nat where
 "MQV_P224_TV15_d1V = 0xac6af8615d1c7529c219ef91b04a2c2886850a077c9fc547bfec5f7c"

definition MQV_P224_TV15_Q1V :: "int point" where 
 "MQV_P224_TV15_Q1V = Point
    0xd2c48b400c205f435fa69ca40cd521de4b8861000c3d6ac268f33a70
    0xce55878d688dedab7f245350fe2eef0da61d82a77dc1dd10c12482be"

definition MQV_P224_TV15_d2V :: nat where
 "MQV_P224_TV15_d2V = 0x0d1e618f276ea0c663fb06edf4c0e1f79e886987bb7b11ed2416c963"

definition MQV_P224_TV15_Q2V :: "int point" where 
 "MQV_P224_TV15_Q2V = Point
    0x0627ef07bc355f439eb7d0d53c74add73d2b4f1fcef8d9ffe9939d4a
    0xcf1aa70fb5132a5b8b83e9bfb5ea1b2ba136c15e75b80b2115bcec46"

definition MQV_P224_TV15_Z :: nat where
  "MQV_P224_TV15_Z = 0x0d7f5e88189fbee30c34187d199678bf6357dd20017ba3c8f5d63e61"
  
lemma MQV_P224_TV15_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV15_d1U MQV_P224_TV15_Q1U 
                       MQV_P224_TV15_d2U MQV_P224_TV15_Q2U 
                       MQV_P224_TV15_Q1V MQV_P224_TV15_Q2V 
    = Some MQV_P224_TV15_Z"
  by eval

lemma MQV_P224_TV15_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV15_d1V MQV_P224_TV15_Q1V 
                       MQV_P224_TV15_d2V MQV_P224_TV15_Q2V 
                       MQV_P224_TV15_Q1U MQV_P224_TV15_Q2U 
    = Some MQV_P224_TV15_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV16›
  
definition MQV_P224_TV16_d1U :: nat where
  "MQV_P224_TV16_d1U = 0xa0684b74ff55a959fc07776c05cbebadc4f96a6507eb800f341d7bbc"

definition MQV_P224_TV16_Q1U :: "int point" where 
  "MQV_P224_TV16_Q1U = Point
    0x4d5994325450a099b9089da2083f5735a69a2c24c596479ade23725a
    0x3abb01f272404798585b4d3d472dfac43a02c9f2dea173e87c6137b2"

definition MQV_P224_TV16_d2U :: nat where
 "MQV_P224_TV16_d2U = 0xb2ddb30fab2512b66a24daeaf9627eb6b4efa69f346d0d6fda307f38"

definition MQV_P224_TV16_Q2U :: "int point" where 
 "MQV_P224_TV16_Q2U = Point
    0xd749acf4cba80639a652344f9e606b30f4b197525434ac964ea68e32
    0xdf3bcbe3a389d8017c6e45ebbe34c04b631672f9663af5053cfbe796"

definition MQV_P224_TV16_d1V :: nat where
 "MQV_P224_TV16_d1V = 0x61a9737edbf4247ea2918b4fbae5f9ffd43b182e28128f264c5b2611"

definition MQV_P224_TV16_Q1V :: "int point" where 
 "MQV_P224_TV16_Q1V = Point
    0xbedbedfcc689479034c54e82afcce491794a56096e7f178b1846c0eb
    0xd0448acd7d23e0bb00ce329f4fbcbe85cfd1381ccd0a1e4d34fa841f"

definition MQV_P224_TV16_d2V :: nat where
 "MQV_P224_TV16_d2V = 0x9f6fa07728da243750f7f24d87b98ede22f63dfff341924fe3bb827a"

definition MQV_P224_TV16_Q2V :: "int point" where 
 "MQV_P224_TV16_Q2V = Point
    0x80a6ee7e6232d77bc2149fdf886dcd8e3230c4ea72dd86e836d9de88
    0x23a7dd4363ffc3a22a4331300b839bc0211d255afaa059d50cee7716"

definition MQV_P224_TV16_Z :: nat where
  "MQV_P224_TV16_Z = 0x2e0123d3e42eff7e055c28f7f17b5cfdbfa50e399e69059c68593917"
  
lemma MQV_P224_TV16_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV16_d1U MQV_P224_TV16_Q1U 
                       MQV_P224_TV16_d2U MQV_P224_TV16_Q2U 
                       MQV_P224_TV16_Q1V MQV_P224_TV16_Q2V 
    = Some MQV_P224_TV16_Z"
  by eval

lemma MQV_P224_TV16_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV16_d1V MQV_P224_TV16_Q1V 
                       MQV_P224_TV16_d2V MQV_P224_TV16_Q2V 
                       MQV_P224_TV16_Q1U MQV_P224_TV16_Q2U 
    = Some MQV_P224_TV16_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV17›
  
definition MQV_P224_TV17_d1U :: nat where
  "MQV_P224_TV17_d1U = 0x6c6b1af3da5ba6f07ab8aaf7e62281309d2b6d4bd6e7af27a715f5b9"

definition MQV_P224_TV17_Q1U :: "int point" where 
  "MQV_P224_TV17_Q1U = Point
    0xcb122bcf29079a9abd2dba23d772a083359f48949cefd63f45baf14c
    0xdc37e8b7d4b813f6e5c9263d5581fcd9ba9e2db898254c44461480ab"

definition MQV_P224_TV17_d2U :: nat where
 "MQV_P224_TV17_d2U = 0x77eb9435d34031af015ca95786b7d51dc40c772ddfb9964d9a59b1e5"

definition MQV_P224_TV17_Q2U :: "int point" where 
 "MQV_P224_TV17_Q2U = Point
    0x71d36f2d87d18a1004a3b474ec7e813b85daa0fb185441a6d6baff15
    0xec8453cea8f60252130586ec5f932f13378216d41e21bfc63e0a4299"

definition MQV_P224_TV17_d1V :: nat where
 "MQV_P224_TV17_d1V = 0x800a4eb2737ce4e5c177b87a2ed061659796f88dfd7be8d85152ed81"

definition MQV_P224_TV17_Q1V :: "int point" where 
 "MQV_P224_TV17_Q1V = Point
    0x7504b5a859a81bda793e7dccaecd212a365215e4dade3527f57819a7
    0xd586fc729556660313c13ce38ab4fb3d1032c1517811cdaa0296fd64"

definition MQV_P224_TV17_d2V :: nat where
 "MQV_P224_TV17_d2V = 0xfbc453372c81b12b9845237425d0ec86d64b63657fd5b0908e03be75"

definition MQV_P224_TV17_Q2V :: "int point" where 
 "MQV_P224_TV17_Q2V = Point
    0x65f035bdef21d65c8ee3ec27e08a77fd2d9ec4c79a8e08911ad719fc
    0x3e407e4c9ec251a44becf487d1c3ce450089d000afe22960d7995747"

definition MQV_P224_TV17_Z :: nat where
  "MQV_P224_TV17_Z = 0x03c59fda17f2b251c71995cc7dfd003737241b17a99362b5a19ec554"
  
lemma MQV_P224_TV17_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV17_d1U MQV_P224_TV17_Q1U 
                       MQV_P224_TV17_d2U MQV_P224_TV17_Q2U 
                       MQV_P224_TV17_Q1V MQV_P224_TV17_Q2V 
    = Some MQV_P224_TV17_Z"
  by eval

lemma MQV_P224_TV17_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV17_d1V MQV_P224_TV17_Q1V 
                       MQV_P224_TV17_d2V MQV_P224_TV17_Q2V 
                       MQV_P224_TV17_Q1U MQV_P224_TV17_Q2U 
    = Some MQV_P224_TV17_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV18›
  
definition MQV_P224_TV18_d1U :: nat where
  "MQV_P224_TV18_d1U = 0xbe6e59d9abeeaf5b7a3efc2e564f5f2e419646e2db91fab5e413160c"

definition MQV_P224_TV18_Q1U :: "int point" where 
  "MQV_P224_TV18_Q1U = Point
    0x2e756aecd5d840bbf76170ad161078c7dc002f55c4895653b021d13b
    0xc783bc0bcf66017ed53aba02906315a9e6e9bc9ef866092182095372"

definition MQV_P224_TV18_d2U :: nat where
 "MQV_P224_TV18_d2U = 0xa25e43755eadbc4bd200df11cb8a65d1631fa94b8a4034574e1ae492"

definition MQV_P224_TV18_Q2U :: "int point" where 
 "MQV_P224_TV18_Q2U = Point
    0x42179394102f7c94cceaea4e98b90062f877d3a4b99cdebdefb44adb
    0xba20727aa3d5e7f73da7e07da62214a5479c0a3250268e11374a5350"

definition MQV_P224_TV18_d1V :: nat where
 "MQV_P224_TV18_d1V = 0x020990b9c313590e9b8cb109e126afe7d291a0cd163f4b91c5e2f02d"

definition MQV_P224_TV18_Q1V :: "int point" where 
 "MQV_P224_TV18_Q1V = Point
    0xbe2bdd32aca353db942fb3675f6b6b548abe194779b16f931c0a0ea8
    0xb1dddb9cd83f5a7b5330ce02ad6d2de25f3fe7e981eb347baf62ff56"

definition MQV_P224_TV18_d2V :: nat where
 "MQV_P224_TV18_d2V = 0x06a8fcdda3df014741adada291a2631ed5719122f2ab2cc9ac593b45"

definition MQV_P224_TV18_Q2V :: "int point" where 
 "MQV_P224_TV18_Q2V = Point
    0xad6908793c608c2df3a57fbf379131f05b7157d2fa6a3cafce7ea3e7
    0x1b499d8e660d4e8c9d7010f6e76eaed80b70a361e8d059ba0daf2dbb"

definition MQV_P224_TV18_Z :: nat where
  "MQV_P224_TV18_Z = 0x01ae35a7ef976fd15d45c7719923af33ecc4c24c8c91b5d719858e2d"
  
lemma MQV_P224_TV18_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV18_d1U MQV_P224_TV18_Q1U 
                       MQV_P224_TV18_d2U MQV_P224_TV18_Q2U 
                       MQV_P224_TV18_Q1V MQV_P224_TV18_Q2V 
    = Some MQV_P224_TV18_Z"
  by eval

lemma MQV_P224_TV18_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV18_d1V MQV_P224_TV18_Q1V 
                       MQV_P224_TV18_d2V MQV_P224_TV18_Q2V 
                       MQV_P224_TV18_Q1U MQV_P224_TV18_Q2U 
    = Some MQV_P224_TV18_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P224_TV19›
  
definition MQV_P224_TV19_d1U :: nat where
  "MQV_P224_TV19_d1U = 0x56fdd849c0065f7a2c59645b6e62b9faed062b05e436105bcc2ee3cb"

definition MQV_P224_TV19_Q1U :: "int point" where 
  "MQV_P224_TV19_Q1U = Point
    0xe3eb3402596d82584c4cfc5dacc29dc672d07f859cd1db0f68e7f8db
    0x5a4f21592e4145abec5959a83d978257ddccf3e3267d9779e75292ea"

definition MQV_P224_TV19_d2U :: nat where
 "MQV_P224_TV19_d2U = 0x709c9e0a75b111187b154191727a64021d95d4ea22ffef3605fe8911"

definition MQV_P224_TV19_Q2U :: "int point" where 
 "MQV_P224_TV19_Q2U = Point
    0x7bc8a58f70a56f80e45e76a88ad164f45417dbd81f1a66209e58b42c
    0x1b42ea009852c86d468d5556f6ce62a247d9c5d4b1bce9063f3414cf"

definition MQV_P224_TV19_d1V :: nat where
 "MQV_P224_TV19_d1V = 0x006105f9ee25c8bb3ac7068395c8b7265a65535aca268573379d75cc"

definition MQV_P224_TV19_Q1V :: "int point" where 
 "MQV_P224_TV19_Q1V = Point
    0xe9f423cc3da1476f10491f1ab635c6d0aefb4a4c72d32bb3510f33d0
    0x2df075712719d6d1eb63c8351156d9a0e91cf8e729574d3aa7310a53"

definition MQV_P224_TV19_d2V :: nat where
 "MQV_P224_TV19_d2V = 0xf765ca637b4706657b77a9469f1eaeac386bb884126046bf15f65b8d"

definition MQV_P224_TV19_Q2V :: "int point" where 
 "MQV_P224_TV19_Q2V = Point
    0xde82ef2e083ca4ecf08347a4450dbc093b663e1d444d8d275ecb0186
    0xeb1cda87b2d050d7e5ef378d41c19707cb76b88b65b41fca05289064"

definition MQV_P224_TV19_Z :: nat where
  "MQV_P224_TV19_Z = 0xf716751f438a9a8786de3430869af2e2c1aa78c877df0ae8c9c47417"
  
lemma MQV_P224_TV19_MQV_check: 
  "SEC1_P224_ECMQVprim MQV_P224_TV19_d1U MQV_P224_TV19_Q1U 
                       MQV_P224_TV19_d2U MQV_P224_TV19_Q2U 
                       MQV_P224_TV19_Q1V MQV_P224_TV19_Q2V 
    = Some MQV_P224_TV19_Z"
  by eval

lemma MQV_P224_TV19_MQV_check': 
  "SEC1_P224_ECMQVprim MQV_P224_TV19_d1V MQV_P224_TV19_Q1V 
                       MQV_P224_TV19_d2V MQV_P224_TV19_Q2V 
                       MQV_P224_TV19_Q1U MQV_P224_TV19_Q2U 
    = Some MQV_P224_TV19_Z"
  by eval  
  
section‹NIST Test Vectors for MQV primitive: Curve P-256›
  
subsubsection‹MQV test vector: MQV_P256_TV00›
  
definition MQV_P256_TV00_d1U :: nat where
  "MQV_P256_TV00_d1U = 0x3a728539f135d9fa3a5fb027e82b22088229f088e6d63111cd663ab41e4edab6"

definition MQV_P256_TV00_Q1U :: "int point" where 
  "MQV_P256_TV00_Q1U = Point
    0x679ee5642729f10d8a2505f54e15118dce08e7612d3cff5945056c77aad36a2d
    0x39bcf2511716db4d0542458c27ae1c83702ac58d651f896cc52b10939e7df850"

definition MQV_P256_TV00_d2U :: nat where
 "MQV_P256_TV00_d2U = 0x400cf51ce8878e04a26dc584801328ae5fb3f5940c2f3354f99e42db3de78035"

definition MQV_P256_TV00_Q2U :: "int point" where 
 "MQV_P256_TV00_Q2U = Point
    0xbf6b27a8631b01855da2737bd471f237733b203b47aa4cbe54d1e3795abad30b
    0xc5949ec3b1c41326b07eebad3cbeda4c63375d323998469b845d9eb0f8f9616b"

definition MQV_P256_TV00_d1V :: nat where
 "MQV_P256_TV00_d1V = 0x5a28a841d842656817d833704069e105bb8bc957ac656bd06cc09d7b7205fc20"

definition MQV_P256_TV00_Q1V :: "int point" where 
 "MQV_P256_TV00_Q1V = Point
    0x1774c26b74f640748830e34f95f1e6f3f21db89786e5f4f3be78d5ea565f0d43
    0x89290856a59b1e5cc49182f17abc2bf9a1684e70fea595b000588bad54cc4dac"

definition MQV_P256_TV00_d2V :: nat where
 "MQV_P256_TV00_d2V = 0xa234543753a4af2539881a93d28cfe1506005d4c1c49b0394587d0fbf41302f2"

definition MQV_P256_TV00_Q2V :: "int point" where 
 "MQV_P256_TV00_Q2V = Point
    0xd4a835653d3228f5a95b98eb4e45210c6f791133e5a9041e3ff942c04f172f9c
    0x59b5763fd3284496daecca60444ee17ad54db5cc057c19723cc1b7a17e64a9f5"

definition MQV_P256_TV00_Z :: nat where
  "MQV_P256_TV00_Z = 0x688e3c7a403a832598cd3aa7e1c5042c97a25e740962745d517f05ba1fe45a39"
  
lemma MQV_P256_TV00_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV00_d1U MQV_P256_TV00_Q1U 
                       MQV_P256_TV00_d2U MQV_P256_TV00_Q2U 
                       MQV_P256_TV00_Q1V MQV_P256_TV00_Q2V 
    = Some MQV_P256_TV00_Z"
  by eval

lemma MQV_P256_TV00_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV00_d1V MQV_P256_TV00_Q1V 
                       MQV_P256_TV00_d2V MQV_P256_TV00_Q2V 
                       MQV_P256_TV00_Q1U MQV_P256_TV00_Q2U 
    = Some MQV_P256_TV00_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV01›
  
definition MQV_P256_TV01_d1U :: nat where
  "MQV_P256_TV01_d1U = 0xb27e7d1ad98f2e6309baa9b3c16d561132df2d4e5528ae35d4fa95324e577a3f"

definition MQV_P256_TV01_Q1U :: "int point" where 
  "MQV_P256_TV01_Q1U = Point
    0x9e887440d6531de1bb905c9bee5ea85c9d648c939a7e2febd78f8b106099ab8c
    0xa274cbdfdd01b20b52ad9bfc1067bcaa439543afeb13f0e6292da1ac8f6bbf82"

definition MQV_P256_TV01_d2U :: nat where
 "MQV_P256_TV01_d2U = 0xa15dc77a36f39c41fad6e8032da057aa25ba5c9ca55b427062b7e59a1f40b683"

definition MQV_P256_TV01_Q2U :: "int point" where 
 "MQV_P256_TV01_Q2U = Point
    0x34c18379f46f2ac91e6b05e349fc99892b4249fb336878b92c9727492aa887f1
    0x6769344dbcdd05a6c1ffd4c7da44a4b87896c45eaee5b0090acdfec6397544f4"

definition MQV_P256_TV01_d1V :: nat where
 "MQV_P256_TV01_d1V = 0xa94b07566c5e216c3022d7a5a932931f2ab79efd3352d699cc1956afdc5d2556"

definition MQV_P256_TV01_Q1V :: "int point" where 
 "MQV_P256_TV01_Q1V = Point
    0xaa235e4e4ea59506c4357eaf7834d0fdce0c68ba6f0b65301f5215dbf5097281
    0x6ca79266fb0ed49502cd9e48dc4bf15e49101f8525719633af8ce75054f9cef1"

definition MQV_P256_TV01_d2V :: nat where
 "MQV_P256_TV01_d2V = 0x60b237081bbe54a9fbee14347a61367dcdcfab4a1b91eb76fbc9f3f7a8dbc35c"

definition MQV_P256_TV01_Q2V :: "int point" where 
 "MQV_P256_TV01_Q2V = Point
    0x0b222885263adc84af7540c4382540914e6ded2844fc7e802cf05312a6bc17ca
    0x12fff2c2681c4807b5291c69051ed9e03c4145ef2bc186b116408cb4db551317"

definition MQV_P256_TV01_Z :: nat where
  "MQV_P256_TV01_Z = 0xc4bffa03bc25d3d592da3dc6e74bd3636f11acc973d67b322a8411ee532d6e0b"
  
lemma MQV_P256_TV01_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV01_d1U MQV_P256_TV01_Q1U 
                       MQV_P256_TV01_d2U MQV_P256_TV01_Q2U 
                       MQV_P256_TV01_Q1V MQV_P256_TV01_Q2V 
    = Some MQV_P256_TV01_Z"
  by eval

lemma MQV_P256_TV01_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV01_d1V MQV_P256_TV01_Q1V 
                       MQV_P256_TV01_d2V MQV_P256_TV01_Q2V 
                       MQV_P256_TV01_Q1U MQV_P256_TV01_Q2U 
    = Some MQV_P256_TV01_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV02›
  
definition MQV_P256_TV02_d1U :: nat where
  "MQV_P256_TV02_d1U = 0x52c960f92f350a5225272a0c0332950bd49e569f28783707e0f2c4615db67200"

definition MQV_P256_TV02_Q1U :: "int point" where 
  "MQV_P256_TV02_Q1U = Point
    0x85e868e605d3f764358b8e38cf34724ed189f3696cfe4c287487fd455e9cca3c
    0x5c93b42bbcf6272c6d5fcf6f9c556334ff02cc2946dc469a7d93dbf7dba80f83"

definition MQV_P256_TV02_d2U :: nat where
 "MQV_P256_TV02_d2U = 0x69d6b1544b4616db78d71131e3b41e918f5f9e17437b9a8f27a0f6e718e0d7e0"

definition MQV_P256_TV02_Q2U :: "int point" where 
 "MQV_P256_TV02_Q2U = Point
    0xe476d7bcb146d9df8b77727a192d3bda34bb85c1a1ab66af0ca2732bbc8da71d
    0x0ef89b2a5eb8647436977eb80376f0cf8e37ee31a9c5f2362ec872507e932301"

definition MQV_P256_TV02_d1V :: nat where
 "MQV_P256_TV02_d1V = 0x423c70143585276f2dd80b46f604c1654a2155bbf4a4fc56c1e569098d0b5b77"

definition MQV_P256_TV02_Q1V :: "int point" where 
 "MQV_P256_TV02_Q1V = Point
    0x703052303c113e787f535c008f5bf993a82cb5a5cea7739ffa34a3ee50d5a7be
    0xd6b8db24b3ff5e5f7edb6b3d2b8e63adf760443b6f6ab7f5e9b76113167e36df"

definition MQV_P256_TV02_d2V :: nat where
 "MQV_P256_TV02_d2V = 0x20892b75d02a34d8760d7fcc234983ecc939178d0fe8370d1057a4f8c188e876"

definition MQV_P256_TV02_Q2V :: "int point" where 
 "MQV_P256_TV02_Q2V = Point
    0xcad63306145f9190712df19e6bdbd92679f11c68e18f5e085d8712266e22ec76
    0x90f98c95163b614b1d5f50ba0c514e2d0c13428938d9139953843d2aedbf7b58"

definition MQV_P256_TV02_Z :: nat where
  "MQV_P256_TV02_Z = 0x0964a08449b8ec2e5668c3bfa238858970b3ee80361dc8e68d72523b09ea575e"
  
lemma MQV_P256_TV02_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV02_d1U MQV_P256_TV02_Q1U 
                       MQV_P256_TV02_d2U MQV_P256_TV02_Q2U 
                       MQV_P256_TV02_Q1V MQV_P256_TV02_Q2V 
    = Some MQV_P256_TV02_Z"
  by eval

lemma MQV_P256_TV02_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV02_d1V MQV_P256_TV02_Q1V 
                       MQV_P256_TV02_d2V MQV_P256_TV02_Q2V 
                       MQV_P256_TV02_Q1U MQV_P256_TV02_Q2U 
    = Some MQV_P256_TV02_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV03›
  
definition MQV_P256_TV03_d1U :: nat where
  "MQV_P256_TV03_d1U = 0x166881e899fda2f3faad9e9682b119a50d94954921ff89f9946f5590459afc18"

definition MQV_P256_TV03_Q1U :: "int point" where 
  "MQV_P256_TV03_Q1U = Point
    0x8977ddad2d49f18dcb76512102147d3432045ead73f52cc964abf55568ecc7d8
    0xe1b225b2675438fda7aeede9c8705f9d527b136c906a6c6ac3d79823c98aa28b"

definition MQV_P256_TV03_d2U :: nat where
 "MQV_P256_TV03_d2U = 0x654b47baa9b877022190fe4779f04ac5013d7df0d6be2dbf1e465896212e782f"

definition MQV_P256_TV03_Q2U :: "int point" where 
 "MQV_P256_TV03_Q2U = Point
    0xa228a690fc1276a2d5ba2f8022ff1244c9e8439706ebc8f76f1929da45766f22
    0x1f4c2cb11f2edbe12827357af2b95789d1a0f49983b2fe712e12b6d87050111f"

definition MQV_P256_TV03_d1V :: nat where
 "MQV_P256_TV03_d1V = 0xeea5f33bb7b5f8661da769883cc6c86db220e5587fd3ef2f0fc1aadfe1020955"

definition MQV_P256_TV03_Q1V :: "int point" where 
 "MQV_P256_TV03_Q1V = Point
    0x87bc4b82b9affb7dc0dae53ad01440db85adb22ea96aaf354fb1b81659043b38
    0x57e56b2930ba348d709fcae5fc299526dfacd0c05126c86505809bbb4fbe6932"

definition MQV_P256_TV03_d2V :: nat where
 "MQV_P256_TV03_d2V = 0xcc6ba145b5fc871991497ec435a81031901bee8aa1fbe1800943809ac98de84d"

definition MQV_P256_TV03_Q2V :: "int point" where 
 "MQV_P256_TV03_Q2V = Point
    0x66988a4e0d21fa44bc1a7d16bf5db30de5b0a00aa87b02fd789b760c6e9049d5
    0xd50f9f348550fb159384f7f7a86f103051bc7ceecc3ebec8bdad37527909692d"

definition MQV_P256_TV03_Z :: nat where
  "MQV_P256_TV03_Z = 0x8eef7092ecd0cb0c1d9d06c4c3fa815974afea5f7db6188eb9fe3997b08d2e74"
  
lemma MQV_P256_TV03_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV03_d1U MQV_P256_TV03_Q1U 
                       MQV_P256_TV03_d2U MQV_P256_TV03_Q2U 
                       MQV_P256_TV03_Q1V MQV_P256_TV03_Q2V 
    = Some MQV_P256_TV03_Z"
  by eval

lemma MQV_P256_TV03_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV03_d1V MQV_P256_TV03_Q1V 
                       MQV_P256_TV03_d2V MQV_P256_TV03_Q2V 
                       MQV_P256_TV03_Q1U MQV_P256_TV03_Q2U 
    = Some MQV_P256_TV03_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV04›
  
definition MQV_P256_TV04_d1U :: nat where
  "MQV_P256_TV04_d1U = 0x4cb3919e1ae177071b78b303ac7aea4b1847f08f00e05674fb6ee73994eba10d"

definition MQV_P256_TV04_Q1U :: "int point" where 
  "MQV_P256_TV04_Q1U = Point
    0xde108768a32854a61373febd855989eb4ec8f31389fb02ce270d1df5e7d7723b
    0xc1acaaae1655e124508ad41683b204bba3e2a1645080a459db5f032aaefe7ed7"

definition MQV_P256_TV04_d2U :: nat where
 "MQV_P256_TV04_d2U = 0x6fc030845e22da811e2080a70bc845b204689c01e0f7242477d703af190fc438"

definition MQV_P256_TV04_Q2U :: "int point" where 
 "MQV_P256_TV04_Q2U = Point
    0x1bc2b5fe3c6cb44096b4394cc7b5e90e8c4ee7b4b8ce23fa4d0778ce99f8a4d3
    0x300bc56e74f0b56371187650dc0d86c58ffe2617278658c5e94095d092790d8b"

definition MQV_P256_TV04_d1V :: nat where
 "MQV_P256_TV04_d1V = 0x48a6b3c3d33242bfc1c902d81e6686e93e87e5504824595898f6eda50328b399"

definition MQV_P256_TV04_Q1V :: "int point" where 
 "MQV_P256_TV04_Q1V = Point
    0xd076e40ec43b605663a4f0a55d620e175443d38a113ef74053689d327143994d
    0xc25d686c2125a8645d4fcb5295c66774022bdfc542ca8f505838cd8816d9216a"

definition MQV_P256_TV04_d2V :: nat where
 "MQV_P256_TV04_d2V = 0xceb3e59b817ca996f23afa412cfa3e24e941423cccfd69af7c3940d8ce89d669"

definition MQV_P256_TV04_Q2V :: "int point" where 
 "MQV_P256_TV04_Q2V = Point
    0x34ebc138cc4f96d0588eb3a93c328819c16613c5112389f5422c4bd0de3d1c9d
    0xc40f51f283f40ef47f850194bb3d99a4f09841f0556768bd4dcc2454bee34ff7"

definition MQV_P256_TV04_Z :: nat where
  "MQV_P256_TV04_Z = 0x27107c6f7b0b727469690d16ce9255f006bdeeb15be4a8d06f21a83c9bc8e4e0"
  
lemma MQV_P256_TV04_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV04_d1U MQV_P256_TV04_Q1U 
                       MQV_P256_TV04_d2U MQV_P256_TV04_Q2U 
                       MQV_P256_TV04_Q1V MQV_P256_TV04_Q2V 
    = Some MQV_P256_TV04_Z"
  by eval

lemma MQV_P256_TV04_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV04_d1V MQV_P256_TV04_Q1V 
                       MQV_P256_TV04_d2V MQV_P256_TV04_Q2V 
                       MQV_P256_TV04_Q1U MQV_P256_TV04_Q2U 
    = Some MQV_P256_TV04_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV05›
  
definition MQV_P256_TV05_d1U :: nat where
  "MQV_P256_TV05_d1U = 0x7e1c1acf75a38faffe394e117eb339cafd8e7b88e018bfd491610fff0591d4b4"

definition MQV_P256_TV05_Q1U :: "int point" where 
  "MQV_P256_TV05_Q1U = Point
    0x33ef41ad1ed1b4c3cbbb1271b15e64ece487187579bc5837a89c9e12f71860e4
    0x1898e18e659a7bf56a00b795d5d16354adaff53652127eef5a6c3b7097945d73"

definition MQV_P256_TV05_d2U :: nat where
 "MQV_P256_TV05_d2U = 0xec6ad6dc203e9a6d9a233bbd57460276e269fdbc43736ebec0d95c727d2573f3"

definition MQV_P256_TV05_Q2U :: "int point" where 
 "MQV_P256_TV05_Q2U = Point
    0x8ac5bdce362551b96db5771947c8dcc965bfe815177280262be0d52850029749
    0x5ff1dd019cc34a58f4c5052612a0f0add6b6284714b2800162445e495ec24d35"

definition MQV_P256_TV05_d1V :: nat where
 "MQV_P256_TV05_d1V = 0x6a21e46b860a00af4f893ab562ddbb68f606ca1665efdf2700e30dae29eb2f16"

definition MQV_P256_TV05_Q1V :: "int point" where 
 "MQV_P256_TV05_Q1V = Point
    0x236cf4a28294f721e08b5a6e5bba6ab48d12400a00eaf8194dd3fcffb8526d8f
    0xa52420519eaec31d76237aa875ae2dfc357036a608ac23a5d565f50d87807775"

definition MQV_P256_TV05_d2V :: nat where
 "MQV_P256_TV05_d2V = 0x13f678d55c08c589e7d094cd96e7b36a9d313f8b282f2af3f4b33a7eaa440616"

definition MQV_P256_TV05_Q2V :: "int point" where 
 "MQV_P256_TV05_Q2V = Point
    0xc54220c662436abd07e4a62cbe2f1d06cfe25285366d2212ac15b833784f7f61
    0xade7729c177e093fb841ca2360e532ca9aa6242ddb499148423bad472d2ef2c0"

definition MQV_P256_TV05_Z :: nat where
  "MQV_P256_TV05_Z = 0x0bf93930e489de71023fad9ffd88c388c4a72ca872943bdfab0c8723c265384a"
  
lemma MQV_P256_TV05_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV05_d1U MQV_P256_TV05_Q1U 
                       MQV_P256_TV05_d2U MQV_P256_TV05_Q2U 
                       MQV_P256_TV05_Q1V MQV_P256_TV05_Q2V 
    = Some MQV_P256_TV05_Z"
  by eval

lemma MQV_P256_TV05_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV05_d1V MQV_P256_TV05_Q1V 
                       MQV_P256_TV05_d2V MQV_P256_TV05_Q2V 
                       MQV_P256_TV05_Q1U MQV_P256_TV05_Q2U 
    = Some MQV_P256_TV05_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV06›
  
definition MQV_P256_TV06_d1U :: nat where
  "MQV_P256_TV06_d1U = 0x0f7baead3f6b222d7714b7532ddf481075425cb948da349c55dcb69d3b0475b0"

definition MQV_P256_TV06_Q1U :: "int point" where 
  "MQV_P256_TV06_Q1U = Point
    0x30f583ad2ea6e14d8d4ca5350b1adcdfa565347c7a9af2e9bef97c1bd402fcc8
    0x41f94974eafbbeafd2748dfc2790f3a9807d3e1fa3b460f0826ac6b756b73c95"

definition MQV_P256_TV06_d2U :: nat where
 "MQV_P256_TV06_d2U = 0xb3e6cedf7a401f63c097b3b432fda07cee0f2b30950fddb37ab95159a4bf422e"

definition MQV_P256_TV06_Q2U :: "int point" where 
 "MQV_P256_TV06_Q2U = Point
    0x8e8afa6ac27a4c56081ac389d8295716c8971a4ba9605d0fe1e93237e4b1bedd
    0x5c8e312528d7485c37faccca11e2b951e2218951e9837888de789d567943b9f4"

definition MQV_P256_TV06_d1V :: nat where
 "MQV_P256_TV06_d1V = 0x62efc808963a9a3dd6a63a247018cb52e7bc054c66e680737811b80b5efecaf9"

definition MQV_P256_TV06_Q1V :: "int point" where 
 "MQV_P256_TV06_Q1V = Point
    0x51234bd1bc1c8ec0ff5700af1fbcee6b69a2b8fe1118aee16b691977c56dd300
    0x06157e5e53c0cc9f97eda80673499b306b6f9095719ad9304e13932ca19e5c9e"

definition MQV_P256_TV06_d2V :: nat where
 "MQV_P256_TV06_d2V = 0xb8c340c0b38743f1bae270950932d2813f1949639449db86ed51e4180ea26a56"

definition MQV_P256_TV06_Q2V :: "int point" where 
 "MQV_P256_TV06_Q2V = Point
    0xbf908d6918629777bba8ff7a309fa425d5d7c5cea7d61fafa64b51dfa20f9035
    0xebc55df317dfc00645e75746da90b3fbb77e53ab3bee144ddfda68365cd57699"

definition MQV_P256_TV06_Z :: nat where
  "MQV_P256_TV06_Z = 0x0653054ed946a4a67b548c19afba900360984ca9281b42e3c86d74913b1dc371"
  
lemma MQV_P256_TV06_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV06_d1U MQV_P256_TV06_Q1U 
                       MQV_P256_TV06_d2U MQV_P256_TV06_Q2U 
                       MQV_P256_TV06_Q1V MQV_P256_TV06_Q2V 
    = Some MQV_P256_TV06_Z"
  by eval

lemma MQV_P256_TV06_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV06_d1V MQV_P256_TV06_Q1V 
                       MQV_P256_TV06_d2V MQV_P256_TV06_Q2V 
                       MQV_P256_TV06_Q1U MQV_P256_TV06_Q2U 
    = Some MQV_P256_TV06_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV07›
  
definition MQV_P256_TV07_d1U :: nat where
  "MQV_P256_TV07_d1U = 0xb3e406c8a0dbdea615842ea1c904295cf7ce7f58a6762b036f62d1c5aa81bbbd"

definition MQV_P256_TV07_Q1U :: "int point" where 
  "MQV_P256_TV07_Q1U = Point
    0x9b3ce04ca508f04852f7f40ddcd162765eadfd68b43ce9e17aeaefd8d33a8295
    0x03b5a5a08e51f03633aa4f4fd36a32e07a0f80555392a099628e420451552a46"

definition MQV_P256_TV07_d2U :: nat where
 "MQV_P256_TV07_d2U = 0x0bbd934c94b03a3f94d89901915bcddb0536acfdd653235c739bb85df9918663"

definition MQV_P256_TV07_Q2U :: "int point" where 
 "MQV_P256_TV07_Q2U = Point
    0xa4af43aba0d12bc634f62d6aada56e9c39bf2820c7770f3a972db033bdb63bfe
    0x5dcfe10dd79c3406d671f42c6c40ca8929e31d372f4dbe25c81bab83961acc21"

definition MQV_P256_TV07_d1V :: nat where
 "MQV_P256_TV07_d1V = 0x8280d32574cc0b81170c037c2fb0bc636ee735c5c2921e2baae3bc9da54e7f37"

definition MQV_P256_TV07_Q1V :: "int point" where 
 "MQV_P256_TV07_Q1V = Point
    0x5b6d6df375b4b00b96c98c95ea229f49d80602ed66a67325d96e9b2b56843a65
    0xd9ff78a1f71dadefcb71d1dbc9ad16dad960a0abb5365204e4505ec2a84e453b"

definition MQV_P256_TV07_d2V :: nat where
 "MQV_P256_TV07_d2V = 0x8da136b5e820db43e45a8226998a11e3101cff0e549958f0b81a97b3e9d2d51c"

definition MQV_P256_TV07_Q2V :: "int point" where 
 "MQV_P256_TV07_Q2V = Point
    0xefec36c348c550a1a7a39b294cc35a7e83f5eaff00db25b102a08a19c3756bc8
    0x65de6535f0c8b54e1283a89a442cbb2d4e4e296762899ae4a1e732cbd90060d8"

definition MQV_P256_TV07_Z :: nat where
  "MQV_P256_TV07_Z = 0x08929e8df6ec48a7a212216282919a6d894c12b1f717cbf5720f26cc5f10be14"
  
lemma MQV_P256_TV07_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV07_d1U MQV_P256_TV07_Q1U 
                       MQV_P256_TV07_d2U MQV_P256_TV07_Q2U 
                       MQV_P256_TV07_Q1V MQV_P256_TV07_Q2V 
    = Some MQV_P256_TV07_Z"
  by eval

lemma MQV_P256_TV07_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV07_d1V MQV_P256_TV07_Q1V 
                       MQV_P256_TV07_d2V MQV_P256_TV07_Q2V 
                       MQV_P256_TV07_Q1U MQV_P256_TV07_Q2U 
    = Some MQV_P256_TV07_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV08›
  
definition MQV_P256_TV08_d1U :: nat where
  "MQV_P256_TV08_d1U = 0x1d75493d633c60157a6764a1acc1cf571e7e8869af1a58203317cbb695c9e699"

definition MQV_P256_TV08_Q1U :: "int point" where 
  "MQV_P256_TV08_Q1U = Point
    0xd54d5ec483b62fe42da9223845332a9e243e78b79148362f2cc4cb9e3099ff42
    0x2a3c05fd4663a2dc44061874b359812c99adb905039d131e5bc0094ca1849822"

definition MQV_P256_TV08_d2U :: nat where
 "MQV_P256_TV08_d2U = 0x87e7a2ebab99bbe0eea8dd113b254ad359f083a16d3be1622a97825a0cfe2056"

definition MQV_P256_TV08_Q2U :: "int point" where 
 "MQV_P256_TV08_Q2U = Point
    0x0bcc34ecfcc9d287587ad831fe08f8a39ef2aac19c0f419b19cb8accdbdc6140
    0xb222f397aafbfd6e66568338fdf0df97c0f8650096b23f06ef8e1edb57157d71"

definition MQV_P256_TV08_d1V :: nat where
 "MQV_P256_TV08_d1V = 0xcc8f9ad008baffd7b558c6384194ecf13bb9fc5d3fd9af7b3e07f6088d26297d"

definition MQV_P256_TV08_Q1V :: "int point" where 
 "MQV_P256_TV08_Q1V = Point
    0xb18e5f7e844b8957b59a9b370699d1bfa5000b717780a11bd5caf6ae2029f78f
    0x80f83cb26477e27a988a6622fce36e547443d85545c4cb312ce7cc4c36b21da7"

definition MQV_P256_TV08_d2V :: nat where
 "MQV_P256_TV08_d2V = 0x0667942a5268ea42fb390b4ffc0b6b072b66caa0f666192ee1d68a987f875525"

definition MQV_P256_TV08_Q2V :: "int point" where 
 "MQV_P256_TV08_Q2V = Point
    0x7acd54317687ffb1fc4cb979f5d4c715979c1abd0b8f2355b46615338a829c1a
    0x5e3e92300e24f7862512cbfba8dbe30db9b990891dada0773659be373baae1b5"

definition MQV_P256_TV08_Z :: nat where
  "MQV_P256_TV08_Z = 0xa9288fea20624107bfc5527295edc110613fa1b39d9d3733bb1a785282b676a2"
  
lemma MQV_P256_TV08_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV08_d1U MQV_P256_TV08_Q1U 
                       MQV_P256_TV08_d2U MQV_P256_TV08_Q2U 
                       MQV_P256_TV08_Q1V MQV_P256_TV08_Q2V 
    = Some MQV_P256_TV08_Z"
  by eval

lemma MQV_P256_TV08_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV08_d1V MQV_P256_TV08_Q1V 
                       MQV_P256_TV08_d2V MQV_P256_TV08_Q2V 
                       MQV_P256_TV08_Q1U MQV_P256_TV08_Q2U 
    = Some MQV_P256_TV08_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV09›
  
definition MQV_P256_TV09_d1U :: nat where
  "MQV_P256_TV09_d1U = 0x12bea85216d4fecdabe1cb4d4c8a74897357ee32dfcda5f03ebe4e7b5fa95454"

definition MQV_P256_TV09_Q1U :: "int point" where 
  "MQV_P256_TV09_Q1U = Point
    0xff3baa6bdd38904e3b84be30579650e3000baacd9bdafe0a9977025769b911bf
    0xf6bcbe2b9dffe92ae93f24a059a78813221e0a4b9ead8fcf1dd86ce60365189d"

definition MQV_P256_TV09_d2U :: nat where
 "MQV_P256_TV09_d2U = 0xf82f508d70c985306c87cb053df4432800d3f3642ae2740db24602de737c12d5"

definition MQV_P256_TV09_Q2U :: "int point" where 
 "MQV_P256_TV09_Q2U = Point
    0x3b827180cd20cbeaa54aacb588d250baaef005396388a3a20eb4214e3d9a65cf
    0xf0cc595006078631350fe975e8b45e0b375ff3fa55a0e7af75a46670f5a20f07"

definition MQV_P256_TV09_d1V :: nat where
 "MQV_P256_TV09_d1V = 0xa5fbae83eb208d96bc3cce0c5975d6e375e451365a1d09a6e526b04db53bcbd5"

definition MQV_P256_TV09_Q1V :: "int point" where 
 "MQV_P256_TV09_Q1V = Point
    0x0e415565f5ca71d0aa32c91563b7910e3ff588dfc9887f9ab0f08dbc94a48da9
    0x13c62ad6587c765663d1e3ba19e2e9ecb1407fc07b0c2f2d21c55742699e393f"

definition MQV_P256_TV09_d2V :: nat where
 "MQV_P256_TV09_d2V = 0x1921439b4165490d933b2171ff9ebcec24d960794d1667b033a1f6211f5bc03e"

definition MQV_P256_TV09_Q2V :: "int point" where 
 "MQV_P256_TV09_Q2V = Point
    0x8cc2f6217e13456812d95cc09559f3cb851892bd9681b1cca85351cf556b8155
    0xafb29f5d727d1fd3fdf22d905fbfee47562a63cd99e1f13fc599e94fba40d2c1"

definition MQV_P256_TV09_Z :: nat where
  "MQV_P256_TV09_Z = 0x044e458061425c36f66230c570c2ed6db86f09fd86c7a38e21376a9f3212d965"
  
lemma MQV_P256_TV09_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV09_d1U MQV_P256_TV09_Q1U 
                       MQV_P256_TV09_d2U MQV_P256_TV09_Q2U 
                       MQV_P256_TV09_Q1V MQV_P256_TV09_Q2V 
    = Some MQV_P256_TV09_Z"
  by eval

lemma MQV_P256_TV09_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV09_d1V MQV_P256_TV09_Q1V 
                       MQV_P256_TV09_d2V MQV_P256_TV09_Q2V 
                       MQV_P256_TV09_Q1U MQV_P256_TV09_Q2U 
    = Some MQV_P256_TV09_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV10›
  
definition MQV_P256_TV10_d1U :: nat where
  "MQV_P256_TV10_d1U = 0xf45a1bd35fcd21644fd636431a80a379c6f2606f3990d5edb1381ae595b825c1"

definition MQV_P256_TV10_Q1U :: "int point" where 
  "MQV_P256_TV10_Q1U = Point
    0xd52136f38a15a01c8748891d233048fea044d21869c675b77f1f63345b9e06fd
    0x480c092296693c3c4e0052cbfc8f5cc021fa9204e77507ab97ce503df756f8e2"

definition MQV_P256_TV10_d2U :: nat where
 "MQV_P256_TV10_d2U = 0x333fdead30a68cfd2dec3aac38d37e6f16c2674f7f809d74d97df429f4c1b7b0"

definition MQV_P256_TV10_Q2U :: "int point" where 
 "MQV_P256_TV10_Q2U = Point
    0xb414fa128861db163e9f5a012f7c3203e7fb03aed585f063a08a94fcc39f37d1
    0xb60d20053385d9d0cda24bea15115b34576a540cb1441f9e6fc7c1c5cab9debd"

definition MQV_P256_TV10_d1V :: nat where
 "MQV_P256_TV10_d1V = 0xf1e19bd4af9303328b22ede6412300fe60f1c3475f445f842875923f439c7690"

definition MQV_P256_TV10_Q1V :: "int point" where 
 "MQV_P256_TV10_Q1V = Point
    0x7e819ab1f72a1722b554ec79d00e876f2fb118822d43cd1350be19a55d15e026
    0x5a8576871a9e4773b22c0d88d6f04f7214350d3fec724e4b5404f407e80e3f45"

definition MQV_P256_TV10_d2V :: nat where
 "MQV_P256_TV10_d2V = 0x9a169daa205de6e21528deaa68f73978bf0fe2b358e3de2dd5c396afdc9dc4f7"

definition MQV_P256_TV10_Q2V :: "int point" where 
 "MQV_P256_TV10_Q2V = Point
    0xb3f019e53f5658c42f91fcc24bdab74dd405b6452500d3b4f4e6efa18075e25c
    0x175af473d7cfa0d5b161f2339c1db11a75aba140a368b1d4471bf6b512e47542"

definition MQV_P256_TV10_Z :: nat where
  "MQV_P256_TV10_Z = 0x15333b96559948c9aca89671adf06fdd9d69173a541b54625a2aaf3ad52e8313"
  
lemma MQV_P256_TV10_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV10_d1U MQV_P256_TV10_Q1U 
                       MQV_P256_TV10_d2U MQV_P256_TV10_Q2U 
                       MQV_P256_TV10_Q1V MQV_P256_TV10_Q2V 
    = Some MQV_P256_TV10_Z"
  by eval

lemma MQV_P256_TV10_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV10_d1V MQV_P256_TV10_Q1V 
                       MQV_P256_TV10_d2V MQV_P256_TV10_Q2V 
                       MQV_P256_TV10_Q1U MQV_P256_TV10_Q2U 
    = Some MQV_P256_TV10_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV11›
  
definition MQV_P256_TV11_d1U :: nat where
  "MQV_P256_TV11_d1U = 0x24406c7ab595c89c0a546a4fe813257f68999fd9fc262e31bcde5f416656b3d7"

definition MQV_P256_TV11_Q1U :: "int point" where 
  "MQV_P256_TV11_Q1U = Point
    0x6e426094adf36f1710b283f2a0ad63bab194d9130337a7508b106be0dd1c79e3
    0xb4af93a2747883000e0b528c12c87483fd44fff254e6d85dd21461193d1de651"

definition MQV_P256_TV11_d2U :: nat where
 "MQV_P256_TV11_d2U = 0xa272a6b30725f0a443b2de8bd1fc51ad39d192b8df005698f966be25bfb03514"

definition MQV_P256_TV11_Q2U :: "int point" where 
 "MQV_P256_TV11_Q2U = Point
    0x61b1daf190f89ac43feb91b65b0f34904f6d83612fa32739ff76c67c1fbc4c18
    0x6e21b51bfae4832d66d6c57530d457b532d01539a7646d8bb93727304dfff6f3"

definition MQV_P256_TV11_d1V :: nat where
 "MQV_P256_TV11_d1V = 0x31061d3da7f083f482a710754c6b4d620736e8fd82f18e96caed93daa54193c9"

definition MQV_P256_TV11_Q1V :: "int point" where 
 "MQV_P256_TV11_Q1V = Point
    0x8f126145960d6ece88bc9eb8001fc5fb6ec0fe21c344d686f7792cdfca43160d
    0x11f52ce6872d74f0fe69e667d0144195bd42f8266d6473e693aeb250e49a891d"

definition MQV_P256_TV11_d2V :: nat where
 "MQV_P256_TV11_d2V = 0xf96c7e5557386e73fc5f2030935758c973a1cb36c9c0578f5fb6dfeb6c6c9c01"

definition MQV_P256_TV11_Q2V :: "int point" where 
 "MQV_P256_TV11_Q2V = Point
    0x07397a6521738d539a5c63e1bd46f1e0ba15fa250dd0a088927548aabdfd2076
    0x89ffc8a437e7c2e26d3317b6e0ba667ddceac5e182153cdf4dba9e8ec45ca4f3"

definition MQV_P256_TV11_Z :: nat where
  "MQV_P256_TV11_Z = 0xef021a2194cb7ff8d6830ff1ff4ae412c8779883e885496a58789a54952558a1"
  
lemma MQV_P256_TV11_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV11_d1U MQV_P256_TV11_Q1U 
                       MQV_P256_TV11_d2U MQV_P256_TV11_Q2U 
                       MQV_P256_TV11_Q1V MQV_P256_TV11_Q2V 
    = Some MQV_P256_TV11_Z"
  by eval

lemma MQV_P256_TV11_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV11_d1V MQV_P256_TV11_Q1V 
                       MQV_P256_TV11_d2V MQV_P256_TV11_Q2V 
                       MQV_P256_TV11_Q1U MQV_P256_TV11_Q2U 
    = Some MQV_P256_TV11_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV12›
  
definition MQV_P256_TV12_d1U :: nat where
  "MQV_P256_TV12_d1U = 0x617a91da5ffef7ea55416c5d8dfa6481aa47a928b2590343c9d68c179e075f6f"

definition MQV_P256_TV12_Q1U :: "int point" where 
  "MQV_P256_TV12_Q1U = Point
    0x930dc67b33951040e70f0d21554bbb49d29def0cdfe7f950b2d5327392ad6050
    0x601976bb62138780e728bdf64e53f7d1dc08bb8c3383b23d218ef3e262c33826"

definition MQV_P256_TV12_d2U :: nat where
 "MQV_P256_TV12_d2U = 0xfa39be83b33686eaa18ca48830c34b9a1a91abb764ffde0f5c9eb92ce7dfe2cc"

definition MQV_P256_TV12_Q2U :: "int point" where 
 "MQV_P256_TV12_Q2U = Point
    0xc62f38597e4988a040f3a6a904304a66916d99a952f6942793bc59bd0523658b
    0xadcc8d04f4e46d48d7f57c94defcfe0714c45d94c648cd1a1dd657eea8769246"

definition MQV_P256_TV12_d1V :: nat where
 "MQV_P256_TV12_d1V = 0x484bd3c097b7891c25023bb66207dfe5d5c41fa8ee28eeec5329ac3d3054d32f"

definition MQV_P256_TV12_Q1V :: "int point" where 
 "MQV_P256_TV12_Q1V = Point
    0x1adf3645c0945703e2bc403b7036880c80b893c9e97823f3bbd08bf3f0ae3f37
    0x68d1b4a2c9bf4f80dea492e9c2d8898cdac040b4b5c0dbbd0c997bfe6b68d7ff"

definition MQV_P256_TV12_d2V :: nat where
 "MQV_P256_TV12_d2V = 0x66f606351f7fbc52e0c3f1497175a473627853fbb8d2b2a1d7a8929ed5a06fa9"

definition MQV_P256_TV12_Q2V :: "int point" where 
 "MQV_P256_TV12_Q2V = Point
    0xf6d7711632e63c457af3875a7fc24f53479c4046fdd6ba658fe9ab87b972065b
    0xd29e47f7b9d30300215431a516cb3ec94c0306ec75ba2a19de4c45b79c584c22"

definition MQV_P256_TV12_Z :: nat where
  "MQV_P256_TV12_Z = 0xcd8d312e8d921f1e365dbc67d04a5a3265d8e267953ca9800894b76044221103"
  
lemma MQV_P256_TV12_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV12_d1U MQV_P256_TV12_Q1U 
                       MQV_P256_TV12_d2U MQV_P256_TV12_Q2U 
                       MQV_P256_TV12_Q1V MQV_P256_TV12_Q2V 
    = Some MQV_P256_TV12_Z"
  by eval

lemma MQV_P256_TV12_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV12_d1V MQV_P256_TV12_Q1V 
                       MQV_P256_TV12_d2V MQV_P256_TV12_Q2V 
                       MQV_P256_TV12_Q1U MQV_P256_TV12_Q2U 
    = Some MQV_P256_TV12_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV13›
  
definition MQV_P256_TV13_d1U :: nat where
  "MQV_P256_TV13_d1U = 0xe6c18947a70bf438b13fb655678da09e620d8dd9d55df42a042bc8e0cc9c7d2a"

definition MQV_P256_TV13_Q1U :: "int point" where 
  "MQV_P256_TV13_Q1U = Point
    0xcae68c1b99e878b2dd2e968a48b0efd7c21e53d515d28c84d649d67dbd1d08f6
    0xe7c5a518c49638a10ac079c9af86d6d4a9371ebca4507f0ded98b2a77f86c836"

definition MQV_P256_TV13_d2U :: nat where
 "MQV_P256_TV13_d2U = 0x3a91ffd58944362ef6b824e3a461dbe8288e2068366ad6495d7f8be8deb32ebe"

definition MQV_P256_TV13_Q2U :: "int point" where 
 "MQV_P256_TV13_Q2U = Point
    0xf44a4df6162236c30fb6c77782e3685847b763b9b8abfae82e147c79f9d3c833
    0xe4d2672f38b6fce352179172895100a80de77a06833c66c7447270173e3db783"

definition MQV_P256_TV13_d1V :: nat where
 "MQV_P256_TV13_d1V = 0x9ce923865b8ca91851b078ce1edeedd1762bace92c0644dbd413ed577e387717"

definition MQV_P256_TV13_Q1V :: "int point" where 
 "MQV_P256_TV13_Q1V = Point
    0xbef984854f781862838aad0108af8ce9db71d1d969daf7281f5f45fb8581d3fc
    0xd08c3bf2d35d9b9344dfba6d1d5d0d46cf549e18a28d0137bc6eaea7ced40eaf"

definition MQV_P256_TV13_d2V :: nat where
 "MQV_P256_TV13_d2V = 0x2df53cfdaf43917308febd88f27c47171dcc623ef41448350a88789100d26b0c"

definition MQV_P256_TV13_Q2V :: "int point" where 
 "MQV_P256_TV13_Q2V = Point
    0x40e69f171d08639f8cac036e0dc14e476739dc438ef6230d0b59a613131cb937
    0xa96e8306affdc8e34e6a3052bf2a97c84da4a91ae4be4f07a26f2e6fb2ad3629"

definition MQV_P256_TV13_Z :: nat where
  "MQV_P256_TV13_Z = 0xe8bad7a353c78eabbaf99ccd2374de21bd7567a167580e243756bd45ab3804bd"
  
lemma MQV_P256_TV13_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV13_d1U MQV_P256_TV13_Q1U 
                       MQV_P256_TV13_d2U MQV_P256_TV13_Q2U 
                       MQV_P256_TV13_Q1V MQV_P256_TV13_Q2V 
    = Some MQV_P256_TV13_Z"
  by eval

lemma MQV_P256_TV13_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV13_d1V MQV_P256_TV13_Q1V 
                       MQV_P256_TV13_d2V MQV_P256_TV13_Q2V 
                       MQV_P256_TV13_Q1U MQV_P256_TV13_Q2U 
    = Some MQV_P256_TV13_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV14›
  
definition MQV_P256_TV14_d1U :: nat where
  "MQV_P256_TV14_d1U = 0xd64c934d05ef37571f4aebd2831389c04cb0d8093208491292653b31840b3a59"

definition MQV_P256_TV14_Q1U :: "int point" where 
  "MQV_P256_TV14_Q1U = Point
    0xf72a114cc876348dc6082ae96ea0455182ddf4c3e71994487c4665fc8dbc4cfa
    0x31f35ae9978c26469b9944e3229ed9958c6c101e0a77071fcadb14529e767998"

definition MQV_P256_TV14_d2U :: nat where
 "MQV_P256_TV14_d2U = 0xcbffc7e152323133ecfbdb557cb33fc65a095385b68121e88f8a8bb205d08afa"

definition MQV_P256_TV14_Q2U :: "int point" where 
 "MQV_P256_TV14_Q2U = Point
    0x7465c547760530fa5fbe8ea2bedbdb6da00ff734955549a0f8366bb2d4272429
    0x0e8e275941490c0334fba37c816b568cb9dcd5b82b15a1e925a2ded666bd65d1"

definition MQV_P256_TV14_d1V :: nat where
 "MQV_P256_TV14_d1V = 0x42ba216c3db6b8fa8e133fbd60ee2a79b3304c1212241907cdaecaa65ef8e21a"

definition MQV_P256_TV14_Q1V :: "int point" where 
 "MQV_P256_TV14_Q1V = Point
    0xef2155bbbb1c58e2c57cd9187009a09b99f30fee101c399e339fd23c6958b273
    0x653e519c0bf1d1d60f8aa4f24b58cddaa1f4ac915c512eaa2288f564de605680"

definition MQV_P256_TV14_d2V :: nat where
 "MQV_P256_TV14_d2V = 0xce185c4af081fe157ba1fa2df71836a49fbde99707f8de6461409dfbf8c97ecc"

definition MQV_P256_TV14_Q2V :: "int point" where 
 "MQV_P256_TV14_Q2V = Point
    0x0cafb764746c282eb011ad7fe8bb297d1f369e684fe6b9bbf47a1d6ba83847f7
    0xe7e45ebdbebe76fb1f7e701f57f08fb72965d1312e85476780cc1bcec5574a00"

definition MQV_P256_TV14_Z :: nat where
  "MQV_P256_TV14_Z = 0x06a7c4ee54e247904d287f536d8d97f3f7af757458008f5ca758859a83b63cdd"
  
lemma MQV_P256_TV14_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV14_d1U MQV_P256_TV14_Q1U 
                       MQV_P256_TV14_d2U MQV_P256_TV14_Q2U 
                       MQV_P256_TV14_Q1V MQV_P256_TV14_Q2V 
    = Some MQV_P256_TV14_Z"
  by eval

lemma MQV_P256_TV14_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV14_d1V MQV_P256_TV14_Q1V 
                       MQV_P256_TV14_d2V MQV_P256_TV14_Q2V 
                       MQV_P256_TV14_Q1U MQV_P256_TV14_Q2U 
    = Some MQV_P256_TV14_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV15›
  
definition MQV_P256_TV15_d1U :: nat where
  "MQV_P256_TV15_d1U = 0x766fbeccea82e4d7493d463fc68e08ab947ea662bef7eab74f9ae424cbcf3c18"

definition MQV_P256_TV15_Q1U :: "int point" where 
  "MQV_P256_TV15_Q1U = Point
    0xb249fda20bb1c2fc1919667e3b58c070782abb3c57f014318f153d2baa32195b
    0x9b45a0a983799ad444cec5638c99adff2d8ed9b897a28e59c941be31958d5aca"

definition MQV_P256_TV15_d2U :: nat where
 "MQV_P256_TV15_d2U = 0x06037b7decf471fea1e4c62a456bad71f46fbf580e522eabbaee07620457b765"

definition MQV_P256_TV15_Q2U :: "int point" where 
 "MQV_P256_TV15_Q2U = Point
    0xad75972e4640c25199c054878dc33eb1b98771ba9ef2259ee8492b9408a2506b
    0x8217f76ce9e3f9cc86ac840b3014bd78ab7c3b34a38cc98ecbbee0de9114c742"

definition MQV_P256_TV15_d1V :: nat where
 "MQV_P256_TV15_d1V = 0x5bf794147a7822bd01610a01c72e345a59da49157f6a42544d63a176b5a9f2a2"

definition MQV_P256_TV15_Q1V :: "int point" where 
 "MQV_P256_TV15_Q1V = Point
    0xfcf225e5016314ad08f5d308baee9ea453eb150977741cebc0fb6d372e8bef9f
    0xdee91f3016d919deeaf3127e7c53bbbaa844af879c0ffed3a5455be9d454f285"

definition MQV_P256_TV15_d2V :: nat where
 "MQV_P256_TV15_d2V = 0xd30e6a1a551fc34eb634c21ef2c9bb188819e5d2e6b4a6dc9e66500d033094c7"

definition MQV_P256_TV15_Q2V :: "int point" where 
 "MQV_P256_TV15_Q2V = Point
    0xa8467e1f1484431a1d0e31524b24f3b7255d59b38e675b6efd6cc9053a9352c9
    0x41cc4d4c80bff794d5b9eaccbaa08c65a4a61257f580ed4a2c70b0e355b069a1"

definition MQV_P256_TV15_Z :: nat where
  "MQV_P256_TV15_Z = 0x0575548fc7ebc102c75eb80d78a16e5d8621db029c9fdcd152c3f9b0c797e476"
  
lemma MQV_P256_TV15_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV15_d1U MQV_P256_TV15_Q1U 
                       MQV_P256_TV15_d2U MQV_P256_TV15_Q2U 
                       MQV_P256_TV15_Q1V MQV_P256_TV15_Q2V 
    = Some MQV_P256_TV15_Z"
  by eval

lemma MQV_P256_TV15_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV15_d1V MQV_P256_TV15_Q1V 
                       MQV_P256_TV15_d2V MQV_P256_TV15_Q2V 
                       MQV_P256_TV15_Q1U MQV_P256_TV15_Q2U 
    = Some MQV_P256_TV15_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV16›
  
definition MQV_P256_TV16_d1U :: nat where
  "MQV_P256_TV16_d1U = 0xf36cd3389c3e762f0c34ed5e2d2e8f433f9459dcd2d8003a7596989c8299cfe2"

definition MQV_P256_TV16_Q1U :: "int point" where 
  "MQV_P256_TV16_Q1U = Point
    0xa0b91f81d9cc493738c4a1f5dea971612ff0e97fbe944702c227d421598f5021
    0x71f2c084de10e941548f2e0586a6757f12fe6f00f632ae0dba2617775b98ca68"

definition MQV_P256_TV16_d2U :: nat where
 "MQV_P256_TV16_d2U = 0x1763c844b0adcadda4f96e15af3a68b9110c286ea1392b4c7ff836dbee71a439"

definition MQV_P256_TV16_Q2U :: "int point" where 
 "MQV_P256_TV16_Q2U = Point
    0x7a10c201daa92ca77e18d378753500ded61795a12d1090133f74c53fd41974e0
    0x2edb6a336b96a745cafe7c298d0309a9d19b04993679222f74dd96129bdcbd0f"

definition MQV_P256_TV16_d1V :: nat where
 "MQV_P256_TV16_d1V = 0x36aacc729e5a42bc18bd7d336f8318f1de3a3011278c0da66c634f9919535e5b"

definition MQV_P256_TV16_Q1V :: "int point" where 
 "MQV_P256_TV16_Q1V = Point
    0xe9434eefaa4bd149b99e55f4003ad35a114589669034d58cf54ff56776f54199
    0x857f5d30b6576ea55dad7f5eb5605b35bf244adb1be4168a6fb057503a41a170"

definition MQV_P256_TV16_d2V :: nat where
 "MQV_P256_TV16_d2V = 0x01627ebe44a15e6021cd482ba678e193edd3f2767abc14e029cf345d375473ed"

definition MQV_P256_TV16_Q2V :: "int point" where 
 "MQV_P256_TV16_Q2V = Point
    0xfad7f0b62c4412531313421166a18efbb580c2e2d7f29f58f4ee1852ce3c1eb6
    0xe53905996ec7818f8915efd9f13c09ea5b20606c81c260887a277e720ea5476b"

definition MQV_P256_TV16_Z :: nat where
  "MQV_P256_TV16_Z = 0xc648ae68287f721b6bd60207a89448d6a59fc4049ca186ac668f9079fe742d11"
  
lemma MQV_P256_TV16_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV16_d1U MQV_P256_TV16_Q1U 
                       MQV_P256_TV16_d2U MQV_P256_TV16_Q2U 
                       MQV_P256_TV16_Q1V MQV_P256_TV16_Q2V 
    = Some MQV_P256_TV16_Z"
  by eval

lemma MQV_P256_TV16_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV16_d1V MQV_P256_TV16_Q1V 
                       MQV_P256_TV16_d2V MQV_P256_TV16_Q2V 
                       MQV_P256_TV16_Q1U MQV_P256_TV16_Q2U 
    = Some MQV_P256_TV16_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV17›
  
definition MQV_P256_TV17_d1U :: nat where
  "MQV_P256_TV17_d1U = 0xe82ad044210a9581c966661d6174cfa32c1acf22c29db25c71283fec7e4be698"

definition MQV_P256_TV17_Q1U :: "int point" where 
  "MQV_P256_TV17_Q1U = Point
    0x69e58cb258201b768edd3db3fab955a424a1c6f20e90dae74f2206b78154a88d
    0x73f364b47a1628e3225884063349a207458ec304bf71faa06cfd0df89d49988e"

definition MQV_P256_TV17_d2U :: nat where
 "MQV_P256_TV17_d2U = 0x113dd9a10e78cea11a8e96b6be4adc98d9b5e1cfc3729ba281110fd77efebe92"

definition MQV_P256_TV17_Q2U :: "int point" where 
 "MQV_P256_TV17_Q2U = Point
    0x4e7a7e8d4ecb07b0041c5165a4d3c6c98b5754ee9dba249a3466240817c9c031
    0x449d99149b39f74316c06901722d61c76764721504b522d94096fdd3de505ec5"

definition MQV_P256_TV17_d1V :: nat where
 "MQV_P256_TV17_d1V = 0x4d0cac157968014c6fc327c404dc480f653714f401ca18f52fe7160f2e7240e9"

definition MQV_P256_TV17_Q1V :: "int point" where 
 "MQV_P256_TV17_Q1V = Point
    0x7c6361b4ee3b80c7d6f5d155210c3e9a3a4550b2bbd109f269e0f1ef86eaaa5c
    0x0d7cc78306c62e42051d2931f6be4f3dfbd5a0b009fc33b125f1fdaca3b10f44"

definition MQV_P256_TV17_d2V :: nat where
 "MQV_P256_TV17_d2V = 0xa0697795bc9898d8cc4e05cf0b97295c5881919f50ee058d809df1288a13c8f9"

definition MQV_P256_TV17_Q2V :: "int point" where 
 "MQV_P256_TV17_Q2V = Point
    0x9c7528573e2056e29480619ab97032becf8e4d3d6ad9268ba134ab3442c61758
    0x6a716b37f934208eaa7e3617bedafd42170879d76ff73a2b2bfa59fa60e86b5f"

definition MQV_P256_TV17_Z :: nat where
  "MQV_P256_TV17_Z = 0x2dd151c22ad5a0653bcf01ef5adc5797faa0bca5536321598e6a8c8920fc27ad"
  
lemma MQV_P256_TV17_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV17_d1U MQV_P256_TV17_Q1U 
                       MQV_P256_TV17_d2U MQV_P256_TV17_Q2U 
                       MQV_P256_TV17_Q1V MQV_P256_TV17_Q2V 
    = Some MQV_P256_TV17_Z"
  by eval

lemma MQV_P256_TV17_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV17_d1V MQV_P256_TV17_Q1V 
                       MQV_P256_TV17_d2V MQV_P256_TV17_Q2V 
                       MQV_P256_TV17_Q1U MQV_P256_TV17_Q2U 
    = Some MQV_P256_TV17_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV18›
  
definition MQV_P256_TV18_d1U :: nat where
  "MQV_P256_TV18_d1U = 0xe06fa0e02920d610a2bb94da1f47292113e4700669eb060b5d4bef9ba331c20f"

definition MQV_P256_TV18_Q1U :: "int point" where 
  "MQV_P256_TV18_Q1U = Point
    0x3b21168580c2df814e6cf01ae16eb2bfeb01def5c7152443474a6f15609c2ed5
    0xeddb634c91a176f0ed15930595cf8b33a99c22bbb3360d6405f6d05cc1236a6e"

definition MQV_P256_TV18_d2U :: nat where
 "MQV_P256_TV18_d2U = 0x3efc55b558d942079217fe8dd0046571fd82004e2d405f08fbb658d41175e8b4"

definition MQV_P256_TV18_Q2U :: "int point" where 
 "MQV_P256_TV18_Q2U = Point
    0x0824503f4ab058e676e06aadfab61ed07f4e67094f758866057920fedccb6123
    0xfb10ad908395cfecbdffbe1e0a94702392ef42ae400b9787c7f0628650267e03"

definition MQV_P256_TV18_d1V :: nat where
 "MQV_P256_TV18_d1V = 0xdf03e4bce5b45ae5f211201f6747bcf50024363ac9fdc94e624af4214aaf0f4c"

definition MQV_P256_TV18_Q1V :: "int point" where 
 "MQV_P256_TV18_Q1V = Point
    0xb82f3eb2e439dc59ca0654d518a10f28571427f986b917613e770042dbdeff29
    0x64ff698fb6be36f8845fd4368241057673e6fd2c7727d389682e16bd033c6d1d"

definition MQV_P256_TV18_d2V :: nat where
 "MQV_P256_TV18_d2V = 0xa80638ef5de90cd55cc3e67823997dee2bf7eff577281fe479e2be6ba9917f25"

definition MQV_P256_TV18_Q2V :: "int point" where 
 "MQV_P256_TV18_Q2V = Point
    0x7baf80f1af176b7b5d43ee3e37e8e9f2aaea6bc2d97f6b6ecccdffe84a166caa
    0xd4e6f86d82deb7dc2ed2e3eddbe3a48dd72a6eab8b8d6cd1b2c218a192769522"

definition MQV_P256_TV18_Z :: nat where
  "MQV_P256_TV18_Z = 0x6551cdc8228fe0a1a267ed15f37a6b437feb1461c37a579c05d1bb861dd73f5f"
  
lemma MQV_P256_TV18_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV18_d1U MQV_P256_TV18_Q1U 
                       MQV_P256_TV18_d2U MQV_P256_TV18_Q2U 
                       MQV_P256_TV18_Q1V MQV_P256_TV18_Q2V 
    = Some MQV_P256_TV18_Z"
  by eval

lemma MQV_P256_TV18_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV18_d1V MQV_P256_TV18_Q1V 
                       MQV_P256_TV18_d2V MQV_P256_TV18_Q2V 
                       MQV_P256_TV18_Q1U MQV_P256_TV18_Q2U 
    = Some MQV_P256_TV18_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P256_TV19›
  
definition MQV_P256_TV19_d1U :: nat where
  "MQV_P256_TV19_d1U = 0xac842b3b4536b4fbfa5fda022cfb0bd7ca8873755e7003848328a59538d359ce"

definition MQV_P256_TV19_Q1U :: "int point" where 
  "MQV_P256_TV19_Q1U = Point
    0x9b2d86bf16c362667ed5b6023d891e79ecc05b94bfb9c39bf94de4b04551cb63
    0x1e19149771b01808dc832ba5a2793f113182cb716141e45fcc5f6befc48b5ada"

definition MQV_P256_TV19_d2U :: nat where
 "MQV_P256_TV19_d2U = 0x601e9c443a0c2951138af96ba4307e5182cd3b1da578a81b6e0d9339459e18ce"

definition MQV_P256_TV19_Q2U :: "int point" where 
 "MQV_P256_TV19_Q2U = Point
    0xa9b137f1ff5f5c16efc634f5f122d436f7e2afc942c47ebb9ff93104f72ea27e
    0x994dc3c7e0c617d4e1a5e962f7d302f107721b45ed5b29f4bf31dd199ce30b0e"

definition MQV_P256_TV19_d1V :: nat where
 "MQV_P256_TV19_d1V = 0x672d260134524d7a2d723858b65a6d9f6da43174739dda044eb3f32f6d5bb631"

definition MQV_P256_TV19_Q1V :: "int point" where 
 "MQV_P256_TV19_Q1V = Point
    0xe4382b5d341459666431abca8b3308e3802b60c67845a14559556945bfdbaca8
    0x01269a18401c01eebf4e176c04b61f08a8f4936d73504e98146f9a43e06fedd9"

definition MQV_P256_TV19_d2V :: nat where
 "MQV_P256_TV19_d2V = 0xbf847f7ebda73b8e1e5ce099ad0baef441ca21d837db5b172e9307b6284056fe"

definition MQV_P256_TV19_Q2V :: "int point" where 
 "MQV_P256_TV19_Q2V = Point
    0xee0e7bd9dabb5c58ffebf6ef0904b0916c7b49c216c97f245a1e6890d2ccc71d
    0x4ef0a3f662027c25d4578b27714be909214ea2a9162d764ba14e5566514f5419"

definition MQV_P256_TV19_Z :: nat where
  "MQV_P256_TV19_Z = 0xbafacc24c7942f470396560b80964014d9ffeef2de43282ee40e9a489b90d6a0"
  
lemma MQV_P256_TV19_MQV_check: 
  "SEC1_P256_ECMQVprim MQV_P256_TV19_d1U MQV_P256_TV19_Q1U 
                       MQV_P256_TV19_d2U MQV_P256_TV19_Q2U 
                       MQV_P256_TV19_Q1V MQV_P256_TV19_Q2V 
    = Some MQV_P256_TV19_Z"
  by eval

lemma MQV_P256_TV19_MQV_check': 
  "SEC1_P256_ECMQVprim MQV_P256_TV19_d1V MQV_P256_TV19_Q1V 
                       MQV_P256_TV19_d2V MQV_P256_TV19_Q2V 
                       MQV_P256_TV19_Q1U MQV_P256_TV19_Q2U 
    = Some MQV_P256_TV19_Z"
  by eval  
  
section‹NIST Test Vectors for MQV primitive: Curve P-384›
  
subsubsection‹MQV test vector: MQV_P384_TV00›
  
definition MQV_P384_TV00_d1U :: nat where
  "MQV_P384_TV00_d1U = 0x33305db8bf8d49945970129b1a977cf353d30811118c2066bc749d36876315420fd13542ed11aa062ad371f0e92ea12b"

definition MQV_P384_TV00_Q1U :: "int point" where 
  "MQV_P384_TV00_Q1U = Point
    0x25dd0644367aa39edcdd1d2d3ac1fde435acfc44029988874f8e143cdbfe2d4cde22c9f0a7f6b3df80ec32ecafccbb1a
    0xd5d3ca28f315d4a1dccbdb4c175ccefefb3625b6120dc073e5ba2c662924df03d0d7db5d6ca54be373f22ba83a12229c"

definition MQV_P384_TV00_d2U :: nat where
 "MQV_P384_TV00_d2U = 0x90751da387bcc90b4c646f92ba2d4c66db35cc8e29f6120b81a1967c189c7ff6f192c69e04bfe1b69604519af6941504"

definition MQV_P384_TV00_Q2U :: "int point" where 
 "MQV_P384_TV00_Q2U = Point
    0xfcb52129a29e2242ad65badbf736b292cf5e0efc6524e5b123b2db7f413fe5e6e0e5605a5dc534e8793030bfe0802c8d
    0x2b8f27af37dc033bff53584a5f5d5ddbfc2c8d0a672bcf9db3388d2144692b299b4df8f0b647ac0d173d738a4204d93d"

definition MQV_P384_TV00_d1V :: nat where
 "MQV_P384_TV00_d1V = 0xe692498288a267355b042a4f39e3acdbd9b341050c5a7b1d1e79e7b32b0355dda8973a4a3727c7b8018c00c2e0defbcd"

definition MQV_P384_TV00_Q1V :: "int point" where 
 "MQV_P384_TV00_Q1V = Point
    0xc6eb6b79142436117a81b91b8d348794f04c6ca4749e548d87e278c1872f64a5be5f554a9cf50f8db9252934b06e7da8
    0x1768c0c31279772aa41bc313713bf94c83b1bf8208013bf55d907f1fb02357c63522ec1d2f53021939e167cea68b2bb6"

definition MQV_P384_TV00_d2V :: nat where
 "MQV_P384_TV00_d2V = 0x9d8d86ac502ff78e1e7572583d76e63f69c5cd0a1320974fede4502b3adcae3bc2d1168c50ba31c569e81f4fca69fa09"

definition MQV_P384_TV00_Q2V :: "int point" where 
 "MQV_P384_TV00_Q2V = Point
    0x378a64380745adc000190f04302162e4c547b09fa7c9b2e52d97bb6f0d7bea0c1dd6f3003f934850568729ff7d9b9c21
    0x215cce1d03669e91b703df0e81d3d4ccb7cf7a5856a188260afb871b2e676b882eb236f62a5009852389ae3c216b7cea"

definition MQV_P384_TV00_Z :: nat where
  "MQV_P384_TV00_Z = 0x0a4c87bdce2f2ac4161e1a3379302205e6257e953927a49d94d140429a7c3d8319e288c809f85f509008512323848eab"
  
lemma MQV_P384_TV00_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV00_d1U MQV_P384_TV00_Q1U 
                       MQV_P384_TV00_d2U MQV_P384_TV00_Q2U 
                       MQV_P384_TV00_Q1V MQV_P384_TV00_Q2V 
    = Some MQV_P384_TV00_Z"
  by eval

lemma MQV_P384_TV00_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV00_d1V MQV_P384_TV00_Q1V 
                       MQV_P384_TV00_d2V MQV_P384_TV00_Q2V 
                       MQV_P384_TV00_Q1U MQV_P384_TV00_Q2U 
    = Some MQV_P384_TV00_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV01›
  
definition MQV_P384_TV01_d1U :: nat where
  "MQV_P384_TV01_d1U = 0x969cc08e31671db82c5e57fd72cfd234a37240f1cb6502fbdba88fc68c40ab3d6caa9aed6b5d1263c26d3163fef1a9fb"

definition MQV_P384_TV01_Q1U :: "int point" where 
  "MQV_P384_TV01_Q1U = Point
    0x858073605c52b466b5baa3c9eac858f41c59889fc0a16dcf06be4ecfa92accfafb93b79683cbacd26405b3e58df4172b
    0x103b42cfa21f1bfdd640fbdd8527fcbb629e63ab2eddcb5eb8cbb436b1f540a34405bed020fb9725ca8e3e3b4d853533"

definition MQV_P384_TV01_d2U :: nat where
 "MQV_P384_TV01_d2U = 0xd8ceda2da5d5be309c8676c6479bbb18f09d7d0df173515b70004a790bce4adb2cafdea079b911760f2c80c697c7c766"

definition MQV_P384_TV01_Q2U :: "int point" where 
 "MQV_P384_TV01_Q2U = Point
    0xb85e3b85c86ce6d0da804bfb93a990b239ea345589ba89d6d83c08492040c50ae650cf0083a224132c0c3859f9bda483
    0x39225c4837c702f1c279ee36e99629fcf209c8372deaa7a49ba174004d9e0ad701bf6f258094f56f6fbdb23914eb16cc"

definition MQV_P384_TV01_d1V :: nat where
 "MQV_P384_TV01_d1V = 0x6cf3d58f8890de56f5085389b349af60f30a2cc1cc5a68c1943d83487a53748555674b97762e1ebde3728f07e82cdfb7"

definition MQV_P384_TV01_Q1V :: "int point" where 
 "MQV_P384_TV01_Q1V = Point
    0xf89008b0e1d161ee4d8143506b2559bced63d8ec694806871d9476f51ec15b2ad8b53c61568f4e3b173e582c14ead2b1
    0x255e960f14dca265ca395faaf58293ede00e69f8dd85a3dada30020bedd57511f4ace7d8771b24c93242ae4cd803f4bc"

definition MQV_P384_TV01_d2V :: nat where
 "MQV_P384_TV01_d2V = 0x0c0c0c45ebee5d298e546e7dabbb7f4f029d54f6f19d0152f763d780549dbf3187bee2ae1e3b8342fe7b3fbd39d4ba30"

definition MQV_P384_TV01_Q2V :: "int point" where 
 "MQV_P384_TV01_Q2V = Point
    0xcc546da698b217a10aaddd6fbe93ca84ddbc8a2f6fa6d2afce23273e6f8a0e10ef730522c93a3827c9cb0fb431dc9b08
    0x28150b8be8cd6b576b27fbd491d257fbda2054ef6f4fe030546012bff7c806bb514f04a619293ac4111ff2d65ed1e0bc"

definition MQV_P384_TV01_Z :: nat where
  "MQV_P384_TV01_Z = 0x52bc24236e8a17d9d5d3fe3ec2d2a8fd488e6c013938586f2412fc3991ddf68b3320707226a2242abe0af2260c1622ca"
  
lemma MQV_P384_TV01_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV01_d1U MQV_P384_TV01_Q1U 
                       MQV_P384_TV01_d2U MQV_P384_TV01_Q2U 
                       MQV_P384_TV01_Q1V MQV_P384_TV01_Q2V 
    = Some MQV_P384_TV01_Z"
  by eval

lemma MQV_P384_TV01_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV01_d1V MQV_P384_TV01_Q1V 
                       MQV_P384_TV01_d2V MQV_P384_TV01_Q2V 
                       MQV_P384_TV01_Q1U MQV_P384_TV01_Q2U 
    = Some MQV_P384_TV01_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV02›
  
definition MQV_P384_TV02_d1U :: nat where
  "MQV_P384_TV02_d1U = 0xa397b71ba549ccfefe86c66312c1fc88dbddddc8d06fdfe053bafaf73a1183f1bc404cdf1218df296ee642a1fa883e91"

definition MQV_P384_TV02_Q1U :: "int point" where 
  "MQV_P384_TV02_Q1U = Point
    0x97daae2c12cdf28cebc5de1bdadd7f54e0db1f28971f7401585165f9f86b0321973639d86822d99671bc33f731c5b41a
    0xc73a705814b46aea0a25627b06b4469e19c1a68736faf2ad1a4c9519eabfa36e7a253bd6eac0285167e467295b0e834e"

definition MQV_P384_TV02_d2U :: nat where
 "MQV_P384_TV02_d2U = 0x34944f989bdabd740d38ee9ceb771518589c972269e4db806c25bfef480d95d8ec1e404a4fffc5f5e2019f3246f9af62"

definition MQV_P384_TV02_Q2U :: "int point" where 
 "MQV_P384_TV02_Q2U = Point
    0x1f87861bfc22069417e7eab975c43a0a586d1362db364b94d249270288da553072a2f95bd6d6488112de626ae3224baf
    0x0ae1dc4ecb0558074a9db65854bb1387a2502f2ddc3f26aae09aef02422ffae08f2168314e0dadb0611248f5b3d2e9a2"

definition MQV_P384_TV02_d1V :: nat where
 "MQV_P384_TV02_d1V = 0x9c300a8eaf837ee0c02d42f551c71c29d1f7a628926158153009d80cc96caea3e4ebd8c4fa6b2f9f07ee8c3606aaf4d9"

definition MQV_P384_TV02_Q1V :: "int point" where 
 "MQV_P384_TV02_Q1V = Point
    0x52d3f43708b7cb397dbad2999dd32fcf24b74866925ac3d2db29a8589637fe9ab0caf2bb6513bacfab8a652343b85cba
    0xbf35cd63b1af642a126bf8d063616b482541f558ed72b13e359df13de8e1eb577ed2c9174828e7e733596826503fda3d"

definition MQV_P384_TV02_d2V :: nat where
 "MQV_P384_TV02_d2V = 0x241b0340c81cd96451c2e244531bbb71f8da6110a28b686df5390816734f10179fe60b83f88795f742f7f8ba2f21d6db"

definition MQV_P384_TV02_Q2V :: "int point" where 
 "MQV_P384_TV02_Q2V = Point
    0xa24be0d19b6ca3cad882aa8ce00b742302551cfcf2f3f965aa757280cda7f65e650fd17a5c5de422195ac677429142b3
    0xa85b96d6c31041417b90402057bddfeb069e74bd500232729b6d1fd84e54ce27a2ef67f6b6e58fae4d40b634280aeeb2"

definition MQV_P384_TV02_Z :: nat where
  "MQV_P384_TV02_Z = 0xcd763a1c494b9387b217b0908cf18c7c3cc7b3afdc83905307c34b1f242492d82bda1abd780b160ad2ba534cd2730084"
  
lemma MQV_P384_TV02_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV02_d1U MQV_P384_TV02_Q1U 
                       MQV_P384_TV02_d2U MQV_P384_TV02_Q2U 
                       MQV_P384_TV02_Q1V MQV_P384_TV02_Q2V 
    = Some MQV_P384_TV02_Z"
  by eval

lemma MQV_P384_TV02_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV02_d1V MQV_P384_TV02_Q1V 
                       MQV_P384_TV02_d2V MQV_P384_TV02_Q2V 
                       MQV_P384_TV02_Q1U MQV_P384_TV02_Q2U 
    = Some MQV_P384_TV02_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV03›
  
definition MQV_P384_TV03_d1U :: nat where
  "MQV_P384_TV03_d1U = 0x951bd95edd147894424bb2dce5619a78cfb9f9aafb00c834d30f268d649b6b00a6a7d015947c08d748f436eb2c3837a2"

definition MQV_P384_TV03_Q1U :: "int point" where 
  "MQV_P384_TV03_Q1U = Point
    0xc7066b9e92bccdcbbc7193c059bf175bde1c952174cdc653d83e985795d83a722a55cfdd45e6bd718bcb13967803286a
    0xb54357c238576297714b22cd2553677a4f6d7611f8b8aa490bfa1bea495f03a4f7796a7eb2a53d7ee69936c558df3b3e"

definition MQV_P384_TV03_d2U :: nat where
 "MQV_P384_TV03_d2U = 0x538e57776900f94db15536656999493ab8e788fc9d36ba96d493842768b3993609c7f8fe0539ea8aa7382f5849142ddd"

definition MQV_P384_TV03_Q2U :: "int point" where 
 "MQV_P384_TV03_Q2U = Point
    0x8443d158373c770d8887517cb1b2c36f09bc4eaebe6ff9ebef3ece7a8480ffa3a22be45a0f238757d530017846e32ab5
    0xdb9619060b644b97133bf9ef988bb6b8bc6ea6509c3572621feadca8059287c1c3f27a12e96bda4dd71a13fc47752ff0"

definition MQV_P384_TV03_d1V :: nat where
 "MQV_P384_TV03_d1V = 0xa5411f9c5b4f0132037dc113fe2042c293f4aa4e689b4cb3e8dca5fb338c989829b9d7405a7cd18e86df8909cf57f745"

definition MQV_P384_TV03_Q1V :: "int point" where 
 "MQV_P384_TV03_Q1V = Point
    0x1d01c6bdf7f19b746062d64d50b40a7cc4125bf67ca5fe853e188346d94d77cb0f3b35f751b972e704a2cb1eebefa6f5
    0x6e5592acfa2340403ee4de01c99550937c1b50e1b842f981ca2b05ac08275f7d839f1bacfc10cb53bb8e18578bb1a2a5"

definition MQV_P384_TV03_d2V :: nat where
 "MQV_P384_TV03_d2V = 0xd8e74352b3f27a6b993a7a45bd0515af0f37b0047f623d538ded927ae82cd8d1c1d703d3f56fd39c72b04c2aa7987f12"

definition MQV_P384_TV03_Q2V :: "int point" where 
 "MQV_P384_TV03_Q2V = Point
    0xf20c016e6a0d70fdc24450c3731e1acfe16ee81f6d4d4bd41d2ca68d23c60318694eba554f65ad2d80a15cc1b10fa664
    0x1c770c30a76c90ece87d14936671209415ae53918e87352bef311684a773f8b36d08ae99ee7f799bec38ade18b6f22d9"

definition MQV_P384_TV03_Z :: nat where
  "MQV_P384_TV03_Z = 0x92e381eadb269e5458136724b31cd7e5230be415fa4754b1b6b06f208798d4b14d97a39492d27d621d13dfd3dae5365d"
  
lemma MQV_P384_TV03_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV03_d1U MQV_P384_TV03_Q1U 
                       MQV_P384_TV03_d2U MQV_P384_TV03_Q2U 
                       MQV_P384_TV03_Q1V MQV_P384_TV03_Q2V 
    = Some MQV_P384_TV03_Z"
  by eval

lemma MQV_P384_TV03_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV03_d1V MQV_P384_TV03_Q1V 
                       MQV_P384_TV03_d2V MQV_P384_TV03_Q2V 
                       MQV_P384_TV03_Q1U MQV_P384_TV03_Q2U 
    = Some MQV_P384_TV03_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV04›
  
definition MQV_P384_TV04_d1U :: nat where
  "MQV_P384_TV04_d1U = 0xbafe123ea3f42e6c4d93e7285822216010ed225f32f090cfd9afecf65d6d1a7eb0b81e860e5f887fbf81d65fffbe8ce5"

definition MQV_P384_TV04_Q1U :: "int point" where 
  "MQV_P384_TV04_Q1U = Point
    0x3415e4b1f7e6b5ba9890f3fe301a65e7947202d4a1e5c9a3520a319a328d810fcfb8420c3e32bd2f804231be6cbf7363
    0xe8c1d459326d1334b235b69eda15269828f7a34f023eae41f7444c1f50c6c0675a5be087dd97ad04abf9788467ee391e"

definition MQV_P384_TV04_d2U :: nat where
 "MQV_P384_TV04_d2U = 0xf7767f6fde7953df2be059d48c1b139f0d23a4de491cbd02ec2abc65ca86dda18c54c07d029686afc05318f8b8645e1f"

definition MQV_P384_TV04_Q2U :: "int point" where 
 "MQV_P384_TV04_Q2U = Point
    0xb492e4729a483b8dcfbbc9d031b99c088a18d3b8ed1f415a97a78bf4fa29ab65a869fdaba7bfb7779a362318e72c2b21
    0xb182579a4daa8c4003088713cba47886a335b97c486b3c6c4f0142cb563330081e581d88d63ec762694fca255a72f3c0"

definition MQV_P384_TV04_d1V :: nat where
 "MQV_P384_TV04_d1V = 0xa7c6137f24e87a679b9c114a4b7b02034517df95f53a3999244ee23ac8e3ba3c7ce3f9f70167a10eb188f44358ce7ae6"

definition MQV_P384_TV04_Q1V :: "int point" where 
 "MQV_P384_TV04_Q1V = Point
    0x14dbfdc1c3670c45475d745613858443c444a424bae52a9bf1537e53ab3ceeb848e20aa200eb22590ade8a09f892fa7e
    0xcc4bff0326fb1a0ca15bb0bea7e3300db3c841e23981a479b6278dcbdf5d640f9d1d9c7f24060cd89dca533a4588a8dd"

definition MQV_P384_TV04_d2V :: nat where
 "MQV_P384_TV04_d2V = 0x7f8c902fd1713125e6bdc32a717d6de8924533ca2296ef726578edad7ba5738e84a1d3bde27ec9f479908c67c7a63827"

definition MQV_P384_TV04_Q2V :: "int point" where 
 "MQV_P384_TV04_Q2V = Point
    0x5e80dcaab77320b55faaaa50b7b0f6b28cb558ac7b7f48e933caa1dddd9f775c7286002e571282abe7a70c03fbcbdc5b
    0x72c2b22162f96c0014f81ae60fa9f135319723ba1d38c33f207b1da510a2758caf7b8cd81bb616b2a669e099d9fdbf37"

definition MQV_P384_TV04_Z :: nat where
  "MQV_P384_TV04_Z = 0xc7a6202b722b3ae165782e1ed57f38a96e1ad65231b229a4fa8343c36b988f7c08b49d1d038d8041492f517ecf4a6dec"
  
lemma MQV_P384_TV04_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV04_d1U MQV_P384_TV04_Q1U 
                       MQV_P384_TV04_d2U MQV_P384_TV04_Q2U 
                       MQV_P384_TV04_Q1V MQV_P384_TV04_Q2V 
    = Some MQV_P384_TV04_Z"
  by eval

lemma MQV_P384_TV04_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV04_d1V MQV_P384_TV04_Q1V 
                       MQV_P384_TV04_d2V MQV_P384_TV04_Q2V 
                       MQV_P384_TV04_Q1U MQV_P384_TV04_Q2U 
    = Some MQV_P384_TV04_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV05›
  
definition MQV_P384_TV05_d1U :: nat where
  "MQV_P384_TV05_d1U = 0xe91a665c8ad862360c74ce95d87e3bb8fb5eb72116fcd6427167c17a9459eb69fa536da50a50e974a188016a2fb4b5fd"

definition MQV_P384_TV05_Q1U :: "int point" where 
  "MQV_P384_TV05_Q1U = Point
    0x19684bfc80569689dd69c2596f5abd56b8fdd1eb1da9c485b775f2567b9483e44a590824a25f68408b0d4aefb1c6c5ec
    0xe9984b7ddbe14ae2ff8871116f6c4d1f982b884991610aee42bf418a41d3b3578461597c42f4c13e643ee8186cf13763"

definition MQV_P384_TV05_d2U :: nat where
 "MQV_P384_TV05_d2U = 0xf9aacde469b637115a97204ff8fc2add2a9d9a08980f3736eb56fa75ff909f0fd01fc22c19cc30649fa2924160e891bc"

definition MQV_P384_TV05_Q2U :: "int point" where 
 "MQV_P384_TV05_Q2U = Point
    0x8a4231f37bae9b438e0f009701402fd49e5badfd81dcae7ee097e6b4cd514d1ed54360c46a89e12420f481a9a9036262
    0x9650e27536577f479c1e3e03d4909bc83c2ee22f34bef4c9e4cb4e28a0208fab1fc64fe9860bed2b3388670d26371be1"

definition MQV_P384_TV05_d1V :: nat where
 "MQV_P384_TV05_d1V = 0xe35bbdc248417b6b47936de71b53e18b3ffd1656249ae57f81b349017550dcc4a6a37aface68db87b2d280655d1bf5e4"

definition MQV_P384_TV05_Q1V :: "int point" where 
 "MQV_P384_TV05_Q1V = Point
    0x550d3c053234004c5df484b47905d9ba84a8b49cfb6aa86133736ce9be2e2669f050f0358c9eedd089b9900f08a60f3e
    0x67e13d6834b40ef4cbe5d190cf6914fd63b39dd97c39df7e20792b8f39d1af4ebeef9cff8ab0b8b9258d8d3fe7fe2641"

definition MQV_P384_TV05_d2V :: nat where
 "MQV_P384_TV05_d2V = 0x1cb15c7a91022ab8b20ffae55999833427e2e7d049bd310e7897829b77b95ff3fb6d012fade05bbde409bafc2ee5a0cb"

definition MQV_P384_TV05_Q2V :: "int point" where 
 "MQV_P384_TV05_Q2V = Point
    0xf9e7dafaa7417f9326c5f5465f55c4ecb5b1810f29993e60826434760b2a01d189fc8dbcdb7832d20aba952606fb2964
    0x39612cb0a7fbc88af07013659b419f3be45321277127c49a38658610fd784e505fcb3149ffd709b406e5148dbd861674"

definition MQV_P384_TV05_Z :: nat where
  "MQV_P384_TV05_Z = 0x089719bdae9d7b46ed3a84b0cbf2dff32e00eac64f2eccb4f638c29f61bfc2192dc382a7152ec35722adfeb56af7f64d"
  
lemma MQV_P384_TV05_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV05_d1U MQV_P384_TV05_Q1U 
                       MQV_P384_TV05_d2U MQV_P384_TV05_Q2U 
                       MQV_P384_TV05_Q1V MQV_P384_TV05_Q2V 
    = Some MQV_P384_TV05_Z"
  by eval

lemma MQV_P384_TV05_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV05_d1V MQV_P384_TV05_Q1V 
                       MQV_P384_TV05_d2V MQV_P384_TV05_Q2V 
                       MQV_P384_TV05_Q1U MQV_P384_TV05_Q2U 
    = Some MQV_P384_TV05_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV06›
  
definition MQV_P384_TV06_d1U :: nat where
  "MQV_P384_TV06_d1U = 0xc83bb4e1c7b6edef808ddaf655511338089a7180d22972242d4b672cbefcdadf5136e6d2e452f4fcdf86ac0fce3e03d8"

definition MQV_P384_TV06_Q1U :: "int point" where 
  "MQV_P384_TV06_Q1U = Point
    0x2390d246c4357f4495e7e850480872e9b347de2a0959fb9e2b0fcd0dce42353819fd4ec724a789ad65e684fa9f16e732
    0xb7f858bea058892f16f49e240fb0e28c3ffc65548bbbbc26da77ef25762a3fe902f95fa39436565ef22233f8cff611b9"

definition MQV_P384_TV06_d2U :: nat where
 "MQV_P384_TV06_d2U = 0x0675ba80ccf9b65b88a252df4f5ecb48b88ddd343f7ea3e7dee4e68a4498c573989ce003244b84c5f2c97d3e5c944ecd"

definition MQV_P384_TV06_Q2U :: "int point" where 
 "MQV_P384_TV06_Q2U = Point
    0xc0ad593b4a93237d071ed8d3ee1ba5de5f36d1ac06ce3e2a76a94438c86b5e675222a0cd7ca847026c1f9324b17ff966
    0x12d95ef66245fff8b03043e1b86fe8d54c67adc5d0397fdbd4a45fee3603d2d00395c75e388ec736b738aec46d08b26b"

definition MQV_P384_TV06_d1V :: nat where
 "MQV_P384_TV06_d1V = 0x0440a32c2ac8831a74c16ed0b3cc99634150bd99821ae7b2a87112557479204a5f43f2e8d2c199672a6e7b47e4877522"

definition MQV_P384_TV06_Q1V :: "int point" where 
 "MQV_P384_TV06_Q1V = Point
    0xfcf679947f06a6ef0cbd44eaf269c688a346e46e2c9afa3e0a4038775dcac580b969ec96243215182ec4d03a4ae22889
    0xf9cf83cba604dabd82f9da0fa553ebeec6f038f5e8f3ca21dc83aa3d501d868a50d2c6396f52306fa7efb291b548d98f"

definition MQV_P384_TV06_d2V :: nat where
 "MQV_P384_TV06_d2V = 0x41de402e9115dd94da8a5b9b7f7a4a65cae124fc0f057e6498246f242509cea1f21cbc238684ae543ec3220248298a84"

definition MQV_P384_TV06_Q2V :: "int point" where 
 "MQV_P384_TV06_Q2V = Point
    0xab1fcde1c54ef5b45ff7484a233ce829f72d1adfac3f7e55251aaa96184cfecfb90c9b9bfe364dedf06d092638570362
    0xc0426f30182b30ff5b2c166767cc11df7de90d25fbc1c16b7145f420088b984ce0333131e1358053f2b5ebc1754d61de"

definition MQV_P384_TV06_Z :: nat where
  "MQV_P384_TV06_Z = 0xc811a85722b3fd678930ab30c46c06c5a3b5197f9e43177c68ef704898d886f8faea6418acc3b1557792fa3bc4ad566d"
  
lemma MQV_P384_TV06_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV06_d1U MQV_P384_TV06_Q1U 
                       MQV_P384_TV06_d2U MQV_P384_TV06_Q2U 
                       MQV_P384_TV06_Q1V MQV_P384_TV06_Q2V 
    = Some MQV_P384_TV06_Z"
  by eval

lemma MQV_P384_TV06_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV06_d1V MQV_P384_TV06_Q1V 
                       MQV_P384_TV06_d2V MQV_P384_TV06_Q2V 
                       MQV_P384_TV06_Q1U MQV_P384_TV06_Q2U 
    = Some MQV_P384_TV06_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV07›
  
definition MQV_P384_TV07_d1U :: nat where
  "MQV_P384_TV07_d1U = 0xa244356494f879537d7dfaac6b1f3801092391f2973eb20767b5995de51237bbc7700bb690c3135cf73893bfa3ab0026"

definition MQV_P384_TV07_Q1U :: "int point" where 
  "MQV_P384_TV07_Q1U = Point
    0x84ca957bfbc4a51b0f53a2a6d2b906331e97a1a006b37376e0e9d4f4b41eabecd329bc3a79d8a6c8844772ef2755271d
    0x71a82b1a5b62e6d873b93f73c438c9a795e484722a80240f5912d6150097e7faaaf8688c8446c671061faa8e4085dea6"

definition MQV_P384_TV07_d2U :: nat where
 "MQV_P384_TV07_d2U = 0xfa5c639720a8c88a3d1062a17dc1d858ef3063128af58df35ed5f3fc170d007cfc6d649a508c39e89b564bdaf08f99e8"

definition MQV_P384_TV07_Q2U :: "int point" where 
 "MQV_P384_TV07_Q2U = Point
    0xc540944fcc0bcc3c8d5b8a6582679041b61013f7bd7648e590b35b3e465b08c9ea12fecaa96c26bdbd851fcc831f0cc5
    0x353d660eec49bab8724ff1ff6cb131775151fd50fa270551e4e81820e80cfde9853ca79227dac43aa252f869a9d8dd6b"

definition MQV_P384_TV07_d1V :: nat where
 "MQV_P384_TV07_d1V = 0xb70e412fbe09a0e00d974ed9df8f9a53ae0d905f423dd6683ff95f7ba61b0ac3e303f4f6fd24008a88994ad799590fcf"

definition MQV_P384_TV07_Q1V :: "int point" where 
 "MQV_P384_TV07_Q1V = Point
    0xb680c137914a74ed8c46230474d32a0a77b526827cf47f7d3185955123e13f2319bd87b8efd3a52b5fb02f926eb86295
    0xe5243310ee4c55da5f182a4124e15ddaf52b1e9c3334c102548e9f1864c48bdb32b868cf1e4e6ebd19acc8ce1a7b728d"

definition MQV_P384_TV07_d2V :: nat where
 "MQV_P384_TV07_d2V = 0x7957d6b13a51621e19c7240d712a56bc39e49d1353ea28d8f77fbd12abac879b308e2107c09bb616dba68f08e3febebc"

definition MQV_P384_TV07_Q2V :: "int point" where 
 "MQV_P384_TV07_Q2V = Point
    0x819a798cf16ff6654f2c2a4ff4e136f1f1a97510c012538cdbd597f4c56656d6c268290ad393d0db5fb3ca74c28cd9f3
    0x48bae40ff1b09ffe9141f9128cf583a46b1a1759caeef05766f2573eefcf9f4c3907d37de9a19c28e29c79455d525481"

definition MQV_P384_TV07_Z :: nat where
  "MQV_P384_TV07_Z = 0x0bec8ace219c023bfa23c9b2964711fa08b047302397c8f8c42aff0fc72ced10b12eec662fa953168e22cecd774d2204"
  
lemma MQV_P384_TV07_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV07_d1U MQV_P384_TV07_Q1U 
                       MQV_P384_TV07_d2U MQV_P384_TV07_Q2U 
                       MQV_P384_TV07_Q1V MQV_P384_TV07_Q2V 
    = Some MQV_P384_TV07_Z"
  by eval

lemma MQV_P384_TV07_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV07_d1V MQV_P384_TV07_Q1V 
                       MQV_P384_TV07_d2V MQV_P384_TV07_Q2V 
                       MQV_P384_TV07_Q1U MQV_P384_TV07_Q2U 
    = Some MQV_P384_TV07_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV08›
  
definition MQV_P384_TV08_d1U :: nat where
  "MQV_P384_TV08_d1U = 0xcf3c2abd0e7ad5159315f372665c54f957ce7bb555621b2a62c1ab69cfb2408b7921f43150ba1757c41dcb736a546205"

definition MQV_P384_TV08_Q1U :: "int point" where 
  "MQV_P384_TV08_Q1U = Point
    0x22cc70f83466bfcb587b90fefa6cf3e942a7d8c8449e56ccfb517747845d3576e54bc00fdba465d234eef9290316be46
    0x4dcd26b97a46a5b976d361ca60ab81d4ee05c9db673fd46c07abfec619786f73492d2583e29bf6906e9ae24ee9146f33"

definition MQV_P384_TV08_d2U :: nat where
 "MQV_P384_TV08_d2U = 0x2a11f35a37f8aec5edab4a9507e50e55b40fb18a61ef5598c0810a050b17ce2792adb5e05bd9ea427a415e1ce90fd782"

definition MQV_P384_TV08_Q2U :: "int point" where 
 "MQV_P384_TV08_Q2U = Point
    0x3675f8c56c8f3cea2562b1200645d213fd2360603f2d8e03cb1a1ba251c5ec4311075515dd6d3c27620e16c2f9fff5c5
    0xb7177f26496f5dd0a634b302633dc75d399c056cc7ef0cbce2090ae4a2330cd660950097864ae13501c83b5604c59007"

definition MQV_P384_TV08_d1V :: nat where
 "MQV_P384_TV08_d1V = 0x674f32a3b97b94087b6d3f14c80235cdc20832f30e9f5574fccbe811c6f766491495838bac803bab6e75e00adf6f5148"

definition MQV_P384_TV08_Q1V :: "int point" where 
 "MQV_P384_TV08_Q1V = Point
    0xe1870679f3b9fbb63468e2d6afeb0904190fa84052446349cf5c0eec289881c59e07865568f4212a3cbaca2dffb1f7ac
    0xa6db638de3bc5d53c1da0611e9b974705e445a8eb521b49ee87e67e321e4e229c1438afcdc970a5a19276f59b9f670de"

definition MQV_P384_TV08_d2V :: nat where
 "MQV_P384_TV08_d2V = 0xf6d844fead95d152aa4c253757fe8080cb0d5cc9b473752a7c80d164fd6e1d0ad1508bbeab100f9a27e4aa35c4e2c7da"

definition MQV_P384_TV08_Q2V :: "int point" where 
 "MQV_P384_TV08_Q2V = Point
    0xc774612cf5a2c0cc67644438231757f1b15f8ccd86d6b915950084eb7e61020163f8244271d3f834ec7248e2ff6659e1
    0xa7c9d5e9393420bcbe592b7af0c644dac903771fe437917ab5a437a0cbca198127b104ea1ecf49bd7b5f39f9ada4070a"

definition MQV_P384_TV08_Z :: nat where
  "MQV_P384_TV08_Z = 0xa12a5344669e6a52154e96b9b4f2edbcbdca6dd4958b347c2b597acf23953febcc64d43dde2e5ddcc7e7a1fd3383e316"
  
lemma MQV_P384_TV08_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV08_d1U MQV_P384_TV08_Q1U 
                       MQV_P384_TV08_d2U MQV_P384_TV08_Q2U 
                       MQV_P384_TV08_Q1V MQV_P384_TV08_Q2V 
    = Some MQV_P384_TV08_Z"
  by eval

lemma MQV_P384_TV08_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV08_d1V MQV_P384_TV08_Q1V 
                       MQV_P384_TV08_d2V MQV_P384_TV08_Q2V 
                       MQV_P384_TV08_Q1U MQV_P384_TV08_Q2U 
    = Some MQV_P384_TV08_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV09›
  
definition MQV_P384_TV09_d1U :: nat where
  "MQV_P384_TV09_d1U = 0x4b27c6ab7a81c5230a8dc09089f2edac246aae4754d7cc61087121d3cd6030b246b605ee0d6b656c01f5ea6821ef1f21"

definition MQV_P384_TV09_Q1U :: "int point" where 
  "MQV_P384_TV09_Q1U = Point
    0x3e526cbe9c0d0e3a9116d4adda6b59ecce1458090776c9b83352636707e9d695d9ddae5c7134b3458da95f8ab2b9adfe
    0x3ef225d6002f0a805c1fa6a5cd0061627d96a1531dca028c3c3dd5877621634337047c4d176d93403e2807bb85d214f8"

definition MQV_P384_TV09_d2U :: nat where
 "MQV_P384_TV09_d2U = 0xbd0f29333fc4de518fcb9e28622cbfe82605d069ea38dc9403e9183d45f530edb3e7b01ec85a0563310c1fe767b6a440"

definition MQV_P384_TV09_Q2U :: "int point" where 
 "MQV_P384_TV09_Q2U = Point
    0xf64553d8665e0d434bf8f72212caead7d1c47efb91e2b208fe415330d0a9a3e2119b4c6eb29e424c25db7465831db1a1
    0x490f436e307b9d4a63612ae71716f5ffd5a80e850a04f47120d0e7657e6078b5a6631d98017c5745823fe27466937e2f"

definition MQV_P384_TV09_d1V :: nat where
 "MQV_P384_TV09_d1V = 0xff41f50cf1ce72ee1e7f82eb68438d17b039b2fca8d27d507b19040b1d25e31df3b8221d7c4e85b24bc9b97d2bf32984"

definition MQV_P384_TV09_Q1V :: "int point" where 
 "MQV_P384_TV09_Q1V = Point
    0x3ee718eaaf7506dadd60f73f64167d67b58a9911820d0d227285c0355a679fd7a09daa8514e0fe72baa055c14988d440
    0x841eddc04d4b3d650fa8fc631728a181c68fcbcc3d0344a7ff1d820b72179f6ee31e95683ee5e0c4d021448c03bbad4e"

definition MQV_P384_TV09_d2V :: nat where
 "MQV_P384_TV09_d2V = 0xf90077d1cb015a0bf5decb290fbbea0a916723a4ead7afa867a31082defa5f68b994dafccdfc820ba2fcefa5dc35f981"

definition MQV_P384_TV09_Q2V :: "int point" where 
 "MQV_P384_TV09_Q2V = Point
    0x202310f9af3aac4565d4723f096c5c0bdda17be735f0b76d01ba91c9dac12731f660d0329ad1e367dcb86b90d6c1be2f
    0xade8e0236dacfe0178ef017770f663082f7a323558fa03d8ce0f64ccfb5e5753b19f1e53e9570df5a8b94d489bbf1f4d"

definition MQV_P384_TV09_Z :: nat where
  "MQV_P384_TV09_Z = 0xe9143adfaf99a7331633deac518cd5f4a626590c5e8247cdb205f384a51854ba8fd82be798e26abf47a1a839773ab254"
  
lemma MQV_P384_TV09_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV09_d1U MQV_P384_TV09_Q1U 
                       MQV_P384_TV09_d2U MQV_P384_TV09_Q2U 
                       MQV_P384_TV09_Q1V MQV_P384_TV09_Q2V 
    = Some MQV_P384_TV09_Z"
  by eval

lemma MQV_P384_TV09_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV09_d1V MQV_P384_TV09_Q1V 
                       MQV_P384_TV09_d2V MQV_P384_TV09_Q2V 
                       MQV_P384_TV09_Q1U MQV_P384_TV09_Q2U 
    = Some MQV_P384_TV09_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV10›
  
definition MQV_P384_TV10_d1U :: nat where
  "MQV_P384_TV10_d1U = 0x301f823d8c9426f64df4dcb6b92f12a3cfb736154f1dabf0f664227b53529353a4426cf4b720aa0ecd6fdc464c125bb4"

definition MQV_P384_TV10_Q1U :: "int point" where 
  "MQV_P384_TV10_Q1U = Point
    0x466236b9cc5454c276aea0033c2a96b39295c64fc19d73854112f401e67491c083403d84ca1d63ec0a9d9330e7a4296f
    0xf204b43f51b0110f3c6936d7adc414df6d851d0dfb95d56572ddbedd836971b6ba99ee5def9619fbb7e9c63e060d342c"

definition MQV_P384_TV10_d2U :: nat where
 "MQV_P384_TV10_d2U = 0x2f350c72476b396cee7d2339f34d8ee94486c789652a926318236b583ee9ddcc52dcfd1499b555da7aeaa5f380d44778"

definition MQV_P384_TV10_Q2U :: "int point" where 
 "MQV_P384_TV10_Q2U = Point
    0xc51d454c79144afa4f1371ec8c0290516ab7a73c415e68f9bc475376cb76f7b9eae7c0a69de03db99e445a3d19b6c8ba
    0xf4321421310a834ff174777335f11ecd7f6589555e8a5713d6055024959df89bf94124115422e5a08d76f24e069bc11f"

definition MQV_P384_TV10_d1V :: nat where
 "MQV_P384_TV10_d1V = 0xc5fc651a53bf6ca36fa3d1b3aa1574dbd3dea132da34f8c787caae4481fb246a97ecd54097d75717cabd5b35fdc88ad6"

definition MQV_P384_TV10_Q1V :: "int point" where 
 "MQV_P384_TV10_Q1V = Point
    0x5736f863ecbc0d4bc1218d44aa2eb74f70df690148aeda9e8e1a48b61ec401451d4aa2eaf6113fab0e3908e6c305d127
    0xa664542f6eb3781ed36c4da30837ee915c6f123346a29512f149117ad810c43cb644c5babce71e177ff87179a8c96454"

definition MQV_P384_TV10_d2V :: nat where
 "MQV_P384_TV10_d2V = 0xb13bae3e0c749e1f19d14ca64a3ae08c685cd6b05909457f3564d2d4566a246bdc8ed78fa19659c5ffdee8b17e14fe04"

definition MQV_P384_TV10_Q2V :: "int point" where 
 "MQV_P384_TV10_Q2V = Point
    0x9ed7d485dfb2425730c576c76195c27192d8e2a606ab2faad5176228697990a3c5317cd9dea7097dd8ba69a828815714
    0x079036b1e32e29987e3905c2a19f342ae9c43440d845faa09fbc1c45a933dd2fd18fd291e9345c341622ace0c36263c3"

definition MQV_P384_TV10_Z :: nat where
  "MQV_P384_TV10_Z = 0x6ccc608d1aabe17a6bdf04477d8fc6b7000d4910788f157ed403434c9667b1e366abec46d424565e17f6462ff80a2134"
  
lemma MQV_P384_TV10_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV10_d1U MQV_P384_TV10_Q1U 
                       MQV_P384_TV10_d2U MQV_P384_TV10_Q2U 
                       MQV_P384_TV10_Q1V MQV_P384_TV10_Q2V 
    = Some MQV_P384_TV10_Z"
  by eval

lemma MQV_P384_TV10_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV10_d1V MQV_P384_TV10_Q1V 
                       MQV_P384_TV10_d2V MQV_P384_TV10_Q2V 
                       MQV_P384_TV10_Q1U MQV_P384_TV10_Q2U 
    = Some MQV_P384_TV10_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV11›
  
definition MQV_P384_TV11_d1U :: nat where
  "MQV_P384_TV11_d1U = 0x8c632fef0b06fa0a7b5b9e92043ddf56476bd01edfa022a14926018f0d44694938a513a4e36f8fd6cdf4dad8f6d7de9e"

definition MQV_P384_TV11_Q1U :: "int point" where 
  "MQV_P384_TV11_Q1U = Point
    0x44aeb71fab0a15412a978879adc7032c2cb7cd4e277b12edc364c322af6e1b2e2c9b4ab2d729b7fe39c456e52141f798
    0x61aebd703c2ac56e5dae448d6835ea2d10f7687e270630c825f168121b898f77a3cff6d42c3e3385dbc7e31daa419091"

definition MQV_P384_TV11_d2U :: nat where
 "MQV_P384_TV11_d2U = 0x2feddc4b980c389a2a1b04e98e4b201d7bca4455cd46137ed3063d20b604e55ec322f7d42753aa5295737dd4014f24b3"

definition MQV_P384_TV11_Q2U :: "int point" where 
 "MQV_P384_TV11_Q2U = Point
    0x630bc6ba5da626a3567f5c6412b1fc37bbc644b911d6c0e12450fc7d06a427ed306a50072ec79fde953aff935664b20d
    0x82fac9aa932eadeee21838617f9be6159dda5f3563e6dadecf3580dfff834cca40294589a23f7ba163f9658ceedb16a6"

definition MQV_P384_TV11_d1V :: nat where
 "MQV_P384_TV11_d1V = 0x46a8948dd443ddd5d9d1d2fa644e9be3d11b61c826a1efec0e7ccca7077ddccc09776f951e9b859c260b297ccc212cee"

definition MQV_P384_TV11_Q1V :: "int point" where 
 "MQV_P384_TV11_Q1V = Point
    0x59309eb93faae86656ff189f92d1dc13ca561c124669c003feb65f5c42000c7f4fca2fee330f22440a5f20ab396cdcfa
    0x4eb30fce62b3f41f53f2d79cc37dca7cb3309562a0cfa6f6e136de84c85be0f6293c4c4f7254f9bceada22c6674727a3"

definition MQV_P384_TV11_d2V :: nat where
 "MQV_P384_TV11_d2V = 0x8973eac128a8825ef6a23fd335ff5a9d0af06364ffe861e7680a075b84c75367b5283dc570114131eee1e4bbad1471b3"

definition MQV_P384_TV11_Q2V :: "int point" where 
 "MQV_P384_TV11_Q2V = Point
    0xc8829cb2c9a215a33d8356f2ce744877351ea24400965dd3138abf3a941f22a4bc12090bbb1b29a92b5fb60230647d3d
    0x9ab72c5a0dc39d8ae5f942d95554de1115d3f49dab68274d5fc3fe2594d810796082a57d9ba98997de0ee562a98f72db"

definition MQV_P384_TV11_Z :: nat where
  "MQV_P384_TV11_Z = 0x0764cf9d1710314bf7e9aedc3825c1a5645bca30dbedcbef5c9fdb026c7789d4296d869f8f7fb2ad77159bececc4b3c7"
  
lemma MQV_P384_TV11_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV11_d1U MQV_P384_TV11_Q1U 
                       MQV_P384_TV11_d2U MQV_P384_TV11_Q2U 
                       MQV_P384_TV11_Q1V MQV_P384_TV11_Q2V 
    = Some MQV_P384_TV11_Z"
  by eval

lemma MQV_P384_TV11_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV11_d1V MQV_P384_TV11_Q1V 
                       MQV_P384_TV11_d2V MQV_P384_TV11_Q2V 
                       MQV_P384_TV11_Q1U MQV_P384_TV11_Q2U 
    = Some MQV_P384_TV11_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV12›
  
definition MQV_P384_TV12_d1U :: nat where
  "MQV_P384_TV12_d1U = 0x6aebdddd5ba774b079bfc555a69a96c4aa9cf58137270f2eb583c933bdc91f3c029dc57929a58cb33fbef5319ab6772e"

definition MQV_P384_TV12_Q1U :: "int point" where 
  "MQV_P384_TV12_Q1U = Point
    0xbc91e5e4bba34ebaebd86e6dee5af21672affe4e63711f122452871568ec18eed4029c4d884f6e5ed5cc613aec9b62cf
    0x55e86ed83d78673ea056e505f551d0696a281beb2564857e1d83078014abdf778d62b043e60b1f3fbbb8fa75962f5458"

definition MQV_P384_TV12_d2U :: nat where
 "MQV_P384_TV12_d2U = 0xdc4d2e3aad5727cca856fef3376dd274a639d9358f5a972a5dad71a0c5a480aa7fb72997da4d0376166cf7beadeb4d05"

definition MQV_P384_TV12_Q2U :: "int point" where 
 "MQV_P384_TV12_Q2U = Point
    0x41b1f80613c2b28fc04ae3fca3eb39c231fab8b9886d73a84f8217366ea98e8b7018ff65647eeda20495e8d26d86cc8c
    0x42aad764c6bac18a06359119b5edc5ee787b457612825fa2200f11c76e11c743044e63d785dc2893097f64efd4647b1e"

definition MQV_P384_TV12_d1V :: nat where
 "MQV_P384_TV12_d1V = 0x82fd5cc5fbe6d5e2f498b57e681badd9ccc56aa343652eeb22c0864abb48f2a020847018463a09d872bf6d7417fdc0cd"

definition MQV_P384_TV12_Q1V :: "int point" where 
 "MQV_P384_TV12_Q1V = Point
    0x1eca4de602f0c670e7f5018d2ad1282b9a6b7af02ff1055c2fdd6456cfca9a5b50ddd85f7270ebc5a4fe814481fcb5ab
    0xc3c586454fe7983f40e06a96808a28c9ec09e1f3a25495e434134cf218a8265e5e11d08d6b9a2675ff593b8ecc6af35c"

definition MQV_P384_TV12_d2V :: nat where
 "MQV_P384_TV12_d2V = 0x5e2cc6f731991616ba72d13cf173668869339165c8d7f80e5ea01cb5d22b15fc0a44e1e325067b37799c122b12b72c24"

definition MQV_P384_TV12_Q2V :: "int point" where 
 "MQV_P384_TV12_Q2V = Point
    0xc167be8a2c4511d9232ebcc746eb6e320b8caf59ef211264e4d22b58f3b79db19c7a480ca53a1b223c9e92f4575b02c1
    0x9d5e380699e7d4752195e17f563006da16b480e90e78a5da8be739920506755b99b6ed986c5b2d5b791995522f955414"

definition MQV_P384_TV12_Z :: nat where
  "MQV_P384_TV12_Z = 0x041de6f64991e2136f28fcd44bdcb213ab6aec666d02b75f7bd83a60f5c6f6b79791e0898374e54ca0d7948dbad6f22c"
  
lemma MQV_P384_TV12_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV12_d1U MQV_P384_TV12_Q1U 
                       MQV_P384_TV12_d2U MQV_P384_TV12_Q2U 
                       MQV_P384_TV12_Q1V MQV_P384_TV12_Q2V 
    = Some MQV_P384_TV12_Z"
  by eval

lemma MQV_P384_TV12_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV12_d1V MQV_P384_TV12_Q1V 
                       MQV_P384_TV12_d2V MQV_P384_TV12_Q2V 
                       MQV_P384_TV12_Q1U MQV_P384_TV12_Q2U 
    = Some MQV_P384_TV12_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV13›
  
definition MQV_P384_TV13_d1U :: nat where
  "MQV_P384_TV13_d1U = 0x80e5ce5f52a4064c09b22d1ea86d969a95ac3ee8785c92e118a3e533d850b0d6c46ce72356fcb2938fd1e91ae2e2cb0d"

definition MQV_P384_TV13_Q1U :: "int point" where 
  "MQV_P384_TV13_Q1U = Point
    0x42ff6ef2968a0f16c69feaae841e1907d1ab422c8cf98a9b535992874781773a5f74fb3263cd39ec816e10461360e994
    0x3bcfa754884231cc085df343f903ecac538a9cf26a889fbaecd917e196baa20cb93a6d07e7fa73d050582bd27d66f495"

definition MQV_P384_TV13_d2U :: nat where
 "MQV_P384_TV13_d2U = 0xa013e38c41b17a5f0a2643f0c8c5bab6999858b79e8f03f9f25a380c2d5798bfa021cce0a85970a945220dea376031a5"

definition MQV_P384_TV13_Q2U :: "int point" where 
 "MQV_P384_TV13_Q2U = Point
    0x81fedb4cdc962085cb0b71e4fa11f847667cdab0f56f3b0294679181595fe030587c0bdbd55d9534c3326738d7c76f92
    0x69e78357b3fb1cbf2a4d8facf645a2b73a711ccbdc450e08c8b1901b17ba022032ac6b346f1b1d9fea3759ea3c73eaae"

definition MQV_P384_TV13_d1V :: nat where
 "MQV_P384_TV13_d1V = 0xc438364e8e28a0dc54481172f06425311a20007fd9a547752041285e3a06d1ab5e77356d7b777d8d657b9a54836ac75a"

definition MQV_P384_TV13_Q1V :: "int point" where 
 "MQV_P384_TV13_Q1V = Point
    0xfaf595c96b2d7378f013fde9f79505a75c3b6648498b446478b227cb2c49d8597c7e2745f807b92b0835f01879d18e1e
    0x77d44219546e56781248edc0d0ff5afa22202ad79a7b3f090a3a9642f2a783d0fd824e8820000eee5624cb2d4bdece65"

definition MQV_P384_TV13_d2V :: nat where
 "MQV_P384_TV13_d2V = 0xc6801083ed130e08a68a13d5f56b7b9181812e87da63cc0a3b73d28451603f7cc80ce8746714af935a1e1e9b867c5493"

definition MQV_P384_TV13_Q2V :: "int point" where 
 "MQV_P384_TV13_Q2V = Point
    0xa05399d9e81e703bad685207435af2917be1cd5bc728f30fb931ab10c172b06c3e2e823532e369050ed8894ec2b81c25
    0xbc9bd96b2e88ac0acd9ac9a07710cd345f7b8ce1269c25dfa5a3820178d94f4057fb4f640a696e42ed08fafd37497b2f"

definition MQV_P384_TV13_Z :: nat where
  "MQV_P384_TV13_Z = 0x08dda40a4409e7f2de83064d1d142c0c440b9331526e4474bdf3df0709c94b95d4f7c248a0e724ad10c991fe957eacba"
  
lemma MQV_P384_TV13_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV13_d1U MQV_P384_TV13_Q1U 
                       MQV_P384_TV13_d2U MQV_P384_TV13_Q2U 
                       MQV_P384_TV13_Q1V MQV_P384_TV13_Q2V 
    = Some MQV_P384_TV13_Z"
  by eval

lemma MQV_P384_TV13_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV13_d1V MQV_P384_TV13_Q1V 
                       MQV_P384_TV13_d2V MQV_P384_TV13_Q2V 
                       MQV_P384_TV13_Q1U MQV_P384_TV13_Q2U 
    = Some MQV_P384_TV13_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV14›
  
definition MQV_P384_TV14_d1U :: nat where
  "MQV_P384_TV14_d1U = 0x39b91edf2d67e2ff758cd0ce9f652614cd456569057eafe540b80264b7b405ada7946e9ab0c9e5f5591901bac9bde0a1"

definition MQV_P384_TV14_Q1U :: "int point" where 
  "MQV_P384_TV14_Q1U = Point
    0x9a46dcc6e2a588d1414fea18b18af8200aeab08912c81d597d9e05047e4678e7950f8aea1bf69ee3816bb6b473d2fa2a
    0x2d0902fd0c121c442e1bbb9e97fce3d656cc236787e00804332ae430777636ede6f34f82b37aa13884700cca47552b27"

definition MQV_P384_TV14_d2U :: nat where
 "MQV_P384_TV14_d2U = 0x14f6e5616d681db1d7414da1a8a5426cf8684f395215d0a1a5bcfd9c53a024151a84eb187d7a77962916e1784c4b46ee"

definition MQV_P384_TV14_Q2U :: "int point" where 
 "MQV_P384_TV14_Q2U = Point
    0x5499dcd4d3e4436e86aedcb9f2ac11f31ee87f6eedde5e7f794ea4f1e63cc9e7b105c77ca4d007c91df7c75b8a91aa54
    0xfb0b3c33db6a46c02b1a1b259589b723844ccb74c71d99d96f9a50eb120f880d70e5fb4809d6d89af31d4365249b1dc8"

definition MQV_P384_TV14_d1V :: nat where
 "MQV_P384_TV14_d1V = 0x80937018598050611d948794d21476f5d750ccae720bae7b9e750c2c491ad5463ac9374d43cc5d08f0c1903378ffa017"

definition MQV_P384_TV14_Q1V :: "int point" where 
 "MQV_P384_TV14_Q1V = Point
    0x6e47f3edcc41f3ffc82a04122824928c5857971c01b87de074b1437b437fb1b7da5bc22e2061f8a102451c4fcf0db945
    0x12e7e6fc58b7cb1c89c6e5553229237e4fc9d3dc8c1d2de532b889c476fdf483663894ae44672e32061e706347363c0f"

definition MQV_P384_TV14_d2V :: nat where
 "MQV_P384_TV14_d2V = 0xe16a92b36cda1ed3d214cddfbc88f4e66f2f198229676feca5888b3fbab80eaf5d239a7403afeb6cbd8b2aa3cb668ebc"

definition MQV_P384_TV14_Q2V :: "int point" where 
 "MQV_P384_TV14_Q2V = Point
    0xc674f700b2ef7a2df01e39cc9714960b11f940d0607a02a7a4812f2f10319bf067e95cfa6d914a16ee169ca4af0ae92b
    0x428fc53ad95948378ad9ec807715109c9e5f393b87c8832c22e7dd2f165a68f5f919a0f8b55cd4ea7aa6f9af2492453c"

definition MQV_P384_TV14_Z :: nat where
  "MQV_P384_TV14_Z = 0x93938d434a51e5af02c99553488357c6757dec9fc7219f74a559c7c5e15f42b2cae7752cff97c9c5fb136e1bc55e94a7"
  
lemma MQV_P384_TV14_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV14_d1U MQV_P384_TV14_Q1U 
                       MQV_P384_TV14_d2U MQV_P384_TV14_Q2U 
                       MQV_P384_TV14_Q1V MQV_P384_TV14_Q2V 
    = Some MQV_P384_TV14_Z"
  by eval

lemma MQV_P384_TV14_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV14_d1V MQV_P384_TV14_Q1V 
                       MQV_P384_TV14_d2V MQV_P384_TV14_Q2V 
                       MQV_P384_TV14_Q1U MQV_P384_TV14_Q2U 
    = Some MQV_P384_TV14_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV15›
  
definition MQV_P384_TV15_d1U :: nat where
  "MQV_P384_TV15_d1U = 0x7c3cc757fe944cbb7c7722f4614538bcd2ed1cdcf53f4204e5c4c061e251a2b3aed3e49c608a7103f4ca048dadfdd2f7"

definition MQV_P384_TV15_Q1U :: "int point" where 
  "MQV_P384_TV15_Q1U = Point
    0xe97f97bb1e19843b76773e094ebf232286bff2cc2f4ba5d075d36888b925e7fb900653a63338a78afbc909b0b75df429
    0x5b7dbb94afbd38d69be26800e6c78a4ef638cad5dcab9d3f62506b588cf6a32ef89444c218ac9d1cbfee5f54621d2e95"

definition MQV_P384_TV15_d2U :: nat where
 "MQV_P384_TV15_d2U = 0xef96c620d63053ac2cc7a2493799a48703ffc4b1fffadcff69b010a8d19a16c0c4ec3855d09402b8e2e92cf0746a1723"

definition MQV_P384_TV15_Q2U :: "int point" where 
 "MQV_P384_TV15_Q2U = Point
    0xe44a63cb098948f0849605bcc7ca7d6060e60340faa8c84da58623e63f24845a759e0760539c03c5dd03adc77e50f0b3
    0x0d781b298cb70905e42b10c557d40830c105c08abdc11453a6c521339c871c98fe824cab6f37427cbff05053f413ef6e"

definition MQV_P384_TV15_d1V :: nat where
 "MQV_P384_TV15_d1V = 0xb3bfc3b0bd0dc79f4d1ef81433b830d40cdcd46a7ea95d14321e5bde233b2f244a2ca2f440a7d38b6fb01eb5f048966a"

definition MQV_P384_TV15_Q1V :: "int point" where 
 "MQV_P384_TV15_Q1V = Point
    0x88c96ce42ea3219e0df044905331e6e7f2f41c6e35566116d677a901bc81d92ac64fc866ad294d346e05a3c3da3f99ee
    0xc6c6d697003dd3f5a4d543eeb84e937cb5ce96a2f880fff263796e4d36daa686c2a798b8206031806f568bf45f652e5d"

definition MQV_P384_TV15_d2V :: nat where
 "MQV_P384_TV15_d2V = 0x252200b5e349d70134bc7ede29eb7969ce69ff592c112d21dfab2b71bef0db32ca7b1d7943088a9c141b379939064c6e"

definition MQV_P384_TV15_Q2V :: "int point" where 
 "MQV_P384_TV15_Q2V = Point
    0xfa80a72bbfc72797f60ae8a2bf5f53c902542ebf2d66bb8f62e8b2753f16b212b9cce96f66ca3db9211ea0e5355ea9f7
    0x747c0a8d7aa0f700cf9d4f733a32fcd1aaf52bcce49e0aac543f502959db474499a3c26797ad35ddbd36ce7523fb14b9"

definition MQV_P384_TV15_Z :: nat where
  "MQV_P384_TV15_Z = 0xe991298e100ffa877edcbf995cf0507d7628aae3eeaca92c281c5023c4809b8f49d5c584c411fa161404527107e8c94d"
  
lemma MQV_P384_TV15_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV15_d1U MQV_P384_TV15_Q1U 
                       MQV_P384_TV15_d2U MQV_P384_TV15_Q2U 
                       MQV_P384_TV15_Q1V MQV_P384_TV15_Q2V 
    = Some MQV_P384_TV15_Z"
  by eval

lemma MQV_P384_TV15_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV15_d1V MQV_P384_TV15_Q1V 
                       MQV_P384_TV15_d2V MQV_P384_TV15_Q2V 
                       MQV_P384_TV15_Q1U MQV_P384_TV15_Q2U 
    = Some MQV_P384_TV15_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV16›
  
definition MQV_P384_TV16_d1U :: nat where
  "MQV_P384_TV16_d1U = 0xbaa40451a992506e32ed1039f4f1a9ddd3fa8448590ca8eb12a377b607d04446e9f7e4314a281c83b1a2df280b49736f"

definition MQV_P384_TV16_Q1U :: "int point" where 
  "MQV_P384_TV16_Q1U = Point
    0x07a6bf31f011d79044e978a12c7d66240f0d2b7b0ccbf7b9499e47dea7617d8512077deda92824c56fd3ea37866b4b6d
    0x2f450d6ded340ed2f6cc504e5399a78ba6be79bc0d200eadb039dc684ffa92a3da5d41ad7c6b0c51e825c67a21852145"

definition MQV_P384_TV16_d2U :: nat where
 "MQV_P384_TV16_d2U = 0xfe059e64764a779a00b4640ec8ce7f2aa3b0c4c537769b1d2bcb6d8520463628eb6dc8068dfbabec72ab1996df3146f0"

definition MQV_P384_TV16_Q2U :: "int point" where 
 "MQV_P384_TV16_Q2U = Point
    0x43b37c996045437fcb4fc01a356cb1258c4171617f475ce97e087ff2105c5e39e33f6a402cc069bd060a8354c8441e13
    0x22d9fd8eace5aad403deb6674e57e28c973245244e7e53676f6f1175b00d94051d4aac705f0b06c031cf9ed4575b6bcd"

definition MQV_P384_TV16_d1V :: nat where
 "MQV_P384_TV16_d1V = 0x2f340f5d90bf53fcc838e37595538fced830c1876f1426b04c8d3a81a5ba93f1f50fb8fcb3fc75e17c12d00349def28a"

definition MQV_P384_TV16_Q1V :: "int point" where 
 "MQV_P384_TV16_Q1V = Point
    0xcab44d55a412f561ba95195cccf1ab5377a5d3efb9694da0e6472ded2f76e3d4457a97b0044589272377237fa5dead19
    0xc0ba5c8d04f2d6599d179a221762d9f19c4427e95ec8ab38f5b6fabc2aab10162e7d4378808868ec71c5b5db292f2f42"

definition MQV_P384_TV16_d2V :: nat where
 "MQV_P384_TV16_d2V = 0xb8353c5cd575dd6f4107eac8f7a53223c148fec95eff8ed6586f54da94f1447bc02001e5cf0829a73a63f2e50ef0a877"

definition MQV_P384_TV16_Q2V :: "int point" where 
 "MQV_P384_TV16_Q2V = Point
    0xa11c21668cd3bc02faef06afe6c9090803b3bd4fad6339f0a85b1669c30065c65a05dd43cb48acec0fa2752750340374
    0x0a412f1ac1b9af30cc141fc4587d5633cd22f06f97deed3f48b9304e991c62fa6174818522974ed24b921a0c056645af"

definition MQV_P384_TV16_Z :: nat where
  "MQV_P384_TV16_Z = 0x02c2c808c63584af00e907f97d2f4c0dca3063aa16609b4254a4b086573f1a55c24ddfa3f95d28667fa558b26be745b8"
  
lemma MQV_P384_TV16_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV16_d1U MQV_P384_TV16_Q1U 
                       MQV_P384_TV16_d2U MQV_P384_TV16_Q2U 
                       MQV_P384_TV16_Q1V MQV_P384_TV16_Q2V 
    = Some MQV_P384_TV16_Z"
  by eval

lemma MQV_P384_TV16_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV16_d1V MQV_P384_TV16_Q1V 
                       MQV_P384_TV16_d2V MQV_P384_TV16_Q2V 
                       MQV_P384_TV16_Q1U MQV_P384_TV16_Q2U 
    = Some MQV_P384_TV16_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV17›
  
definition MQV_P384_TV17_d1U :: nat where
  "MQV_P384_TV17_d1U = 0x2cac960c6738cb8eff52286153a2c3d9a019f822d29b2f1559bf93d7a291666ce86dada549054c76d214e1213f338af1"

definition MQV_P384_TV17_Q1U :: "int point" where 
  "MQV_P384_TV17_Q1U = Point
    0x67fd18d968abbf2cd1f0f25143f2730d6c4e866b7c01eb9934bf6961007df89d7931b5b66fd714f1d0f03df34f7cf44f
    0x07beba2577ee5d5a8728f4e3f169fc125b83b49727aee1d91f09a9511ccfa228be9cdbc13cd2b532ecc91020e0dbae3f"

definition MQV_P384_TV17_d2U :: nat where
 "MQV_P384_TV17_d2U = 0xe4f92cfca732714b9b4a5f46de5e0f9e2d4adf653ac9df36ed477695d0d9db8c9d834f10502d7075163ef4b6669a5122"

definition MQV_P384_TV17_Q2U :: "int point" where 
 "MQV_P384_TV17_Q2U = Point
    0x9d1ddd2044c6b80f7e1ce20cbb4525cfa30e86484d95d48167a05f5816ba4c7d8b16a71858d5d482040b0df735122ff2
    0x0add065d9fe748960fdc6dc6397930e66a134cde6cb8ee066f6255452d7e9e80a1a61a13070745bc6865bfd8014e378f"

definition MQV_P384_TV17_d1V :: nat where
 "MQV_P384_TV17_d1V = 0x946b4024647d29ddf829c9e3980ca55e2622746edc618484af16515f70683037d2f1bb83c15f96a3342e3c79b2a74925"

definition MQV_P384_TV17_Q1V :: "int point" where 
 "MQV_P384_TV17_Q1V = Point
    0x5b1ca6477a40778df6e268599c0fe4064344eab38b4946855821d4eb73f8ab36c57e9c71b4f0d88fe7f5b250535a6f3a
    0xb5cba0a23f873f1f9fd4a9af17a2afa96cb6b7da2bd0e6b3a893895b99b91998707b31c24f3ea0f4b8cee61383bf2c3d"

definition MQV_P384_TV17_d2V :: nat where
 "MQV_P384_TV17_d2V = 0x9923744ecc8e00664d97af238759ed81d65fbcd7d2b7a4b7966fd70ec6d927cada84f040c94209d6c98dfe8eb7c41fcb"

definition MQV_P384_TV17_Q2V :: "int point" where 
 "MQV_P384_TV17_Q2V = Point
    0xc45c3c8cf90e389eca5516984d65094cb7f4d22c97ff3438d7d6d6907d15e0d8fb8a9a9a50083791b35ed6a389675f2c
    0x73bd39a15df2e5985729acada3e909ab95f077b6cc38bf3c75d83fe9a050ba8941c6d560e09999b795c96a49cae2156e"

definition MQV_P384_TV17_Z :: nat where
  "MQV_P384_TV17_Z = 0xa5cac2de3739f6e88c34809b5e7d61952bbc327100692fb315404293c02cca9e249b37883462fc5e2ce4fe71d29facfb"
  
lemma MQV_P384_TV17_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV17_d1U MQV_P384_TV17_Q1U 
                       MQV_P384_TV17_d2U MQV_P384_TV17_Q2U 
                       MQV_P384_TV17_Q1V MQV_P384_TV17_Q2V 
    = Some MQV_P384_TV17_Z"
  by eval

lemma MQV_P384_TV17_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV17_d1V MQV_P384_TV17_Q1V 
                       MQV_P384_TV17_d2V MQV_P384_TV17_Q2V 
                       MQV_P384_TV17_Q1U MQV_P384_TV17_Q2U 
    = Some MQV_P384_TV17_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV18›
  
definition MQV_P384_TV18_d1U :: nat where
  "MQV_P384_TV18_d1U = 0x7d14a2b32fbaa341138b141ff3d2852276bfeed77a9bc4769adfedf1e0adb95384fb21457485da87db6a18bbc5ba9802"

definition MQV_P384_TV18_Q1U :: "int point" where 
  "MQV_P384_TV18_Q1U = Point
    0x1442f08c6da2427b7d041c0ca0b0283b529be3b9a272162b0f8c5ef8d70c2d27565c6fb75f2b113d839664d67fd16445
    0x0a7045d4ec647ff91a0dcacbbbffbdf72ae30745ad67dd2a30365bc5c50dcc469987f4c546d7567d8441aaa912032bda"

definition MQV_P384_TV18_d2U :: nat where
 "MQV_P384_TV18_d2U = 0xbe84ac8ec10e0f0afa9808d2439503e5885b297729130f5980df388364a4e7d8002283047704640b929709561041848b"

definition MQV_P384_TV18_Q2U :: "int point" where 
 "MQV_P384_TV18_Q2U = Point
    0xb68c4a79e01f89fcc0a80b78d117c6d22c75350aab639b2682733660b1fc531138a9fb3059be54f38c52b457cd2b910c
    0xac303224014cb03a779cef72f561c5743f5788f22f0ced3e443b357a11aa004103d83cab096b793b1d83a04c8cbea583"

definition MQV_P384_TV18_d1V :: nat where
 "MQV_P384_TV18_d1V = 0x9ed69ef3622a749d80d3821731c9319c0a067de83a8d08c2add9772890ab111b9571b4ca2ca56fadc32fb05fb79c0424"

definition MQV_P384_TV18_Q1V :: "int point" where 
 "MQV_P384_TV18_Q1V = Point
    0xb102107c91222ee6adc00e5baeaef92b6d07af2c55e8be35b6cba484221efa4b31b212af27eaecebd02e7ce191941059
    0xa8516b2cbc60810de5e1776793f5cefcb206122f0a8c4131a03218e941f78758fa7f11c0c8d8414fd99efa3e0a7519ac"

definition MQV_P384_TV18_d2V :: nat where
 "MQV_P384_TV18_d2V = 0x80c478dd26af49ff2d3eed03a6f2ac952ec65bc9bd83ca025bfddfdbd2d09b2577040a0f91bd229322297dfb326f0adf"

definition MQV_P384_TV18_Q2V :: "int point" where 
 "MQV_P384_TV18_Q2V = Point
    0x2746cce3c5d6b0dbedc53c7306364a8fed9cb6fa8186c42777293e5a65186d61b96f6b7653dfe1bd041ef198c780f9b4
    0x2306ef70a0c9befec1d475354e6c0e9124b624ca11d5de0b444e1c8c1c7cded4c53eada145cf70fd58ff8e54473fa525"

definition MQV_P384_TV18_Z :: nat where
  "MQV_P384_TV18_Z = 0x6f74195f4ee9a0625d8cbd2bd9a9db7b0d482e61242b87d20b238b129892200f3d9546c237fbe795983406534668bee7"
  
lemma MQV_P384_TV18_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV18_d1U MQV_P384_TV18_Q1U 
                       MQV_P384_TV18_d2U MQV_P384_TV18_Q2U 
                       MQV_P384_TV18_Q1V MQV_P384_TV18_Q2V 
    = Some MQV_P384_TV18_Z"
  by eval

lemma MQV_P384_TV18_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV18_d1V MQV_P384_TV18_Q1V 
                       MQV_P384_TV18_d2V MQV_P384_TV18_Q2V 
                       MQV_P384_TV18_Q1U MQV_P384_TV18_Q2U 
    = Some MQV_P384_TV18_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P384_TV19›
  
definition MQV_P384_TV19_d1U :: nat where
  "MQV_P384_TV19_d1U = 0xdc21d5dffdb18138e1fc673822817f8bd82bacd582e6a06ee0e52153dd2f3329796e3e190da7649fe00f7a77c78d0891"

definition MQV_P384_TV19_Q1U :: "int point" where 
  "MQV_P384_TV19_Q1U = Point
    0x4ac52337e744f88014b40a91badd3f291ffc8a32d0a4538471c83a264f3bc0a852d9d7cb193a0e5e48fd89c8c47a5623
    0xddcb94b4a36149083d5fff3ef45f0ffb8081345347da11931caf9271295a300f527135f397e2780277223b5f27de6f8b"

definition MQV_P384_TV19_d2U :: nat where
 "MQV_P384_TV19_d2U = 0x2e418d540e5cdd8cd259bcb5f44d0338f3a1e6c7507b317064b007801d3fda79795068f7979f02076c6ae1465712c6dc"

definition MQV_P384_TV19_Q2U :: "int point" where 
 "MQV_P384_TV19_Q2U = Point
    0x8e280f074dde3a0ed8d8f26c41471caf176333c3f4609cbe532ee4005d7a1962150b5075aba51befa55c8bfc4d37cc4c
    0x798ff850f8a23f330c79d4c07e5c182e1c4eb80bdf314903bc9effce087b13b57394efddcaa0f1ebc8cf1d2eb0a5a740"

definition MQV_P384_TV19_d1V :: nat where
 "MQV_P384_TV19_d1V = 0x9c6ca9843cf0cfc73e6c140e4c551ddbc4a601744309c92b31d0a48da50cb5ee08465d669cb798def5cf0d2538234ae2"

definition MQV_P384_TV19_Q1V :: "int point" where 
 "MQV_P384_TV19_Q1V = Point
    0x9936d20d11f51eaa77c519eb9ab0d49b463bbb719439bb3bbf4e1a4437d0850d81576816bca05c9ebef40e62fc8f78ac
    0x52c5afe64c46c9291a771d1cdeffae7cdf7abbf2da4416301c0889390c06160b47f38f292e6abd06aaaafc0af4502881"

definition MQV_P384_TV19_d2V :: nat where
 "MQV_P384_TV19_d2V = 0x95035eba40cbf54f2e05b3481cfb77a8c5362d9e7ac57e230de09e8bb7218f4f1c33f53989ebd43c1e7dfd36be20fd47"

definition MQV_P384_TV19_Q2V :: "int point" where 
 "MQV_P384_TV19_Q2V = Point
    0xebea484e0c5eb51aaa2a7d38cf9af905fc3c3f0cd22634fe72726c177ad671e94cab02aa19737d5f08a10b359f11e051
    0x8b78d226b632fe0b15d666d3f914b893a2f401b02ecada2408de63613e0a34738c86ea766ec1ea4b9ab2eed2491dc0e9"

definition MQV_P384_TV19_Z :: nat where
  "MQV_P384_TV19_Z = 0x0c1e19e627c2c37e43325c2c5f62700db4877db99af83962258c5cf15fc6a946470141825a9314a5a1b734f1e68db82f"
  
lemma MQV_P384_TV19_MQV_check: 
  "SEC1_P384_ECMQVprim MQV_P384_TV19_d1U MQV_P384_TV19_Q1U 
                       MQV_P384_TV19_d2U MQV_P384_TV19_Q2U 
                       MQV_P384_TV19_Q1V MQV_P384_TV19_Q2V 
    = Some MQV_P384_TV19_Z"
  by eval

lemma MQV_P384_TV19_MQV_check': 
  "SEC1_P384_ECMQVprim MQV_P384_TV19_d1V MQV_P384_TV19_Q1V 
                       MQV_P384_TV19_d2V MQV_P384_TV19_Q2V 
                       MQV_P384_TV19_Q1U MQV_P384_TV19_Q2U 
    = Some MQV_P384_TV19_Z"
  by eval  
  
section‹NIST Test Vectors for MQV primitive: Curve P-521›
  
subsubsection‹MQV test vector: MQV_P521_TV00›
  
definition MQV_P521_TV00_d1U :: nat where
  "MQV_P521_TV00_d1U = 0x000001556ca827de266fc16b2cc6f57560253cd559ec364a61007954919cb262c4a3fd323178e8a57718c759c0ba673153939cb6ff26a3154858cd65afdf9c16f9dc9dcc"

definition MQV_P521_TV00_Q1U :: "int point" where 
  "MQV_P521_TV00_Q1U = Point
    0x000000d2bf6fa6fadac4df1076d532dd31e6ba5cf0b74531d5afd83daad263de2913114b98741f21766034f1111f1af018e1e50530f739e5977cc611fadbe9015fb31f7e
    0x00000091da4014a0511735033e26a163dd52e6a26cae4663decdbdcfebca48e88fb00422040a19e423def031ae30ad287871b5ff96d588d9dddd372fad687f97c29b5c6b"

definition MQV_P521_TV00_d2U :: nat where
 "MQV_P521_TV00_d2U = 0x000001e20b1242ab087f508f202e76b242bec17da80fa0ee700f4e85cc6823a0eea725e9df669ba0e13fdd061733d2d7a3a830e78c9ab03731f2d5d5e22364b74a29db82"

definition MQV_P521_TV00_Q2U :: "int point" where 
 "MQV_P521_TV00_Q2U = Point
    0x0000014bff23a380874370b80f8e0a83e9760950cc630a848e28fbf21d04e017f3560a2e4c98a99bf9584c1415b5d2ec46aed809ab50ee2e1f60191093bc31b5555d6db5
    0x00000030e70aa5f8c09e1d60b3bad0e74f563ed1ada247f01024e6314aeb4c1a0350f2761ebc36d84f29dd1e5b8abae771705527f2338b442ed7100b88529f508ec14992"

definition MQV_P521_TV00_d1V :: nat where
 "MQV_P521_TV00_d1V = 0x0000007fa41b895eae7a9ba00337ba82a4bdb8716881c49bcce617fbcbe55d95649577bcbdf73f8a2bd6a39ccaab59e0dcbe0dd0b75900682ba3d2528d9846f51cbf36b1"

definition MQV_P521_TV00_Q1V :: "int point" where 
 "MQV_P521_TV00_Q1V = Point
    0x00000056d82971c3b39cfc34978f5ebb2a8b59c34c511358966ce55c66239368984c5bcc0e5424eebde2516ab410d1db77a932ccd35ee12577316d703a0c1d318b46f341
    0x00000156fe4cc32974467b0ecb25747f133887aef4656c035fd2d5b597c5c9141b4441b78d26604076947e2790c4e906588ea78aec31b0c9ccf2909111d5213c9eda42ce"

definition MQV_P521_TV00_d2V :: nat where
 "MQV_P521_TV00_d2V = 0x000000c809ef97b5e2e8b65e652148bdbb08a04875f98cda02b02a86e7b0a81f4e6db1bf15489968cdeedcfffe61eec68c3ef3021ec1643d9721dd70fda323a9a23cfcfe"

definition MQV_P521_TV00_Q2V :: "int point" where 
 "MQV_P521_TV00_Q2V = Point
    0x0000010a0ebe9aebfebea6a565613f8c98127facfc4857e974685e3a0152d0ca4e725b00ee69c0fab3e1c41b1cb0e135a861824d29dd01d25800c978db1790fea56fccb9
    0x00000054e7d413ef5eea10168139821d37920729a4d99de35abdd142242dee54e65af4c32c442b91b31248e6b9c5f87b58f8f92a285a8867dabed97f8bb3e78bb134ff6e"

definition MQV_P521_TV00_Z :: nat where
  "MQV_P521_TV00_Z = 0x00d11edceafeba18633061e60bd6052244ec6f7931a5e44649f33c77da5261952aa5854a22769724c2e2861c30ec465f7c91129852f900bf7163aa401ebecb6d4ec6"
  
lemma MQV_P521_TV00_MQV_check: 
  "SEC1_P521_ECMQVprim MQV_P521_TV00_d1U MQV_P521_TV00_Q1U 
                       MQV_P521_TV00_d2U MQV_P521_TV00_Q2U 
                       MQV_P521_TV00_Q1V MQV_P521_TV00_Q2V 
    = Some MQV_P521_TV00_Z"
  by eval

lemma MQV_P521_TV00_MQV_check': 
  "SEC1_P521_ECMQVprim MQV_P521_TV00_d1V MQV_P521_TV00_Q1V 
                       MQV_P521_TV00_d2V MQV_P521_TV00_Q2V 
                       MQV_P521_TV00_Q1U MQV_P521_TV00_Q2U 
    = Some MQV_P521_TV00_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P521_TV01›
  
definition MQV_P521_TV01_d1U :: nat where
  "MQV_P521_TV01_d1U = 0x0000012dcfef6e3fd1793d7a7ec2f28301f704b5a56704aee5f91fe81d8f5f8d10018efe16ebffdead756298cb6a63773d07ef5ab4cc3f7f16a8162a4e9d5cc4e529b6c8"

definition MQV_P521_TV01_Q1U :: "int point" where 
  "MQV_P521_TV01_Q1U = Point
    0x00000098f3d5deff51bd2faef9ad8790f26b53a443dbdd881d860ba8ed40aab327fbaca4a4b840af37489ca532dc4aa8eed27da0251bbc5f6adb2b53cb3b0d28449a3732
    0x00000114a46135a73704241abb505d6e22eb0857020832ab607992d6906855eda38d7ea5433a6408a1a51ea11eded56d6847ef57067d4572889d471b42bd5c75d6fb43f7"

definition MQV_P521_TV01_d2U :: nat where
 "MQV_P521_TV01_d2U = 0x0000018eaa15774115d28508f2843cb6908d0bbd0f8dadfb5bbee037dad10e0bd3292af4a352bec6366d828258e40c07d071ecba643d9f6167d23f677e4da774344f643e"

definition MQV_P521_TV01_Q2U :: "int point" where 
 "MQV_P521_TV01_Q2U = Point
    0x000001f97b5b68c410118b18c4e67caa2182fcdbeb55a41da8e6ea8bd0a132b7bc747f1718a80d0844c185449c6d524e6dda0fc486924e9e5ecfe12282495f846a29065b
    0x000001338e3fd18ac833ecd2e0811ca4842802917402dac25f7bc28cc4af162b0a10449268dfaf3382f64580111dd20966b22efe1a805ae7aeef809e4fa04a378c54e4a4"

definition MQV_P521_TV01_d1V :: nat where
 "MQV_P521_TV01_d1V = 0x000001b27fb1fd58966173e0fb0c78ec101420e45cb8bf755293c921192230275cf502730bd6addff3bc5c05d06591b0389e6d00c63fb46c97a28412b944fe88dbae45a3"

definition MQV_P521_TV01_Q1V :: "int point" where 
 "MQV_P521_TV01_Q1V = Point
    0x00000004c0217db22d90f63228c913318087233a3b304c5271e8062ed060dfceca1afe889a28816bead8a97a769b03899baa3acb7ee48712d3b9c4b082e642c60b81d4d4
    0x0000005788f183903574bc9a969798bd078b916a1965dd86c5d6308ea691a336fe09bf1a09c7cdce669b2566d8a7f3d999ce150134d49ac90a0e552ac2a3915b3f3ca4e7"

definition MQV_P521_TV01_d2V :: nat where
 "MQV_P521_TV01_d2V = 0x0000001d64fb6c3b0e4614b1c66f75079f35459c2d1af0bec8df95e7dcc46aa7922ef1b78ba2b617a5de47bade6371f94672af5a541ee0aee1435404eb170cadeea651b8"

definition MQV_P521_TV01_Q2V :: "int point" where 
 "MQV_P521_TV01_Q2V = Point
    0x0000009d834d1a8ed366499d9e06704fd6bb4fe6ad15eb7a63202de325b81619b4d48857f2eae092e6b95cc70c90dc75d3525f9cc2a4c2adff2795be3309d90735598f4e
    0x0000012e5feae563fc4553a7e81ae45a395bedde351c1ad0401270e51c9d156c9b12d8176dd7e882bdab262b2ecab331b99ee5300843e0054967eacf726c5c84392c9262"

definition MQV_P521_TV01_Z :: nat where
  "MQV_P521_TV01_Z = 0x01e271aee4ef15ef35d769218aca51a9dde6928236ba7fefa78966c410c00a68d8d94ffcc51f507d18361943b2e71b93f8d24b729408af72b51beda4453368db5bef"
  
lemma MQV_P521_TV01_MQV_check: 
  "SEC1_P521_ECMQVprim MQV_P521_TV01_d1U MQV_P521_TV01_Q1U 
                       MQV_P521_TV01_d2U MQV_P521_TV01_Q2U 
                       MQV_P521_TV01_Q1V MQV_P521_TV01_Q2V 
    = Some MQV_P521_TV01_Z"
  by eval

lemma MQV_P521_TV01_MQV_check': 
  "SEC1_P521_ECMQVprim MQV_P521_TV01_d1V MQV_P521_TV01_Q1V 
                       MQV_P521_TV01_d2V MQV_P521_TV01_Q2V 
                       MQV_P521_TV01_Q1U MQV_P521_TV01_Q2U 
    = Some MQV_P521_TV01_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P521_TV02›
  
definition MQV_P521_TV02_d1U :: nat where
  "MQV_P521_TV02_d1U = 0x0000019c5f633c204f9b24f326af2a3349b3dbc0e9682432827a599a53eace10f62827eb231091e91fb658de6718fb0e39610c9d2b67c72c1ce103dcd9bd628b72a0664f"

definition MQV_P521_TV02_Q1U :: "int point" where 
  "MQV_P521_TV02_Q1U = Point
    0x000001d69335a754c6a1a7110e24f32eced832ac1c8f93c35dd63a445a362b10ea76e6c1b5934f91686fba13ecb8abad0134ccc11cb802e86f9893f26abaee530a23e6c9
    0x000000c0b1de2926b80533ae75bf7f81e19649ce9d9e3dadbd7a2b289131a3ebd4b89e40f33341073aefa206e0283cf3dc0d0eb8652d1e0c83a876ef7e20dc5e11b57094"

definition MQV_P521_TV02_d2U :: nat where
 "MQV_P521_TV02_d2U = 0x000000c79051e04e2f2ad1a8199d151cc7079bebee16808a52b65a3cbcffca84133d8b3dc63213c0964237763d718fbd8a3ce06ea6629a00d5b14b719c24da3f0e62b649"

definition MQV_P521_TV02_Q2U :: "int point" where 
 "MQV_P521_TV02_Q2U = Point
    0x0000004cb6af1990ef879114f6e3946c6a575dae258ccc490b076892f52b1f9e71e89cbeb5d4d92513ec9f3531738de9ca1833751a6c50b7c4e9c5e9cc3727a40e9cdc5f
    0x000000c5594121df96740e9a9df2b258324caa91cb33ea5ccb5c83a4c9c844056666ba63fad0ee4eb467514513535e2756d6821d0ca4f1688467b26c2804d9c4f02afd2c"

definition MQV_P521_TV02_d1V :: nat where
 "MQV_P521_TV02_d1V = 0x000000118ff885d79f57edf468a8311b93f5cdd8a686af1e482fe8a6376a815b1703fe15c966c959063b59aff080c326a1d3a5a2aa471de2b72734546499de50db46e8e3"

definition MQV_P521_TV02_Q1V :: "int point" where 
 "MQV_P521_TV02_Q1V = Point
    0x000000a83747b7ecf2b636bae3af05e1c1e8b35e3772a7031d508a1094c51dc40bda0615fe00842387c40fe6d706d5d0efc6d69076e7c64e696ef255a01322d63373b83a
    0x0000009f56bf21bcf83388f581c79a0b6aad96e8cf013bf585d254ecd9b4d7fd3e6207ac46679dddb54d2501d65c2c55c86b667c2e477cad30b3f6fb3c00c2c7e3565fd0"

definition MQV_P521_TV02_d2V :: nat where
 "MQV_P521_TV02_d2V = 0x0000013dd24785b6ebe32242de300b41341df41277a55c094eee60954039efded09a64879bebb247ebef60cbda3f331ad11ec569152bfc6620d5b06e03fb3bda88314b78"

definition MQV_P521_TV02_Q2V :: "int point" where 
 "MQV_P521_TV02_Q2V = Point
    0x00000144d5b5db2e3143acac483b6873f5b584210cd478f0aec0937333004e71fd1eae27916c8dde2c13500c2725476e66670b195cabecde6e8ca8ee0f22447f17643777
    0x00000002049988dd7b661567ae2e3f6b5d9da2e8bdcb1785481bbb3a356aa79ae476865d10d4bbd6ebe9757bd1bb784ce9d32957671eb913e6373b683b8d77ede26f3cd8"

definition MQV_P521_TV02_Z :: nat where
  "MQV_P521_TV02_Z = 0x01cabca01f2a8ff77bb41f2606f0e0e61b9ec54acb37b988993c47fc916c35213a471edb157505eb055c9cdec324351a0c8eb2b84942f86cf153dd36e197df472ee5"
  
lemma MQV_P521_TV02_MQV_check: 
  "SEC1_P521_ECMQVprim MQV_P521_TV02_d1U MQV_P521_TV02_Q1U 
                       MQV_P521_TV02_d2U MQV_P521_TV02_Q2U 
                       MQV_P521_TV02_Q1V MQV_P521_TV02_Q2V 
    = Some MQV_P521_TV02_Z"
  by eval

lemma MQV_P521_TV02_MQV_check': 
  "SEC1_P521_ECMQVprim MQV_P521_TV02_d1V MQV_P521_TV02_Q1V 
                       MQV_P521_TV02_d2V MQV_P521_TV02_Q2V 
                       MQV_P521_TV02_Q1U MQV_P521_TV02_Q2U 
    = Some MQV_P521_TV02_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P521_TV03›
  
definition MQV_P521_TV03_d1U :: nat where
  "MQV_P521_TV03_d1U = 0x00000156d47dbdbe8383bbfb0a8f0e9eeb0e7837fed8baeed9ee3dc071a92a892b6117c96cdc329f2caeb9a49f7f4c2c0677b7721a0539899be6de5df3d3b70c635e5f16"

definition MQV_P521_TV03_Q1U :: "int point" where 
  "MQV_P521_TV03_Q1U = Point
    0x000000f63ab48fb7606b8cc5bb80a580901cb2a6e6f13f26eddc2408ce64b19872cf334b3e51ab12b771813a944d4e5af4c0e9db570bf606888ef5c5df03a8f78d25c39e
    0x000000c7fce3a597b7551dd01059ce10de965e8bf4b9ecff423eca9d01fa91a79184e0f1ba28258e49d4d0eacd5bb4389b7737282c8a66123c607295babe081bc95aad8d"

definition MQV_P521_TV03_d2U :: nat where
 "MQV_P521_TV03_d2U = 0x000000aa7917c6f7062a3c88f30f1f276e1ec67c3e6ad5d77afbb43cbbbf816ce110d4783aa53f33ab63522d8d7d62d3098604f4cfaea795ebd053e6127b28bf67de033f"

definition MQV_P521_TV03_Q2U :: "int point" where 
 "MQV_P521_TV03_Q2U = Point
    0x000001f01454c3d8e601b5310504e6e80277970536887bf3fe30350709c5035344acf8fcb10425d9c9e0c28d134503de74a863411d6518ed11376fe0864f445cfb2bec90
    0x000001fe6ab40a9120b238b0e284f9276691177390c7fb9b2f2aa54796257ec2153c4dd27941f53bb6ebdc74737234208bc27aefa06e6dc40cb356ac7cd185ac5de8d7d6"

definition MQV_P521_TV03_d1V :: nat where
 "MQV_P521_TV03_d1V = 0x00000138ecb2d7abcc6ee17dce64f4aadb299172ff77a37cad1f90f0b428d33be55452872c30a21473b0ed27bfbb499ea5cb485fddd8d864240ade8f2bb02cdafba60e5c"

definition MQV_P521_TV03_Q1V :: "int point" where 
 "MQV_P521_TV03_Q1V = Point
    0x000001a4315cfd7019d614596a8a04e7efae2ea00122814ab9638c3c6c9f030b95921fd23b945c199837768123c4cf823f2506264ab5ca63512902f795d3686abfaf6bfd
    0x00000045624ac91ef2e2e9c3e79a6f3eb218f30b91e2c5c1f4dc7d6a0bcd686801399a62cb2217af4e7d2150194101787c5bd7f1cd350c9ff7c6c13f7c9bcfa3c555d590"

definition MQV_P521_TV03_d2V :: nat where
 "MQV_P521_TV03_d2V = 0x0000019ca7d40015013034f6173d2bb5c05eeb2b7a9f3463dfbd3b3dfa3dd1c5c59378b7e0b64956499544c28373ddf71c5e43b5a67c9efde2256875980e11ff0d685068"

definition MQV_P521_TV03_Q2V :: "int point" where 
 "MQV_P521_TV03_Q2V = Point
    0x000000152061b93964526ce2915aca337d0730d9c8acf7cb56398b325fc12fa66646fb0853cc0bb221bac6397dd5221090f6eac3746eb442cb931364b85f709e3e420682
    0x000000ca2739ed55335a661a794ebb18ebd11fa55021c0a3a83533ff01cbc8f1638f5e8e9b162c37bee86ed875758dbda3b2e2e05280afced5c033029655f9abd13756a8"

definition MQV_P521_TV03_Z :: nat where
  "MQV_P521_TV03_Z = 0x00f7f3fbf1472b352ddb4a692f19f32f0d96370e9a87dd4de3dc22be589f12722fe4ce86d0da862d0433816bd8335ebeba131e6b82a52acc0eca965dcd965a710007"
  
lemma MQV_P521_TV03_MQV_check: 
  "SEC1_P521_ECMQVprim MQV_P521_TV03_d1U MQV_P521_TV03_Q1U 
                       MQV_P521_TV03_d2U MQV_P521_TV03_Q2U 
                       MQV_P521_TV03_Q1V MQV_P521_TV03_Q2V 
    = Some MQV_P521_TV03_Z"
  by eval

lemma MQV_P521_TV03_MQV_check': 
  "SEC1_P521_ECMQVprim MQV_P521_TV03_d1V MQV_P521_TV03_Q1V 
                       MQV_P521_TV03_d2V MQV_P521_TV03_Q2V 
                       MQV_P521_TV03_Q1U MQV_P521_TV03_Q2U 
    = Some MQV_P521_TV03_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P521_TV04›
  
definition MQV_P521_TV04_d1U :: nat where
  "MQV_P521_TV04_d1U = 0x000001c61402a5be55d47e14482e13b3b69c9d90085666353cd3021038aa698094e85eb6f1132ec0a9d1cbc145c5871b306a70d68d79c97e50c03b3faefb3546d97f6446"

definition MQV_P521_TV04_Q1U :: "int point" where 
  "MQV_P521_TV04_Q1U = Point
    0x0000006b8b42b1b54f85ecbb0ecd935ded9de9d45e6aea31b73f50553355f6c3328535adb483edb0606dfb5ed8c13b8dcfd65b64f5c4263d71f23b69b4495c23fa5f0d57
    0x000001b5213064307bde786af75e2b82653412ebe46ca4bcf81689c54994e9248c8c3212af487f66362a1f71986cdad0acb8202bd076f842e86797f945cab5a5b9ea0979"

definition MQV_P521_TV04_d2U :: nat where
 "MQV_P521_TV04_d2U = 0x000001879bf36fd845ab896405c18ad4db20ee516cebc3e7b8d36b37714c170eb4fd4ce3b111d61f94e7be2706425265e9109d19eb6547e1da4722d1d43fa80d679afdc7"

definition MQV_P521_TV04_Q2U :: "int point" where 
 "MQV_P521_TV04_Q2U = Point
    0x0000014e80746ca37816a1365d9ded6e2c637461a06650f44e3a4298598f91d354fe233eb21a4a08051c6a2dafcd83dd2da7d8b1f7655dc20fd4e359a5c87bfd6a4f6e74
    0x000001aeac0c689f6ba6af210c7533cafb9ffb77b689217e06a0dfa13f8be31cbe668250c2f8f4479c69b9690b36bd1a6b7523ad2db67f27e69d4e086bb4f02f7949a3af"

definition MQV_P521_TV04_d1V :: nat where
 "MQV_P521_TV04_d1V = 0x0000002473029216bcf260adcdfaf398082b2f022d81277c5e3ed797e8ad61bc39195a1abf157c00eec7c866abf50b7d973480041c7e03df3ba8a5d7b4cfd98ff1ff4736"

definition MQV_P521_TV04_Q1V :: "int point" where 
 "MQV_P521_TV04_Q1V = Point
    0x0000004250c84feba21739cc9afdf33770b63c74ef614b538adbdcf538aedbb7fbcdeeb01cad215f261e35bdbeeee500dacb405ae8867778b3ce985e64550dec7d14e2ed
    0x0000013b9b75310d602c9de9d76205e7fb6a516c95e33e8533b89b792ee62c2202bbd52d78a2c919edda713d98aff76b98d1037a673df49965e93e1d073fe5b6584bd1c5"

definition MQV_P521_TV04_d2V :: nat where
 "MQV_P521_TV04_d2V = 0x0000003bd9e451a03a4562260f75701645ef58a0418218d87adca392969ba83ce352bcc354abc7bec2405949b6e793b351c88d23381ae132f8f262034cd0a54ad0f36e0b"

definition MQV_P521_TV04_Q2V :: "int point" where 
 "MQV_P521_TV04_Q2V = Point
    0x000001f8bd69d823a8fe683a5af31a230f49ebf2147f5420e30c30f83e0af3dcbbe238d6d50be9ef3bd37454f0f14c90c08b9bb51b50d646a5fbe17152fd091ec03bceed
    0x0000006838262011a5bda6bbdd3587ab0b16709166e41937c242dcc4745ad651b386f9fc956a78674cf8d6400937ca4662c583bf77ce2d1bdaeccf056943264e865b374f"

definition MQV_P521_TV04_Z :: nat where
  "MQV_P521_TV04_Z = 0x0166507f4fdf0546fed858bd83514505ea9a80da82376c6ad1f4d5607783bc53ee48faa69418d7711e8439554654bafd759b565137a4cf8560fe3b720b3a0e327478"
  
lemma MQV_P521_TV04_MQV_check: 
  "SEC1_P521_ECMQVprim MQV_P521_TV04_d1U MQV_P521_TV04_Q1U 
                       MQV_P521_TV04_d2U MQV_P521_TV04_Q2U 
                       MQV_P521_TV04_Q1V MQV_P521_TV04_Q2V 
    = Some MQV_P521_TV04_Z"
  by eval

lemma MQV_P521_TV04_MQV_check': 
  "SEC1_P521_ECMQVprim MQV_P521_TV04_d1V MQV_P521_TV04_Q1V 
                       MQV_P521_TV04_d2V MQV_P521_TV04_Q2V 
                       MQV_P521_TV04_Q1U MQV_P521_TV04_Q2U 
    = Some MQV_P521_TV04_Z"
  by eval  
  
subsubsection‹MQV test vector: MQV_P521_TV05›
  
definition MQV_P521_TV05_d1U :: nat where
  "MQV_P521_TV05_d1U = 0x000001fb96662139e349acabcf76788d1fb002918cdf52ded53b2f525f3e607abf40e585b2be6cb63504c2de12ea31cce858b3b294c12a3c80ac7043c675e9882e8fe742"

definition MQV_P521_TV05_Q1U :: "int point" where 
  "MQV_P521_TV05_Q1U = Point
    0x00000001fdaa925d9d0559a6aa7f555fac4f7d977266a5c68b5db2c2f026b6a01d09fd5252bc2a4281996ee9b47a606dc748e1c8b213da98d12add6f598bf21f53cde130
    0x000001423c6ba6b28aa6a1d253c0c3d56a2b9042575c219709bfbcd27bff96975a9b7b5b51652bab71f476f2837dc5b4825e2d3a7034e47e5b472ef27f397b07de356fd9"

definition MQV_P521_TV05_d2U :: nat where
 "MQV_P521_TV05_d2U = 0x0000015d27e33d483aeb401e93930bfc0c75d376e55df95149721f25d06b8fdfc50d359cf596a70806557bed3c50d0633bc8cce2ed679c6fda080e8981c03c1f92f51275"

definition MQV_P521_TV05_Q2U :: "int point" where 
 "MQV_P521_TV05_Q2U = Point
    0x000000ed5263ac51d4115ceb237f36e126b8050cef792eba2373895f4ea7f8db3c483dd3675360db52e92fe3e883b8c9dbfc1fc73dfd1d909f4975b6565eeb28fa68d3ea
    0x000000a99794c555c8a27ffaa51e0febaf78946be2d2ca27005c702ecd83efb63c5dc479fbee2e21817e94fefbf72ba890c6374dd01020ec13b7db67be501026402a08a1"

definition MQV_P521_TV05_d1V :: nat where
 "MQV_P521_TV05_d1V = 0x0000013be34124bd262f21b768be8d92ee57cd519b7516a7847b35bd775726f77b6289f8d2c9f175f9eec3a52540137b1478fa67f8ee1b84afeeb2d65d3f0bcb2efd0fae"

definition MQV_P521_TV05_Q1V :: "int point" where 
 "MQV_P521_TV05_Q1V = Point
    0x000000773267d742ea5e1e6e90ceb0a2d8a43e04e1a8dd0c71d15243e5f75184d3d2b338879e8e63cd24278075fd8cf946f77b1c9a381be502cebefc34c12ca2cb7d294d
    0x000000efae1d7764701597c36a47940fc49da7901c281e524386cd2eb5302ad96b688da69eca1bfcf2f8f546d7e3b24fec10d8116b8449a86a49aabcf1dff70106442627"

definition MQV_P521_TV05_d2V :: nat where
 "MQV_P521_TV05_d2V = 0x00000159acf275822991be33713a7164a45332067103a2d10d0f1eca38b1b9ff423e89b782dbdc8bb6cc7407d66f735d055d428070446fb858a8c4fdb3e58892181914b3"

definition MQV_P521_TV05_Q2V :: "int point" where 
 "MQV_P521_TV05_Q2V = Point
    0x000000ce2a266c943326ca5ab837ce65c0aeeef9d1d85c1060ac7bca34762339d4f1060621c6120646ceb7fd405e990a90115a5ae6cd6990009de1d5676aa31e864965d6
    0x000000e0c0b736e43ccf05d07c7fffc04f4bd78aa31e20debe95b0ec0c7d955947e933cf8d921c0f0c80284cc1b141220a5096648b5a2d48624d6bcf79d67586af0e95c4"

definition MQV_P521_TV05_Z :: nat where
  "MQV_P521_TV05_Z = 0x002b4413aca0f93b551c2fe8ce560d6ce9c5e6fb9755426814e6cc6364282a7923a34bba1b55526ea8dc8b01c556c1ce42e2faa3e2ca1c56201409ad466da539db0c"
  
lemma MQV_P521_TV05_MQV_check: 
  "SEC1_P521_ECMQVprim MQV_P521_TV05_d1U MQV_P521_TV05_Q1U 
                       MQV_P521_TV05_d2U MQV_P521_TV05_Q2U 
                       MQV_P521_TV05_Q1V MQV_P521_TV05_Q2V 
    = Some MQV_P521_TV05_Z"
  by eval

lemma MQV_P521_TV05_MQV_check': 
  "SEC1_P521_ECMQVprim MQV_P521_TV05_d1V MQV_P521_TV05_Q1V 
                       MQV_P521_TV05_d2V MQV_P521_TV05_Q2V 
                       MQV_P521_TV05_Q1U MQV_P521_TV05_Q2U 
    = Some MQV_P521_TV05_Z"
  by eval  
  



end