src/HOL/SPARK/Examples/RIPEMD-160/K_R.thy
author wenzelm
Sat, 07 Sep 2013 15:10:33 +0200
changeset 53456 d12be8f62285
parent 41561 d1318f3c86ba
child 56798 939e88e79724
permissions -rw-r--r--
Build_Dialog based on System_Dialog; avoid hopping between threads;

(*  Title:      HOL/SPARK/Examples/RIPEMD-160/K_R.thy
    Author:     Fabian Immler, TU Muenchen

Verification of the RIPEMD-160 hash function
*)

theory K_R
imports RMD_Specification
begin

spark_open "rmd/k_r.siv"

spark_vc function_k_r_6
  using assms by (simp add: K'_def)

spark_vc function_k_r_7
proof-
  from H1 have "16 <= nat j" by simp
  moreover from H2 have "nat j <= 31" by simp
  ultimately show ?thesis by (simp add: K'_def)
qed

spark_vc function_k_r_8
proof -
  from H1 have "32 <= nat j" by simp
  moreover from H2 have "nat j <= 47" by simp
  ultimately show ?thesis by (simp add: K'_def)
qed

spark_vc function_k_r_9
proof -
  from H1 have "48 <= nat j" by simp
  moreover from H2 have "nat j <= 63" by simp
  ultimately show ?thesis by (simp add: K'_def)
qed

spark_vc function_k_r_10
proof -
  from H6 have "64 <= nat j" by simp
  moreover from H2 have "nat j <= 79" by simp
  ultimately show ?thesis by (simp add: K'_def)
qed

spark_end

end