| author | nipkow | 
| Fri, 16 Dec 2022 18:19:23 +0100 | |
| changeset 76651 | 0cc3679999d9 | 
| parent 74097 | 6d7be1227d02 | 
| permissions | -rw-r--r-- | 
| 41561 | 1 | (* Title: HOL/SPARK/Examples/RIPEMD-160/RMD.thy | 
| 2 | Author: Fabian Immler, TU Muenchen | |
| 3 | ||
| 4 | Verification of the RIPEMD-160 hash function | |
| 5 | *) | |
| 6 | ||
| 7 | theory RMD | |
| 72515 
c7038c397ae3
moved most material from session HOL-Word to Word_Lib in the AFP
 haftmann parents: 
67407diff
changeset | 8 | imports "HOL-Library.Word" | 
| 41561 | 9 | begin | 
| 10 | ||
| 74097 | 11 | unbundle bit_operations_syntax | 
| 41561 | 12 | |
| 67407 | 13 | \<comment> \<open>all operations are defined on 32-bit words\<close> | 
| 41561 | 14 | |
| 41587 | 15 | type_synonym word32 = "32 word" | 
| 16 | type_synonym byte = "8 word" | |
| 17 | type_synonym perm = "nat \<Rightarrow> nat" | |
| 18 | type_synonym chain = "word32 * word32 * word32 * word32 * word32" | |
| 19 | type_synonym block = "nat \<Rightarrow> word32" | |
| 20 | type_synonym message = "nat \<Rightarrow> block" | |
| 41561 | 21 | |
| 67407 | 22 | \<comment> \<open>nonlinear functions at bit level\<close> | 
| 41561 | 23 | |
| 24 | definition f::"[nat, word32, word32, word32] => word32" | |
| 25 | where | |
| 26 | "f j x y z = | |
| 27 | (if ( 0 <= j & j <= 15) then x XOR y XOR z | |
| 28 | else if (16 <= j & j <= 31) then (x AND y) OR (NOT x AND z) | |
| 29 | else if (32 <= j & j <= 47) then (x OR NOT y) XOR z | |
| 30 | else if (48 <= j & j <= 63) then (x AND z) OR (y AND NOT z) | |
| 31 | else if (64 <= j & j <= 79) then x XOR (y OR NOT z) | |
| 32 | else 0)" | |
| 33 | ||
| 34 | ||
| 67407 | 35 | \<comment> \<open>added constants (hexadecimal)\<close> | 
| 41561 | 36 | |
| 37 | definition K::"nat => word32" | |
| 38 | where | |
| 39 | "K j = | |
| 40 | (if ( 0 <= j & j <= 15) then 0x00000000 | |
| 41 | else if (16 <= j & j <= 31) then 0x5A827999 | |
| 42 | else if (32 <= j & j <= 47) then 0x6ED9EBA1 | |
| 43 | else if (48 <= j & j <= 63) then 0x8F1BBCDC | |
| 44 | else if (64 <= j & j <= 79) then 0xA953FD4E | |
| 45 | else 0)" | |
| 46 | ||
| 47 | definition K'::"nat => word32" | |
| 48 | where | |
| 49 | "K' j = | |
| 50 | (if ( 0 <= j & j <= 15) then 0x50A28BE6 | |
| 51 | else if (16 <= j & j <= 31) then 0x5C4DD124 | |
| 52 | else if (32 <= j & j <= 47) then 0x6D703EF3 | |
| 53 | else if (48 <= j & j <= 63) then 0x7A6D76E9 | |
| 54 | else if (64 <= j & j <= 79) then 0x00000000 | |
| 55 | else 0)" | |
| 56 | ||
| 57 | ||
| 67407 | 58 | \<comment> \<open>selection of message word\<close> | 
| 41561 | 59 | |
| 60 | definition r_list :: "nat list" | |
| 61 | where "r_list = [ | |
| 62 | 0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, | |
| 63 | 7, 4, 13, 1, 10, 6, 15, 3, 12, 0, 9, 5, 2, 14, 11, 8, | |
| 64 | 3, 10, 14, 4, 9, 15, 8, 1, 2, 7, 0, 6, 13, 11, 5, 12, | |
| 65 | 1, 9, 11, 10, 0, 8, 12, 4, 13, 3, 7, 15, 14, 5, 6, 2, | |
| 66 | 4, 0, 5, 9, 7, 12, 2, 10, 14, 1, 3, 8, 11, 6, 15, 13 | |
| 67 | ]" | |
| 68 | ||
| 69 | definition r'_list :: "nat list" | |
| 70 | where "r'_list = [ | |
| 71 | 5, 14, 7, 0, 9, 2, 11, 4, 13, 6, 15, 8, 1, 10, 3, 12, | |
| 72 | 6, 11, 3, 7, 0, 13, 5, 10, 14, 15, 8, 12, 4, 9, 1, 2, | |
| 73 | 15, 5, 1, 3, 7, 14, 6, 9, 11, 8, 12, 2, 10, 0, 4, 13, | |
| 74 | 8, 6, 4, 1, 3, 11, 15, 0, 5, 12, 2, 13, 9, 7, 10, 14, | |
| 75 | 12, 15, 10, 4, 1, 5, 8, 7, 6, 2, 13, 14, 0, 3, 9, 11 | |
| 76 | ]" | |
| 77 | ||
| 78 | definition r :: perm | |
| 79 | where "r j = r_list ! j" | |
| 80 | ||
| 81 | definition r' :: perm | |
| 82 | where "r' j = r'_list ! j" | |
| 83 | ||
| 84 | ||
| 67407 | 85 | \<comment> \<open>amount for rotate left (rol)\<close> | 
| 41561 | 86 | |
| 87 | definition s_list :: "nat list" | |
| 88 | where "s_list = [ | |
| 89 | 11, 14, 15, 12, 5, 8, 7, 9, 11, 13, 14, 15, 6, 7, 9, 8, | |
| 90 | 7, 6, 8, 13, 11, 9, 7, 15, 7, 12, 15, 9, 11, 7, 13, 12, | |
| 91 | 11, 13, 6, 7, 14, 9, 13, 15, 14, 8, 13, 6, 5, 12, 7, 5, | |
| 92 | 11, 12, 14, 15, 14, 15, 9, 8, 9, 14, 5, 6, 8, 6, 5, 12, | |
| 93 | 9, 15, 5, 11, 6, 8, 13, 12, 5, 12, 13, 14, 11, 8, 5, 6 | |
| 94 | ]" | |
| 95 | ||
| 96 | definition s'_list :: "nat list" | |
| 97 | where "s'_list = [ | |
| 98 | 8, 9, 9, 11, 13, 15, 15, 5, 7, 7, 8, 11, 14, 14, 12, 6, | |
| 99 | 9, 13, 15, 7, 12, 8, 9, 11, 7, 7, 12, 7, 6, 15, 13, 11, | |
| 100 | 9, 7, 15, 11, 8, 6, 6, 14, 12, 13, 5, 14, 13, 13, 7, 5, | |
| 101 | 15, 5, 8, 11, 14, 14, 6, 14, 6, 9, 12, 9, 12, 5, 15, 8, | |
| 102 | 8, 5, 12, 9, 12, 5, 14, 6, 8, 13, 6, 5, 15, 13, 11, 11 | |
| 103 | ]" | |
| 104 | ||
| 105 | definition s :: perm | |
| 106 | where "s j = s_list ! j" | |
| 107 | ||
| 108 | definition s' :: perm | |
| 109 | where "s' j = s'_list ! j" | |
| 110 | ||
| 111 | ||
| 67407 | 112 | \<comment> \<open>Initial value (hexadecimal)\<close> | 
| 41561 | 113 | |
| 114 | definition h0_0::word32 where "h0_0 = 0x67452301" | |
| 115 | definition h1_0::word32 where "h1_0 = 0xEFCDAB89" | |
| 116 | definition h2_0::word32 where "h2_0 = 0x98BADCFE" | |
| 117 | definition h3_0::word32 where "h3_0 = 0x10325476" | |
| 118 | definition h4_0::word32 where "h4_0 = 0xC3D2E1F0" | |
| 119 | definition h_0::chain where "h_0 = (h0_0, h1_0, h2_0, h3_0, h4_0)" | |
| 120 | ||
| 121 | ||
| 122 | definition step_l :: | |
| 123 | "[ block, | |
| 124 | chain, | |
| 125 | nat | |
| 126 | ] => chain" | |
| 127 | where | |
| 128 | "step_l X c j = | |
| 129 | (let (A, B, C, D, E) = c in | |
| 67407 | 130 | (\<comment> \<open>\<open>A:\<close>\<close> E, | 
| 131 | \<comment> \<open>\<open>B:\<close>\<close> word_rotl (s j) (A + f j B C D + X (r j) + K j) + E, | |
| 132 | \<comment> \<open>\<open>C:\<close>\<close> B, | |
| 133 | \<comment> \<open>\<open>D:\<close>\<close> word_rotl 10 C, | |
| 134 | \<comment> \<open>\<open>E:\<close>\<close> D))" | |
| 41561 | 135 | |
| 136 | definition step_r :: | |
| 137 | "[ block, | |
| 138 | chain, | |
| 139 | nat | |
| 140 | ] \<Rightarrow> chain" | |
| 141 | where | |
| 142 | "step_r X c' j = | |
| 143 | (let (A', B', C', D', E') = c' in | |
| 67407 | 144 | (\<comment> \<open>\<open>A':\<close>\<close> E', | 
| 145 | \<comment> \<open>\<open>B':\<close>\<close> word_rotl (s' j) (A' + f (79 - j) B' C' D' + X (r' j) + K' j) + E', | |
| 146 | \<comment> \<open>\<open>C':\<close>\<close> B', | |
| 147 | \<comment> \<open>\<open>D':\<close>\<close> word_rotl 10 C', | |
| 148 | \<comment> \<open>\<open>E':\<close>\<close> D'))" | |
| 41561 | 149 | |
| 150 | definition step_both :: | |
| 151 | "[ block, chain * chain, nat ] \<Rightarrow> chain * chain" | |
| 152 | where | |
| 153 | "step_both X cc j = (case cc of (c, c') \<Rightarrow> | |
| 154 | (step_l X c j, step_r X c' j))" | |
| 155 | ||
| 156 | definition steps::"[ block, chain * chain, nat] \<Rightarrow> chain * chain" | |
| 157 | where "steps X cc i = foldl (step_both X) cc [0..<i]" | |
| 158 | ||
| 159 | definition round::"[ block, chain ] \<Rightarrow> chain" | |
| 160 | where "round X h = | |
| 161 | (let (h0, h1, h2, h3, h4) = h in | |
| 162 | let ((A, B, C, D, E), (A', B', C', D', E')) = steps X (h, h) 80 in | |
| 67407 | 163 | (\<comment> \<open>\<open>h0:\<close>\<close> h1 + C + D', | 
| 164 | \<comment> \<open>\<open>h1:\<close>\<close> h2 + D + E', | |
| 165 | \<comment> \<open>\<open>h2:\<close>\<close> h3 + E + A', | |
| 166 | \<comment> \<open>\<open>h3:\<close>\<close> h4 + A + B', | |
| 167 | \<comment> \<open>\<open>h4:\<close>\<close> h0 + B + C'))" | |
| 41561 | 168 | |
| 169 | definition rmd_body::"[ message, chain, nat ] => chain" | |
| 170 | where | |
| 171 | "rmd_body X h i = round (X i) h" | |
| 172 | ||
| 173 | definition rounds::"message \<Rightarrow> chain \<Rightarrow> nat \<Rightarrow> chain" | |
| 174 | where | |
| 175 | "rounds X h i = foldl (rmd_body X) h_0 [0..<i]" | |
| 176 | ||
| 177 | definition rmd :: "message \<Rightarrow> nat \<Rightarrow> chain" | |
| 178 | where | |
| 179 | "rmd X len = rounds X h_0 len" | |
| 180 | ||
| 62390 | 181 | end |