| author | wenzelm | 
| Fri, 05 Nov 2021 12:11:30 +0100 | |
| changeset 74697 | c492c8efcab4 | 
| parent 67407 | dbaa38bd223a | 
| permissions | -rw-r--r-- | 
| 41561 | 1 | (* Title: HOL/SPARK/Examples/RIPEMD-160/RMD_Specification.thy | 
| 2 | Author: Fabian Immler, TU Muenchen | |
| 3 | ||
| 4 | Verification of the RIPEMD-160 hash function | |
| 5 | *) | |
| 6 | ||
| 7 | theory RMD_Specification | |
| 66992 
69673025292e
less global theories -- avoid confusion about special cases;
 wenzelm parents: 
66453diff
changeset | 8 | imports RMD "HOL-SPARK.SPARK" | 
| 41561 | 9 | begin | 
| 10 | ||
| 67407 | 11 | \<comment> \<open>bit operations\<close> | 
| 41561 | 12 | |
| 13 | abbreviation rotate_left :: "int \<Rightarrow> int \<Rightarrow> int" where | |
| 14 | "rotate_left i w == uint (word_rotl (nat i) (word_of_int w::word32))" | |
| 15 | ||
| 16 | spark_proof_functions | |
| 17 | wordops__rotate_left = rotate_left | |
| 18 | ||
| 19 | ||
| 67407 | 20 | \<comment> \<open>Conversions for proof functions\<close> | 
| 41561 | 21 | abbreviation k_l_spec :: " int => int " where | 
| 22 | "k_l_spec j == uint (K (nat j))" | |
| 23 | abbreviation k_r_spec :: " int => int " where | |
| 24 | "k_r_spec j == uint (K' (nat j))" | |
| 25 | abbreviation r_l_spec :: " int => int " where | |
| 26 | "r_l_spec j == int (r (nat j))" | |
| 27 | abbreviation r_r_spec :: " int => int " where | |
| 28 | "r_r_spec j == int (r' (nat j))" | |
| 29 | abbreviation s_l_spec :: " int => int " where | |
| 30 | "s_l_spec j == int (s (nat j))" | |
| 31 | abbreviation s_r_spec :: " int => int " where | |
| 32 | "s_r_spec j == int (s' (nat j))" | |
| 33 | abbreviation f_spec :: "int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int" where | |
| 34 | "f_spec j x y z == | |
| 35 | uint (f (nat j) (word_of_int x::word32) (word_of_int y) (word_of_int z))" | |
| 36 | ||
| 37 | spark_proof_functions | |
| 38 | k_l_spec = k_l_spec | |
| 39 | k_r_spec = k_r_spec | |
| 40 | r_l_spec = r_l_spec | |
| 41 | r_r_spec = r_r_spec | |
| 42 | s_l_spec = s_l_spec | |
| 43 | s_r_spec = s_r_spec | |
| 44 | f_spec = f_spec | |
| 45 | ||
| 46 | end |