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