Added new SPARK verification environment.
authorberghofe
Sat, 15 Jan 2011 12:35:29 +0100
changeset 41561 d1318f3c86ba
parent 41558 236cd8f07f7b
child 41562 90fb3d7474df
Added new SPARK verification environment.
src/HOL/SPARK/Examples/Gcd/Gcd.adb
src/HOL/SPARK/Examples/Gcd/Gcd.ads
src/HOL/SPARK/Examples/Gcd/Greatest_Common_Divisor.thy
src/HOL/SPARK/Examples/Gcd/greatest_common_divisor/g_c_d.fdl
src/HOL/SPARK/Examples/Gcd/greatest_common_divisor/g_c_d.rls
src/HOL/SPARK/Examples/Gcd/greatest_common_divisor/g_c_d.siv
src/HOL/SPARK/Examples/Liseq/Liseq.adb
src/HOL/SPARK/Examples/Liseq/Liseq.ads
src/HOL/SPARK/Examples/Liseq/Longest_Increasing_Subsequence.thy
src/HOL/SPARK/Examples/Liseq/liseq/liseq_length.fdl
src/HOL/SPARK/Examples/Liseq/liseq/liseq_length.rls
src/HOL/SPARK/Examples/Liseq/liseq/liseq_length.siv
src/HOL/SPARK/Examples/README
src/HOL/SPARK/Examples/RIPEMD-160/F.thy
src/HOL/SPARK/Examples/RIPEMD-160/Hash.thy
src/HOL/SPARK/Examples/RIPEMD-160/K_L.thy
src/HOL/SPARK/Examples/RIPEMD-160/K_R.thy
src/HOL/SPARK/Examples/RIPEMD-160/RMD.thy
src/HOL/SPARK/Examples/RIPEMD-160/RMD_Lemmas.thy
src/HOL/SPARK/Examples/RIPEMD-160/RMD_Specification.thy
src/HOL/SPARK/Examples/RIPEMD-160/R_L.thy
src/HOL/SPARK/Examples/RIPEMD-160/R_R.thy
src/HOL/SPARK/Examples/RIPEMD-160/Round.thy
src/HOL/SPARK/Examples/RIPEMD-160/S_L.thy
src/HOL/SPARK/Examples/RIPEMD-160/S_R.thy
src/HOL/SPARK/Examples/RIPEMD-160/rmd.adb
src/HOL/SPARK/Examples/RIPEMD-160/rmd.ads
src/HOL/SPARK/Examples/RIPEMD-160/rmd/f.fdl
src/HOL/SPARK/Examples/RIPEMD-160/rmd/f.rls
src/HOL/SPARK/Examples/RIPEMD-160/rmd/f.siv
src/HOL/SPARK/Examples/RIPEMD-160/rmd/hash.fdl
src/HOL/SPARK/Examples/RIPEMD-160/rmd/hash.rls
src/HOL/SPARK/Examples/RIPEMD-160/rmd/hash.siv
src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_l.fdl
src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_l.rls
src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_l.siv
src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_r.fdl
src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_r.rls
src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_r.siv
src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_l.fdl
src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_l.rls
src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_l.siv
src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_r.fdl
src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_r.rls
src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_r.siv
src/HOL/SPARK/Examples/RIPEMD-160/rmd/round.fdl
src/HOL/SPARK/Examples/RIPEMD-160/rmd/round.rls
src/HOL/SPARK/Examples/RIPEMD-160/rmd/round.siv
src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_l.fdl
src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_l.rls
src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_l.siv
src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_r.fdl
src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_r.rls
src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_r.siv
src/HOL/SPARK/Examples/RIPEMD-160/shadow/interfaces.ads
src/HOL/SPARK/Examples/RIPEMD-160/wordops.adb
src/HOL/SPARK/Examples/RIPEMD-160/wordops.ads
src/HOL/SPARK/Examples/ROOT.ML
src/HOL/SPARK/Examples/Sqrt/Sqrt.adb
src/HOL/SPARK/Examples/Sqrt/Sqrt.ads
src/HOL/SPARK/Examples/Sqrt/Sqrt.thy
src/HOL/SPARK/Examples/Sqrt/sqrt/isqrt.fdl
src/HOL/SPARK/Examples/Sqrt/sqrt/isqrt.rls
src/HOL/SPARK/Examples/Sqrt/sqrt/isqrt.siv
src/HOL/SPARK/ROOT.ML
src/HOL/SPARK/SPARK.thy
src/HOL/SPARK/SPARK_Setup.thy
src/HOL/SPARK/Tools/fdl_lexer.ML
src/HOL/SPARK/Tools/fdl_parser.ML
src/HOL/SPARK/Tools/spark_commands.ML
src/HOL/SPARK/Tools/spark_vcs.ML
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Gcd/Gcd.adb	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,17 @@
+package body Greatest_Common_Divisor
+is
+
+   procedure G_C_D(M, N: in Natural; G: out Natural)
+   is
+      C, D, R: Integer;
+   begin
+      C := M; D := N;
+      while D /= 0 loop
+         --# assert C >= 0 and D > 0 and Gcd(C, D) = Gcd(M, N);
+         R := C rem D;
+         C := D; D := R;
+      end loop;
+      G := C;
+   end G_C_D;
+
+end Greatest_Common_Divisor;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Gcd/Gcd.ads	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,11 @@
+package Greatest_Common_Divisor
+is
+
+   --# function Gcd(A, B: Natural) return Natural;
+
+   procedure G_C_D(M, N: in Natural; G: out Natural);
+     --# derives G from M, N;
+     --# pre M >= 0 and N > 0;
+     --# post G = Gcd(M,N);
+
+end Greatest_Common_Divisor;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Gcd/Greatest_Common_Divisor.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,36 @@
+(*  Title:      HOL/SPARK/Examples/Gcd/Greatest_Common_Divisor.thy
+    Author:     Stefan Berghofer
+    Copyright:  secunet Security Networks AG
+*)
+
+theory Greatest_Common_Divisor
+imports SPARK GCD
+begin
+
+spark_proof_functions
+  gcd = "gcd :: int \<Rightarrow> int \<Rightarrow> int"
+
+spark_open "greatest_common_divisor/g_c_d.siv"
+
+spark_vc procedure_g_c_d_4
+proof -
+  from `0 < d` have "0 \<le> c mod d" by (rule pos_mod_sign)
+  with `0 \<le> c` `0 < d` `c - c sdiv d * d \<noteq> 0` show ?C1
+    by (simp add: sdiv_pos_pos zmod_zdiv_equality')
+next
+  from `0 \<le> c` `0 < d` `gcd c d = gcd m n` show ?C2
+    by (simp add: sdiv_pos_pos zmod_zdiv_equality' gcd_non_0_int)
+qed
+
+spark_vc procedure_g_c_d_11
+proof -
+  from `0 \<le> c` `0 < d` `c - c sdiv d * d = 0`
+  have "d dvd c"
+    by (auto simp add: sdiv_pos_pos dvd_def mult_ac)
+  with `0 < d` `gcd c d = gcd m n` show ?C1
+    by simp
+qed
+
+spark_end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Gcd/greatest_common_divisor/g_c_d.fdl	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,32 @@
+           {*******************************************************}
+                               {FDL Declarations}
+    {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039}
+             {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.}
+           {*******************************************************}
+
+
+                        {DATE : 29-NOV-2010 14:30:10.98}
+
+                  {procedure Greatest_Common_Divisor.G_C_D}
+
+
+title procedure g_c_d;
+
+  function round__(real) : integer;
+  const natural__base__first : integer = pending; 
+  const natural__base__last : integer = pending; 
+  const integer__base__first : integer = pending; 
+  const integer__base__last : integer = pending; 
+  const natural__first : integer = pending; 
+  const natural__last : integer = pending; 
+  const natural__size : integer = pending; 
+  const integer__first : integer = pending; 
+  const integer__last : integer = pending; 
+  const integer__size : integer = pending; 
+  var m : integer;
+  var n : integer;
+  var c : integer;
+  var d : integer;
+  function gcd(integer, integer) : integer;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Gcd/greatest_common_divisor/g_c_d.rls	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,27 @@
+           /*********************************************************/
+                           /*Proof Rule Declarations*/
+    /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/
+             /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/
+           /*********************************************************/
+
+
+                        /*DATE : 29-NOV-2010 14:30:10.98*/
+
+                  /*procedure Greatest_Common_Divisor.G_C_D*/
+
+
+rule_family g_c_d_rules:
+     X      requires [X:any] &
+     X <= Y requires [X:ire, Y:ire] &
+     X >= Y requires [X:ire, Y:ire].
+
+g_c_d_rules(1): integer__size >= 0 may_be_deduced.
+g_c_d_rules(2): integer__first may_be_replaced_by -2147483648.
+g_c_d_rules(3): integer__last may_be_replaced_by 2147483647.
+g_c_d_rules(4): integer__base__first may_be_replaced_by -2147483648.
+g_c_d_rules(5): integer__base__last may_be_replaced_by 2147483647.
+g_c_d_rules(6): natural__size >= 0 may_be_deduced.
+g_c_d_rules(7): natural__first may_be_replaced_by 0.
+g_c_d_rules(8): natural__last may_be_replaced_by 2147483647.
+g_c_d_rules(9): natural__base__first may_be_replaced_by -2147483648.
+g_c_d_rules(10): natural__base__last may_be_replaced_by 2147483647.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Gcd/greatest_common_divisor/g_c_d.siv	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,117 @@
+*****************************************************************************
+                       Semantic Analysis of SPARK Text
+    Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+             Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+*****************************************************************************
+
+
+CREATED 29-NOV-2010, 14:30:10  SIMPLIFIED 29-NOV-2010, 14:30:11
+
+SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+
+procedure Greatest_Common_Divisor.G_C_D
+
+
+
+
+For path(s) from start to run-time check associated with statement of line 8:
+
+procedure_g_c_d_1.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 8:
+
+procedure_g_c_d_2.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to assertion of line 10:
+
+procedure_g_c_d_3.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 10 to assertion of line 10:
+
+procedure_g_c_d_4.
+H1:    c >= 0 .
+H2:    d > 0 .
+H3:    gcd(c, d) = gcd(m, n) .
+H4:    m >= 0 .
+H5:    m <= 2147483647 .
+H6:    n <= 2147483647 .
+H7:    n > 0 .
+H8:    c <= 2147483647 .
+H9:    d <= 2147483647 .
+H10:   c - c div d * d >= - 2147483648 .
+H11:   c - c div d * d <= 2147483647 .
+H12:   c - c div d * d <> 0 .
+H13:   integer__size >= 0 .
+H14:   natural__size >= 0 .
+       ->
+C1:    c - c div d * d > 0 .
+C2:    gcd(d, c - c div d * d) = gcd(m, n) .
+
+
+For path(s) from assertion of line 10 to run-time check associated with 
+          statement of line 11:
+
+procedure_g_c_d_5.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 10 to run-time check associated with 
+          statement of line 12:
+
+procedure_g_c_d_6.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 10 to run-time check associated with 
+          statement of line 12:
+
+procedure_g_c_d_7.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 14:
+
+procedure_g_c_d_8.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 10 to run-time check associated with 
+          statement of line 14:
+
+procedure_g_c_d_9.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to finish:
+
+procedure_g_c_d_10.
+*** true .   /* contradiction within hypotheses. */
+
+
+
+For path(s) from assertion of line 10 to finish:
+
+procedure_g_c_d_11.
+H1:    c >= 0 .
+H2:    d > 0 .
+H3:    gcd(c, d) = gcd(m, n) .
+H4:    m >= 0 .
+H5:    m <= 2147483647 .
+H6:    n <= 2147483647 .
+H7:    n > 0 .
+H8:    c <= 2147483647 .
+H9:    d <= 2147483647 .
+H10:   c - c div d * d = 0 .
+H11:   integer__size >= 0 .
+H12:   natural__size >= 0 .
+       ->
+C1:    d = gcd(m, n) .
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Liseq/Liseq.adb	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,55 @@
+-------------------------------------------------------------------------------
+-- Longest increasing subsequence of an array of integers
+-------------------------------------------------------------------------------
+
+package body Liseq is
+
+   procedure Liseq_length(A: in Vector; L: in out Vector; maxi: out Integer)
+   is
+      maxj,i,j,pmax : Integer;
+   begin
+      L(0) := 1;
+      pmax := 0;
+      maxi := 1;
+      i := 1;
+      while i <= L'Last
+        --# assert
+        --#   (for all i2 in Integer range 0 .. i-1 =>
+        --#      (L(i2) = Liseq_ends_at(A, i2))) and
+        --#   L(pmax) = maxi and L(pmax) = Liseq_prfx(A, i) and
+        --#   0 <= i and i <= L'Last+1 and 0 <= pmax and pmax < i and
+        --#   A'First = 0 and L'First = 0 and A'Last = L'Last and A'Last < Integer'Last;
+      loop
+         if A(i) < A(pmax) then
+            maxj := 0;
+            j := 0;
+            while j < i
+              --# assert
+              --#  (for all i2 in Integer range 0 .. i-1 =>
+              --#     (L(i2) = Liseq_ends_at(A, i2))) and
+              --#  L(pmax) = maxi and L(pmax) = Liseq_prfx(A, I) and
+              --#  0 <= i and i <= L'Last and 0 <= pmax and pmax < i and
+              --#  0 <= j and j <= i and
+              --#  maxj = Max_ext (A, i, j) and
+              --#  A'First = 0 and L'First = 0 and A'Last = L'Last and A'Last < Integer'Last;
+            loop
+               if (A(j) <= A(i) and
+                     maxj < L(j)) then
+                  maxj := L(j);
+               end if;
+               j := j+1;
+            end loop;
+            L(i) := maxj+1;
+            if L(i) > maxi then
+               maxi := maxi+1;
+               pmax := i;
+            end if;
+         else
+            maxi := maxi+1;
+            L(i) := maxi;
+            pmax := i;
+         end if;
+         i := i+1;
+      end loop;
+   end Liseq_length;
+end Liseq;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Liseq/Liseq.ads	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,18 @@
+-------------------------------------------------------------------------------
+-- Longest increasing subsequence of an array of integers
+-------------------------------------------------------------------------------
+
+package Liseq is
+
+   type Vector is array (Integer range <>) of Integer;
+
+   --# function Liseq_prfx(A: Vector; i: Integer) return Integer;
+   --# function Liseq_ends_at(A: Vector; i: Integer) return Integer;
+   --# function Max_ext(A: Vector; i, j: Integer) return Integer;
+
+   procedure Liseq_length(A: in Vector; L: in out Vector; maxi: out Integer);
+   --# derives maxi, L from A, L;
+   --# pre A'First = 0 and L'First = 0 and A'Last = L'Last and A'Last < Integer'Last;
+   --# post maxi = Liseq_prfx (A, A'Last+1);
+
+end Liseq;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Liseq/Longest_Increasing_Subsequence.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,665 @@
+(*  Title:      HOL/SPARK/Examples/Liseq/Longest_Increasing_Subsequence.thy
+    Author:     Stefan Berghofer
+    Copyright:  secunet Security Networks AG
+*)
+
+theory Longest_Increasing_Subsequence
+imports SPARK
+begin
+
+text {*
+Set of all increasing subsequences in a prefix of an array
+*}
+
+definition iseq :: "(nat \<Rightarrow> 'a::linorder) \<Rightarrow> nat \<Rightarrow> nat set set" where
+  "iseq xs l = {is. (\<forall>i\<in>is. i < l) \<and>
+     (\<forall>i\<in>is. \<forall>j\<in>is. i \<le> j \<longrightarrow> xs i \<le> xs j)}"
+
+text {*
+Length of longest increasing subsequence in a prefix of an array
+*}
+
+definition liseq :: "(nat \<Rightarrow> 'a::linorder) \<Rightarrow> nat \<Rightarrow> nat" where
+  "liseq xs i = Max (card ` iseq xs i)"
+
+text {*
+Length of longest increasing subsequence ending at a particular position
+*}
+
+definition liseq' :: "(nat \<Rightarrow> 'a::linorder) \<Rightarrow> nat \<Rightarrow> nat" where
+  "liseq' xs i = Max (card ` (iseq xs (Suc i) \<inter> {is. Max is = i}))"  
+
+lemma iseq_finite: "finite (iseq xs i)"
+  apply (simp add: iseq_def)
+  apply (rule finite_subset [OF _
+    finite_Collect_subsets [of "{j. j < i}"]])
+  apply auto
+  done
+
+lemma iseq_finite': "is \<in> iseq xs i \<Longrightarrow> finite is"
+  by (auto simp add: iseq_def bounded_nat_set_is_finite)
+
+lemma iseq_singleton: "i < l \<Longrightarrow> {i} \<in> iseq xs l"
+  by (simp add: iseq_def)
+
+lemma iseq_trivial: "{} \<in> iseq xs i"
+  by (simp add: iseq_def)
+
+lemma iseq_nonempty: "iseq xs i \<noteq> {}"
+  by (auto intro: iseq_trivial)
+
+lemma liseq'_ge1: "1 \<le> liseq' xs x"
+  apply (simp add: liseq'_def)
+  apply (subgoal_tac "iseq xs (Suc x) \<inter> {is. Max is = x} \<noteq> {}")
+  apply (simp add: Max_ge_iff iseq_finite)
+  apply (rule_tac x="{x}" in bexI)
+  apply (auto intro: iseq_singleton)
+  done
+
+lemma liseq_expand:
+  assumes R: "\<And>is. liseq xs i = card is \<Longrightarrow> is \<in> iseq xs i \<Longrightarrow>
+    (\<And>js. js \<in> iseq xs i \<Longrightarrow> card js \<le> card is) \<Longrightarrow> P"
+  shows "P"
+proof -
+  have "Max (card ` iseq xs i) \<in> card ` iseq xs i"
+    by (rule Max_in) (simp_all add: iseq_finite iseq_nonempty)
+  then obtain js where js: "liseq xs i = card js" and "js \<in> iseq xs i"
+    by (rule imageE) (simp add: liseq_def)
+  moreover {
+    fix js'
+    assume "js' \<in> iseq xs i"
+    then have "card js' \<le> card js"
+      by (simp add: js [symmetric] liseq_def iseq_finite iseq_trivial)
+  }
+  ultimately show ?thesis by (rule R)
+qed
+
+lemma liseq'_expand:
+  assumes R: "\<And>is. liseq' xs i = card is \<Longrightarrow> is \<in> iseq xs (Suc i) \<Longrightarrow>
+    finite is \<Longrightarrow> Max is = i \<Longrightarrow>
+    (\<And>js. js \<in> iseq xs (Suc i) \<Longrightarrow> Max js = i \<Longrightarrow> card js \<le> card is) \<Longrightarrow>
+    is \<noteq> {} \<Longrightarrow> P"
+  shows "P"
+proof -
+  have "Max (card ` (iseq xs (Suc i) \<inter> {is. Max is = i})) \<in>
+    card ` (iseq xs (Suc i) \<inter> {is. Max is = i})"
+    by (auto simp add: iseq_finite intro!: iseq_singleton Max_in)
+  then obtain js where js: "liseq' xs i = card js" and "js \<in> iseq xs (Suc i)"
+    and "finite js" and "Max js = i"
+    by (auto simp add: liseq'_def intro: iseq_finite')
+  moreover {
+    fix js'
+    assume "js' \<in> iseq xs (Suc i)" "Max js' = i"
+    then have "card js' \<le> card js"
+      by (auto simp add: js [symmetric] liseq'_def iseq_finite intro!: iseq_singleton)
+  }
+  note max = this
+  moreover have "card {i} \<le> card js"
+    by (rule max) (simp_all add: iseq_singleton)
+  then have "js \<noteq> {}" by auto
+  ultimately show ?thesis by (rule R)
+qed
+
+lemma liseq'_ge:
+  "j = card js \<Longrightarrow> js \<in> iseq xs (Suc i) \<Longrightarrow> Max js = i \<Longrightarrow>
+  js \<noteq> {} \<Longrightarrow> j \<le> liseq' xs i"
+  by (simp add: liseq'_def iseq_finite)
+
+lemma liseq'_eq:
+  "j = card js \<Longrightarrow> js \<in> iseq xs (Suc i) \<Longrightarrow> Max js = i \<Longrightarrow>
+  js \<noteq> {} \<Longrightarrow> (\<And>js'. js' \<in> iseq xs (Suc i) \<Longrightarrow> Max js' = i \<Longrightarrow> finite js' \<Longrightarrow>
+    js' \<noteq> {} \<Longrightarrow> card js' \<le> card js) \<Longrightarrow>
+  j = liseq' xs i"
+  by (fastsimp simp add: liseq'_def iseq_finite
+    intro: Max_eqI [symmetric])
+
+lemma liseq_ge:
+  "j = card js \<Longrightarrow> js \<in> iseq xs i \<Longrightarrow> j \<le> liseq xs i"
+  by (auto simp add: liseq_def iseq_finite)
+
+lemma liseq_eq:
+  "j = card js \<Longrightarrow> js \<in> iseq xs i \<Longrightarrow>
+  (\<And>js'. js' \<in> iseq xs i \<Longrightarrow> finite js' \<Longrightarrow>
+    js' \<noteq> {} \<Longrightarrow> card js' \<le> card js) \<Longrightarrow>
+  j = liseq xs i"
+  by (fastsimp simp add: liseq_def iseq_finite
+    intro: Max_eqI [symmetric])
+
+lemma max_notin: "finite xs \<Longrightarrow> Max xs < x \<Longrightarrow> x \<notin> xs"
+  by (cases "xs = {}") auto
+
+lemma iseq_insert:
+  "xs (Max is) \<le> xs i \<Longrightarrow> is \<in> iseq xs i \<Longrightarrow>
+  is \<union> {i} \<in> iseq xs (Suc i)"
+  apply (frule iseq_finite')
+  apply (cases "is = {}")
+  apply (auto simp add: iseq_def)
+  apply (rule order_trans [of _ "xs (Max is)"])
+  apply auto
+  apply (thin_tac "\<forall>a\<in>is. a < i")
+  apply (drule_tac x=ia in bspec)
+  apply assumption
+  apply (drule_tac x="Max is" in bspec)
+  apply (auto intro: Max_in)
+  done
+
+lemma iseq_diff: "is \<in> iseq xs (Suc (Max is)) \<Longrightarrow>
+  is - {Max is} \<in> iseq xs (Suc (Max (is - {Max is})))"
+  apply (frule iseq_finite')
+  apply (simp add: iseq_def less_Suc_eq_le)
+  done
+
+lemma iseq_butlast:
+  assumes "js \<in> iseq xs (Suc i)" and "js \<noteq> {}"
+  and "Max js \<noteq> i"
+  shows "js \<in> iseq xs i"
+proof -
+  from assms have fin: "finite js"
+    by (simp add: iseq_finite')
+  with assms have "Max js \<in> js"
+    by auto
+  with assms have "Max js < i"
+    by (auto simp add: iseq_def)
+  with fin assms have "\<forall>j\<in>js. j < i"
+    by (simp add: Max_less_iff)
+  with assms show ?thesis
+    by (simp add: iseq_def)
+qed
+
+lemma iseq_mono: "is \<in> iseq xs i \<Longrightarrow> i \<le> j \<Longrightarrow> is \<in> iseq xs j"
+  by (auto simp add: iseq_def)
+
+lemma diff_nonempty:
+  assumes "1 < card is"
+  shows "is - {i} \<noteq> {}"
+proof -
+  from assms have fin: "finite is" by (auto intro: card_ge_0_finite)
+  with assms fin have "card is - 1 \<le> card (is - {i})"
+    by (simp add: card_Diff_singleton_if)
+  with assms have "0 < card (is - {i})" by simp
+  then show ?thesis by (simp add: card_gt_0_iff)
+qed
+
+lemma Max_diff:
+  assumes "1 < card is"
+  shows "Max (is - {Max is}) < Max is"
+proof -
+  from assms have "finite is" by (auto intro: card_ge_0_finite)
+  moreover from assms have "is - {Max is} \<noteq> {}"
+    by (rule diff_nonempty)
+  ultimately show ?thesis using assms
+    apply (auto simp add: not_less)
+    apply (subgoal_tac "a \<le> Max is")
+    apply auto
+    done
+qed
+
+lemma iseq_nth: "js \<in> iseq xs l \<Longrightarrow> 1 < card js \<Longrightarrow>
+  xs (Max (js - {Max js})) \<le> xs (Max js)"
+  apply (auto simp add: iseq_def)
+  apply (subgoal_tac "Max (js - {Max js}) \<in> js")
+  apply (thin_tac "\<forall>i\<in>js. i < l")
+  apply (drule_tac x="Max (js - {Max js})" in bspec)
+  apply assumption
+  apply (drule_tac x="Max js" in bspec)
+  using card_gt_0_iff [of js]
+  apply simp
+  using Max_diff [of js]
+  apply simp
+  using Max_in [of "js - {Max js}", OF _ diff_nonempty] card_gt_0_iff [of js]
+  apply auto
+  done
+
+lemma card_leq1_singleton:
+  assumes "finite xs" "xs \<noteq> {}" "card xs \<le> 1"
+  obtains x where "xs = {x}"
+  using assms
+  by induct simp_all
+
+lemma longest_iseq1:
+  "liseq' xs i =
+   Max ({0} \<union> {liseq' xs j |j. j < i \<and> xs j \<le> xs i}) + 1"
+proof -
+  have "Max ({0} \<union> {liseq' xs j |j. j < i \<and> xs j \<le> xs i}) = liseq' xs i - 1"
+  proof (rule Max_eqI)
+    fix y
+    assume "y \<in> {0} \<union> {liseq' xs j |j. j < i \<and> xs j \<le> xs i}"
+    then show "y \<le> liseq' xs i - 1"
+    proof
+      assume "y \<in> {liseq' xs j |j. j < i \<and> xs j \<le> xs i}"
+      then obtain j where j: "j < i" "xs j \<le> xs i" "y = liseq' xs j"
+        by auto
+      have "liseq' xs j + 1 \<le> liseq' xs i"
+      proof (rule liseq'_expand)
+        fix "is"
+        assume H: "liseq' xs j = card is" "is \<in> iseq xs (Suc j)"
+          "finite is" "Max is = j" "is \<noteq> {}"
+        from H j have "card is + 1 = card (is \<union> {i})"
+          by (simp add: card_insert max_notin)
+        moreover {
+          from H j have "xs (Max is) \<le> xs i" by simp
+          moreover from `j < i` have "Suc j \<le> i" by simp
+          with `is \<in> iseq xs (Suc j)` have "is \<in> iseq xs i"
+            by (rule iseq_mono)
+          ultimately have "is \<union> {i} \<in> iseq xs (Suc i)"
+          by (rule iseq_insert)
+        } moreover from H j have "Max (is \<union> {i}) = i" by simp
+        moreover have "is \<union> {i} \<noteq> {}" by simp
+        ultimately have "card is + 1 \<le> liseq' xs i"
+          by (rule liseq'_ge)
+        with H show ?thesis by simp
+      qed
+      with j show "y \<le> liseq' xs i - 1"
+        by simp
+    qed simp
+  next
+    have "liseq' xs i \<le> 1 \<or>
+      (\<exists>j. liseq' xs i - 1 = liseq' xs j \<and> j < i \<and> xs j \<le> xs i)"
+    proof (rule liseq'_expand)
+      fix "is"
+      assume H: "liseq' xs i = card is" "is \<in> iseq xs (Suc i)"
+        "finite is" "Max is = i" "is \<noteq> {}"
+      assume R: "\<And>js. js \<in> iseq xs (Suc i) \<Longrightarrow> Max js = i \<Longrightarrow>
+        card js \<le> card is"
+      show ?thesis
+      proof (cases "card is \<le> 1")
+        case True with H show ?thesis by simp
+      next
+        case False
+        then have "1 < card is" by simp
+        then have "Max (is - {Max is}) < Max is"
+          by (rule Max_diff)
+        from `is \<in> iseq xs (Suc i)` `1 < card is`
+        have "xs (Max (is - {Max is})) \<le> xs (Max is)"
+          by (rule iseq_nth)
+        have "card is - 1 = liseq' xs (Max (is - {i}))"
+        proof (rule liseq'_eq)
+          from `Max is = i` [symmetric] `finite is` `is \<noteq> {}`
+          show "card is - 1 = card (is - {i})" by simp
+        next
+          from `is \<in> iseq xs (Suc i)` `Max is = i` [symmetric]
+          show "is - {i} \<in> iseq xs (Suc (Max (is - {i})))"
+            by simp (rule iseq_diff)
+        next
+          from `1 < card is`
+          show "is - {i} \<noteq> {}" by (rule diff_nonempty)
+        next
+          fix js
+          assume "js \<in> iseq xs (Suc (Max (is - {i})))"
+            "Max js = Max (is - {i})" "finite js" "js \<noteq> {}"
+          from `xs (Max (is - {Max is})) \<le> xs (Max is)`
+            `Max js = Max (is - {i})` `Max is = i`
+          have "xs (Max js) \<le> xs i" by simp
+          moreover from `Max is = i` `Max (is - {Max is}) < Max is`
+          have "Suc (Max (is - {i})) \<le> i"
+            by simp
+          with `js \<in> iseq xs (Suc (Max (is - {i})))`
+          have "js \<in> iseq xs i"
+            by (rule iseq_mono)
+          ultimately have "js \<union> {i} \<in> iseq xs (Suc i)"
+            by (rule iseq_insert)
+          moreover from `js \<noteq> {}` `finite js` `Max js = Max (is - {i})`
+            `Max is = i` [symmetric] `Max (is - {Max is}) < Max is`
+          have "Max (js \<union> {i}) = i"
+            by simp
+          ultimately have "card (js \<union> {i}) \<le> card is" by (rule R)
+          moreover from `Max is = i` [symmetric] `finite js`
+            `Max (is - {Max is}) < Max is` `Max js = Max (is - {i})`
+          have "i \<notin> js" by (simp add: max_notin)
+          with `finite js`
+          have "card (js \<union> {i}) = card ((js \<union> {i}) - {i}) + 1"
+            by simp
+          ultimately show "card js \<le> card (is - {i})"
+            using `i \<notin> js` `Max is = i` [symmetric] `is \<noteq> {}` `finite is`
+            by simp
+        qed simp
+        with H `Max (is - {Max is}) < Max is`
+          `xs (Max (is - {Max is})) \<le> xs (Max is)`
+        show ?thesis by auto
+      qed
+    qed
+    then show "liseq' xs i - 1 \<in> {0} \<union>
+      {liseq' xs j |j. j < i \<and> xs j \<le> xs i}" by simp
+  qed simp
+  moreover have "1 \<le> liseq' xs i" by (rule liseq'_ge1)
+  ultimately show ?thesis by simp
+qed
+
+lemma longest_iseq2': "liseq xs i < liseq' xs i \<Longrightarrow>
+  liseq xs (Suc i) = liseq' xs i"
+  apply (rule_tac xs=xs and i=i in liseq'_expand)
+  apply simp
+  apply (rule liseq_eq [symmetric])
+  apply (rule refl)
+  apply assumption
+  apply (case_tac "Max js' = i")
+  apply simp
+  apply (drule_tac js=js' in iseq_butlast)
+  apply assumption+
+  apply (drule_tac js=js' in liseq_ge [OF refl])
+  apply simp
+  done
+
+lemma longest_iseq2: "liseq xs i < liseq' xs i \<Longrightarrow>
+  liseq xs i + 1 = liseq' xs i"
+  apply (rule_tac xs=xs and i=i in liseq'_expand)
+  apply simp
+  apply (rule_tac xs=xs and i=i in liseq_expand)
+  apply (drule_tac s="Max is" in sym)
+  apply simp
+  apply (case_tac "card is \<le> 1")
+  apply simp
+  apply (drule iseq_diff)
+  apply (drule_tac i="Suc (Max (is - {Max is}))" and j="Max is" in iseq_mono)
+  apply (simp add: less_eq_Suc_le [symmetric])
+  apply (rule Max_diff)
+  apply simp
+  apply (drule_tac x="is - {Max is}" in meta_spec,
+    drule meta_mp, assumption)
+  apply simp
+  done
+
+lemma longest_iseq3:
+  "liseq xs j = liseq' xs i \<Longrightarrow> xs i \<le> xs j \<Longrightarrow> i < j \<Longrightarrow>
+  liseq xs (Suc j) = liseq xs j + 1"
+  apply (rule_tac xs=xs and i=j in liseq_expand)
+  apply simp
+  apply (rule_tac xs=xs and i=i in liseq'_expand)
+  apply simp
+  apply (rule_tac js="isa \<union> {j}" in liseq_eq [symmetric])
+  apply (simp add: card_insert card_Diff_singleton_if max_notin)
+  apply (rule iseq_insert)
+  apply simp
+  apply (erule iseq_mono)
+  apply simp
+  apply (case_tac "j = Max js'")
+  apply simp
+  apply (drule iseq_diff)
+  apply (drule_tac x="js' - {j}" in meta_spec)
+  apply (drule meta_mp)
+  apply simp
+  apply (case_tac "card js' \<le> 1")
+  apply (erule_tac xs=js' in card_leq1_singleton)
+  apply assumption+
+  apply (simp add: iseq_trivial)
+  apply (erule iseq_mono)
+  apply (simp add: less_eq_Suc_le [symmetric])
+  apply (rule Max_diff)
+  apply simp
+  apply (rule le_diff_iff [THEN iffD1, of 1])
+  apply (simp add: card_0_eq [symmetric] del: card_0_eq)
+  apply (simp add: card_insert)
+  apply (subgoal_tac "card (js' - {j}) = card js' - 1")
+  apply (simp add: card_insert card_Diff_singleton_if max_notin)
+  apply (frule_tac A=js' in Max_in)
+  apply assumption
+  apply (simp add: card_Diff_singleton_if)
+  apply (drule_tac js=js' in iseq_butlast)
+  apply assumption
+  apply (erule not_sym)
+  apply (drule_tac x=js' in meta_spec)
+  apply (drule meta_mp)
+  apply assumption
+  apply (simp add: card_insert_disjoint max_notin)
+  done
+
+lemma longest_iseq4:
+  "liseq xs j = liseq' xs i \<Longrightarrow> xs i \<le> xs j \<Longrightarrow> i < j \<Longrightarrow>
+  liseq' xs j = liseq' xs i + 1"
+  apply (rule_tac xs=xs and i=j in liseq_expand)
+  apply simp
+  apply (rule_tac xs=xs and i=i in liseq'_expand)
+  apply simp
+  apply (rule_tac js="isa \<union> {j}" in liseq'_eq [symmetric])
+  apply (simp add: card_insert card_Diff_singleton_if max_notin)
+  apply (rule iseq_insert)
+  apply simp
+  apply (erule iseq_mono)
+  apply simp
+  apply simp
+  apply simp
+  apply (drule_tac s="Max js'" in sym)
+  apply simp
+  apply (drule iseq_diff)
+  apply (drule_tac x="js' - {j}" in meta_spec)
+  apply (drule meta_mp)
+  apply simp
+  apply (case_tac "card js' \<le> 1")
+  apply (erule_tac xs=js' in card_leq1_singleton)
+  apply assumption+
+  apply (simp add: iseq_trivial)
+  apply (erule iseq_mono)
+  apply (simp add: less_eq_Suc_le [symmetric])
+  apply (rule Max_diff)
+  apply simp
+  apply (rule le_diff_iff [THEN iffD1, of 1])
+  apply (simp add: card_0_eq [symmetric] del: card_0_eq)
+  apply (simp add: card_insert)
+  apply (subgoal_tac "card (js' - {j}) = card js' - 1")
+  apply (simp add: card_insert card_Diff_singleton_if max_notin)
+  apply (frule_tac A=js' in Max_in)
+  apply assumption
+  apply (simp add: card_Diff_singleton_if)
+  done
+
+lemma longest_iseq5: "liseq' xs i \<le> liseq xs i \<Longrightarrow>
+  liseq xs (Suc i) = liseq xs i"
+  apply (rule_tac i=i and xs=xs in liseq'_expand)
+  apply simp
+  apply (rule_tac xs=xs and i=i in liseq_expand)
+  apply simp
+  apply (rule liseq_eq [symmetric])
+  apply (rule refl)
+  apply (erule iseq_mono)
+  apply simp
+  apply (case_tac "Max js' = i")
+  apply (drule_tac x=js' in meta_spec)
+  apply simp
+  apply (drule iseq_butlast, assumption, assumption)
+  apply simp
+  done
+
+lemma liseq_empty: "liseq xs 0 = 0"
+  apply (rule_tac js="{}" in liseq_eq [symmetric])
+  apply simp
+  apply (rule iseq_trivial)
+  apply (simp add: iseq_def)
+  done
+
+lemma liseq'_singleton: "liseq' xs 0 = 1"
+  by (simp add: longest_iseq1 [of _ 0])
+
+lemma liseq_singleton: "liseq xs (Suc 0) = Suc 0"
+  by (simp add: longest_iseq2' liseq_empty liseq'_singleton)
+
+lemma liseq'_Suc_unfold:
+  "A j \<le> x \<Longrightarrow>
+   (insert 0 {liseq' A j' |j'. j' < Suc j \<and> A j' \<le> x}) =
+   (insert 0 {liseq' A j' |j'. j' < j \<and> A j' \<le> x}) \<union>
+   {liseq' A j}"
+  by (auto simp add: less_Suc_eq)
+
+lemma liseq'_Suc_unfold':
+  "\<not> (A j \<le> x) \<Longrightarrow>
+   {liseq' A j' |j'. j' < Suc j \<and> A j' \<le> x} =
+   {liseq' A j' |j'. j' < j \<and> A j' \<le> x}"
+  by (auto simp add: less_Suc_eq)
+
+lemma iseq_card_limit:
+  assumes "is \<in> iseq A i"
+  shows "card is \<le> i"
+proof -
+  from assms have "is \<subseteq> {0..<i}"
+    by (auto simp add: iseq_def)
+  with finite_atLeastLessThan have "card is \<le> card {0..<i}"
+    by (rule card_mono)
+  with card_atLeastLessThan show ?thesis by simp
+qed
+
+lemma liseq_limit: "liseq A i \<le> i"
+  by (rule_tac xs=A and i=i in liseq_expand)
+    (simp add: iseq_card_limit)
+
+lemma liseq'_limit: "liseq' A i \<le> i + 1"
+  by (rule_tac xs=A and i=i in liseq'_expand)
+    (simp add: iseq_card_limit)
+
+definition max_ext :: "(nat \<Rightarrow> 'a::linorder) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" where
+  "max_ext A i j = Max ({0} \<union> {liseq' A j' |j'. j' < j \<and> A j' \<le> A i})"
+
+lemma max_ext_limit: "max_ext A i j \<le> j"
+  apply (auto simp add: max_ext_def)
+  apply (drule Suc_leI)
+  apply (cut_tac i=j' and A=A in liseq'_limit)
+  apply simp
+  done
+
+
+text {* Proof functions *}
+
+abbreviation (input)
+  "arr_conv a \<equiv> (\<lambda>n. a (int n))"
+
+lemma idx_conv_suc:
+  "0 \<le> i \<Longrightarrow> nat (i + 1) = nat i + 1"
+  by simp
+
+abbreviation liseq_ends_at' :: "(int \<Rightarrow> 'a::linorder) \<Rightarrow> int \<Rightarrow> int" where
+  "liseq_ends_at' A i \<equiv> int (liseq' (\<lambda>l. A (int l)) (nat i))"
+
+abbreviation liseq_prfx' :: "(int \<Rightarrow> 'a::linorder) \<Rightarrow> int \<Rightarrow> int" where
+  "liseq_prfx' A i \<equiv> int (liseq (\<lambda>l. A (int l)) (nat i))"
+
+abbreviation max_ext' :: "(int \<Rightarrow> 'a::linorder) \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int" where
+  "max_ext' A i j \<equiv> int (max_ext (\<lambda>l. A (int l)) (nat i) (nat j))"
+
+spark_proof_functions
+  liseq_ends_at = "liseq_ends_at' :: (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> int"
+  liseq_prfx = "liseq_prfx' :: (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> int"
+  max_ext = "max_ext' :: (int \<Rightarrow> int) \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int"
+
+
+text {* The verification conditions *}
+
+spark_open "liseq/liseq_length.siv"
+
+spark_vc procedure_liseq_length_5
+  by (simp_all add: liseq_singleton liseq'_singleton)
+
+spark_vc procedure_liseq_length_6
+proof -
+  from H1 H2 H3 H4
+  have eq: "liseq (arr_conv a) (nat i) =
+    liseq' (arr_conv a) (nat pmax)"
+    by simp
+  from H14 H3 H4
+  have pmax1: "arr_conv a (nat pmax) \<le> arr_conv a (nat i)"
+    by simp
+  from H3 H4 have pmax2: "nat pmax < nat i"
+    by simp
+  {
+    fix i2
+    assume i2: "0 \<le> i2" "i2 \<le> i"
+    have "(l(i := l pmax + 1)) i2 =
+      int (liseq' (arr_conv a) (nat i2))"
+    proof (cases "i2 = i")
+      case True
+      from eq pmax1 pmax2 have "liseq' (arr_conv a) (nat i) =
+        liseq' (arr_conv a) (nat pmax) + 1"
+        by (rule longest_iseq4)
+      with True H1 H3 H4 show ?thesis
+        by simp
+    next
+      case False
+      with H1 i2 show ?thesis
+        by simp
+    qed
+  }
+  then show ?C1 by simp
+  from eq pmax1 pmax2
+  have "liseq (arr_conv a) (Suc (nat i)) =
+    liseq (arr_conv a) (nat i) + 1"
+    by (rule longest_iseq3)
+ with H2 H3 H4 show ?C2
+    by (simp add: idx_conv_suc)
+qed
+
+spark_vc procedure_liseq_length_7
+proof -
+  from H1 show ?C1
+    by (simp add: max_ext_def longest_iseq1 [of _ "nat i"])
+  from H6
+  have m: "max_ext (arr_conv a) (nat i) (nat i) + 1 =
+    liseq' (arr_conv a) (nat i)"
+    by (simp add: max_ext_def longest_iseq1 [of _ "nat i"])
+  with H2 H18
+  have gt: "liseq (arr_conv a) (nat i) < liseq' (arr_conv a) (nat i)"
+    by simp
+  then have "liseq' (arr_conv a) (nat i) = liseq (arr_conv a) (nat i) + 1"
+    by (rule longest_iseq2 [symmetric])
+  with H2 m show ?C2 by simp
+  from gt have "liseq (arr_conv a) (Suc (nat i)) = liseq' (arr_conv a) (nat i)"
+    by (rule longest_iseq2')
+  with m H6 show ?C3 by (simp add: idx_conv_suc)
+qed
+
+spark_vc procedure_liseq_length_8
+proof -
+  {
+    fix i2
+    assume i2: "0 \<le> i2" "i2 \<le> i"
+    have "(l(i := max_ext' a i i + 1)) i2 =
+      int (liseq' (arr_conv a) (nat i2))"
+    proof (cases "i2 = i")
+      case True
+      with H1 show ?thesis
+        by (simp add: max_ext_def longest_iseq1 [of _ "nat i"])
+    next
+      case False
+      with H1 i2 show ?thesis by simp
+    qed
+  }
+  then show ?C1 by simp
+  from H2 H6 H18
+  have "liseq' (arr_conv a) (nat i) \<le> liseq (arr_conv a) (nat i)"
+    by (simp add: max_ext_def longest_iseq1 [of _ "nat i"])
+  then have "liseq (arr_conv a) (Suc (nat i)) = liseq (arr_conv a) (nat i)"
+    by (rule longest_iseq5)
+  with H2 H6 show ?C2 by (simp add: idx_conv_suc)
+qed
+
+spark_vc procedure_liseq_length_12
+  by (simp add: max_ext_def)
+
+spark_vc procedure_liseq_length_13
+  using H1 H6 H13 H21 H22
+  by (simp add: max_ext_def
+    idx_conv_suc liseq'_Suc_unfold max_def del: Max_less_iff)
+
+spark_vc procedure_liseq_length_14
+  using H1 H6 H13 H21
+  by (cases "a j \<le> a i")
+    (simp_all add: max_ext_def
+      idx_conv_suc liseq'_Suc_unfold liseq'_Suc_unfold')
+
+spark_vc procedure_liseq_length_19
+  using H3 H4 H5 H8 H9
+  apply (rule_tac y="int (nat i)" in order_trans)
+  apply (cut_tac A="arr_conv a" and i="nat i" and j="nat i" in max_ext_limit)
+  apply simp_all
+  done
+
+spark_vc procedure_liseq_length_23
+  using H2 H3 H4 H7 H8 H11
+  apply (rule_tac y="int (nat i)" in order_trans)
+  apply (cut_tac A="arr_conv a" and i="nat i" in liseq_limit)
+  apply simp_all
+  done
+
+spark_vc procedure_liseq_length_29
+  using H2 H3 H8 H13
+  by (simp add: add1_zle_eq [symmetric])
+
+spark_end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Liseq/liseq/liseq_length.fdl	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,37 @@
+           {*******************************************************}
+                               {FDL Declarations}
+    {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039}
+             {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.}
+           {*******************************************************}
+
+
+                        {DATE : 29-NOV-2010 14:30:13.02}
+
+                        {procedure Liseq.Liseq_length}
+
+
+title procedure liseq_length;
+
+  function round__(real) : integer;
+  type vector = array [integer] of integer;
+  const integer__base__first : integer = pending; 
+  const integer__base__last : integer = pending; 
+  const l__index__subtype__1__first : integer = pending; 
+  const l__index__subtype__1__last : integer = pending; 
+  const a__index__subtype__1__first : integer = pending; 
+  const a__index__subtype__1__last : integer = pending; 
+  const integer__first : integer = pending; 
+  const integer__last : integer = pending; 
+  const integer__size : integer = pending; 
+  var a : vector;
+  var l : vector;
+  var maxi : integer;
+  var maxj : integer;
+  var i : integer;
+  var j : integer;
+  var pmax : integer;
+  function liseq_prfx(vector, integer) : integer;
+  function liseq_ends_at(vector, integer) : integer;
+  function max_ext(vector, integer, integer) : integer;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Liseq/liseq/liseq_length.rls	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,34 @@
+           /*********************************************************/
+                           /*Proof Rule Declarations*/
+    /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/
+             /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/
+           /*********************************************************/
+
+
+                        /*DATE : 29-NOV-2010 14:30:13.02*/
+
+                        /*procedure Liseq.Liseq_length*/
+
+
+rule_family liseq_length_rules:
+     X      requires [X:any] &
+     X <= Y requires [X:ire, Y:ire] &
+     X >= Y requires [X:ire, Y:ire].
+
+liseq_length_rules(1): integer__size >= 0 may_be_deduced.
+liseq_length_rules(2): integer__first may_be_replaced_by -2147483648.
+liseq_length_rules(3): integer__last may_be_replaced_by 2147483647.
+liseq_length_rules(4): integer__base__first may_be_replaced_by -2147483648.
+liseq_length_rules(5): integer__base__last may_be_replaced_by 2147483647.
+liseq_length_rules(6): a__index__subtype__1__first >= integer__first may_be_deduced.
+liseq_length_rules(7): a__index__subtype__1__last <= integer__last may_be_deduced.
+liseq_length_rules(8): a__index__subtype__1__first <= 
+     a__index__subtype__1__last may_be_deduced.
+liseq_length_rules(9): a__index__subtype__1__last >= integer__first may_be_deduced.
+liseq_length_rules(10): a__index__subtype__1__first <= integer__last may_be_deduced.
+liseq_length_rules(11): l__index__subtype__1__first >= integer__first may_be_deduced.
+liseq_length_rules(12): l__index__subtype__1__last <= integer__last may_be_deduced.
+liseq_length_rules(13): l__index__subtype__1__first <= 
+     l__index__subtype__1__last may_be_deduced.
+liseq_length_rules(14): l__index__subtype__1__last >= integer__first may_be_deduced.
+liseq_length_rules(15): l__index__subtype__1__first <= integer__last may_be_deduced.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Liseq/liseq/liseq_length.siv	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,547 @@
+*****************************************************************************
+                       Semantic Analysis of SPARK Text
+    Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+             Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+*****************************************************************************
+
+
+CREATED 29-NOV-2010, 14:30:13  SIMPLIFIED 29-NOV-2010, 14:30:13
+
+SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+
+procedure Liseq.Liseq_length
+
+
+
+
+For path(s) from start to run-time check associated with statement of line 11:
+
+procedure_liseq_length_1.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 12:
+
+procedure_liseq_length_2.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 13:
+
+procedure_liseq_length_3.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 14:
+
+procedure_liseq_length_4.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to assertion of line 15:
+
+procedure_liseq_length_5.
+H1:    a__index__subtype__1__first = 0 .
+H2:    l__index__subtype__1__first = 0 .
+H3:    a__index__subtype__1__last = l__index__subtype__1__last .
+H4:    a__index__subtype__1__last < 2147483647 .
+H5:    for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 
+          <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) 
+          and element(a, [i___1]) <= 2147483647) .
+H6:    for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 
+          <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) 
+          and element(l, [i___1]) <= 2147483647) .
+H7:    0 <= l__index__subtype__1__last .
+H8:    integer__size >= 0 .
+H9:    a__index__subtype__1__first <= a__index__subtype__1__last .
+H10:   l__index__subtype__1__first <= l__index__subtype__1__last .
+H11:   a__index__subtype__1__first >= - 2147483648 .
+H12:   a__index__subtype__1__last >= - 2147483648 .
+H13:   l__index__subtype__1__first >= - 2147483648 .
+H14:   l__index__subtype__1__last >= - 2147483648 .
+H15:   a__index__subtype__1__last <= 2147483647 .
+H16:   a__index__subtype__1__first <= 2147483647 .
+H17:   l__index__subtype__1__last <= 2147483647 .
+H18:   l__index__subtype__1__first <= 2147483647 .
+       ->
+C1:    for_all(i2_ : integer, 0 <= i2_ and i2_ <= 0 -> element(update(l, [0], 1)
+          , [i2_]) = liseq_ends_at(a, i2_)) .
+C2:    1 = liseq_prfx(a, 1) .
+
+
+For path(s) from assertion of line 15 to assertion of line 15:
+
+procedure_liseq_length_6.
+H1:    for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = 
+          liseq_ends_at(a, i2_)) .
+H2:    element(l, [pmax]) = liseq_prfx(a, i) .
+H3:    0 <= pmax .
+H4:    pmax < i .
+H5:    a__index__subtype__1__first = 0 .
+H6:    l__index__subtype__1__first = 0 .
+H7:    a__index__subtype__1__last = l__index__subtype__1__last .
+H8:    a__index__subtype__1__last < 2147483647 .
+H9:    for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 
+          <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) 
+          and element(a, [i___1]) <= 2147483647) .
+H10:   for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 
+          <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) 
+          and element(l, [i___1]) <= 2147483647) .
+H11:   i <= l__index__subtype__1__last .
+H12:   pmax >= a__index__subtype__1__first .
+H13:   i <= a__index__subtype__1__last .
+H14:   element(a, [pmax]) <= element(a, [i]) .
+H15:   element(l, [pmax]) >= - 2147483648 .
+H16:   element(l, [pmax]) <= 2147483646 .
+H17:   i >= l__index__subtype__1__first .
+H18:   i <= 2147483646 .
+H19:   integer__size >= 0 .
+H20:   a__index__subtype__1__first <= a__index__subtype__1__last .
+H21:   l__index__subtype__1__first <= l__index__subtype__1__last .
+H22:   a__index__subtype__1__first >= - 2147483648 .
+H23:   a__index__subtype__1__last >= - 2147483648 .
+H24:   l__index__subtype__1__first >= - 2147483648 .
+H25:   l__index__subtype__1__last >= - 2147483648 .
+H26:   a__index__subtype__1__last <= 2147483647 .
+H27:   a__index__subtype__1__first <= 2147483647 .
+H28:   l__index__subtype__1__last <= 2147483647 .
+H29:   l__index__subtype__1__first <= 2147483647 .
+       ->
+C1:    for_all(i2_ : integer, 0 <= i2_ and i2_ <= i -> element(update(l, [i], 
+          element(l, [pmax]) + 1), [i2_]) = liseq_ends_at(a, i2_)) .
+C2:    element(l, [pmax]) + 1 = liseq_prfx(a, i + 1) .
+
+
+For path(s) from assertion of line 26 to assertion of line 15:
+
+procedure_liseq_length_7.
+H1:    for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = 
+          liseq_ends_at(a, i2_)) .
+H2:    element(l, [pmax]) = liseq_prfx(a, i) .
+H3:    i <= l__index__subtype__1__last .
+H4:    0 <= pmax .
+H5:    pmax < i .
+H6:    0 <= i .
+H7:    a__index__subtype__1__first = 0 .
+H8:    l__index__subtype__1__first = 0 .
+H9:    a__index__subtype__1__last = l__index__subtype__1__last .
+H10:   a__index__subtype__1__last < 2147483647 .
+H11:   for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 
+          <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) 
+          and element(a, [i___1]) <= 2147483647) .
+H12:   for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 
+          <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) 
+          and element(l, [i___1]) <= 2147483647) .
+H13:   i <= 2147483647 .
+H14:   max_ext(a, i, i) >= - 2147483648 .
+H15:   max_ext(a, i, i) <= 2147483646 .
+H16:   i >= l__index__subtype__1__first .
+H17:   element(l, [pmax]) >= - 2147483648 .
+H18:   max_ext(a, i, i) + 1 > element(l, [pmax]) .
+H19:   element(l, [pmax]) <= 2147483646 .
+H20:   i <= 2147483646 .
+H21:   integer__size >= 0 .
+H22:   a__index__subtype__1__first <= a__index__subtype__1__last .
+H23:   l__index__subtype__1__first <= l__index__subtype__1__last .
+H24:   a__index__subtype__1__first >= - 2147483648 .
+H25:   a__index__subtype__1__last >= - 2147483648 .
+H26:   l__index__subtype__1__first >= - 2147483648 .
+H27:   l__index__subtype__1__last >= - 2147483648 .
+H28:   a__index__subtype__1__last <= 2147483647 .
+H29:   a__index__subtype__1__first <= 2147483647 .
+H30:   l__index__subtype__1__last <= 2147483647 .
+H31:   l__index__subtype__1__first <= 2147483647 .
+       ->
+C1:    for_all(i2_ : integer, 0 <= i2_ and i2_ <= i -> element(update(l, [i], 
+          max_ext(a, i, i) + 1), [i2_]) = liseq_ends_at(a, i2_)) .
+C2:    max_ext(a, i, i) + 1 = element(l, [pmax]) + 1 .
+C3:    max_ext(a, i, i) + 1 = liseq_prfx(a, i + 1) .
+
+
+procedure_liseq_length_8.
+H1:    for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = 
+          liseq_ends_at(a, i2_)) .
+H2:    element(l, [pmax]) = liseq_prfx(a, i) .
+H3:    i <= l__index__subtype__1__last .
+H4:    0 <= pmax .
+H5:    pmax < i .
+H6:    0 <= i .
+H7:    a__index__subtype__1__first = 0 .
+H8:    l__index__subtype__1__first = 0 .
+H9:    a__index__subtype__1__last = l__index__subtype__1__last .
+H10:   a__index__subtype__1__last < 2147483647 .
+H11:   for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 
+          <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) 
+          and element(a, [i___1]) <= 2147483647) .
+H12:   for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 
+          <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) 
+          and element(l, [i___1]) <= 2147483647) .
+H13:   i <= 2147483647 .
+H14:   max_ext(a, i, i) >= - 2147483648 .
+H15:   max_ext(a, i, i) <= 2147483646 .
+H16:   i >= l__index__subtype__1__first .
+H17:   element(l, [pmax]) <= 2147483647 .
+H18:   max_ext(a, i, i) + 1 <= element(l, [pmax]) .
+H19:   i <= 2147483646 .
+H20:   integer__size >= 0 .
+H21:   a__index__subtype__1__first <= a__index__subtype__1__last .
+H22:   l__index__subtype__1__first <= l__index__subtype__1__last .
+H23:   a__index__subtype__1__first >= - 2147483648 .
+H24:   a__index__subtype__1__last >= - 2147483648 .
+H25:   l__index__subtype__1__first >= - 2147483648 .
+H26:   l__index__subtype__1__last >= - 2147483648 .
+H27:   a__index__subtype__1__last <= 2147483647 .
+H28:   a__index__subtype__1__first <= 2147483647 .
+H29:   l__index__subtype__1__last <= 2147483647 .
+H30:   l__index__subtype__1__first <= 2147483647 .
+       ->
+C1:    for_all(i2_ : integer, 0 <= i2_ and i2_ <= i -> element(update(l, [i], 
+          max_ext(a, i, i) + 1), [i2_]) = liseq_ends_at(a, i2_)) .
+C2:    element(l, [pmax]) = liseq_prfx(a, i + 1) .
+
+
+For path(s) from assertion of line 15 to run-time check associated with 
+          statement of line 23:
+
+procedure_liseq_length_9.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 15 to run-time check associated with 
+          statement of line 24:
+
+procedure_liseq_length_10.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 15 to run-time check associated with 
+          statement of line 25:
+
+procedure_liseq_length_11.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 15 to assertion of line 26:
+
+procedure_liseq_length_12.
+H1:    for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = 
+          liseq_ends_at(a, i2_)) .
+H2:    element(l, [pmax]) = liseq_prfx(a, i) .
+H3:    0 <= pmax .
+H4:    pmax < i .
+H5:    a__index__subtype__1__first = 0 .
+H6:    l__index__subtype__1__first = 0 .
+H7:    a__index__subtype__1__last = l__index__subtype__1__last .
+H8:    a__index__subtype__1__last < 2147483647 .
+H9:    for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 
+          <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) 
+          and element(a, [i___1]) <= 2147483647) .
+H10:   for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 
+          <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) 
+          and element(l, [i___1]) <= 2147483647) .
+H11:   i <= l__index__subtype__1__last .
+H12:   pmax <= 2147483647 .
+H13:   pmax >= a__index__subtype__1__first .
+H14:   i <= a__index__subtype__1__last .
+H15:   element(a, [i]) < element(a, [pmax]) .
+H16:   integer__size >= 0 .
+H17:   a__index__subtype__1__first <= a__index__subtype__1__last .
+H18:   l__index__subtype__1__first <= l__index__subtype__1__last .
+H19:   a__index__subtype__1__first >= - 2147483648 .
+H20:   a__index__subtype__1__last >= - 2147483648 .
+H21:   l__index__subtype__1__first >= - 2147483648 .
+H22:   l__index__subtype__1__last >= - 2147483648 .
+H23:   a__index__subtype__1__last <= 2147483647 .
+H24:   a__index__subtype__1__first <= 2147483647 .
+H25:   l__index__subtype__1__last <= 2147483647 .
+H26:   l__index__subtype__1__first <= 2147483647 .
+       ->
+C1:    0 = max_ext(a, i, 0) .
+
+
+For path(s) from assertion of line 26 to assertion of line 26:
+
+procedure_liseq_length_13.
+H1:    for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = 
+          liseq_ends_at(a, i2_)) .
+H2:    element(l, [pmax]) = liseq_prfx(a, i) .
+H3:    i <= l__index__subtype__1__last .
+H4:    0 <= pmax .
+H5:    pmax < i .
+H6:    0 <= j .
+H7:    a__index__subtype__1__first = 0 .
+H8:    l__index__subtype__1__first = 0 .
+H9:    a__index__subtype__1__last = l__index__subtype__1__last .
+H10:   a__index__subtype__1__last < 2147483647 .
+H11:   for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 
+          <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) 
+          and element(a, [i___1]) <= 2147483647) .
+H12:   for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 
+          <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) 
+          and element(l, [i___1]) <= 2147483647) .
+H13:   j < i .
+H14:   max_ext(a, i, j) >= - 2147483648 .
+H15:   max_ext(a, i, j) <= 2147483647 .
+H16:   j >= l__index__subtype__1__first .
+H17:   i >= a__index__subtype__1__first .
+H18:   i <= a__index__subtype__1__last .
+H19:   j >= a__index__subtype__1__first .
+H20:   j <= a__index__subtype__1__last .
+H21:   element(a, [j]) <= element(a, [i]) .
+H22:   max_ext(a, i, j) < element(l, [j]) .
+H23:   element(l, [j]) >= - 2147483648 .
+H24:   element(l, [j]) <= 2147483647 .
+H25:   j <= 2147483646 .
+H26:   integer__size >= 0 .
+H27:   a__index__subtype__1__first <= a__index__subtype__1__last .
+H28:   l__index__subtype__1__first <= l__index__subtype__1__last .
+H29:   a__index__subtype__1__first >= - 2147483648 .
+H30:   a__index__subtype__1__last >= - 2147483648 .
+H31:   l__index__subtype__1__first >= - 2147483648 .
+H32:   l__index__subtype__1__last >= - 2147483648 .
+H33:   a__index__subtype__1__last <= 2147483647 .
+H34:   a__index__subtype__1__first <= 2147483647 .
+H35:   l__index__subtype__1__last <= 2147483647 .
+H36:   l__index__subtype__1__first <= 2147483647 .
+       ->
+C1:    element(l, [j]) = max_ext(a, i, j + 1) .
+
+
+procedure_liseq_length_14.
+H1:    for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = 
+          liseq_ends_at(a, i2_)) .
+H2:    element(l, [pmax]) = liseq_prfx(a, i) .
+H3:    i <= l__index__subtype__1__last .
+H4:    0 <= pmax .
+H5:    pmax < i .
+H6:    0 <= j .
+H7:    a__index__subtype__1__first = 0 .
+H8:    l__index__subtype__1__first = 0 .
+H9:    a__index__subtype__1__last = l__index__subtype__1__last .
+H10:   a__index__subtype__1__last < 2147483647 .
+H11:   for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 
+          <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) 
+          and element(a, [i___1]) <= 2147483647) .
+H12:   for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 
+          <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) 
+          and element(l, [i___1]) <= 2147483647) .
+H13:   j < i .
+H14:   max_ext(a, i, j) >= - 2147483648 .
+H15:   max_ext(a, i, j) <= 2147483647 .
+H16:   j >= l__index__subtype__1__first .
+H17:   i >= a__index__subtype__1__first .
+H18:   i <= a__index__subtype__1__last .
+H19:   j >= a__index__subtype__1__first .
+H20:   j <= a__index__subtype__1__last .
+H21:   element(a, [i]) < element(a, [j]) or element(l, [j]) <= max_ext(a, i, j) 
+          .
+H22:   j <= 2147483646 .
+H23:   integer__size >= 0 .
+H24:   a__index__subtype__1__first <= a__index__subtype__1__last .
+H25:   l__index__subtype__1__first <= l__index__subtype__1__last .
+H26:   a__index__subtype__1__first >= - 2147483648 .
+H27:   a__index__subtype__1__last >= - 2147483648 .
+H28:   l__index__subtype__1__first >= - 2147483648 .
+H29:   l__index__subtype__1__last >= - 2147483648 .
+H30:   a__index__subtype__1__last <= 2147483647 .
+H31:   a__index__subtype__1__first <= 2147483647 .
+H32:   l__index__subtype__1__last <= 2147483647 .
+H33:   l__index__subtype__1__first <= 2147483647 .
+       ->
+C1:    max_ext(a, i, j) = max_ext(a, i, j + 1) .
+
+
+For path(s) from assertion of line 26 to run-time check associated with 
+          statement of line 36:
+
+procedure_liseq_length_15.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 26 to run-time check associated with 
+          statement of line 38:
+
+procedure_liseq_length_16.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 26 to run-time check associated with 
+          statement of line 40:
+
+procedure_liseq_length_17.
+*** true .          /* all conclusions proved */
+
+
+procedure_liseq_length_18.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 26 to run-time check associated with 
+          statement of line 42:
+
+procedure_liseq_length_19.
+H1:    for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = 
+          liseq_ends_at(a, i2_)) .
+H2:    element(l, [pmax]) = liseq_prfx(a, i) .
+H3:    i <= l__index__subtype__1__last .
+H4:    0 <= pmax .
+H5:    pmax < i .
+H6:    a__index__subtype__1__first = 0 .
+H7:    l__index__subtype__1__first = 0 .
+H8:    a__index__subtype__1__last = l__index__subtype__1__last .
+H9:    a__index__subtype__1__last < 2147483647 .
+H10:   for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 
+          <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) 
+          and element(a, [i___1]) <= 2147483647) .
+H11:   for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 
+          <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) 
+          and element(l, [i___1]) <= 2147483647) .
+H12:   i <= 2147483647 .
+H13:   max_ext(a, i, i) >= - 2147483648 .
+H14:   max_ext(a, i, i) <= 2147483647 .
+H15:   integer__size >= 0 .
+H16:   a__index__subtype__1__first <= a__index__subtype__1__last .
+H17:   l__index__subtype__1__first <= l__index__subtype__1__last .
+H18:   a__index__subtype__1__first >= - 2147483648 .
+H19:   a__index__subtype__1__last >= - 2147483648 .
+H20:   l__index__subtype__1__first >= - 2147483648 .
+H21:   l__index__subtype__1__last >= - 2147483648 .
+H22:   a__index__subtype__1__last <= 2147483647 .
+H23:   a__index__subtype__1__first <= 2147483647 .
+H24:   l__index__subtype__1__last <= 2147483647 .
+H25:   l__index__subtype__1__first <= 2147483647 .
+       ->
+C1:    max_ext(a, i, i) <= 2147483646 .
+
+
+For path(s) from assertion of line 26 to run-time check associated with 
+          statement of line 43:
+
+procedure_liseq_length_20.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 26 to run-time check associated with 
+          statement of line 44:
+
+procedure_liseq_length_21.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 26 to run-time check associated with 
+          statement of line 45:
+
+procedure_liseq_length_22.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 15 to run-time check associated with 
+          statement of line 48:
+
+procedure_liseq_length_23.
+H1:    for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = 
+          liseq_ends_at(a, i2_)) .
+H2:    element(l, [pmax]) = liseq_prfx(a, i) .
+H3:    0 <= pmax .
+H4:    pmax < i .
+H5:    a__index__subtype__1__first = 0 .
+H6:    l__index__subtype__1__first = 0 .
+H7:    a__index__subtype__1__last = l__index__subtype__1__last .
+H8:    a__index__subtype__1__last < 2147483647 .
+H9:    for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 
+          <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) 
+          and element(a, [i___1]) <= 2147483647) .
+H10:   for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 
+          <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) 
+          and element(l, [i___1]) <= 2147483647) .
+H11:   i <= l__index__subtype__1__last .
+H12:   pmax <= 2147483647 .
+H13:   pmax >= a__index__subtype__1__first .
+H14:   i <= a__index__subtype__1__last .
+H15:   element(a, [pmax]) <= element(a, [i]) .
+H16:   element(l, [pmax]) >= - 2147483648 .
+H17:   element(l, [pmax]) <= 2147483647 .
+H18:   integer__size >= 0 .
+H19:   a__index__subtype__1__first <= a__index__subtype__1__last .
+H20:   l__index__subtype__1__first <= l__index__subtype__1__last .
+H21:   a__index__subtype__1__first >= - 2147483648 .
+H22:   a__index__subtype__1__last >= - 2147483648 .
+H23:   l__index__subtype__1__first >= - 2147483648 .
+H24:   l__index__subtype__1__last >= - 2147483648 .
+H25:   a__index__subtype__1__last <= 2147483647 .
+H26:   a__index__subtype__1__first <= 2147483647 .
+H27:   l__index__subtype__1__last <= 2147483647 .
+H28:   l__index__subtype__1__first <= 2147483647 .
+       ->
+C1:    element(l, [pmax]) <= 2147483646 .
+
+
+For path(s) from assertion of line 15 to run-time check associated with 
+          statement of line 49:
+
+procedure_liseq_length_24.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 15 to run-time check associated with 
+          statement of line 50:
+
+procedure_liseq_length_25.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 15 to run-time check associated with 
+          statement of line 52:
+
+procedure_liseq_length_26.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 26 to run-time check associated with 
+          statement of line 52:
+
+procedure_liseq_length_27.
+*** true .          /* all conclusions proved */
+
+
+procedure_liseq_length_28.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 15 to finish:
+
+procedure_liseq_length_29.
+H1:    for_all(i2_ : integer, 0 <= i2_ and i2_ <= i - 1 -> element(l, [i2_]) = 
+          liseq_ends_at(a, i2_)) .
+H2:    element(l, [pmax]) = liseq_prfx(a, i) .
+H3:    i <= l__index__subtype__1__last + 1 .
+H4:    0 <= pmax .
+H5:    pmax < i .
+H6:    a__index__subtype__1__first = 0 .
+H7:    l__index__subtype__1__first = 0 .
+H8:    a__index__subtype__1__last = l__index__subtype__1__last .
+H9:    a__index__subtype__1__last < 2147483647 .
+H10:   for_all(i___1 : integer, a__index__subtype__1__first <= i___1 and i___1 
+          <= a__index__subtype__1__last -> - 2147483648 <= element(a, [i___1]) 
+          and element(a, [i___1]) <= 2147483647) .
+H11:   for_all(i___1 : integer, l__index__subtype__1__first <= i___1 and i___1 
+          <= l__index__subtype__1__last -> - 2147483648 <= element(l, [i___1]) 
+          and element(l, [i___1]) <= 2147483647) .
+H12:   i <= 2147483647 .
+H13:   l__index__subtype__1__last < i .
+H14:   integer__size >= 0 .
+H15:   a__index__subtype__1__first <= a__index__subtype__1__last .
+H16:   l__index__subtype__1__first <= l__index__subtype__1__last .
+H17:   a__index__subtype__1__first >= - 2147483648 .
+H18:   a__index__subtype__1__last >= - 2147483648 .
+H19:   l__index__subtype__1__first >= - 2147483648 .
+H20:   l__index__subtype__1__last >= - 2147483648 .
+H21:   a__index__subtype__1__last <= 2147483647 .
+H22:   a__index__subtype__1__first <= 2147483647 .
+H23:   l__index__subtype__1__last <= 2147483647 .
+H24:   l__index__subtype__1__first <= 2147483647 .
+       ->
+C1:    element(l, [pmax]) = liseq_prfx(a, a__index__subtype__1__last + 1) .
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/README	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,3 @@
+The copyright notice contained in the *.siv, *.fdl, and *.rls files in
+the example subdirectories refers to the tools that have been used to
+generate the files, but not to the files themselves.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/F.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,66 @@
+(*  Title:      HOL/SPARK/Examples/RIPEMD-160/F.thy
+    Author:     Fabian Immler, TU Muenchen
+
+Verification of the RIPEMD-160 hash function
+*)
+
+theory F
+imports RMD_Specification
+begin
+
+spark_open "rmd/f.siv"
+
+spark_vc function_f_2
+  using assms by simp_all
+
+spark_vc function_f_3
+  using assms by simp_all
+
+spark_vc function_f_4
+  using assms by simp_all
+
+spark_vc function_f_5
+  using assms by simp_all
+
+spark_vc function_f_6
+proof -
+  from H8 have "nat j <= 15" by simp
+  with assms show ?thesis
+    by (simp add: f_def bwsimps int_word_uint int_mod_eq')
+qed
+
+spark_vc function_f_7
+proof -
+  from H7 have "16 <= nat j" by simp
+  moreover from H8 have "nat j <= 31" by simp
+  ultimately show ?thesis using assms
+    by (simp add: f_def bwsimps int_word_uint int_mod_eq')
+qed
+
+spark_vc function_f_8
+proof -
+  from H7 have "32 <= nat j" by simp
+  moreover from H8 have "nat j <= 47" by simp
+  ultimately show ?thesis using assms
+    by (simp add: f_def bwsimps int_word_uint int_mod_eq')
+qed
+
+spark_vc function_f_9
+proof -
+  from H7 have "48 <= nat j" by simp
+  moreover from H8 have   "nat j <= 63" by simp
+  ultimately show ?thesis using assms
+    by (simp add: f_def bwsimps int_word_uint int_mod_eq')
+qed
+
+spark_vc function_f_10
+proof -
+  from H2 have "nat j <= 79" by simp
+  moreover from H12 have "64 <= nat j" by simp
+  ultimately show ?thesis using assms
+    by (simp add: f_def bwsimps int_word_uint int_mod_eq')
+qed
+
+spark_end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/Hash.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,100 @@
+(*  Title:      HOL/SPARK/Examples/RIPEMD-160/Hash.thy
+    Author:     Fabian Immler, TU Muenchen
+
+Verification of the RIPEMD-160 hash function
+*)
+
+theory Hash
+imports RMD_Specification
+begin
+
+spark_open "rmd/hash.siv"
+
+abbreviation from_chain :: "chain \<Rightarrow> RMD.chain" where
+  "from_chain c \<equiv> (
+    word_of_int (h0 c),
+    word_of_int (h1 c),
+    word_of_int (h2 c),
+    word_of_int (h3 c),
+    word_of_int (h4 c))"
+
+abbreviation to_chain :: "RMD.chain \<Rightarrow> chain" where
+  "to_chain c \<equiv>
+    (let (h0, h1, h2, h3, h4) = c in
+      (|h0 = uint h0,
+        h1 = uint h1,
+        h2 = uint h2,
+        h3 = uint h3,
+        h4 = uint h4|))"
+
+abbreviation round' :: "chain \<Rightarrow> block \<Rightarrow> chain" where
+  "round' c b == to_chain (round (\<lambda>n. word_of_int (b (int n))) (from_chain c))"
+
+abbreviation rounds' :: "chain \<Rightarrow> int \<Rightarrow> message \<Rightarrow> chain" where
+  "rounds' h i X ==
+     to_chain (rounds
+      (\<lambda>n. \<lambda>m. word_of_int (X (int n) (int m)))
+      (from_chain h)
+      (nat i))"
+
+abbreviation rmd_hash :: "message \<Rightarrow> int \<Rightarrow> chain" where
+  "rmd_hash X i == to_chain (rmd
+    (\<lambda>n. \<lambda>m. word_of_int (X (int n) (int m)))
+    (nat i))"
+
+spark_proof_functions
+  round_spec = round'
+  rounds = rounds'
+  rmd_hash = rmd_hash
+
+spark_vc function_hash_12
+  using H1 H6
+  by (simp add:
+    rounds_def rmd_body_def round_def
+    h_0_def h0_0_def h1_0_def h2_0_def h3_0_def h4_0_def)
+
+
+lemma rounds_step:
+  assumes "0 <= i"
+  shows "rounds X b (Suc i) = round (X i) (rounds X b i)"
+  by (simp add: rounds_def rmd_body_def)
+
+lemma from_to_id: "from_chain (to_chain C) = C"
+proof (cases C)
+  fix a b c d e f::word32
+  assume "C = (a, b, c, d, e)"
+  thus ?thesis by (cases a) simp
+qed
+
+lemma steps_to_steps':
+  "round X (foldl a b c) = round X (from_chain (to_chain (foldl a b c)))"
+  unfolding from_to_id ..
+
+lemma rounds'_step:
+  assumes "0 <= i"
+  shows "rounds' c (i + 1) x = round' (rounds' c i x) (x i)"
+proof -
+  have makesuc: "nat (i + 1) = Suc (nat i)" using assms by simp
+  show ?thesis using assms
+    by (simp add: makesuc rounds_def rmd_body_def steps_to_steps')
+qed
+
+
+spark_vc function_hash_13
+proof -
+  have loop_suc: "loop__1__i + 2 = (loop__1__i + 1) + 1" by simp
+  have "0 <= loop__1__i + 1" using `0 <= loop__1__i` by simp
+  show ?thesis
+    unfolding loop_suc
+    unfolding rounds'_step[OF `0 <= loop__1__i + 1`]
+    unfolding H1[symmetric]
+    unfolding H18 ..
+qed
+
+
+spark_vc function_hash_17
+  unfolding rmd_def H1 rounds_def ..
+
+spark_end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/K_L.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,46 @@
+(*  Title:      HOL/SPARK/Examples/RIPEMD-160/K_L.thy
+    Author:     Fabian Immler, TU Muenchen
+
+Verification of the RIPEMD-160 hash function
+*)
+
+theory K_L
+imports RMD_Specification
+begin
+
+spark_open "rmd/k_l.siv"
+
+spark_vc function_k_l_6
+  using assms by (simp add: K_def)
+
+spark_vc function_k_l_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_l_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_l_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_l_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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/K_R.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,46 @@
+(*  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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/RMD.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,180 @@
+(*  Title:      HOL/SPARK/Examples/RIPEMD-160/RMD.thy
+    Author:     Fabian Immler, TU Muenchen
+
+Verification of the RIPEMD-160 hash function
+*)
+
+theory RMD
+imports Word
+begin
+
+
+(* all operations are defined on 32-bit words *)
+
+types word32 = "32 word"
+types byte = "8 word"
+types perm = "nat \<Rightarrow> nat"
+types chain = "word32 * word32 * word32 * word32 * word32"
+types block = "nat \<Rightarrow> word32"
+types message = "nat \<Rightarrow> block"
+
+(* nonlinear functions at bit level *)
+
+definition f::"[nat, word32, word32, word32] => word32"
+where
+"f j x y z =
+      (if ( 0 <= j & j <= 15) then x XOR y XOR z
+  else if (16 <= j & j <= 31) then (x AND y) OR (NOT x AND z)
+  else if (32 <= j & j <= 47) then (x OR NOT y) XOR z
+  else if (48 <= j & j <= 63) then (x AND z) OR (y AND NOT z)
+  else if (64 <= j & j <= 79) then x XOR (y OR NOT z)
+  else 0)"
+
+
+(* added constants (hexadecimal) *)
+
+definition K::"nat => word32"
+where
+"K j =
+      (if ( 0 <= j & j <= 15) then 0x00000000
+  else if (16 <= j & j <= 31) then 0x5A827999
+  else if (32 <= j & j <= 47) then 0x6ED9EBA1
+  else if (48 <= j & j <= 63) then 0x8F1BBCDC
+  else if (64 <= j & j <= 79) then 0xA953FD4E
+  else 0)"
+
+definition K'::"nat => word32"
+where
+"K' j =
+      (if ( 0 <= j & j <= 15) then 0x50A28BE6
+  else if (16 <= j & j <= 31) then 0x5C4DD124
+  else if (32 <= j & j <= 47) then 0x6D703EF3
+  else if (48 <= j & j <= 63) then 0x7A6D76E9
+  else if (64 <= j & j <= 79) then 0x00000000
+  else 0)"
+
+
+(* selection of message word *)
+
+definition r_list :: "nat list"
+  where "r_list = [
+  0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15,
+  7, 4, 13, 1, 10, 6, 15, 3, 12, 0, 9, 5, 2, 14, 11, 8,
+  3, 10, 14, 4, 9, 15, 8, 1, 2, 7, 0, 6, 13, 11, 5, 12,
+  1, 9, 11, 10, 0, 8, 12, 4, 13, 3, 7, 15, 14, 5, 6, 2,
+  4, 0, 5, 9, 7, 12, 2, 10, 14, 1, 3, 8, 11, 6, 15, 13
+  ]"
+
+definition r'_list :: "nat list"
+  where "r'_list = [
+  5, 14, 7, 0, 9, 2, 11, 4, 13, 6, 15, 8, 1, 10, 3, 12,
+  6, 11, 3, 7, 0, 13, 5, 10, 14, 15, 8, 12, 4, 9, 1, 2,
+  15, 5, 1, 3, 7, 14, 6, 9, 11, 8, 12, 2, 10, 0, 4, 13,
+  8, 6, 4, 1, 3, 11, 15, 0, 5, 12, 2, 13, 9, 7, 10, 14,
+  12, 15, 10, 4, 1, 5, 8, 7, 6, 2, 13, 14, 0, 3, 9, 11
+  ]"
+
+definition r :: perm
+  where "r j = r_list ! j"
+
+definition r' :: perm
+  where "r' j = r'_list ! j"
+
+
+(* amount for rotate left (rol) *)
+
+definition s_list :: "nat list"
+  where "s_list = [
+  11, 14, 15, 12, 5, 8, 7, 9, 11, 13, 14, 15, 6, 7, 9, 8,
+  7, 6, 8, 13, 11, 9, 7, 15, 7, 12, 15, 9, 11, 7, 13, 12,
+  11, 13, 6, 7, 14, 9, 13, 15, 14, 8, 13, 6, 5, 12, 7, 5,
+  11, 12, 14, 15, 14, 15, 9, 8, 9, 14, 5, 6, 8, 6, 5, 12,
+  9, 15, 5, 11, 6, 8, 13, 12, 5, 12, 13, 14, 11, 8, 5, 6
+  ]"
+
+definition s'_list :: "nat list"
+  where "s'_list = [
+  8, 9, 9, 11, 13, 15, 15, 5, 7, 7, 8, 11, 14, 14, 12, 6,
+  9, 13, 15, 7, 12, 8, 9, 11, 7, 7, 12, 7, 6, 15, 13, 11,
+  9, 7, 15, 11, 8, 6, 6, 14, 12, 13, 5, 14, 13, 13, 7, 5,
+  15, 5, 8, 11, 14, 14, 6, 14, 6, 9, 12, 9, 12, 5, 15, 8,
+  8, 5, 12, 9, 12, 5, 14, 6, 8, 13, 6, 5, 15, 13, 11, 11
+  ]"
+
+definition s :: perm
+  where "s j = s_list ! j"
+
+definition s' :: perm
+  where "s' j = s'_list ! j"
+
+
+(* Initial value (hexadecimal *)
+
+definition h0_0::word32 where "h0_0 = 0x67452301"
+definition h1_0::word32 where "h1_0 = 0xEFCDAB89"
+definition h2_0::word32 where "h2_0 = 0x98BADCFE"
+definition h3_0::word32 where "h3_0 = 0x10325476"
+definition h4_0::word32 where "h4_0 = 0xC3D2E1F0"
+definition h_0::chain where "h_0 = (h0_0, h1_0, h2_0, h3_0, h4_0)"
+
+
+definition step_l ::
+  "[ block,
+     chain,
+     nat
+  ] => chain"
+  where
+  "step_l X c j =
+    (let (A, B, C, D, E) = c in
+    ((* A *) E,
+     (* B *) word_rotl (s j) (A + f j B C D + X (r j) + K j) + E,
+     (* C *) B,
+     (* D *) word_rotl 10 C,
+     (* E *) D))"
+
+definition step_r ::
+  "[ block,
+     chain,
+     nat
+   ] \<Rightarrow> chain"
+where
+  "step_r X c' j =
+    (let (A', B', C', D', E') = c' in
+    ((* A' *) E',
+     (* B' *) word_rotl (s' j) (A' + f (79 - j) B' C' D' + X (r' j) + K' j) + E',
+     (* C' *) B',
+     (* D' *) word_rotl 10 C',
+     (* E' *) D'))"
+
+definition step_both ::
+  "[ block, chain * chain, nat ] \<Rightarrow> chain * chain"
+  where
+  "step_both X cc j = (case cc of (c, c') \<Rightarrow>
+  (step_l X c j, step_r X c' j))"
+
+definition steps::"[ block, chain * chain, nat] \<Rightarrow> chain * chain"
+  where "steps X cc i = foldl (step_both X) cc [0..<i]"
+
+definition round::"[ block, chain ] \<Rightarrow> chain"
+  where "round X h =
+    (let (h0, h1, h2, h3, h4) = h in
+     let ((A, B, C, D, E), (A', B', C', D', E')) = steps X (h, h) 80 in
+      ((* h0 *) h1 + C + D',
+       (* h1 *) h2 + D + E',
+       (* h2 *) h3 + E + A',
+       (* h3 *) h4 + A + B',
+       (* h4 *) h0 + B + C'))"
+
+definition rmd_body::"[ message, chain, nat ] => chain"
+where
+  "rmd_body X h i = round (X i) h"
+
+definition rounds::"message \<Rightarrow> chain \<Rightarrow> nat \<Rightarrow> chain"
+where
+  "rounds X h i = foldl (rmd_body X) h_0 [0..<i]"
+
+definition rmd :: "message \<Rightarrow> nat \<Rightarrow> chain"
+where
+  "rmd X len = rounds X h_0 len"
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/RMD_Lemmas.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,27 @@
+(*  Title:      HOL/SPARK/Examples/RIPEMD-160/RMD_Lemmas.thy
+    Author:     Fabian Immler, TU Muenchen
+
+Verification of the RIPEMD-160 hash function
+*)
+
+theory RMD_Lemmas
+imports Main
+begin
+
+definition "fun_of_list i xs g j =
+  (if j < i \<or> i + int (length xs) \<le> j then g j else xs ! nat (j - i))"
+
+lemma fun_of_list_Nil [simp]: "fun_of_list i [] g = g"
+  by (auto simp add: fun_eq_iff fun_of_list_def)
+
+lemma fun_of_list_Cons [simp]:
+  "fun_of_list i (x # xs) g = fun_of_list (i + 1) xs (g(i:=x))"
+  by (auto simp add: fun_eq_iff fun_of_list_def nth_Cons'
+    nat_diff_distrib [of "int (Suc 0)", simplified, symmetric]
+    diff_diff_eq)
+
+lemma nth_fun_of_list_eq:
+  "0 \<le> i \<Longrightarrow> i < int (length xs) \<Longrightarrow> xs ! nat i = fun_of_list 0 xs g i"
+  by (simp add: fun_of_list_def)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/RMD_Specification.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,47 @@
+(*  Title:      HOL/SPARK/Examples/RIPEMD-160/RMD_Specification.thy
+    Author:     Fabian Immler, TU Muenchen
+
+Verification of the RIPEMD-160 hash function
+*)
+
+theory RMD_Specification
+imports RMD SPARK
+
+begin
+
+(* bit operations *)
+
+abbreviation rotate_left :: "int \<Rightarrow> int \<Rightarrow> int" where
+  "rotate_left i w == uint (word_rotl (nat i) (word_of_int w::word32))"
+
+spark_proof_functions
+  wordops__rotate_left = rotate_left
+
+
+(* Conversions for proof functions *)
+abbreviation k_l_spec :: " int => int " where
+  "k_l_spec j == uint (K (nat j))"
+abbreviation k_r_spec :: " int => int " where
+  "k_r_spec j == uint (K' (nat j))"
+abbreviation r_l_spec :: " int => int " where
+  "r_l_spec j == int (r (nat j))"
+abbreviation r_r_spec :: " int => int " where
+  "r_r_spec j == int (r' (nat j))"
+abbreviation s_l_spec :: " int => int " where
+  "s_l_spec j == int (s (nat j))"
+abbreviation s_r_spec :: " int => int " where
+  "s_r_spec j == int (s' (nat j))"
+abbreviation f_spec :: "int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int \<Rightarrow> int" where
+  "f_spec j x y z ==
+    uint (f (nat j) (word_of_int x::word32) (word_of_int y) (word_of_int z))"
+
+spark_proof_functions
+  k_l_spec = k_l_spec
+  k_r_spec = k_r_spec
+  r_l_spec = r_l_spec
+  r_r_spec = r_r_spec
+  s_l_spec = s_l_spec
+  s_r_spec = s_r_spec
+  f_spec = f_spec
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/R_L.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,30 @@
+(*  Title:      HOL/SPARK/Examples/RIPEMD-160/R_L.thy
+    Author:     Fabian Immler, TU Muenchen
+
+Verification of the RIPEMD-160 hash function
+*)
+
+theory R_L
+imports RMD_Specification RMD_Lemmas
+begin
+
+spark_open "rmd/r_l.siv"
+
+spark_vc function_r_l_2
+proof -
+  from `0 \<le> j` `j \<le> 79`
+  show C: ?C1
+    by (simp add: r_def r_list_def nth_map [symmetric, of _ _ int] del: fun_upd_apply)
+      (simp add: nth_fun_of_list_eq [of _ _ undefined] del: fun_upd_apply)
+  from C show ?C2 by simp
+  have "list_all (\<lambda>n. int n \<le> 15) r_list"
+    by (simp add: r_list_def)
+  moreover have "length r_list = 80"
+    by (simp add: r_list_def)
+  ultimately show ?C3 unfolding C using `j \<le> 79`
+    by (simp add: r_def list_all_length)
+qed
+
+spark_end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/R_R.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,30 @@
+(*  Title:      HOL/SPARK/Examples/RIPEMD-160/R_R.thy
+    Author:     Fabian Immler, TU Muenchen
+
+Verification of the RIPEMD-160 hash function
+*)
+
+theory R_R
+imports RMD_Specification RMD_Lemmas
+begin
+
+spark_open "rmd/r_r.siv"
+
+spark_vc function_r_r_2
+proof -
+  from `0 \<le> j` `j \<le> 79`
+  show C: ?C1
+    by (simp add: r'_def r'_list_def nth_map [symmetric, of _ _ int] del: fun_upd_apply)
+      (simp add: nth_fun_of_list_eq [of _ _ undefined] del: fun_upd_apply)
+  from C show ?C2 by simp
+  have "list_all (\<lambda>n. int n \<le> 15) r'_list"
+    by (simp add: r'_list_def)
+  moreover have "length r'_list = 80"
+    by (simp add: r'_list_def)
+  ultimately show ?C3 unfolding C using `j \<le> 79`
+    by (simp add: r'_def list_all_length)
+qed
+
+spark_end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/Round.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,465 @@
+(*  Title:      HOL/SPARK/Examples/RIPEMD-160/Round.thy
+    Author:     Fabian Immler, TU Muenchen
+
+Verification of the RIPEMD-160 hash function
+*)
+
+theory Round
+imports RMD_Specification
+begin
+
+spark_open "rmd/round.siv"
+
+abbreviation from_chain :: "chain \<Rightarrow> RMD.chain" where
+  "from_chain c \<equiv> (
+    word_of_int (h0 c),
+    word_of_int (h1 c),
+    word_of_int (h2 c),
+    word_of_int (h3 c),
+    word_of_int (h4 c))"
+
+abbreviation from_chain_pair :: "chain_pair \<Rightarrow> RMD.chain \<times> RMD.chain" where
+  "from_chain_pair cc \<equiv> (
+    from_chain (left cc),
+    from_chain (right cc))"
+
+abbreviation to_chain :: "RMD.chain \<Rightarrow> chain" where
+  "to_chain c \<equiv>
+    (let (h0, h1, h2, h3, h4) = c in
+      (|h0 = uint h0,
+        h1 = uint h1,
+        h2 = uint h2,
+        h3 = uint h3,
+        h4 = uint h4|))"
+
+abbreviation to_chain_pair :: "RMD.chain \<times> RMD.chain \<Rightarrow> chain_pair" where
+  "to_chain_pair c == (let (c1, c2) = c in
+    (| left = to_chain c1,
+      right = to_chain c2 |))"
+
+abbreviation steps' :: "chain_pair \<Rightarrow> int \<Rightarrow> block \<Rightarrow> chain_pair" where
+  "steps' cc i b == to_chain_pair (steps
+    (\<lambda>n. word_of_int (b (int n)))
+    (from_chain_pair cc)
+    (nat i))"
+
+abbreviation round_spec :: "chain \<Rightarrow> block \<Rightarrow> chain" where
+  "round_spec c b == to_chain (round (\<lambda>n. word_of_int (b (int n))) (from_chain c))"
+
+spark_proof_functions
+  steps = steps'
+  round_spec = round_spec
+
+lemma uint_word_of_int_id:
+  assumes "0 <= (x::int)"
+  assumes "x <= 4294967295"
+  shows"uint(word_of_int x::word32) = x"
+  unfolding int_word_uint
+  using assms
+  by (simp add:int_mod_eq')
+
+lemma steps_step: "steps X cc (Suc i) = step_both X (steps X cc i) i"
+  unfolding steps_def
+  by (induct i) simp_all
+
+lemma from_to_id: "from_chain_pair (to_chain_pair CC) = CC"
+proof (cases CC)
+  fix a::RMD.chain
+  fix b c d e f::word32
+  assume "CC = (a, b, c, d, e, f)"
+  thus ?thesis by (cases a) simp
+qed
+
+lemma steps_to_steps':
+  "F A (steps X cc i) B =
+   F A (from_chain_pair (to_chain_pair (steps X cc i))) B"
+  unfolding from_to_id ..
+
+lemma steps'_step:
+  assumes "0 <= i"
+  shows
+  "steps' cc (i + 1) X = to_chain_pair (
+     step_both
+       (\<lambda>n. word_of_int (X (int n)))
+       (from_chain_pair (steps' cc i X))
+       (nat i))"
+proof -
+  have "nat (i + 1) = Suc (nat i)" using assms by simp
+  show ?thesis
+    unfolding `nat (i + 1) = Suc (nat i)` steps_step steps_to_steps'
+    ..
+qed 
+
+lemma step_from_hyp:
+  assumes
+  step_hyp:
+  "\<lparr>left =
+      \<lparr>h0 = a, h1 = b, h2 = c, h3 = d, h4 = e\<rparr>,
+    right =
+      \<lparr>h0 = a', h1 = b', h2 = c', h3 = d', h4 = e'\<rparr>\<rparr> =
+   steps'
+     (\<lparr>left =
+         \<lparr>h0 = a_0, h1 = b_0, h2 = c_0,
+          h3 = d_0, h4 = e_0\<rparr>,
+       right =
+         \<lparr>h0 = a_0, h1 = b_0, h2 = c_0,
+          h3 = d_0, h4 = e_0\<rparr>\<rparr>)
+     j x"
+  assumes "a <= 4294967295" (is "_ <= ?M")
+  assumes                "b  <= ?M" and "c  <= ?M" and "d  <= ?M" and "e  <= ?M"
+  assumes "a' <= ?M" and "b' <= ?M" and "c' <= ?M" and "d' <= ?M" and "e' <= ?M"
+  assumes "0 <= a " and "0 <= b " and "0 <= c " and "0 <= d " and "0 <= e "
+  assumes "0 <= a'" and "0 <= b'" and "0 <= c'" and "0 <= d'" and "0 <= e'"
+  assumes "0 <= x (r_l_spec j)" and "x (r_l_spec j) <= ?M"
+  assumes "0 <= x (r_r_spec j)" and "x (r_r_spec j) <= ?M"
+  assumes "0 <= j" and "j <= 79"
+  shows
+  "\<lparr>left =
+      \<lparr>h0 = e,
+         h1 =
+           (rotate_left (s_l_spec j)
+             ((((a + f_spec j b c d) mod 4294967296 +
+                x (r_l_spec j)) mod
+               4294967296 +
+               k_l_spec j) mod
+              4294967296) +
+            e) mod
+           4294967296,
+         h2 = b, h3 = rotate_left 10 c,
+         h4 = d\<rparr>,
+      right =
+        \<lparr>h0 = e',
+           h1 =
+             (rotate_left (s_r_spec j)
+               ((((a' + f_spec (79 - j) b' c' d') mod
+                  4294967296 +
+                  x (r_r_spec j)) mod
+                 4294967296 +
+                 k_r_spec j) mod
+                4294967296) +
+              e') mod
+             4294967296,
+           h2 = b', h3 = rotate_left 10 c',
+           h4 = d'\<rparr>\<rparr> =
+   steps'
+    (\<lparr>left =
+        \<lparr>h0 = a_0, h1 = b_0, h2 = c_0,
+           h3 = d_0, h4 = e_0\<rparr>,
+        right =
+          \<lparr>h0 = a_0, h1 = b_0, h2 = c_0,
+             h3 = d_0, h4 = e_0\<rparr>\<rparr>)
+    (j + 1) x"
+  using step_hyp
+proof -
+  let ?MM = 4294967296
+  have AL: "uint(word_of_int e::word32) = e"
+    by (rule uint_word_of_int_id[OF `0 <= e` `e <= ?M`])
+  have CL: "uint(word_of_int b::word32) = b"
+    by (rule uint_word_of_int_id[OF `0 <= b` `b <= ?M`])
+  have DL: "True" ..
+  have EL: "uint(word_of_int d::word32) = d"
+    by (rule uint_word_of_int_id[OF `0 <= d` `d <= ?M`])
+  have AR: "uint(word_of_int e'::word32) = e'"
+    by (rule uint_word_of_int_id[OF `0 <= e'` `e' <= ?M`])
+  have CR: "uint(word_of_int b'::word32) = b'"
+    by (rule uint_word_of_int_id[OF `0 <= b'` `b' <= ?M`])
+  have DR: "True" ..
+  have ER: "uint(word_of_int d'::word32) = d'"
+    by (rule uint_word_of_int_id[OF `0 <= d'` `d' <= ?M`])
+  have BL:
+    "(uint
+      (word_rotl (s (nat j))
+        ((word_of_int::int\<Rightarrow>word32)
+          ((((a + f_spec j b c d) mod ?MM +
+             x (r_l_spec j)) mod ?MM +
+            k_l_spec j) mod ?MM))) +
+        e) mod ?MM
+    =
+    uint
+              (word_rotl (s (nat j))
+                (word_of_int a +
+                 f (nat j) (word_of_int b)
+                  (word_of_int c) (word_of_int d) +
+                 word_of_int (x (r_l_spec j)) +
+                 K (nat j)) +
+               word_of_int e)"
+    (is "(uint (word_rotl _ (_ ((((_ + ?F) mod _ + ?X) mod _ + _) mod _))) + _) mod _ = _")
+  proof -
+    have "a mod ?MM = a" using `0 <= a` `a <= ?M`
+      by (simp add: int_mod_eq')
+    have "?X mod ?MM = ?X" using `0 <= ?X` `?X <= ?M`
+      by (simp add: int_mod_eq')
+    have "e mod ?MM = e" using `0 <= e` `e <= ?M`
+      by (simp add: int_mod_eq')
+    have "(?MM::int) = 2 ^ len_of TYPE(32)" by simp
+    show ?thesis
+      unfolding
+        word_add_alt
+        uint_word_of_int_id[OF `0 <= a` `a <= ?M`]
+        uint_word_of_int_id[OF `0 <= ?X` `?X <= ?M`]
+        int_word_uint
+      unfolding `?MM = 2 ^ len_of TYPE(32)`
+      unfolding word_uint.Abs_norm
+      by (simp add:
+        `a mod ?MM = a`
+        `e mod ?MM = e`
+        `?X mod ?MM = ?X`)
+  qed
+
+  have BR:
+    "(uint
+      (word_rotl (s' (nat j))
+        ((word_of_int::int\<Rightarrow>word32)
+          ((((a' + f_spec (79 - j) b' c' d') mod ?MM +
+             x (r_r_spec j)) mod ?MM +
+            k_r_spec j) mod ?MM))) +
+        e') mod ?MM
+    =
+    uint
+              (word_rotl (s' (nat j))
+                (word_of_int a' +
+                 f (79 - nat j) (word_of_int b')
+                  (word_of_int c') (word_of_int d') +
+                 word_of_int (x (r_r_spec j)) +
+                 K' (nat j)) +
+               word_of_int e')"
+    (is "(uint (word_rotl _ (_ ((((_ + ?F) mod _ + ?X) mod _ + _) mod _))) + _) mod _ = _")
+  proof -
+    have "a' mod ?MM = a'" using `0 <= a'` `a' <= ?M`
+      by (simp add: int_mod_eq')
+    have "?X mod ?MM = ?X" using `0 <= ?X` `?X <= ?M`
+      by (simp add: int_mod_eq')
+    have "e' mod ?MM = e'" using `0 <= e'` `e' <= ?M`
+      by (simp add: int_mod_eq')
+    have "(?MM::int) = 2 ^ len_of TYPE(32)" by simp
+    have nat_transfer: "79 - nat j = nat (79 - j)"
+      using nat_diff_distrib `0 <= j`  `j <= 79`
+      by simp
+    show ?thesis
+      unfolding
+        word_add_alt
+        uint_word_of_int_id[OF `0 <= a'` `a' <= ?M`]
+        uint_word_of_int_id[OF `0 <= ?X` `?X <= ?M`]
+        int_word_uint
+        nat_transfer
+      unfolding `?MM = 2 ^ len_of TYPE(32)`
+      unfolding word_uint.Abs_norm
+      by (simp add: 
+        `a' mod ?MM = a'`
+        `e' mod ?MM = e'`
+        `?X mod ?MM = ?X`)
+  qed
+
+  show ?thesis
+    unfolding steps'_step[OF `0 <= j`] step_hyp[symmetric]
+      step_both_def step_r_def step_l_def
+    by (simp add: AL BL CL DL EL AR BR CR DR ER)
+qed
+
+spark_vc procedure_round_61
+proof -
+  let ?M = "4294967295::int"
+  have step_hyp:
+  "\<lparr>left =
+      \<lparr>h0 = ca, h1 = cb, h2 = cc,
+         h3 = cd, h4 = ce\<rparr>,
+      right =
+        \<lparr>h0 = ca, h1 = cb, h2 = cc,
+           h3 = cd, h4 = ce\<rparr>\<rparr> =
+   steps'
+    (\<lparr>left =
+        \<lparr>h0 = ca, h1 = cb, h2 = cc,
+           h3 = cd, h4 = ce\<rparr>,
+        right =
+          \<lparr>h0 = ca, h1 = cb, h2 = cc,
+             h3 = cd, h4 = ce\<rparr>\<rparr>)
+    0 x"
+    unfolding steps_def 
+    by (simp add:
+      uint_word_of_int_id[OF `0 <= ca` `ca <= ?M`]
+      uint_word_of_int_id[OF `0 <= cb` `cb <= ?M`]
+      uint_word_of_int_id[OF `0 <= cc` `cc <= ?M`]
+      uint_word_of_int_id[OF `0 <= cd` `cd <= ?M`]
+      uint_word_of_int_id[OF `0 <= ce` `ce <= ?M`])
+  let ?rotate_arg_l =
+    "((((ca + f 0 cb cc cd) smod 4294967296 +
+        x (r_l 0)) smod 4294967296 + k_l 0) smod 4294967296)"
+  let ?rotate_arg_r =
+    "((((ca + f 79 cb cc cd) smod 4294967296 +
+        x (r_r 0)) smod 4294967296 + k_r 0) smod 4294967296)"
+  note returns =
+    `wordops__rotate (s_l 0) ?rotate_arg_l =
+     rotate_left (s_l 0) ?rotate_arg_l`
+    `wordops__rotate (s_r 0) ?rotate_arg_r =
+     rotate_left (s_r 0) ?rotate_arg_r`
+    `wordops__rotate 10 cc = rotate_left 10 cc`
+    `f 0 cb cc cd = f_spec 0 cb cc cd`
+    `f 79 cb cc cd = f_spec 79 cb cc cd`
+    `k_l 0 = k_l_spec 0`
+    `k_r 0 = k_r_spec 0`
+    `r_l 0 = r_l_spec 0`
+    `r_r 0 = r_r_spec 0`
+    `s_l 0 = s_l_spec 0`
+    `s_r 0 = s_r_spec 0`
+
+  note x_borders = `\<forall>i. 0 \<le> i \<and> i \<le> 15 \<longrightarrow> 0 \<le> x i \<and> x i \<le> ?M`
+
+  from `0 <= r_l 0` `r_l 0 <= 15` x_borders
+  have "0 \<le> x (r_l 0)" by blast
+  hence x_lower: "0 <= x (r_l_spec 0)" unfolding returns .
+
+  from `0 <= r_l 0` `r_l 0 <= 15` x_borders
+  have "x (r_l 0) <= ?M" by blast
+  hence x_upper: "x (r_l_spec 0) <= ?M" unfolding returns .
+
+  from `0 <= r_r 0` `r_r 0 <= 15` x_borders
+  have "0 \<le> x (r_r 0)" by blast
+  hence x_lower': "0 <= x (r_r_spec 0)" unfolding returns .
+
+  from `0 <= r_r 0` `r_r 0 <= 15` x_borders
+  have "x (r_r 0) <= ?M" by blast
+  hence x_upper': "x (r_r_spec 0) <= ?M" unfolding returns .
+
+  have "0 <= (0::int)" by simp
+  have "0 <= (79::int)" by simp
+  note step_from_hyp [OF
+    step_hyp
+    H2 H4 H6 H8 H10 H2 H4 H6 H8 H10 (* upper bounds *)
+    H1 H3 H5 H7 H9  H1 H3 H5 H7 H9  (* lower bounds *)
+  ]
+  from this[OF x_lower x_upper x_lower' x_upper' `0 <= 0` `0 <= 79`]
+    `0 \<le> ca` `0 \<le> ce` x_lower x_lower'
+  show ?thesis unfolding returns(1) returns(2) unfolding returns
+    by (simp add: smod_pos_pos)
+qed
+
+spark_vc procedure_round_62
+proof -
+  let ?M = "4294967295::int"
+  let ?rotate_arg_l =
+    "((((cla + f (loop__1__j + 1) clb clc cld) smod 4294967296 +
+         x (r_l (loop__1__j + 1))) smod 4294967296 +
+         k_l (loop__1__j + 1)) smod 4294967296)"
+  let ?rotate_arg_r =
+    "((((cra + f (79 - (loop__1__j + 1)) crb crc crd) smod
+         4294967296 + x (r_r (loop__1__j + 1))) smod 4294967296 +
+         k_r (loop__1__j + 1)) smod 4294967296)"
+
+  have s: "78 - loop__1__j = (79 - (loop__1__j + 1))" by simp
+  note returns =
+    `wordops__rotate (s_l (loop__1__j + 1)) ?rotate_arg_l =
+     rotate_left (s_l (loop__1__j + 1)) ?rotate_arg_l`
+    `wordops__rotate (s_r (loop__1__j + 1)) ?rotate_arg_r =
+     rotate_left (s_r (loop__1__j + 1)) ?rotate_arg_r`
+    `f (loop__1__j + 1) clb clc cld =
+     f_spec (loop__1__j + 1) clb clc cld`
+    `f (78 - loop__1__j) crb crc crd =
+     f_spec (78 - loop__1__j) crb crc crd`[simplified s]
+    `wordops__rotate 10 clc = rotate_left 10 clc`
+    `wordops__rotate 10 crc = rotate_left 10 crc`
+    `k_l (loop__1__j + 1) = k_l_spec (loop__1__j + 1)`
+    `k_r (loop__1__j + 1) = k_r_spec (loop__1__j + 1)`
+    `r_l (loop__1__j + 1) = r_l_spec (loop__1__j + 1)`
+    `r_r (loop__1__j + 1) = r_r_spec (loop__1__j + 1)`
+    `s_l (loop__1__j + 1) = s_l_spec (loop__1__j + 1)`
+    `s_r (loop__1__j + 1) = s_r_spec (loop__1__j + 1)`
+
+  note x_borders = `\<forall>i. 0 \<le> i \<and> i \<le> 15 \<longrightarrow> 0 \<le> x i \<and> x i \<le> ?M`
+
+  from `0 <= r_l (loop__1__j + 1)` `r_l (loop__1__j + 1) <= 15` x_borders
+  have "0 \<le> x (r_l (loop__1__j + 1))" by blast
+  hence x_lower: "0 <= x (r_l_spec (loop__1__j + 1))" unfolding returns .
+
+  from `0 <= r_l (loop__1__j + 1)` `r_l (loop__1__j + 1) <= 15` x_borders
+  have "x (r_l (loop__1__j + 1)) <= ?M" by blast
+  hence x_upper: "x (r_l_spec (loop__1__j + 1)) <= ?M" unfolding returns .
+
+  from `0 <= r_r (loop__1__j + 1)` `r_r (loop__1__j + 1) <= 15` x_borders
+  have "0 \<le> x (r_r (loop__1__j + 1))" by blast
+  hence x_lower': "0 <= x (r_r_spec (loop__1__j + 1))" unfolding returns .
+
+  from `0 <= r_r (loop__1__j + 1)` `r_r (loop__1__j + 1) <= 15` x_borders
+  have "x (r_r (loop__1__j + 1)) <= ?M" by blast
+  hence x_upper': "x (r_r_spec (loop__1__j + 1)) <= ?M" unfolding returns .
+
+  from `0 <= loop__1__j` have "0 <= loop__1__j + 1" by simp
+  from `loop__1__j <= 78` have "loop__1__j + 1 <= 79" by simp
+
+  have "loop__1__j + 1 + 1 = loop__1__j + 2" by simp
+
+  note step_from_hyp[OF H1
+    `cla <= ?M`
+    `clb <= ?M`
+    `clc <= ?M`
+    `cld <= ?M`
+    `cle <= ?M`
+    `cra <= ?M`
+    `crb <= ?M`
+    `crc <= ?M`
+    `crd <= ?M`
+    `cre <= ?M`
+    
+    `0 <= cla`
+    `0 <= clb`
+    `0 <= clc`
+    `0 <= cld`
+    `0 <= cle`
+    `0 <= cra`
+    `0 <= crb`
+    `0 <= crc`
+    `0 <= crd`
+    `0 <= cre`]
+  from this[OF
+    x_lower x_upper x_lower' x_upper'
+    `0 <= loop__1__j + 1` `loop__1__j + 1 <= 79`]
+    `0 \<le> cla` `0 \<le> cle` `0 \<le> cra` `0 \<le> cre` x_lower x_lower'
+  show ?thesis unfolding `loop__1__j + 1 + 1 = loop__1__j + 2`
+    unfolding returns(1) returns(2) unfolding returns
+    by (simp add: smod_pos_pos)
+qed
+
+spark_vc procedure_round_76
+proof -
+  let ?M = "4294967295 :: int"
+  let ?INIT_CHAIN =
+     "\<lparr>h0 = ca_init, h1 = cb_init,
+         h2 = cc_init, h3 = cd_init,
+         h4 = ce_init\<rparr>"
+  have steps_to_steps':
+    "steps
+       (\<lambda>n\<Colon>nat. word_of_int (x (int n)))
+       (from_chain ?INIT_CHAIN, from_chain ?INIT_CHAIN)
+       80 =
+    from_chain_pair (
+      steps'
+      (\<lparr>left = ?INIT_CHAIN, right = ?INIT_CHAIN\<rparr>)
+      80
+      x)"
+    unfolding from_to_id by simp
+  from
+    `0 \<le> ca_init` `ca_init \<le> ?M`
+    `0 \<le> cb_init` `cb_init \<le> ?M`
+    `0 \<le> cc_init` `cc_init \<le> ?M`
+    `0 \<le> cd_init` `cd_init \<le> ?M`
+    `0 \<le> ce_init` `ce_init \<le> ?M`
+    `0 \<le> cla` `cla \<le> ?M`
+    `0 \<le> clb` `clb \<le> ?M`
+    `0 \<le> clc` `clc \<le> ?M`
+    `0 \<le> cld` `cld \<le> ?M`
+    `0 \<le> cle` `cle \<le> ?M`
+    `0 \<le> cra` `cra \<le> ?M`
+    `0 \<le> crb` `crb \<le> ?M`
+    `0 \<le> crc` `crc \<le> ?M`
+    `0 \<le> crd` `crd \<le> ?M`
+    `0 \<le> cre` `cre \<le> ?M`
+  show ?thesis
+    unfolding round_def
+    unfolding steps_to_steps'
+    unfolding H1[symmetric]
+    by (simp add: uint_word_ariths(2) rdmods smod_pos_pos
+      uint_word_of_int_id)
+qed
+
+spark_end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/S_L.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,30 @@
+(*  Title:      HOL/SPARK/Examples/RIPEMD-160/S_L.thy
+    Author:     Fabian Immler, TU Muenchen
+
+Verification of the RIPEMD-160 hash function
+*)
+
+theory S_L
+imports RMD_Specification RMD_Lemmas
+begin
+
+spark_open "rmd/s_l.siv"
+
+spark_vc function_s_l_2
+proof -
+  from `0 \<le> j` `j \<le> 79`
+  show C: ?C1
+    by (simp add: s_def s_list_def nth_map [symmetric, of _ _ int] del: fun_upd_apply)
+      (simp add: nth_fun_of_list_eq [of _ _ undefined] del: fun_upd_apply)
+  from C show ?C2 by simp
+  have "list_all (\<lambda>n. int n \<le> 15) s_list"
+    by (simp add: s_list_def)
+  moreover have "length s_list = 80"
+    by (simp add: s_list_def)
+  ultimately show ?C3 unfolding C using `j \<le> 79`
+    by (simp add: s_def list_all_length)
+qed
+
+spark_end
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/S_R.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,30 @@
+(*  Title:      HOL/SPARK/Examples/RIPEMD-160/S_R.thy
+    Author:     Fabian Immler, TU Muenchen
+
+Verification of the RIPEMD-160 hash function
+*)
+
+theory S_R
+imports RMD_Specification RMD_Lemmas
+begin
+
+spark_open "rmd/s_r.siv"
+
+spark_vc function_s_r_2
+proof -
+  from `0 \<le> j` `j \<le> 79`
+  show C: ?C1
+    by (simp add: s'_def s'_list_def nth_map [symmetric, of _ _ int] del: fun_upd_apply)
+      (simp add: nth_fun_of_list_eq [of _ _ undefined] del: fun_upd_apply)
+  from C show ?C2 by simp
+  have "list_all (\<lambda>n. int n \<le> 15) s'_list"
+    by (simp add: s'_list_def)
+  moreover have "length s'_list = 80"
+    by (simp add: s'_list_def)
+  ultimately show ?C3 unfolding C using `j \<le> 79`
+    by (simp add: s'_def list_all_length)
+qed
+
+spark_end
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd.adb	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,187 @@
+package body RMD is
+
+
+
+   function F(J : Round_Index; X,Y,Z : Word) return Word
+   is
+      Result: Word;
+   begin
+      if     0 <= J and J <= 15 then Result := X xor Y xor Z;
+      elsif 16 <= J and J <= 31 then Result := (X and Y) or (not X and Z);
+      elsif 32 <= J and J <= 47 then Result := (X or not Y) xor Z;
+      elsif 48 <= J and J <= 63 then Result := (X and Z) or (Y and not Z);
+      else                           Result := X xor (Y or not Z);
+      end if;
+      return Result;
+   end F;
+
+
+
+   function K_L(J : Round_Index) return Word
+   is
+      K: Word;
+   begin
+      if     0 <= J and J <= 15 then K := 16#0000_0000#;
+      elsif 16 <= J and J <= 31 then K := 16#5A82_7999#;
+      elsif 32 <= J and J <= 47 then K := 16#6ED9_EBA1#;
+      elsif 48 <= J and J <= 63 then K := 16#8F1B_BCDC#;
+      else                           K := 16#A953_FD4E#;
+      end if;
+      return K;
+   end K_L;
+
+
+   function K_R(J : Round_Index) return Word
+   is
+      K: Word;
+   begin
+      if     0 <= J and J <= 15 then K := 16#50A2_8BE6#;
+      elsif 16 <= J and J <= 31 then K := 16#5C4D_D124#;
+      elsif 32 <= J and J <= 47 then K := 16#6D70_3EF3#;
+      elsif 48 <= J and J <= 63 then K := 16#7A6D_76E9#;
+      else                           K := 16#0000_0000#;
+      end if;
+      return K;
+   end K_R;
+
+
+
+   function R_L(J : Round_Index) return Block_Index
+   is
+      R_Values : constant Block_Permutation := Block_Permutation'
+        (0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15,
+         7, 4, 13, 1, 10, 6, 15, 3, 12, 0, 9, 5, 2, 14, 11, 8,
+         3, 10, 14, 4, 9, 15, 8, 1, 2, 7, 0, 6, 13, 11, 5, 12,
+         1, 9, 11, 10, 0, 8, 12, 4, 13, 3, 7, 15, 14, 5, 6, 2,
+         4, 0, 5, 9, 7, 12, 2, 10, 14, 1, 3, 8, 11, 6, 15, 13);
+      --# for R_Values declare rule;
+   begin
+      return R_Values(J);
+   end R_L;
+
+
+   function R_R(J : Round_Index) return Block_Index
+   is
+      R_Values : constant Block_Permutation := Block_Permutation'
+        (5, 14, 7, 0, 9, 2, 11, 4, 13, 6, 15, 8, 1, 10, 3, 12,
+         6, 11, 3, 7, 0, 13, 5, 10, 14, 15, 8, 12, 4, 9, 1, 2,
+         15, 5, 1, 3, 7, 14, 6, 9, 11, 8, 12, 2, 10, 0, 4, 13,
+         8, 6, 4, 1, 3, 11, 15, 0, 5, 12, 2, 13, 9, 7, 10, 14,
+         12, 15, 10, 4, 1, 5, 8, 7, 6, 2, 13, 14, 0, 3, 9, 11);
+      --# for R_Values declare rule;
+   begin
+      return R_Values(J);
+   end R_R;
+
+
+   function S_L(J : Round_Index) return Rotate_Amount
+   is
+      S_Values : constant Rotate_Definition := Rotate_Definition'
+        (11, 14, 15, 12, 5, 8, 7, 9, 11, 13, 14, 15, 6, 7, 9, 8,
+         7, 6, 8, 13, 11, 9, 7, 15, 7, 12, 15, 9, 11, 7, 13, 12,
+         11, 13, 6, 7, 14, 9, 13, 15, 14, 8, 13, 6, 5, 12, 7, 5,
+         11, 12, 14, 15, 14, 15, 9, 8, 9, 14, 5, 6, 8, 6, 5, 12,
+         9, 15, 5, 11, 6, 8, 13, 12, 5, 12, 13, 14, 11, 8, 5, 6);
+      --# for S_Values declare rule;
+   begin
+      return S_Values(J);
+   end S_L;
+
+
+   function S_R(J : Round_Index) return Rotate_Amount
+   is
+      S_Values : constant Rotate_Definition := Rotate_Definition'
+        (8, 9, 9, 11, 13, 15, 15, 5, 7, 7, 8, 11, 14, 14, 12, 6,
+         9, 13, 15, 7, 12, 8, 9, 11, 7, 7, 12, 7, 6, 15, 13, 11,
+         9, 7, 15, 11, 8, 6, 6, 14, 12, 13, 5, 14, 13, 13, 7, 5,
+         15, 5, 8, 11, 14, 14, 6, 14, 6, 9, 12, 9, 12, 5, 15, 8,
+         8, 5, 12, 9, 12, 5, 14, 6, 8, 13, 6, 5, 15, 13, 11, 11);
+      --# for S_Values declare rule;
+   begin
+      return S_Values(J);
+   end S_R;
+
+
+
+   procedure Round(CA, CB, CC, CD, CE : in out Word; X : in Block)
+   is
+      CLA, CLB, CLC, CLD, CLE, CRA, CRB, CRC, CRD, CRE : Word;
+      T : Word;
+   begin
+      CLA := CA;
+      CLB := CB;
+      CLC := CC;
+      CLD := CD;
+      CLE := CE;
+      CRA := CA;
+      CRB := CB;
+      CRC := CC;
+      CRD := CD;
+      CRE := CE;
+      for J in Round_Index range 0..79
+      loop
+         -- left
+         T := Wordops.Rotate(S_L(J),
+                             CLA +
+                               F(J, CLB, CLC, CLD) +
+                               X(R_L(J)) +
+                               K_L(J)) +
+           CLE;
+         CLA := CLE;
+         CLE := CLD;
+         CLD := Wordops.Rotate(10, CLC);
+         CLC := CLB;
+         CLB := T;
+         -- right
+         T := Wordops.Rotate(S_R(J),
+                             CRA +
+                               F(79 - J, CRB, CRC, CRD) +
+                               X(R_R(J)) +
+                               K_R(J)) +
+           CRE;
+         CRA := CRE;
+         CRE := CRD;
+         CRD := Wordops.Rotate(10, CRC);
+         CRC := CRB;
+         CRB := T;
+         --# assert Chain_Pair'(Chain'(CLA, CLB, CLC, CLD, CLE),
+         --#                    Chain'(CRA, CRB, CRC, CRD, CRE)) =
+         --#   steps(Chain_Pair'(Chain'(CA~, CB~, CC~, CD~, CE~),
+         --#                    Chain'(CA~, CB~, CC~, CD~, CE~)), J + 1, X)
+         --# and CA = CA~ and CB = CB~ and CC = CC~ and CD = CD~ and CE = CE~;
+      end loop;
+      T    := CB + CLC + CRD;
+      CB := CC + CLD + CRE;
+      CC := CD + CLE + CRA;
+      CD := CE + CLA + CRB;
+      CE := CA + CLB + CRC;
+      CA := T;
+   end Round;
+
+   function Hash(X : Message) return Chain
+   is
+      CA_Init : constant Word := 16#6745_2301#;
+      CB_Init : constant Word := 16#EFCD_AB89#;
+      CC_Init : constant Word := 16#98BA_DCFE#;
+      CD_Init : constant Word := 16#1032_5476#;
+      CE_Init : constant Word := 16#C3D2_E1F0#;
+      CA, CB, CC, CD, CE : Word;
+   begin
+      CA := CA_Init;
+      CB := CB_Init;
+      CC := CC_Init;
+      CD := CD_Init;
+      CE := CE_Init;
+      for I in Message_Index range X'First..X'Last
+      loop
+         Round(CA, CB, CC, CD, CE, X(I));
+         --# assert Chain'(CA, CB, CC, CD, CE) = rounds(
+         --#    Chain'(CA_Init, CB_Init, CC_Init, CD_Init, CE_Init),
+         --#    I + 1,
+         --#    X);
+      end loop;
+      return Chain'(CA, CB, CC, CD, CE);
+   end Hash;
+
+end RMD;
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd.ads	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,92 @@
+with Wordops;
+use type Wordops.Word;
+--# inherit Wordops;
+
+package RMD
+is
+
+   -- Types
+
+   subtype Word is Wordops.Word;
+
+   type Chain is
+     record
+        H0, H1, H2, H3, H4 : Word;
+     end record;
+
+   type Block_Index is range 0..15;
+   type Block is array(Block_Index) of Word;
+
+   type Message_Index is range 0..2**32;
+   type Message is array(Message_Index range <>) of Block;
+
+   -- Isabelle specification
+
+   --# function rmd_hash(X : Message; L : Message_Index) return Chain;
+
+   function Hash(X : Message) return Chain;
+   --# pre X'First = 0;
+   --# return rmd_hash(X, X'Last + 1);
+
+private
+
+   -- Types
+
+   type Round_Index is range 0..79;
+
+   type Chain_Pair is
+      record
+         Left, Right : Chain;
+      end record;
+
+   type Block_Permutation is array(Round_Index) of Block_Index;
+
+   subtype Rotate_Amount is Wordops.Rotate_Amount;
+   type Rotate_Definition is array(Round_Index) of Rotate_Amount;
+
+
+   -- Isabelle proof functions
+
+   --# function f_spec(J : Round_Index; X,Y,Z : Word) return Word;
+   --# function K_l_spec(J : Round_Index) return Word;
+   --# function K_r_spec(J : Round_Index) return Word;
+   --# function r_l_spec(J : Round_Index) return Block_Index;
+   --# function r_r_spec(J : Round_Index) return Block_Index;
+   --# function s_l_spec(J : Round_Index) return Rotate_Amount;
+   --# function s_r_spec(J : Round_Index) return Rotate_Amount;
+   --# function steps(CS : Chain_Pair; I : Round_Index; B : Block)
+   --#    return Chain_Pair;
+   --# function round_spec(C : Chain; B : Block) return Chain;
+   --# function rounds(C : Chain; I : Message_Index; X : Message)
+   --#    return Chain;
+
+
+   -- Spark Implementation
+
+   function F(J : Round_Index; X,Y,Z : Word) return Word;
+   --# return f_spec(J, X, Y, Z);
+
+   function K_L(J : Round_Index) return Word;
+   --# return K_l_spec(J);
+
+   function K_R(J : Round_Index) return Word;
+   --# return K_r_spec(J);
+
+   function R_L(J : Round_Index) return Block_Index;
+   --# return r_l_spec(J);
+
+   function R_R(J : Round_Index) return Block_Index;
+   --# return r_r_spec(J);
+
+   function S_L(J : Round_Index) return Rotate_Amount;
+   --# return s_l_spec(J);
+
+   function S_R(J : Round_Index) return Rotate_Amount;
+   --# return s_r_spec(J);
+
+   procedure Round(CA, CB, CC, CD, CE : in out Word; X: in Block);
+   --# derives CA, CB, CC, CD, CE from X, CA, CB, CC, CD, CE;
+   --# post Chain'(CA, CB, CC, CD, CE) =
+   --#        round_spec(Chain'(CA~, CB~, CC~, CD~, CE~), X);
+
+end RMD;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/f.fdl	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,41 @@
+           {*******************************************************}
+                               {FDL Declarations}
+    {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039}
+             {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.}
+           {*******************************************************}
+
+
+                        {DATE : 29-NOV-2010 14:30:19.73}
+
+                               {function RMD.F}
+
+
+title function f;
+
+  function round__(real) : integer;
+  type interfaces__unsigned_32 = integer;
+  type round_index = integer;
+  const round_index__base__first : integer = pending; 
+  const round_index__base__last : integer = pending; 
+  const word__base__first : integer = pending; 
+  const word__base__last : integer = pending; 
+  const interfaces__unsigned_32__base__first : integer = pending; 
+  const interfaces__unsigned_32__base__last : integer = pending; 
+  const round_index__first : integer = pending; 
+  const round_index__last : integer = pending; 
+  const round_index__size : integer = pending; 
+  const word__first : integer = pending; 
+  const word__last : integer = pending; 
+  const word__modulus : integer = pending; 
+  const word__size : integer = pending; 
+  const interfaces__unsigned_32__first : integer = pending; 
+  const interfaces__unsigned_32__last : integer = pending; 
+  const interfaces__unsigned_32__modulus : integer = pending; 
+  const interfaces__unsigned_32__size : integer = pending; 
+  var j : integer;
+  var x : integer;
+  var y : integer;
+  var z : integer;
+  function f_spec(integer, integer, integer, integer) : integer;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/f.rls	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,35 @@
+           /*********************************************************/
+                           /*Proof Rule Declarations*/
+    /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/
+             /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/
+           /*********************************************************/
+
+
+                        /*DATE : 29-NOV-2010 14:30:19.73*/
+
+                               /*function RMD.F*/
+
+
+rule_family f_rules:
+     X      requires [X:any] &
+     X <= Y requires [X:ire, Y:ire] &
+     X >= Y requires [X:ire, Y:ire].
+
+f_rules(1): interfaces__unsigned_32__size >= 0 may_be_deduced.
+f_rules(2): interfaces__unsigned_32__first may_be_replaced_by 0.
+f_rules(3): interfaces__unsigned_32__last may_be_replaced_by 4294967295.
+f_rules(4): interfaces__unsigned_32__base__first may_be_replaced_by 0.
+f_rules(5): interfaces__unsigned_32__base__last may_be_replaced_by 4294967295.
+f_rules(6): interfaces__unsigned_32__modulus may_be_replaced_by 4294967296.
+f_rules(7): word__size >= 0 may_be_deduced.
+f_rules(8): word__first may_be_replaced_by 0.
+f_rules(9): word__last may_be_replaced_by 4294967295.
+f_rules(10): word__base__first may_be_replaced_by 0.
+f_rules(11): word__base__last may_be_replaced_by 4294967295.
+f_rules(12): word__modulus may_be_replaced_by 4294967296.
+f_rules(13): round_index__size >= 0 may_be_deduced.
+f_rules(14): round_index__first may_be_replaced_by 0.
+f_rules(15): round_index__last may_be_replaced_by 79.
+f_rules(16): round_index__base__first <= round_index__base__last may_be_deduced.
+f_rules(17): round_index__base__first <= round_index__first may_be_deduced.
+f_rules(18): round_index__base__last >= round_index__last may_be_deduced.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/f.siv	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,228 @@
+*****************************************************************************
+                       Semantic Analysis of SPARK Text
+    Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+             Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+*****************************************************************************
+
+
+CREATED 29-NOV-2010, 14:30:19  SIMPLIFIED 29-NOV-2010, 14:30:28
+
+SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+
+function RMD.F
+
+
+
+
+For path(s) from start to run-time check associated with statement of line 9:
+
+function_f_1.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 10:
+
+function_f_2.
+H1:    x >= 0 .
+H2:    x <= 4294967295 .
+H3:    y >= 0 .
+H4:    y <= 4294967295 .
+H5:    z >= 0 .
+H6:    z <= 4294967295 .
+H7:    16 <= j .
+H8:    j <= 31 .
+H9:    interfaces__unsigned_32__size >= 0 .
+H10:   word__size >= 0 .
+H11:   round_index__size >= 0 .
+H12:   round_index__base__first <= round_index__base__last .
+H13:   round_index__base__first <= 0 .
+H14:   round_index__base__last >= 79 .
+       ->
+C1:    bit__or(bit__and(x, y), bit__and(4294967295 - x, z)) >= 0 .
+C2:    bit__or(bit__and(x, y), bit__and(4294967295 - x, z)) <= 4294967295 .
+
+
+For path(s) from start to run-time check associated with statement of line 11:
+
+function_f_3.
+H1:    x >= 0 .
+H2:    x <= 4294967295 .
+H3:    y >= 0 .
+H4:    y <= 4294967295 .
+H5:    z >= 0 .
+H6:    z <= 4294967295 .
+H7:    32 <= j .
+H8:    j <= 47 .
+H9:    interfaces__unsigned_32__size >= 0 .
+H10:   word__size >= 0 .
+H11:   round_index__size >= 0 .
+H12:   round_index__base__first <= round_index__base__last .
+H13:   round_index__base__first <= 0 .
+H14:   round_index__base__last >= 79 .
+       ->
+C1:    bit__xor(bit__or(x, 4294967295 - y), z) >= 0 .
+C2:    bit__xor(bit__or(x, 4294967295 - y), z) <= 4294967295 .
+
+
+For path(s) from start to run-time check associated with statement of line 12:
+
+function_f_4.
+H1:    x >= 0 .
+H2:    x <= 4294967295 .
+H3:    y >= 0 .
+H4:    y <= 4294967295 .
+H5:    z >= 0 .
+H6:    z <= 4294967295 .
+H7:    48 <= j .
+H8:    j <= 63 .
+H9:    interfaces__unsigned_32__size >= 0 .
+H10:   word__size >= 0 .
+H11:   round_index__size >= 0 .
+H12:   round_index__base__first <= round_index__base__last .
+H13:   round_index__base__first <= 0 .
+H14:   round_index__base__last >= 79 .
+       ->
+C1:    bit__or(bit__and(x, z), bit__and(y, 4294967295 - z)) >= 0 .
+C2:    bit__or(bit__and(x, z), bit__and(y, 4294967295 - z)) <= 4294967295 .
+
+
+For path(s) from start to run-time check associated with statement of line 13:
+
+function_f_5.
+H1:    j >= 0 .
+H2:    j <= 79 .
+H3:    x >= 0 .
+H4:    x <= 4294967295 .
+H5:    y >= 0 .
+H6:    y <= 4294967295 .
+H7:    z >= 0 .
+H8:    z <= 4294967295 .
+H9:    15 < j .
+H10:   31 < j .
+H11:   47 < j .
+H12:   63 < j .
+H13:   interfaces__unsigned_32__size >= 0 .
+H14:   word__size >= 0 .
+H15:   round_index__size >= 0 .
+H16:   round_index__base__first <= round_index__base__last .
+H17:   round_index__base__first <= 0 .
+H18:   round_index__base__last >= 79 .
+       ->
+C1:    bit__xor(x, bit__or(y, 4294967295 - z)) >= 0 .
+C2:    bit__xor(x, bit__or(y, 4294967295 - z)) <= 4294967295 .
+
+
+For path(s) from start to finish:
+
+function_f_6.
+H1:    j >= 0 .
+H2:    x >= 0 .
+H3:    x <= 4294967295 .
+H4:    y >= 0 .
+H5:    y <= 4294967295 .
+H6:    z >= 0 .
+H7:    z <= 4294967295 .
+H8:    j <= 15 .
+H9:    bit__xor(x, bit__xor(y, z)) >= 0 .
+H10:   bit__xor(x, bit__xor(y, z)) <= 4294967295 .
+H11:   interfaces__unsigned_32__size >= 0 .
+H12:   word__size >= 0 .
+H13:   round_index__size >= 0 .
+H14:   round_index__base__first <= round_index__base__last .
+H15:   round_index__base__first <= 0 .
+H16:   round_index__base__last >= 79 .
+       ->
+C1:    bit__xor(x, bit__xor(y, z)) = f_spec(j, x, y, z) .
+
+
+function_f_7.
+H1:    x >= 0 .
+H2:    x <= 4294967295 .
+H3:    y >= 0 .
+H4:    y <= 4294967295 .
+H5:    z >= 0 .
+H6:    z <= 4294967295 .
+H7:    16 <= j .
+H8:    j <= 31 .
+H9:    bit__or(bit__and(x, y), bit__and(4294967295 - x, z)) >= 0 .
+H10:   bit__or(bit__and(x, y), bit__and(4294967295 - x, z)) <= 4294967295 .
+H11:   interfaces__unsigned_32__size >= 0 .
+H12:   word__size >= 0 .
+H13:   round_index__size >= 0 .
+H14:   round_index__base__first <= round_index__base__last .
+H15:   round_index__base__first <= 0 .
+H16:   round_index__base__last >= 79 .
+       ->
+C1:    bit__or(bit__and(x, y), bit__and(4294967295 - x, z)) = f_spec(j, x, y, z)
+           .
+
+
+function_f_8.
+H1:    x >= 0 .
+H2:    x <= 4294967295 .
+H3:    y >= 0 .
+H4:    y <= 4294967295 .
+H5:    z >= 0 .
+H6:    z <= 4294967295 .
+H7:    32 <= j .
+H8:    j <= 47 .
+H9:    bit__xor(bit__or(x, 4294967295 - y), z) >= 0 .
+H10:   bit__xor(bit__or(x, 4294967295 - y), z) <= 4294967295 .
+H11:   interfaces__unsigned_32__size >= 0 .
+H12:   word__size >= 0 .
+H13:   round_index__size >= 0 .
+H14:   round_index__base__first <= round_index__base__last .
+H15:   round_index__base__first <= 0 .
+H16:   round_index__base__last >= 79 .
+       ->
+C1:    bit__xor(bit__or(x, 4294967295 - y), z) = f_spec(j, x, y, z) .
+
+
+function_f_9.
+H1:    x >= 0 .
+H2:    x <= 4294967295 .
+H3:    y >= 0 .
+H4:    y <= 4294967295 .
+H5:    z >= 0 .
+H6:    z <= 4294967295 .
+H7:    48 <= j .
+H8:    j <= 63 .
+H9:    bit__or(bit__and(x, z), bit__and(y, 4294967295 - z)) >= 0 .
+H10:   bit__or(bit__and(x, z), bit__and(y, 4294967295 - z)) <= 4294967295 .
+H11:   interfaces__unsigned_32__size >= 0 .
+H12:   word__size >= 0 .
+H13:   round_index__size >= 0 .
+H14:   round_index__base__first <= round_index__base__last .
+H15:   round_index__base__first <= 0 .
+H16:   round_index__base__last >= 79 .
+       ->
+C1:    bit__or(bit__and(x, z), bit__and(y, 4294967295 - z)) = f_spec(j, x, y, z)
+           .
+
+
+function_f_10.
+H1:    j >= 0 .
+H2:    j <= 79 .
+H3:    x >= 0 .
+H4:    x <= 4294967295 .
+H5:    y >= 0 .
+H6:    y <= 4294967295 .
+H7:    z >= 0 .
+H8:    z <= 4294967295 .
+H9:    15 < j .
+H10:   31 < j .
+H11:   47 < j .
+H12:   63 < j .
+H13:   bit__xor(x, bit__or(y, 4294967295 - z)) >= 0 .
+H14:   bit__xor(x, bit__or(y, 4294967295 - z)) <= 4294967295 .
+H15:   interfaces__unsigned_32__size >= 0 .
+H16:   word__size >= 0 .
+H17:   round_index__size >= 0 .
+H18:   round_index__base__first <= round_index__base__last .
+H19:   round_index__base__first <= 0 .
+H20:   round_index__base__last >= 79 .
+       ->
+C1:    bit__xor(x, bit__or(y, 4294967295 - z)) = f_spec(j, x, y, z) .
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/hash.fdl	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,74 @@
+           {*******************************************************}
+                               {FDL Declarations}
+    {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039}
+             {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.}
+           {*******************************************************}
+
+
+                        {DATE : 29-NOV-2010 14:30:20.17}
+
+                             {function RMD.Hash}
+
+
+title function hash;
+
+  function round__(real) : integer;
+  type interfaces__unsigned_32 = integer;
+  type block_index = integer;
+  type message_index = integer;
+  type chain = record
+        h0 : integer;
+        h1 : integer;
+        h2 : integer;
+        h3 : integer;
+        h4 : integer
+     end;
+  type block = array [integer] of integer;
+  type message = array [integer] of block;
+  const ca_init : integer = pending;
+  const cb_init : integer = pending;
+  const cc_init : integer = pending;
+  const cd_init : integer = pending;
+  const ce_init : integer = pending;
+  const message_index__base__first : integer = pending; 
+  const message_index__base__last : integer = pending; 
+  const block_index__base__first : integer = pending; 
+  const block_index__base__last : integer = pending; 
+  const word__base__first : integer = pending; 
+  const word__base__last : integer = pending; 
+  const interfaces__unsigned_32__base__first : integer = pending; 
+  const interfaces__unsigned_32__base__last : integer = pending; 
+  const x__index__subtype__1__first : integer = pending; 
+  const x__index__subtype__1__last : integer = pending; 
+  const message_index__first : integer = pending; 
+  const message_index__last : integer = pending; 
+  const message_index__size : integer = pending; 
+  const block_index__first : integer = pending; 
+  const block_index__last : integer = pending; 
+  const block_index__size : integer = pending; 
+  const chain__size : integer = pending; 
+  const word__first : integer = pending; 
+  const word__last : integer = pending; 
+  const word__modulus : integer = pending; 
+  const word__size : integer = pending; 
+  const interfaces__unsigned_32__first : integer = pending; 
+  const interfaces__unsigned_32__last : integer = pending; 
+  const interfaces__unsigned_32__modulus : integer = pending; 
+  const interfaces__unsigned_32__size : integer = pending; 
+  var x : message;
+  var ca : integer;
+  var cb : integer;
+  var cc : integer;
+  var cd : integer;
+  var ce : integer;
+  var loop__1__i : integer;
+  function rmd_hash(message, integer) : chain;
+  function round_spec(chain, block) : chain;
+  function rounds(chain, integer, message) : chain;
+  var ce__1 : integer;
+  var cd__1 : integer;
+  var cc__1 : integer;
+  var cb__1 : integer;
+  var ca__1 : integer;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/hash.rls	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,61 @@
+           /*********************************************************/
+                           /*Proof Rule Declarations*/
+    /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/
+             /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/
+           /*********************************************************/
+
+
+                        /*DATE : 29-NOV-2010 14:30:20.17*/
+
+                             /*function RMD.Hash*/
+
+
+rule_family hash_rules:
+     X      requires [X:any] &
+     X <= Y requires [X:ire, Y:ire] &
+     X >= Y requires [X:ire, Y:ire].
+
+hash_rules(1): ca_init may_be_replaced_by 1732584193.
+hash_rules(2): cb_init may_be_replaced_by 4023233417.
+hash_rules(3): cc_init may_be_replaced_by 2562383102.
+hash_rules(4): cd_init may_be_replaced_by 271733878.
+hash_rules(5): ce_init may_be_replaced_by 3285377520.
+hash_rules(6): interfaces__unsigned_32__size >= 0 may_be_deduced.
+hash_rules(7): interfaces__unsigned_32__first may_be_replaced_by 0.
+hash_rules(8): interfaces__unsigned_32__last may_be_replaced_by 4294967295.
+hash_rules(9): interfaces__unsigned_32__base__first may_be_replaced_by 0.
+hash_rules(10): interfaces__unsigned_32__base__last may_be_replaced_by 4294967295.
+hash_rules(11): interfaces__unsigned_32__modulus may_be_replaced_by 4294967296.
+hash_rules(12): word__size >= 0 may_be_deduced.
+hash_rules(13): word__first may_be_replaced_by 0.
+hash_rules(14): word__last may_be_replaced_by 4294967295.
+hash_rules(15): word__base__first may_be_replaced_by 0.
+hash_rules(16): word__base__last may_be_replaced_by 4294967295.
+hash_rules(17): word__modulus may_be_replaced_by 4294967296.
+hash_rules(18): chain__size >= 0 may_be_deduced.
+hash_rules(19): A = B may_be_deduced_from
+     [goal(checktype(A,chain)),
+      goal(checktype(B,chain)),
+      fld_h0(A) = fld_h0(B),
+      fld_h1(A) = fld_h1(B),
+      fld_h2(A) = fld_h2(B),
+      fld_h3(A) = fld_h3(B),
+      fld_h4(A) = fld_h4(B)].
+hash_rules(20): block_index__size >= 0 may_be_deduced.
+hash_rules(21): block_index__first may_be_replaced_by 0.
+hash_rules(22): block_index__last may_be_replaced_by 15.
+hash_rules(23): block_index__base__first <= block_index__base__last may_be_deduced.
+hash_rules(24): block_index__base__first <= block_index__first may_be_deduced.
+hash_rules(25): block_index__base__last >= block_index__last may_be_deduced.
+hash_rules(26): message_index__size >= 0 may_be_deduced.
+hash_rules(27): message_index__first may_be_replaced_by 0.
+hash_rules(28): message_index__last may_be_replaced_by 4294967296.
+hash_rules(29): message_index__base__first <= message_index__base__last may_be_deduced.
+hash_rules(30): message_index__base__first <= message_index__first may_be_deduced.
+hash_rules(31): message_index__base__last >= message_index__last may_be_deduced.
+hash_rules(32): x__index__subtype__1__first >= message_index__first may_be_deduced.
+hash_rules(33): x__index__subtype__1__last <= message_index__last may_be_deduced.
+hash_rules(34): x__index__subtype__1__first <= 
+     x__index__subtype__1__last may_be_deduced.
+hash_rules(35): x__index__subtype__1__last >= message_index__first may_be_deduced.
+hash_rules(36): x__index__subtype__1__first <= message_index__last may_be_deduced.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/hash.siv	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,240 @@
+*****************************************************************************
+                       Semantic Analysis of SPARK Text
+    Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+             Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+*****************************************************************************
+
+
+CREATED 29-NOV-2010, 14:30:20  SIMPLIFIED 29-NOV-2010, 14:30:20
+
+SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+
+function RMD.Hash
+
+
+
+
+For path(s) from start to run-time check associated with statement of line 170:
+
+function_hash_1.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 171:
+
+function_hash_2.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 172:
+
+function_hash_3.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 173:
+
+function_hash_4.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 174:
+
+function_hash_5.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 175:
+
+function_hash_6.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 175:
+
+function_hash_7.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 177:
+
+function_hash_8.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 178 to run-time check associated with 
+          statement of line 177:
+
+function_hash_9.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 177:
+
+function_hash_10.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 178 to run-time check associated with 
+          statement of line 177:
+
+function_hash_11.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to assertion of line 178:
+
+function_hash_12.
+H1:    x__index__subtype__1__first = 0 .
+H2:    for_all(i___2 : integer, 0 <= i___2 and i___2 <= 15 -> for_all(i___1 : 
+          integer, x__index__subtype__1__first <= i___1 and i___1 <= 
+          x__index__subtype__1__last -> 0 <= element(element(x, [i___1]), [
+          i___2]) and element(element(x, [i___1]), [i___2]) <= 4294967295)) .
+H3:    x__index__subtype__1__last >= 0 .
+H4:    x__index__subtype__1__last <= 4294967296 .
+H5:    x__index__subtype__1__first <= x__index__subtype__1__last .
+H6:    mk__chain(h0 := ca__1, h1 := cb__1, h2 := cc__1, h3 := cd__1, h4 := 
+          ce__1) = round_spec(mk__chain(h0 := 1732584193, h1 := 4023233417, h2 
+          := 2562383102, h3 := 271733878, h4 := 3285377520), element(x, [
+          x__index__subtype__1__first])) .
+H7:    ca__1 >= 0 .
+H8:    ca__1 <= 4294967295 .
+H9:    cb__1 >= 0 .
+H10:   cb__1 <= 4294967295 .
+H11:   cc__1 >= 0 .
+H12:   cc__1 <= 4294967295 .
+H13:   cd__1 >= 0 .
+H14:   cd__1 <= 4294967295 .
+H15:   ce__1 >= 0 .
+H16:   ce__1 <= 4294967295 .
+       ->
+C1:    mk__chain(h0 := ca__1, h1 := cb__1, h2 := cc__1, h3 := cd__1, h4 := 
+          ce__1) = rounds(mk__chain(h0 := 1732584193, h1 := 4023233417, h2 := 
+          2562383102, h3 := 271733878, h4 := 3285377520), 
+          x__index__subtype__1__first + 1, x) .
+
+
+For path(s) from assertion of line 178 to assertion of line 178:
+
+function_hash_13.
+H1:    mk__chain(h0 := ca, h1 := cb, h2 := cc, h3 := cd, h4 := ce) = rounds(
+          mk__chain(h0 := 1732584193, h1 := 4023233417, h2 := 2562383102, h3 := 
+          271733878, h4 := 3285377520), loop__1__i + 1, x) .
+H2:    for_all(i___2 : integer, 0 <= i___2 and i___2 <= 15 -> for_all(i___1 : 
+          integer, x__index__subtype__1__first <= i___1 and i___1 <= 
+          x__index__subtype__1__last -> 0 <= element(element(x, [i___1]), [
+          i___2]) and element(element(x, [i___1]), [i___2]) <= 4294967295)) .
+H3:    x__index__subtype__1__first = 0 .
+H4:    loop__1__i >= 0 .
+H5:    loop__1__i <= 4294967296 .
+H6:    loop__1__i >= x__index__subtype__1__first .
+H7:    ca >= 0 .
+H8:    ca <= 4294967295 .
+H9:    cb >= 0 .
+H10:   cb <= 4294967295 .
+H11:   cc >= 0 .
+H12:   cc <= 4294967295 .
+H13:   cd >= 0 .
+H14:   cd <= 4294967295 .
+H15:   ce >= 0 .
+H16:   ce <= 4294967295 .
+H17:   loop__1__i + 1 <= x__index__subtype__1__last .
+H18:   mk__chain(h0 := ca__1, h1 := cb__1, h2 := cc__1, h3 := cd__1, h4 := 
+          ce__1) = round_spec(mk__chain(h0 := ca, h1 := cb, h2 := cc, h3 := cd, 
+          h4 := ce), element(x, [loop__1__i + 1])) .
+H19:   ca__1 >= 0 .
+H20:   ca__1 <= 4294967295 .
+H21:   cb__1 >= 0 .
+H22:   cb__1 <= 4294967295 .
+H23:   cc__1 >= 0 .
+H24:   cc__1 <= 4294967295 .
+H25:   cd__1 >= 0 .
+H26:   cd__1 <= 4294967295 .
+H27:   ce__1 >= 0 .
+H28:   ce__1 <= 4294967295 .
+H29:   interfaces__unsigned_32__size >= 0 .
+H30:   word__size >= 0 .
+H31:   chain__size >= 0 .
+H32:   block_index__size >= 0 .
+H33:   block_index__base__first <= block_index__base__last .
+H34:   message_index__size >= 0 .
+H35:   message_index__base__first <= message_index__base__last .
+H36:   x__index__subtype__1__first <= x__index__subtype__1__last .
+H37:   block_index__base__first <= 0 .
+H38:   block_index__base__last >= 15 .
+H39:   message_index__base__first <= 0 .
+H40:   x__index__subtype__1__first >= 0 .
+H41:   x__index__subtype__1__last >= 0 .
+H42:   message_index__base__last >= 4294967296 .
+H43:   x__index__subtype__1__last <= 4294967296 .
+H44:   x__index__subtype__1__first <= 4294967296 .
+       ->
+C1:    mk__chain(h0 := ca__1, h1 := cb__1, h2 := cc__1, h3 := cd__1, h4 := 
+          ce__1) = rounds(mk__chain(h0 := 1732584193, h1 := 4023233417, h2 := 
+          2562383102, h3 := 271733878, h4 := 3285377520), loop__1__i + 2, x) .
+
+
+For path(s) from start to run-time check associated with statement of line 183:
+
+function_hash_14.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 178 to run-time check associated with 
+          statement of line 183:
+
+function_hash_15.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to finish:
+
+function_hash_16.
+*** true .   /* contradiction within hypotheses. */
+
+
+
+For path(s) from assertion of line 178 to finish:
+
+function_hash_17.
+H1:    mk__chain(h0 := ca, h1 := cb, h2 := cc, h3 := cd, h4 := ce) = rounds(
+          mk__chain(h0 := 1732584193, h1 := 4023233417, h2 := 2562383102, h3 := 
+          271733878, h4 := 3285377520), x__index__subtype__1__last + 1, x) .
+H2:    for_all(i___2 : integer, 0 <= i___2 and i___2 <= 15 -> for_all(i___1 : 
+          integer, x__index__subtype__1__first <= i___1 and i___1 <= 
+          x__index__subtype__1__last -> 0 <= element(element(x, [i___1]), [
+          i___2]) and element(element(x, [i___1]), [i___2]) <= 4294967295)) .
+H3:    x__index__subtype__1__first = 0 .
+H4:    x__index__subtype__1__last >= 0 .
+H5:    x__index__subtype__1__last <= 4294967296 .
+H6:    x__index__subtype__1__last >= x__index__subtype__1__first .
+H7:    ca >= 0 .
+H8:    ca <= 4294967295 .
+H9:    cb >= 0 .
+H10:   cb <= 4294967295 .
+H11:   cc >= 0 .
+H12:   cc <= 4294967295 .
+H13:   cd >= 0 .
+H14:   cd <= 4294967295 .
+H15:   ce >= 0 .
+H16:   ce <= 4294967295 .
+H17:   interfaces__unsigned_32__size >= 0 .
+H18:   word__size >= 0 .
+H19:   chain__size >= 0 .
+H20:   block_index__size >= 0 .
+H21:   block_index__base__first <= block_index__base__last .
+H22:   message_index__size >= 0 .
+H23:   message_index__base__first <= message_index__base__last .
+H24:   x__index__subtype__1__first <= x__index__subtype__1__last .
+H25:   block_index__base__first <= 0 .
+H26:   block_index__base__last >= 15 .
+H27:   message_index__base__first <= 0 .
+H28:   x__index__subtype__1__first >= 0 .
+H29:   message_index__base__last >= 4294967296 .
+H30:   x__index__subtype__1__first <= 4294967296 .
+       ->
+C1:    mk__chain(h0 := ca, h1 := cb, h2 := cc, h3 := cd, h4 := ce) = rmd_hash(
+          x, x__index__subtype__1__last + 1) .
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_l.fdl	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,38 @@
+           {*******************************************************}
+                               {FDL Declarations}
+    {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039}
+             {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.}
+           {*******************************************************}
+
+
+                        {DATE : 29-NOV-2010 14:30:19.74}
+
+                              {function RMD.K_L}
+
+
+title function k_l;
+
+  function round__(real) : integer;
+  type interfaces__unsigned_32 = integer;
+  type round_index = integer;
+  const round_index__base__first : integer = pending; 
+  const round_index__base__last : integer = pending; 
+  const word__base__first : integer = pending; 
+  const word__base__last : integer = pending; 
+  const interfaces__unsigned_32__base__first : integer = pending; 
+  const interfaces__unsigned_32__base__last : integer = pending; 
+  const round_index__first : integer = pending; 
+  const round_index__last : integer = pending; 
+  const round_index__size : integer = pending; 
+  const word__first : integer = pending; 
+  const word__last : integer = pending; 
+  const word__modulus : integer = pending; 
+  const word__size : integer = pending; 
+  const interfaces__unsigned_32__first : integer = pending; 
+  const interfaces__unsigned_32__last : integer = pending; 
+  const interfaces__unsigned_32__modulus : integer = pending; 
+  const interfaces__unsigned_32__size : integer = pending; 
+  var j : integer;
+  function k_l_spec(integer) : integer;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_l.rls	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,35 @@
+           /*********************************************************/
+                           /*Proof Rule Declarations*/
+    /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/
+             /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/
+           /*********************************************************/
+
+
+                        /*DATE : 29-NOV-2010 14:30:19.74*/
+
+                              /*function RMD.K_L*/
+
+
+rule_family k_l_rules:
+     X      requires [X:any] &
+     X <= Y requires [X:ire, Y:ire] &
+     X >= Y requires [X:ire, Y:ire].
+
+k_l_rules(1): interfaces__unsigned_32__size >= 0 may_be_deduced.
+k_l_rules(2): interfaces__unsigned_32__first may_be_replaced_by 0.
+k_l_rules(3): interfaces__unsigned_32__last may_be_replaced_by 4294967295.
+k_l_rules(4): interfaces__unsigned_32__base__first may_be_replaced_by 0.
+k_l_rules(5): interfaces__unsigned_32__base__last may_be_replaced_by 4294967295.
+k_l_rules(6): interfaces__unsigned_32__modulus may_be_replaced_by 4294967296.
+k_l_rules(7): word__size >= 0 may_be_deduced.
+k_l_rules(8): word__first may_be_replaced_by 0.
+k_l_rules(9): word__last may_be_replaced_by 4294967295.
+k_l_rules(10): word__base__first may_be_replaced_by 0.
+k_l_rules(11): word__base__last may_be_replaced_by 4294967295.
+k_l_rules(12): word__modulus may_be_replaced_by 4294967296.
+k_l_rules(13): round_index__size >= 0 may_be_deduced.
+k_l_rules(14): round_index__first may_be_replaced_by 0.
+k_l_rules(15): round_index__last may_be_replaced_by 79.
+k_l_rules(16): round_index__base__first <= round_index__base__last may_be_deduced.
+k_l_rules(17): round_index__base__first <= round_index__first may_be_deduced.
+k_l_rules(18): round_index__base__last >= round_index__last may_be_deduced.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_l.siv	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,118 @@
+*****************************************************************************
+                       Semantic Analysis of SPARK Text
+    Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+             Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+*****************************************************************************
+
+
+CREATED 29-NOV-2010, 14:30:19  SIMPLIFIED 29-NOV-2010, 14:30:28
+
+SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+
+function RMD.K_L
+
+
+
+
+For path(s) from start to run-time check associated with statement of line 24:
+
+function_k_l_1.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 25:
+
+function_k_l_2.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 26:
+
+function_k_l_3.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 27:
+
+function_k_l_4.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 28:
+
+function_k_l_5.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to finish:
+
+function_k_l_6.
+H1:    j >= 0 .
+H2:    j <= 15 .
+H3:    interfaces__unsigned_32__size >= 0 .
+H4:    word__size >= 0 .
+H5:    round_index__size >= 0 .
+H6:    round_index__base__first <= round_index__base__last .
+H7:    round_index__base__first <= 0 .
+H8:    round_index__base__last >= 79 .
+       ->
+C1:    0 = k_l_spec(j) .
+
+
+function_k_l_7.
+H1:    16 <= j .
+H2:    j <= 31 .
+H3:    interfaces__unsigned_32__size >= 0 .
+H4:    word__size >= 0 .
+H5:    round_index__size >= 0 .
+H6:    round_index__base__first <= round_index__base__last .
+H7:    round_index__base__first <= 0 .
+H8:    round_index__base__last >= 79 .
+       ->
+C1:    1518500249 = k_l_spec(j) .
+
+
+function_k_l_8.
+H1:    32 <= j .
+H2:    j <= 47 .
+H3:    interfaces__unsigned_32__size >= 0 .
+H4:    word__size >= 0 .
+H5:    round_index__size >= 0 .
+H6:    round_index__base__first <= round_index__base__last .
+H7:    round_index__base__first <= 0 .
+H8:    round_index__base__last >= 79 .
+       ->
+C1:    1859775393 = k_l_spec(j) .
+
+
+function_k_l_9.
+H1:    48 <= j .
+H2:    j <= 63 .
+H3:    interfaces__unsigned_32__size >= 0 .
+H4:    word__size >= 0 .
+H5:    round_index__size >= 0 .
+H6:    round_index__base__first <= round_index__base__last .
+H7:    round_index__base__first <= 0 .
+H8:    round_index__base__last >= 79 .
+       ->
+C1:    2400959708 = k_l_spec(j) .
+
+
+function_k_l_10.
+H1:    j >= 0 .
+H2:    j <= 79 .
+H3:    15 < j .
+H4:    31 < j .
+H5:    47 < j .
+H6:    63 < j .
+H7:    interfaces__unsigned_32__size >= 0 .
+H8:    word__size >= 0 .
+H9:    round_index__size >= 0 .
+H10:   round_index__base__first <= round_index__base__last .
+H11:   round_index__base__first <= 0 .
+H12:   round_index__base__last >= 79 .
+       ->
+C1:    2840853838 = k_l_spec(j) .
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_r.fdl	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,38 @@
+           {*******************************************************}
+                               {FDL Declarations}
+    {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039}
+             {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.}
+           {*******************************************************}
+
+
+                        {DATE : 29-NOV-2010 14:30:19.76}
+
+                              {function RMD.K_R}
+
+
+title function k_r;
+
+  function round__(real) : integer;
+  type interfaces__unsigned_32 = integer;
+  type round_index = integer;
+  const round_index__base__first : integer = pending; 
+  const round_index__base__last : integer = pending; 
+  const word__base__first : integer = pending; 
+  const word__base__last : integer = pending; 
+  const interfaces__unsigned_32__base__first : integer = pending; 
+  const interfaces__unsigned_32__base__last : integer = pending; 
+  const round_index__first : integer = pending; 
+  const round_index__last : integer = pending; 
+  const round_index__size : integer = pending; 
+  const word__first : integer = pending; 
+  const word__last : integer = pending; 
+  const word__modulus : integer = pending; 
+  const word__size : integer = pending; 
+  const interfaces__unsigned_32__first : integer = pending; 
+  const interfaces__unsigned_32__last : integer = pending; 
+  const interfaces__unsigned_32__modulus : integer = pending; 
+  const interfaces__unsigned_32__size : integer = pending; 
+  var j : integer;
+  function k_r_spec(integer) : integer;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_r.rls	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,35 @@
+           /*********************************************************/
+                           /*Proof Rule Declarations*/
+    /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/
+             /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/
+           /*********************************************************/
+
+
+                        /*DATE : 29-NOV-2010 14:30:19.76*/
+
+                              /*function RMD.K_R*/
+
+
+rule_family k_r_rules:
+     X      requires [X:any] &
+     X <= Y requires [X:ire, Y:ire] &
+     X >= Y requires [X:ire, Y:ire].
+
+k_r_rules(1): interfaces__unsigned_32__size >= 0 may_be_deduced.
+k_r_rules(2): interfaces__unsigned_32__first may_be_replaced_by 0.
+k_r_rules(3): interfaces__unsigned_32__last may_be_replaced_by 4294967295.
+k_r_rules(4): interfaces__unsigned_32__base__first may_be_replaced_by 0.
+k_r_rules(5): interfaces__unsigned_32__base__last may_be_replaced_by 4294967295.
+k_r_rules(6): interfaces__unsigned_32__modulus may_be_replaced_by 4294967296.
+k_r_rules(7): word__size >= 0 may_be_deduced.
+k_r_rules(8): word__first may_be_replaced_by 0.
+k_r_rules(9): word__last may_be_replaced_by 4294967295.
+k_r_rules(10): word__base__first may_be_replaced_by 0.
+k_r_rules(11): word__base__last may_be_replaced_by 4294967295.
+k_r_rules(12): word__modulus may_be_replaced_by 4294967296.
+k_r_rules(13): round_index__size >= 0 may_be_deduced.
+k_r_rules(14): round_index__first may_be_replaced_by 0.
+k_r_rules(15): round_index__last may_be_replaced_by 79.
+k_r_rules(16): round_index__base__first <= round_index__base__last may_be_deduced.
+k_r_rules(17): round_index__base__first <= round_index__first may_be_deduced.
+k_r_rules(18): round_index__base__last >= round_index__last may_be_deduced.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/k_r.siv	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,118 @@
+*****************************************************************************
+                       Semantic Analysis of SPARK Text
+    Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+             Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+*****************************************************************************
+
+
+CREATED 29-NOV-2010, 14:30:19  SIMPLIFIED 29-NOV-2010, 14:30:21
+
+SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+
+function RMD.K_R
+
+
+
+
+For path(s) from start to run-time check associated with statement of line 38:
+
+function_k_r_1.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 39:
+
+function_k_r_2.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 40:
+
+function_k_r_3.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 41:
+
+function_k_r_4.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 42:
+
+function_k_r_5.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to finish:
+
+function_k_r_6.
+H1:    j >= 0 .
+H2:    j <= 15 .
+H3:    interfaces__unsigned_32__size >= 0 .
+H4:    word__size >= 0 .
+H5:    round_index__size >= 0 .
+H6:    round_index__base__first <= round_index__base__last .
+H7:    round_index__base__first <= 0 .
+H8:    round_index__base__last >= 79 .
+       ->
+C1:    1352829926 = k_r_spec(j) .
+
+
+function_k_r_7.
+H1:    16 <= j .
+H2:    j <= 31 .
+H3:    interfaces__unsigned_32__size >= 0 .
+H4:    word__size >= 0 .
+H5:    round_index__size >= 0 .
+H6:    round_index__base__first <= round_index__base__last .
+H7:    round_index__base__first <= 0 .
+H8:    round_index__base__last >= 79 .
+       ->
+C1:    1548603684 = k_r_spec(j) .
+
+
+function_k_r_8.
+H1:    32 <= j .
+H2:    j <= 47 .
+H3:    interfaces__unsigned_32__size >= 0 .
+H4:    word__size >= 0 .
+H5:    round_index__size >= 0 .
+H6:    round_index__base__first <= round_index__base__last .
+H7:    round_index__base__first <= 0 .
+H8:    round_index__base__last >= 79 .
+       ->
+C1:    1836072691 = k_r_spec(j) .
+
+
+function_k_r_9.
+H1:    48 <= j .
+H2:    j <= 63 .
+H3:    interfaces__unsigned_32__size >= 0 .
+H4:    word__size >= 0 .
+H5:    round_index__size >= 0 .
+H6:    round_index__base__first <= round_index__base__last .
+H7:    round_index__base__first <= 0 .
+H8:    round_index__base__last >= 79 .
+       ->
+C1:    2053994217 = k_r_spec(j) .
+
+
+function_k_r_10.
+H1:    j >= 0 .
+H2:    j <= 79 .
+H3:    15 < j .
+H4:    31 < j .
+H5:    47 < j .
+H6:    63 < j .
+H7:    interfaces__unsigned_32__size >= 0 .
+H8:    word__size >= 0 .
+H9:    round_index__size >= 0 .
+H10:   round_index__base__first <= round_index__base__last .
+H11:   round_index__base__first <= 0 .
+H12:   round_index__base__last >= 79 .
+       ->
+C1:    0 = k_r_spec(j) .
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_l.fdl	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,33 @@
+           {*******************************************************}
+                               {FDL Declarations}
+    {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039}
+             {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.}
+           {*******************************************************}
+
+
+                        {DATE : 29-NOV-2010 14:30:19.77}
+
+                              {function RMD.R_L}
+
+
+title function r_l;
+
+  function round__(real) : integer;
+  type block_index = integer;
+  type round_index = integer;
+  type block_permutation = array [integer] of integer;
+  const r_values : block_permutation = pending;
+  const round_index__base__first : integer = pending; 
+  const round_index__base__last : integer = pending; 
+  const block_index__base__first : integer = pending; 
+  const block_index__base__last : integer = pending; 
+  const round_index__first : integer = pending; 
+  const round_index__last : integer = pending; 
+  const round_index__size : integer = pending; 
+  const block_index__first : integer = pending; 
+  const block_index__last : integer = pending; 
+  const block_index__size : integer = pending; 
+  var j : integer;
+  function r_l_spec(integer) : integer;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_l.rls	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,75 @@
+           /*********************************************************/
+                           /*Proof Rule Declarations*/
+    /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/
+             /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/
+           /*********************************************************/
+
+
+                        /*DATE : 29-NOV-2010 14:30:19.77*/
+
+                              /*function RMD.R_L*/
+
+
+rule_family r_l_rules:
+     X      requires [X:any] &
+     X <= Y requires [X:ire, Y:ire] &
+     X >= Y requires [X:ire, Y:ire].
+
+r_l_rules(1): block_index__first <= element(r_values, [I]) may_be_deduced_from [0 <= I, I <= 79].
+r_l_rules(2): element(r_values, [I]) <= block_index__last may_be_deduced_from [0 <= I, I <= 79].
+r_l_rules(3): r_values may_be_replaced_by 
+           mk__block_permutation([round_index__first] := 0, [
+           round_index__first + 1] := 1, [round_index__first + 2] := 2, [
+           round_index__first + 3] := 3, [round_index__first + 4] := 4, [
+           round_index__first + 5] := 5, [round_index__first + 6] := 6, [
+           round_index__first + 7] := 7, [round_index__first + 8] := 8, [
+           round_index__first + 9] := 9, [round_index__first + 10] := 
+           10, [round_index__first + 11] := 11, [
+           round_index__first + 12] := 12, [round_index__first + 13] := 
+           13, [round_index__first + 14] := 14, [
+           round_index__first + 15] := 15, [round_index__first + 16] := 
+           7, [round_index__first + 17] := 4, [round_index__first + 18] := 
+           13, [round_index__first + 19] := 1, [round_index__first + 20] := 
+           10, [round_index__first + 21] := 6, [round_index__first + 22] := 
+           15, [round_index__first + 23] := 3, [round_index__first + 24] := 
+           12, [round_index__first + 25] := 0, [round_index__first + 26] := 
+           9, [round_index__first + 27] := 5, [round_index__first + 28] := 
+           2, [round_index__first + 29] := 14, [round_index__first + 30] := 
+           11, [round_index__first + 31] := 8, [round_index__first + 32] := 
+           3, [round_index__first + 33] := 10, [round_index__first + 34] := 
+           14, [round_index__first + 35] := 4, [round_index__first + 36] := 
+           9, [round_index__first + 37] := 15, [round_index__first + 38] := 
+           8, [round_index__first + 39] := 1, [round_index__first + 40] := 
+           2, [round_index__first + 41] := 7, [round_index__first + 42] := 
+           0, [round_index__first + 43] := 6, [round_index__first + 44] := 
+           13, [round_index__first + 45] := 11, [
+           round_index__first + 46] := 5, [round_index__first + 47] := 
+           12, [round_index__first + 48] := 1, [round_index__first + 49] := 
+           9, [round_index__first + 50] := 11, [round_index__first + 51] := 
+           10, [round_index__first + 52] := 0, [round_index__first + 53] := 
+           8, [round_index__first + 54] := 12, [round_index__first + 55] := 
+           4, [round_index__first + 56] := 13, [round_index__first + 57] := 
+           3, [round_index__first + 58] := 7, [round_index__first + 59] := 
+           15, [round_index__first + 60] := 14, [
+           round_index__first + 61] := 5, [round_index__first + 62] := 
+           6, [round_index__first + 63] := 2, [round_index__first + 64] := 
+           4, [round_index__first + 65] := 0, [round_index__first + 66] := 
+           5, [round_index__first + 67] := 9, [round_index__first + 68] := 
+           7, [round_index__first + 69] := 12, [round_index__first + 70] := 
+           2, [round_index__first + 71] := 10, [round_index__first + 72] := 
+           14, [round_index__first + 73] := 1, [round_index__first + 74] := 
+           3, [round_index__first + 75] := 8, [round_index__first + 76] := 
+           11, [round_index__first + 77] := 6, [round_index__first + 78] := 
+           15, [round_index__first + 79] := 13).
+r_l_rules(4): block_index__size >= 0 may_be_deduced.
+r_l_rules(5): block_index__first may_be_replaced_by 0.
+r_l_rules(6): block_index__last may_be_replaced_by 15.
+r_l_rules(7): block_index__base__first <= block_index__base__last may_be_deduced.
+r_l_rules(8): block_index__base__first <= block_index__first may_be_deduced.
+r_l_rules(9): block_index__base__last >= block_index__last may_be_deduced.
+r_l_rules(10): round_index__size >= 0 may_be_deduced.
+r_l_rules(11): round_index__first may_be_replaced_by 0.
+r_l_rules(12): round_index__last may_be_replaced_by 79.
+r_l_rules(13): round_index__base__first <= round_index__base__last may_be_deduced.
+r_l_rules(14): round_index__base__first <= round_index__first may_be_deduced.
+r_l_rules(15): round_index__base__last >= round_index__last may_be_deduced.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_l.siv	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,81 @@
+*****************************************************************************
+                       Semantic Analysis of SPARK Text
+    Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+             Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+*****************************************************************************
+
+
+CREATED 29-NOV-2010, 14:30:19  SIMPLIFIED 29-NOV-2010, 14:30:28
+
+SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+
+function RMD.R_L
+
+
+
+
+For path(s) from start to run-time check associated with statement of line 59:
+
+function_r_l_1.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to finish:
+
+function_r_l_2.
+H1:    j >= 0 .
+H2:    j <= 79 .
+H3:    block_index__size >= 0 .
+H4:    block_index__base__first <= block_index__base__last .
+H5:    round_index__size >= 0 .
+H6:    round_index__base__first <= round_index__base__last .
+H7:    block_index__base__first <= 0 .
+H8:    block_index__base__last >= 15 .
+H9:    round_index__base__first <= 0 .
+H10:   round_index__base__last >= 79 .
+       ->
+C1:    element(mk__block_permutation([0] := 0, [1] := 1, [2] := 2, [3] := 3, [4]
+           := 4, [5] := 5, [6] := 6, [7] := 7, [8] := 8, [9] := 9, [10] := 10, [
+          11] := 11, [12] := 12, [13] := 13, [14] := 14, [15] := 15, [16] := 7, 
+          [17] := 4, [18] := 13, [19] := 1, [20] := 10, [21] := 6, [22] := 15, [
+          23] := 3, [24] := 12, [25] := 0, [26] := 9, [27] := 5, [28] := 2, [29]
+           := 14, [30] := 11, [31] := 8, [32] := 3, [33] := 10, [34] := 14, [35]
+           := 4, [36] := 9, [37] := 15, [38] := 8, [39] := 1, [40] := 2, [41] 
+          := 7, [42] := 0, [43] := 6, [44] := 13, [45] := 11, [46] := 5, [47] 
+          := 12, [48] := 1, [49] := 9, [50] := 11, [51] := 10, [52] := 0, [53] 
+          := 8, [54] := 12, [55] := 4, [56] := 13, [57] := 3, [58] := 7, [59] 
+          := 15, [60] := 14, [61] := 5, [62] := 6, [63] := 2, [64] := 4, [65] 
+          := 0, [66] := 5, [67] := 9, [68] := 7, [69] := 12, [70] := 2, [71] := 
+          10, [72] := 14, [73] := 1, [74] := 3, [75] := 8, [76] := 11, [77] := 
+          6, [78] := 15, [79] := 13), [j]) = r_l_spec(j) .
+C2:    element(mk__block_permutation([0] := 0, [1] := 1, [2] := 2, [3] := 3, [4]
+           := 4, [5] := 5, [6] := 6, [7] := 7, [8] := 8, [9] := 9, [10] := 10, [
+          11] := 11, [12] := 12, [13] := 13, [14] := 14, [15] := 15, [16] := 7, 
+          [17] := 4, [18] := 13, [19] := 1, [20] := 10, [21] := 6, [22] := 15, [
+          23] := 3, [24] := 12, [25] := 0, [26] := 9, [27] := 5, [28] := 2, [29]
+           := 14, [30] := 11, [31] := 8, [32] := 3, [33] := 10, [34] := 14, [35]
+           := 4, [36] := 9, [37] := 15, [38] := 8, [39] := 1, [40] := 2, [41] 
+          := 7, [42] := 0, [43] := 6, [44] := 13, [45] := 11, [46] := 5, [47] 
+          := 12, [48] := 1, [49] := 9, [50] := 11, [51] := 10, [52] := 0, [53] 
+          := 8, [54] := 12, [55] := 4, [56] := 13, [57] := 3, [58] := 7, [59] 
+          := 15, [60] := 14, [61] := 5, [62] := 6, [63] := 2, [64] := 4, [65] 
+          := 0, [66] := 5, [67] := 9, [68] := 7, [69] := 12, [70] := 2, [71] := 
+          10, [72] := 14, [73] := 1, [74] := 3, [75] := 8, [76] := 11, [77] := 
+          6, [78] := 15, [79] := 13), [j]) >= 0 .
+C3:    element(mk__block_permutation([0] := 0, [1] := 1, [2] := 2, [3] := 3, [4]
+           := 4, [5] := 5, [6] := 6, [7] := 7, [8] := 8, [9] := 9, [10] := 10, [
+          11] := 11, [12] := 12, [13] := 13, [14] := 14, [15] := 15, [16] := 7, 
+          [17] := 4, [18] := 13, [19] := 1, [20] := 10, [21] := 6, [22] := 15, [
+          23] := 3, [24] := 12, [25] := 0, [26] := 9, [27] := 5, [28] := 2, [29]
+           := 14, [30] := 11, [31] := 8, [32] := 3, [33] := 10, [34] := 14, [35]
+           := 4, [36] := 9, [37] := 15, [38] := 8, [39] := 1, [40] := 2, [41] 
+          := 7, [42] := 0, [43] := 6, [44] := 13, [45] := 11, [46] := 5, [47] 
+          := 12, [48] := 1, [49] := 9, [50] := 11, [51] := 10, [52] := 0, [53] 
+          := 8, [54] := 12, [55] := 4, [56] := 13, [57] := 3, [58] := 7, [59] 
+          := 15, [60] := 14, [61] := 5, [62] := 6, [63] := 2, [64] := 4, [65] 
+          := 0, [66] := 5, [67] := 9, [68] := 7, [69] := 12, [70] := 2, [71] := 
+          10, [72] := 14, [73] := 1, [74] := 3, [75] := 8, [76] := 11, [77] := 
+          6, [78] := 15, [79] := 13), [j]) <= 15 .
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_r.fdl	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,33 @@
+           {*******************************************************}
+                               {FDL Declarations}
+    {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039}
+             {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.}
+           {*******************************************************}
+
+
+                        {DATE : 29-NOV-2010 14:30:19.81}
+
+                              {function RMD.R_R}
+
+
+title function r_r;
+
+  function round__(real) : integer;
+  type block_index = integer;
+  type round_index = integer;
+  type block_permutation = array [integer] of integer;
+  const r_values : block_permutation = pending;
+  const round_index__base__first : integer = pending; 
+  const round_index__base__last : integer = pending; 
+  const block_index__base__first : integer = pending; 
+  const block_index__base__last : integer = pending; 
+  const round_index__first : integer = pending; 
+  const round_index__last : integer = pending; 
+  const round_index__size : integer = pending; 
+  const block_index__first : integer = pending; 
+  const block_index__last : integer = pending; 
+  const block_index__size : integer = pending; 
+  var j : integer;
+  function r_r_spec(integer) : integer;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_r.rls	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,76 @@
+           /*********************************************************/
+                           /*Proof Rule Declarations*/
+    /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/
+             /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/
+           /*********************************************************/
+
+
+                        /*DATE : 29-NOV-2010 14:30:19.81*/
+
+                              /*function RMD.R_R*/
+
+
+rule_family r_r_rules:
+     X      requires [X:any] &
+     X <= Y requires [X:ire, Y:ire] &
+     X >= Y requires [X:ire, Y:ire].
+
+r_r_rules(1): block_index__first <= element(r_values, [I]) may_be_deduced_from [0 <= I, I <= 79].
+r_r_rules(2): element(r_values, [I]) <= block_index__last may_be_deduced_from [0 <= I, I <= 79].
+r_r_rules(3): r_values may_be_replaced_by 
+           mk__block_permutation([round_index__first] := 5, [
+           round_index__first + 1] := 14, [round_index__first + 2] := 7, [
+           round_index__first + 3] := 0, [round_index__first + 4] := 9, [
+           round_index__first + 5] := 2, [round_index__first + 6] := 11, [
+           round_index__first + 7] := 4, [round_index__first + 8] := 13, [
+           round_index__first + 9] := 6, [round_index__first + 10] := 
+           15, [round_index__first + 11] := 8, [round_index__first + 12] := 
+           1, [round_index__first + 13] := 10, [round_index__first + 14] := 
+           3, [round_index__first + 15] := 12, [round_index__first + 16] := 
+           6, [round_index__first + 17] := 11, [round_index__first + 18] := 
+           3, [round_index__first + 19] := 7, [round_index__first + 20] := 
+           0, [round_index__first + 21] := 13, [round_index__first + 22] := 
+           5, [round_index__first + 23] := 10, [round_index__first + 24] := 
+           14, [round_index__first + 25] := 15, [
+           round_index__first + 26] := 8, [round_index__first + 27] := 
+           12, [round_index__first + 28] := 4, [round_index__first + 29] := 
+           9, [round_index__first + 30] := 1, [round_index__first + 31] := 
+           2, [round_index__first + 32] := 15, [round_index__first + 33] := 
+           5, [round_index__first + 34] := 1, [round_index__first + 35] := 
+           3, [round_index__first + 36] := 7, [round_index__first + 37] := 
+           14, [round_index__first + 38] := 6, [round_index__first + 39] := 
+           9, [round_index__first + 40] := 11, [round_index__first + 41] := 
+           8, [round_index__first + 42] := 12, [round_index__first + 43] := 
+           2, [round_index__first + 44] := 10, [round_index__first + 45] := 
+           0, [round_index__first + 46] := 4, [round_index__first + 47] := 
+           13, [round_index__first + 48] := 8, [round_index__first + 49] := 
+           6, [round_index__first + 50] := 4, [round_index__first + 51] := 
+           1, [round_index__first + 52] := 3, [round_index__first + 53] := 
+           11, [round_index__first + 54] := 15, [
+           round_index__first + 55] := 0, [round_index__first + 56] := 
+           5, [round_index__first + 57] := 12, [round_index__first + 58] := 
+           2, [round_index__first + 59] := 13, [round_index__first + 60] := 
+           9, [round_index__first + 61] := 7, [round_index__first + 62] := 
+           10, [round_index__first + 63] := 14, [
+           round_index__first + 64] := 12, [round_index__first + 65] := 
+           15, [round_index__first + 66] := 10, [
+           round_index__first + 67] := 4, [round_index__first + 68] := 
+           1, [round_index__first + 69] := 5, [round_index__first + 70] := 
+           8, [round_index__first + 71] := 7, [round_index__first + 72] := 
+           6, [round_index__first + 73] := 2, [round_index__first + 74] := 
+           13, [round_index__first + 75] := 14, [
+           round_index__first + 76] := 0, [round_index__first + 77] := 
+           3, [round_index__first + 78] := 9, [round_index__first + 79] := 
+           11).
+r_r_rules(4): block_index__size >= 0 may_be_deduced.
+r_r_rules(5): block_index__first may_be_replaced_by 0.
+r_r_rules(6): block_index__last may_be_replaced_by 15.
+r_r_rules(7): block_index__base__first <= block_index__base__last may_be_deduced.
+r_r_rules(8): block_index__base__first <= block_index__first may_be_deduced.
+r_r_rules(9): block_index__base__last >= block_index__last may_be_deduced.
+r_r_rules(10): round_index__size >= 0 may_be_deduced.
+r_r_rules(11): round_index__first may_be_replaced_by 0.
+r_r_rules(12): round_index__last may_be_replaced_by 79.
+r_r_rules(13): round_index__base__first <= round_index__base__last may_be_deduced.
+r_r_rules(14): round_index__base__first <= round_index__first may_be_deduced.
+r_r_rules(15): round_index__base__last >= round_index__last may_be_deduced.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/r_r.siv	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,81 @@
+*****************************************************************************
+                       Semantic Analysis of SPARK Text
+    Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+             Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+*****************************************************************************
+
+
+CREATED 29-NOV-2010, 14:30:19  SIMPLIFIED 29-NOV-2010, 14:30:21
+
+SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+
+function RMD.R_R
+
+
+
+
+For path(s) from start to run-time check associated with statement of line 73:
+
+function_r_r_1.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to finish:
+
+function_r_r_2.
+H1:    j >= 0 .
+H2:    j <= 79 .
+H3:    block_index__size >= 0 .
+H4:    block_index__base__first <= block_index__base__last .
+H5:    round_index__size >= 0 .
+H6:    round_index__base__first <= round_index__base__last .
+H7:    block_index__base__first <= 0 .
+H8:    block_index__base__last >= 15 .
+H9:    round_index__base__first <= 0 .
+H10:   round_index__base__last >= 79 .
+       ->
+C1:    element(mk__block_permutation([0] := 5, [1] := 14, [2] := 7, [3] := 0, [
+          4] := 9, [5] := 2, [6] := 11, [7] := 4, [8] := 13, [9] := 6, [10] := 
+          15, [11] := 8, [12] := 1, [13] := 10, [14] := 3, [15] := 12, [16] := 
+          6, [17] := 11, [18] := 3, [19] := 7, [20] := 0, [21] := 13, [22] := 
+          5, [23] := 10, [24] := 14, [25] := 15, [26] := 8, [27] := 12, [28] := 
+          4, [29] := 9, [30] := 1, [31] := 2, [32] := 15, [33] := 5, [34] := 1, 
+          [35] := 3, [36] := 7, [37] := 14, [38] := 6, [39] := 9, [40] := 11, [
+          41] := 8, [42] := 12, [43] := 2, [44] := 10, [45] := 0, [46] := 4, [
+          47] := 13, [48] := 8, [49] := 6, [50] := 4, [51] := 1, [52] := 3, [53]
+           := 11, [54] := 15, [55] := 0, [56] := 5, [57] := 12, [58] := 2, [59] 
+          := 13, [60] := 9, [61] := 7, [62] := 10, [63] := 14, [64] := 12, [65] 
+          := 15, [66] := 10, [67] := 4, [68] := 1, [69] := 5, [70] := 8, [71] 
+          := 7, [72] := 6, [73] := 2, [74] := 13, [75] := 14, [76] := 0, [77] 
+          := 3, [78] := 9, [79] := 11), [j]) = r_r_spec(j) .
+C2:    element(mk__block_permutation([0] := 5, [1] := 14, [2] := 7, [3] := 0, [
+          4] := 9, [5] := 2, [6] := 11, [7] := 4, [8] := 13, [9] := 6, [10] := 
+          15, [11] := 8, [12] := 1, [13] := 10, [14] := 3, [15] := 12, [16] := 
+          6, [17] := 11, [18] := 3, [19] := 7, [20] := 0, [21] := 13, [22] := 
+          5, [23] := 10, [24] := 14, [25] := 15, [26] := 8, [27] := 12, [28] := 
+          4, [29] := 9, [30] := 1, [31] := 2, [32] := 15, [33] := 5, [34] := 1, 
+          [35] := 3, [36] := 7, [37] := 14, [38] := 6, [39] := 9, [40] := 11, [
+          41] := 8, [42] := 12, [43] := 2, [44] := 10, [45] := 0, [46] := 4, [
+          47] := 13, [48] := 8, [49] := 6, [50] := 4, [51] := 1, [52] := 3, [53]
+           := 11, [54] := 15, [55] := 0, [56] := 5, [57] := 12, [58] := 2, [59] 
+          := 13, [60] := 9, [61] := 7, [62] := 10, [63] := 14, [64] := 12, [65] 
+          := 15, [66] := 10, [67] := 4, [68] := 1, [69] := 5, [70] := 8, [71] 
+          := 7, [72] := 6, [73] := 2, [74] := 13, [75] := 14, [76] := 0, [77] 
+          := 3, [78] := 9, [79] := 11), [j]) >= 0 .
+C3:    element(mk__block_permutation([0] := 5, [1] := 14, [2] := 7, [3] := 0, [
+          4] := 9, [5] := 2, [6] := 11, [7] := 4, [8] := 13, [9] := 6, [10] := 
+          15, [11] := 8, [12] := 1, [13] := 10, [14] := 3, [15] := 12, [16] := 
+          6, [17] := 11, [18] := 3, [19] := 7, [20] := 0, [21] := 13, [22] := 
+          5, [23] := 10, [24] := 14, [25] := 15, [26] := 8, [27] := 12, [28] := 
+          4, [29] := 9, [30] := 1, [31] := 2, [32] := 15, [33] := 5, [34] := 1, 
+          [35] := 3, [36] := 7, [37] := 14, [38] := 6, [39] := 9, [40] := 11, [
+          41] := 8, [42] := 12, [43] := 2, [44] := 10, [45] := 0, [46] := 4, [
+          47] := 13, [48] := 8, [49] := 6, [50] := 4, [51] := 1, [52] := 3, [53]
+           := 11, [54] := 15, [55] := 0, [56] := 5, [57] := 12, [58] := 2, [59] 
+          := 13, [60] := 9, [61] := 7, [62] := 10, [63] := 14, [64] := 12, [65] 
+          := 15, [66] := 10, [67] := 4, [68] := 1, [69] := 5, [70] := 8, [71] 
+          := 7, [72] := 6, [73] := 2, [74] := 13, [75] := 14, [76] := 0, [77] 
+          := 3, [78] := 9, [79] := 11), [j]) <= 15 .
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/round.fdl	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,112 @@
+           {*******************************************************}
+                               {FDL Declarations}
+    {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039}
+             {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.}
+           {*******************************************************}
+
+
+                        {DATE : 29-NOV-2010 14:30:19.87}
+
+                            {procedure RMD.Round}
+
+
+title procedure round;
+
+  function round__(real) : integer;
+  type interfaces__unsigned_32 = integer;
+  type block_index = integer;
+  type round_index = integer;
+  type chain = record
+        h0 : integer;
+        h1 : integer;
+        h2 : integer;
+        h3 : integer;
+        h4 : integer
+     end;
+  type block = array [integer] of integer;
+  type chain_pair = record
+        left : chain;
+        right : chain
+     end;
+  const rotate_amount__base__first : integer = pending; 
+  const rotate_amount__base__last : integer = pending; 
+  const round_index__base__first : integer = pending; 
+  const round_index__base__last : integer = pending; 
+  const block_index__base__first : integer = pending; 
+  const block_index__base__last : integer = pending; 
+  const word__base__first : integer = pending; 
+  const word__base__last : integer = pending; 
+  const wordops__rotate_amount__base__first : integer = pending; 
+  const wordops__rotate_amount__base__last : integer = pending; 
+  const wordops__word__base__first : integer = pending; 
+  const wordops__word__base__last : integer = pending; 
+  const interfaces__unsigned_32__base__first : integer = pending; 
+  const interfaces__unsigned_32__base__last : integer = pending; 
+  const integer__base__first : integer = pending; 
+  const integer__base__last : integer = pending; 
+  const rotate_amount__first : integer = pending; 
+  const rotate_amount__last : integer = pending; 
+  const rotate_amount__size : integer = pending; 
+  const chain_pair__size : integer = pending; 
+  const round_index__first : integer = pending; 
+  const round_index__last : integer = pending; 
+  const round_index__size : integer = pending; 
+  const block_index__first : integer = pending; 
+  const block_index__last : integer = pending; 
+  const block_index__size : integer = pending; 
+  const chain__size : integer = pending; 
+  const word__first : integer = pending; 
+  const word__last : integer = pending; 
+  const word__modulus : integer = pending; 
+  const word__size : integer = pending; 
+  const wordops__rotate_amount__first : integer = pending; 
+  const wordops__rotate_amount__last : integer = pending; 
+  const wordops__rotate_amount__size : integer = pending; 
+  const wordops__word__first : integer = pending; 
+  const wordops__word__last : integer = pending; 
+  const wordops__word__modulus : integer = pending; 
+  const wordops__word__size : integer = pending; 
+  const interfaces__unsigned_32__first : integer = pending; 
+  const interfaces__unsigned_32__last : integer = pending; 
+  const interfaces__unsigned_32__modulus : integer = pending; 
+  const interfaces__unsigned_32__size : integer = pending; 
+  const integer__first : integer = pending; 
+  const integer__last : integer = pending; 
+  const integer__size : integer = pending; 
+  var ca : integer;
+  var cb : integer;
+  var cc : integer;
+  var cd : integer;
+  var ce : integer;
+  var x : block;
+  var cla : integer;
+  var clb : integer;
+  var clc : integer;
+  var cld : integer;
+  var cle : integer;
+  var cra : integer;
+  var crb : integer;
+  var crc : integer;
+  var crd : integer;
+  var cre : integer;
+  var loop__1__j : integer;
+  function wordops__rotate_left(integer, integer) : integer;
+  function wordops__rotate(integer, integer) : integer;
+  function f_spec(integer, integer, integer, integer) : integer;
+  function k_l_spec(integer) : integer;
+  function k_r_spec(integer) : integer;
+  function r_l_spec(integer) : integer;
+  function r_r_spec(integer) : integer;
+  function s_l_spec(integer) : integer;
+  function s_r_spec(integer) : integer;
+  function steps(chain_pair, integer, block) : chain_pair;
+  function round_spec(chain, block) : chain;
+  function f(integer, integer, integer, integer) : integer;
+  function k_l(integer) : integer;
+  function k_r(integer) : integer;
+  function r_l(integer) : integer;
+  function r_r(integer) : integer;
+  function s_l(integer) : integer;
+  function s_r(integer) : integer;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/round.rls	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,77 @@
+           /*********************************************************/
+                           /*Proof Rule Declarations*/
+    /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/
+             /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/
+           /*********************************************************/
+
+
+                        /*DATE : 29-NOV-2010 14:30:19.87*/
+
+                            /*procedure RMD.Round*/
+
+
+rule_family round_rules:
+     X      requires [X:any] &
+     X <= Y requires [X:ire, Y:ire] &
+     X >= Y requires [X:ire, Y:ire].
+
+round_rules(1): integer__size >= 0 may_be_deduced.
+round_rules(2): integer__first may_be_replaced_by -2147483648.
+round_rules(3): integer__last may_be_replaced_by 2147483647.
+round_rules(4): integer__base__first may_be_replaced_by -2147483648.
+round_rules(5): integer__base__last may_be_replaced_by 2147483647.
+round_rules(6): interfaces__unsigned_32__size >= 0 may_be_deduced.
+round_rules(7): interfaces__unsigned_32__first may_be_replaced_by 0.
+round_rules(8): interfaces__unsigned_32__last may_be_replaced_by 4294967295.
+round_rules(9): interfaces__unsigned_32__base__first may_be_replaced_by 0.
+round_rules(10): interfaces__unsigned_32__base__last may_be_replaced_by 4294967295.
+round_rules(11): interfaces__unsigned_32__modulus may_be_replaced_by 4294967296.
+round_rules(12): wordops__word__size >= 0 may_be_deduced.
+round_rules(13): wordops__word__first may_be_replaced_by 0.
+round_rules(14): wordops__word__last may_be_replaced_by 4294967295.
+round_rules(15): wordops__word__base__first may_be_replaced_by 0.
+round_rules(16): wordops__word__base__last may_be_replaced_by 4294967295.
+round_rules(17): wordops__word__modulus may_be_replaced_by 4294967296.
+round_rules(18): wordops__rotate_amount__size >= 0 may_be_deduced.
+round_rules(19): wordops__rotate_amount__first may_be_replaced_by 0.
+round_rules(20): wordops__rotate_amount__last may_be_replaced_by 15.
+round_rules(21): wordops__rotate_amount__base__first may_be_replaced_by -2147483648.
+round_rules(22): wordops__rotate_amount__base__last may_be_replaced_by 2147483647.
+round_rules(23): word__size >= 0 may_be_deduced.
+round_rules(24): word__first may_be_replaced_by 0.
+round_rules(25): word__last may_be_replaced_by 4294967295.
+round_rules(26): word__base__first may_be_replaced_by 0.
+round_rules(27): word__base__last may_be_replaced_by 4294967295.
+round_rules(28): word__modulus may_be_replaced_by 4294967296.
+round_rules(29): chain__size >= 0 may_be_deduced.
+round_rules(30): A = B may_be_deduced_from
+     [goal(checktype(A,chain)),
+      goal(checktype(B,chain)),
+      fld_h0(A) = fld_h0(B),
+      fld_h1(A) = fld_h1(B),
+      fld_h2(A) = fld_h2(B),
+      fld_h3(A) = fld_h3(B),
+      fld_h4(A) = fld_h4(B)].
+round_rules(31): block_index__size >= 0 may_be_deduced.
+round_rules(32): block_index__first may_be_replaced_by 0.
+round_rules(33): block_index__last may_be_replaced_by 15.
+round_rules(34): block_index__base__first <= block_index__base__last may_be_deduced.
+round_rules(35): block_index__base__first <= block_index__first may_be_deduced.
+round_rules(36): block_index__base__last >= block_index__last may_be_deduced.
+round_rules(37): round_index__size >= 0 may_be_deduced.
+round_rules(38): round_index__first may_be_replaced_by 0.
+round_rules(39): round_index__last may_be_replaced_by 79.
+round_rules(40): round_index__base__first <= round_index__base__last may_be_deduced.
+round_rules(41): round_index__base__first <= round_index__first may_be_deduced.
+round_rules(42): round_index__base__last >= round_index__last may_be_deduced.
+round_rules(43): chain_pair__size >= 0 may_be_deduced.
+round_rules(44): A = B may_be_deduced_from
+     [goal(checktype(A,chain_pair)),
+      goal(checktype(B,chain_pair)),
+      fld_left(A) = fld_left(B),
+      fld_right(A) = fld_right(B)].
+round_rules(45): rotate_amount__size >= 0 may_be_deduced.
+round_rules(46): rotate_amount__first may_be_replaced_by 0.
+round_rules(47): rotate_amount__last may_be_replaced_by 15.
+round_rules(48): rotate_amount__base__first may_be_replaced_by -2147483648.
+round_rules(49): rotate_amount__base__last may_be_replaced_by 2147483647.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/round.siv	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,834 @@
+*****************************************************************************
+                       Semantic Analysis of SPARK Text
+    Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+             Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+*****************************************************************************
+
+
+CREATED 29-NOV-2010, 14:30:19  SIMPLIFIED 29-NOV-2010, 14:30:21
+
+SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+
+procedure RMD.Round
+
+
+
+
+For path(s) from start to run-time check associated with statement of line 111:
+
+procedure_round_1.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 112:
+
+procedure_round_2.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 113:
+
+procedure_round_3.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 114:
+
+procedure_round_4.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 115:
+
+procedure_round_5.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 116:
+
+procedure_round_6.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 117:
+
+procedure_round_7.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 118:
+
+procedure_round_8.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 119:
+
+procedure_round_9.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 120:
+
+procedure_round_10.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 121:
+
+procedure_round_11.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 121:
+
+procedure_round_12.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 124:
+
+procedure_round_13.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 124:
+
+procedure_round_14.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 124:
+
+procedure_round_15.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 124:
+
+procedure_round_16.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 124:
+
+procedure_round_17.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 124:
+
+procedure_round_18.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 124:
+
+procedure_round_19.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 124:
+
+procedure_round_20.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 124:
+
+procedure_round_21.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 124:
+
+procedure_round_22.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 124:
+
+procedure_round_23.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 124:
+
+procedure_round_24.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 130:
+
+procedure_round_25.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 130:
+
+procedure_round_26.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 131:
+
+procedure_round_27.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 131:
+
+procedure_round_28.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 132:
+
+procedure_round_29.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 132:
+
+procedure_round_30.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 132:
+
+procedure_round_31.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 132:
+
+procedure_round_32.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 133:
+
+procedure_round_33.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 133:
+
+procedure_round_34.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 134:
+
+procedure_round_35.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 134:
+
+procedure_round_36.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 136:
+
+procedure_round_37.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 136:
+
+procedure_round_38.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 136:
+
+procedure_round_39.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 136:
+
+procedure_round_40.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 136:
+
+procedure_round_41.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 136:
+
+procedure_round_42.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 136:
+
+procedure_round_43.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 136:
+
+procedure_round_44.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 136:
+
+procedure_round_45.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 136:
+
+procedure_round_46.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 136:
+
+procedure_round_47.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 136:
+
+procedure_round_48.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 142:
+
+procedure_round_49.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 142:
+
+procedure_round_50.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 143:
+
+procedure_round_51.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 143:
+
+procedure_round_52.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 144:
+
+procedure_round_53.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 144:
+
+procedure_round_54.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 144:
+
+procedure_round_55.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 144:
+
+procedure_round_56.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 145:
+
+procedure_round_57.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 145:
+
+procedure_round_58.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 146:
+
+procedure_round_59.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 146:
+
+procedure_round_60.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to assertion of line 147:
+
+procedure_round_61.
+H1:    ca >= 0 .
+H2:    ca <= 4294967295 .
+H3:    cb >= 0 .
+H4:    cb <= 4294967295 .
+H5:    cc >= 0 .
+H6:    cc <= 4294967295 .
+H7:    cd >= 0 .
+H8:    cd <= 4294967295 .
+H9:    ce >= 0 .
+H10:   ce <= 4294967295 .
+H11:   for_all(i___1 : integer, 0 <= i___1 and i___1 <= 15 -> 0 <= element(x, [
+          i___1]) and element(x, [i___1]) <= 4294967295) .
+H12:   s_l(0) >= 0 .
+H13:   s_l(0) <= 15 .
+H14:   s_l(0) = s_l_spec(0) .
+H15:   f(0, cb, cc, cd) >= 0 .
+H16:   f(0, cb, cc, cd) <= 4294967295 .
+H17:   f(0, cb, cc, cd) = f_spec(0, cb, cc, cd) .
+H18:   r_l(0) >= 0 .
+H19:   r_l(0) <= 15 .
+H20:   r_l(0) = r_l_spec(0) .
+H21:   k_l(0) >= 0 .
+H22:   k_l(0) <= 4294967295 .
+H23:   k_l(0) = k_l_spec(0) .
+H24:   (((ca + f(0, cb, cc, cd)) mod 4294967296 + element(x, [r_l(0)])) mod 
+          4294967296 + k_l(0)) mod 4294967296 >= 0 .
+H25:   (((ca + f(0, cb, cc, cd)) mod 4294967296 + element(x, [r_l(0)])) mod 
+          4294967296 + k_l(0)) mod 4294967296 <= 4294967295 .
+H26:   wordops__rotate(s_l(0), (((ca + f(0, cb, cc, cd)) mod 4294967296 + 
+          element(x, [r_l(0)])) mod 4294967296 + k_l(0)) mod 4294967296) >= 0 .
+H27:   wordops__rotate(s_l(0), (((ca + f(0, cb, cc, cd)) mod 4294967296 + 
+          element(x, [r_l(0)])) mod 4294967296 + k_l(0)) mod 4294967296) <= 
+          4294967295 .
+H28:   wordops__rotate(s_l(0), (((ca + f(0, cb, cc, cd)) mod 4294967296 + 
+          element(x, [r_l(0)])) mod 4294967296 + k_l(0)) mod 4294967296) = 
+          wordops__rotate_left(s_l(0), (((ca + f(0, cb, cc, cd)) mod 4294967296 
+          + element(x, [r_l(0)])) mod 4294967296 + k_l(0)) mod 4294967296) .
+H29:   (wordops__rotate(s_l(0), (((ca + f(0, cb, cc, cd)) mod 4294967296 + 
+          element(x, [r_l(0)])) mod 4294967296 + k_l(0)) mod 4294967296) + ce) 
+          mod 4294967296 >= 0 .
+H30:   (wordops__rotate(s_l(0), (((ca + f(0, cb, cc, cd)) mod 4294967296 + 
+          element(x, [r_l(0)])) mod 4294967296 + k_l(0)) mod 4294967296) + ce) 
+          mod 4294967296 <= 4294967295 .
+H31:   wordops__rotate(10, cc) >= 0 .
+H32:   wordops__rotate(10, cc) <= 4294967295 .
+H33:   wordops__rotate(10, cc) = wordops__rotate_left(10, cc) .
+H34:   s_r(0) >= 0 .
+H35:   s_r(0) <= 15 .
+H36:   s_r(0) = s_r_spec(0) .
+H37:   79 >= round_index__base__first .
+H38:   79 <= round_index__base__last .
+H39:   f(79, cb, cc, cd) >= 0 .
+H40:   f(79, cb, cc, cd) <= 4294967295 .
+H41:   f(79, cb, cc, cd) = f_spec(79, cb, cc, cd) .
+H42:   r_r(0) >= 0 .
+H43:   r_r(0) <= 15 .
+H44:   r_r(0) = r_r_spec(0) .
+H45:   k_r(0) >= 0 .
+H46:   k_r(0) <= 4294967295 .
+H47:   k_r(0) = k_r_spec(0) .
+H48:   (((ca + f(79, cb, cc, cd)) mod 4294967296 + element(x, [r_r(0)])) mod 
+          4294967296 + k_r(0)) mod 4294967296 >= 0 .
+H49:   (((ca + f(79, cb, cc, cd)) mod 4294967296 + element(x, [r_r(0)])) mod 
+          4294967296 + k_r(0)) mod 4294967296 <= 4294967295 .
+H50:   wordops__rotate(s_r(0), (((ca + f(79, cb, cc, cd)) mod 4294967296 + 
+          element(x, [r_r(0)])) mod 4294967296 + k_r(0)) mod 4294967296) >= 0 .
+H51:   wordops__rotate(s_r(0), (((ca + f(79, cb, cc, cd)) mod 4294967296 + 
+          element(x, [r_r(0)])) mod 4294967296 + k_r(0)) mod 4294967296) <= 
+          4294967295 .
+H52:   wordops__rotate(s_r(0), (((ca + f(79, cb, cc, cd)) mod 4294967296 + 
+          element(x, [r_r(0)])) mod 4294967296 + k_r(0)) mod 4294967296) = 
+          wordops__rotate_left(s_r(0), (((ca + f(79, cb, cc, cd)) mod 
+          4294967296 + element(x, [r_r(0)])) mod 4294967296 + k_r(0)) mod 
+          4294967296) .
+H53:   (wordops__rotate(s_r(0), (((ca + f(79, cb, cc, cd)) mod 4294967296 + 
+          element(x, [r_r(0)])) mod 4294967296 + k_r(0)) mod 4294967296) + ce) 
+          mod 4294967296 >= 0 .
+H54:   (wordops__rotate(s_r(0), (((ca + f(79, cb, cc, cd)) mod 4294967296 + 
+          element(x, [r_r(0)])) mod 4294967296 + k_r(0)) mod 4294967296) + ce) 
+          mod 4294967296 <= 4294967295 .
+H55:   integer__size >= 0 .
+H56:   interfaces__unsigned_32__size >= 0 .
+H57:   wordops__word__size >= 0 .
+H58:   wordops__rotate_amount__size >= 0 .
+H59:   word__size >= 0 .
+H60:   chain__size >= 0 .
+H61:   block_index__size >= 0 .
+H62:   block_index__base__first <= block_index__base__last .
+H63:   round_index__size >= 0 .
+H64:   round_index__base__first <= round_index__base__last .
+H65:   chain_pair__size >= 0 .
+H66:   rotate_amount__size >= 0 .
+H67:   block_index__base__first <= 0 .
+H68:   block_index__base__last >= 15 .
+H69:   round_index__base__first <= 0 .
+H70:   round_index__base__last >= 79 .
+       ->
+C1:    mk__chain_pair(left := mk__chain(h0 := ce, h1 := (wordops__rotate(s_l(0)
+          , (((ca + f(0, cb, cc, cd)) mod 4294967296 + element(x, [r_l(0)])) 
+          mod 4294967296 + k_l(0)) mod 4294967296) + ce) mod 4294967296, h2 := 
+          cb, h3 := wordops__rotate(10, cc), h4 := cd), right := mk__chain(h0 
+          := ce, h1 := (wordops__rotate(s_r(0), (((ca + f(79, cb, cc, cd)) mod 
+          4294967296 + element(x, [r_r(0)])) mod 4294967296 + k_r(0)) mod 
+          4294967296) + ce) mod 4294967296, h2 := cb, h3 := wordops__rotate(10, 
+          cc), h4 := cd)) = steps(mk__chain_pair(left := mk__chain(h0 := ca, h1 
+          := cb, h2 := cc, h3 := cd, h4 := ce), right := mk__chain(h0 := ca, h1 
+          := cb, h2 := cc, h3 := cd, h4 := ce)), 1, x) .
+
+
+For path(s) from assertion of line 147 to assertion of line 147:
+
+procedure_round_62.
+H1:    mk__chain_pair(left := mk__chain(h0 := cla, h1 := clb, h2 := clc, h3 := 
+          cld, h4 := cle), right := mk__chain(h0 := cra, h1 := crb, h2 := crc, 
+          h3 := crd, h4 := cre)) = steps(mk__chain_pair(left := mk__chain(h0 := 
+          ca~, h1 := cb~, h2 := cc~, h3 := cd~, h4 := ce~), right := mk__chain(
+          h0 := ca~, h1 := cb~, h2 := cc~, h3 := cd~, h4 := ce~)), loop__1__j + 
+          1, x) .
+H2:    ca~ >= 0 .
+H3:    ca~ <= 4294967295 .
+H4:    cb~ >= 0 .
+H5:    cb~ <= 4294967295 .
+H6:    cc~ >= 0 .
+H7:    cc~ <= 4294967295 .
+H8:    cd~ >= 0 .
+H9:    cd~ <= 4294967295 .
+H10:   ce~ >= 0 .
+H11:   ce~ <= 4294967295 .
+H12:   for_all(i___1 : integer, 0 <= i___1 and i___1 <= 15 -> 0 <= element(x, [
+          i___1]) and element(x, [i___1]) <= 4294967295) .
+H13:   loop__1__j >= 0 .
+H14:   loop__1__j <= 78 .
+H15:   s_l(loop__1__j + 1) >= 0 .
+H16:   s_l(loop__1__j + 1) <= 15 .
+H17:   s_l(loop__1__j + 1) = s_l_spec(loop__1__j + 1) .
+H18:   cla >= 0 .
+H19:   cla <= 4294967295 .
+H20:   clb >= 0 .
+H21:   clb <= 4294967295 .
+H22:   clc >= 0 .
+H23:   clc <= 4294967295 .
+H24:   cld >= 0 .
+H25:   cld <= 4294967295 .
+H26:   f(loop__1__j + 1, clb, clc, cld) >= 0 .
+H27:   f(loop__1__j + 1, clb, clc, cld) <= 4294967295 .
+H28:   f(loop__1__j + 1, clb, clc, cld) = f_spec(loop__1__j + 1, clb, clc, cld) 
+          .
+H29:   r_l(loop__1__j + 1) >= 0 .
+H30:   r_l(loop__1__j + 1) <= 15 .
+H31:   r_l(loop__1__j + 1) = r_l_spec(loop__1__j + 1) .
+H32:   k_l(loop__1__j + 1) >= 0 .
+H33:   k_l(loop__1__j + 1) <= 4294967295 .
+H34:   k_l(loop__1__j + 1) = k_l_spec(loop__1__j + 1) .
+H35:   (((cla + f(loop__1__j + 1, clb, clc, cld)) mod 4294967296 + element(x, [
+          r_l(loop__1__j + 1)])) mod 4294967296 + k_l(loop__1__j + 1)) mod 
+          4294967296 >= 0 .
+H36:   (((cla + f(loop__1__j + 1, clb, clc, cld)) mod 4294967296 + element(x, [
+          r_l(loop__1__j + 1)])) mod 4294967296 + k_l(loop__1__j + 1)) mod 
+          4294967296 <= 4294967295 .
+H37:   wordops__rotate(s_l(loop__1__j + 1), (((cla + f(loop__1__j + 1, clb, 
+          clc, cld)) mod 4294967296 + element(x, [r_l(loop__1__j + 1)])) mod 
+          4294967296 + k_l(loop__1__j + 1)) mod 4294967296) >= 0 .
+H38:   wordops__rotate(s_l(loop__1__j + 1), (((cla + f(loop__1__j + 1, clb, 
+          clc, cld)) mod 4294967296 + element(x, [r_l(loop__1__j + 1)])) mod 
+          4294967296 + k_l(loop__1__j + 1)) mod 4294967296) <= 4294967295 .
+H39:   wordops__rotate(s_l(loop__1__j + 1), (((cla + f(loop__1__j + 1, clb, 
+          clc, cld)) mod 4294967296 + element(x, [r_l(loop__1__j + 1)])) mod 
+          4294967296 + k_l(loop__1__j + 1)) mod 4294967296) = 
+          wordops__rotate_left(s_l(loop__1__j + 1), (((cla + f(loop__1__j + 1, 
+          clb, clc, cld)) mod 4294967296 + element(x, [r_l(loop__1__j + 1)])) 
+          mod 4294967296 + k_l(loop__1__j + 1)) mod 4294967296) .
+H40:   cle >= 0 .
+H41:   cle <= 4294967295 .
+H42:   (wordops__rotate(s_l(loop__1__j + 1), (((cla + f(loop__1__j + 1, clb, 
+          clc, cld)) mod 4294967296 + element(x, [r_l(loop__1__j + 1)])) mod 
+          4294967296 + k_l(loop__1__j + 1)) mod 4294967296) + cle) mod 
+          4294967296 >= 0 .
+H43:   (wordops__rotate(s_l(loop__1__j + 1), (((cla + f(loop__1__j + 1, clb, 
+          clc, cld)) mod 4294967296 + element(x, [r_l(loop__1__j + 1)])) mod 
+          4294967296 + k_l(loop__1__j + 1)) mod 4294967296) + cle) mod 
+          4294967296 <= 4294967295 .
+H44:   wordops__rotate(10, clc) >= 0 .
+H45:   wordops__rotate(10, clc) <= 4294967295 .
+H46:   wordops__rotate(10, clc) = wordops__rotate_left(10, clc) .
+H47:   s_r(loop__1__j + 1) >= 0 .
+H48:   s_r(loop__1__j + 1) <= 15 .
+H49:   s_r(loop__1__j + 1) = s_r_spec(loop__1__j + 1) .
+H50:   cra >= 0 .
+H51:   cra <= 4294967295 .
+H52:   crb >= 0 .
+H53:   crb <= 4294967295 .
+H54:   crc >= 0 .
+H55:   crc <= 4294967295 .
+H56:   crd >= 0 .
+H57:   crd <= 4294967295 .
+H58:   79 - (loop__1__j + 1) >= round_index__base__first .
+H59:   79 - (loop__1__j + 1) <= round_index__base__last .
+H60:   f(79 - (loop__1__j + 1), crb, crc, crd) >= 0 .
+H61:   f(79 - (loop__1__j + 1), crb, crc, crd) <= 4294967295 .
+H62:   f(78 - loop__1__j, crb, crc, crd) = f_spec(78 - loop__1__j, crb, crc, 
+          crd) .
+H63:   r_r(loop__1__j + 1) >= 0 .
+H64:   r_r(loop__1__j + 1) <= 15 .
+H65:   r_r(loop__1__j + 1) = r_r_spec(loop__1__j + 1) .
+H66:   k_r(loop__1__j + 1) >= 0 .
+H67:   k_r(loop__1__j + 1) <= 4294967295 .
+H68:   k_r(loop__1__j + 1) = k_r_spec(loop__1__j + 1) .
+H69:   (((cra + f(79 - (loop__1__j + 1), crb, crc, crd)) mod 4294967296 + 
+          element(x, [r_r(loop__1__j + 1)])) mod 4294967296 + k_r(loop__1__j + 
+          1)) mod 4294967296 >= 0 .
+H70:   (((cra + f(79 - (loop__1__j + 1), crb, crc, crd)) mod 4294967296 + 
+          element(x, [r_r(loop__1__j + 1)])) mod 4294967296 + k_r(loop__1__j + 
+          1)) mod 4294967296 <= 4294967295 .
+H71:   wordops__rotate(s_r(loop__1__j + 1), (((cra + f(79 - (loop__1__j + 1), 
+          crb, crc, crd)) mod 4294967296 + element(x, [r_r(loop__1__j + 1)])) 
+          mod 4294967296 + k_r(loop__1__j + 1)) mod 4294967296) >= 0 .
+H72:   wordops__rotate(s_r(loop__1__j + 1), (((cra + f(79 - (loop__1__j + 1), 
+          crb, crc, crd)) mod 4294967296 + element(x, [r_r(loop__1__j + 1)])) 
+          mod 4294967296 + k_r(loop__1__j + 1)) mod 4294967296) <= 4294967295 .
+H73:   wordops__rotate(s_r(loop__1__j + 1), (((cra + f(79 - (loop__1__j + 1), 
+          crb, crc, crd)) mod 4294967296 + element(x, [r_r(loop__1__j + 1)])) 
+          mod 4294967296 + k_r(loop__1__j + 1)) mod 4294967296) = 
+          wordops__rotate_left(s_r(loop__1__j + 1), (((cra + f(79 - (loop__1__j 
+          + 1), crb, crc, crd)) mod 4294967296 + element(x, [r_r(loop__1__j + 1)
+          ])) mod 4294967296 + k_r(loop__1__j + 1)) mod 4294967296) .
+H74:   cre >= 0 .
+H75:   cre <= 4294967295 .
+H76:   (wordops__rotate(s_r(loop__1__j + 1), (((cra + f(79 - (loop__1__j + 1), 
+          crb, crc, crd)) mod 4294967296 + element(x, [r_r(loop__1__j + 1)])) 
+          mod 4294967296 + k_r(loop__1__j + 1)) mod 4294967296) + cre) mod 
+          4294967296 >= 0 .
+H77:   (wordops__rotate(s_r(loop__1__j + 1), (((cra + f(79 - (loop__1__j + 1), 
+          crb, crc, crd)) mod 4294967296 + element(x, [r_r(loop__1__j + 1)])) 
+          mod 4294967296 + k_r(loop__1__j + 1)) mod 4294967296) + cre) mod 
+          4294967296 <= 4294967295 .
+H78:   wordops__rotate(10, crc) >= 0 .
+H79:   wordops__rotate(10, crc) <= 4294967295 .
+H80:   wordops__rotate(10, crc) = wordops__rotate_left(10, crc) .
+H81:   integer__size >= 0 .
+H82:   interfaces__unsigned_32__size >= 0 .
+H83:   wordops__word__size >= 0 .
+H84:   wordops__rotate_amount__size >= 0 .
+H85:   word__size >= 0 .
+H86:   chain__size >= 0 .
+H87:   block_index__size >= 0 .
+H88:   block_index__base__first <= block_index__base__last .
+H89:   round_index__size >= 0 .
+H90:   round_index__base__first <= round_index__base__last .
+H91:   chain_pair__size >= 0 .
+H92:   rotate_amount__size >= 0 .
+H93:   block_index__base__first <= 0 .
+H94:   block_index__base__last >= 15 .
+H95:   round_index__base__first <= 0 .
+H96:   round_index__base__last >= 79 .
+       ->
+C1:    mk__chain_pair(left := mk__chain(h0 := cle, h1 := (wordops__rotate(s_l(
+          loop__1__j + 1), (((cla + f(loop__1__j + 1, clb, clc, cld)) mod 
+          4294967296 + element(x, [r_l(loop__1__j + 1)])) mod 4294967296 + k_l(
+          loop__1__j + 1)) mod 4294967296) + cle) mod 4294967296, h2 := clb, h3 
+          := wordops__rotate(10, clc), h4 := cld), right := mk__chain(h0 := 
+          cre, h1 := (wordops__rotate(s_r(loop__1__j + 1), (((cra + f(79 - (
+          loop__1__j + 1), crb, crc, crd)) mod 4294967296 + element(x, [r_r(
+          loop__1__j + 1)])) mod 4294967296 + k_r(loop__1__j + 1)) mod 
+          4294967296) + cre) mod 4294967296, h2 := crb, h3 := wordops__rotate(
+          10, crc), h4 := crd)) = steps(mk__chain_pair(left := mk__chain(h0 := 
+          ca~, h1 := cb~, h2 := cc~, h3 := cd~, h4 := ce~), right := mk__chain(
+          h0 := ca~, h1 := cb~, h2 := cc~, h3 := cd~, h4 := ce~)), loop__1__j + 
+          2, x) .
+
+
+For path(s) from start to run-time check associated with statement of line 153:
+
+procedure_round_63.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 153:
+
+procedure_round_64.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 154:
+
+procedure_round_65.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 154:
+
+procedure_round_66.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 155:
+
+procedure_round_67.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 155:
+
+procedure_round_68.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 156:
+
+procedure_round_69.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 156:
+
+procedure_round_70.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 157:
+
+procedure_round_71.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 157:
+
+procedure_round_72.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to run-time check associated with statement of line 158:
+
+procedure_round_73.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 147 to run-time check associated with 
+          statement of line 158:
+
+procedure_round_74.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to finish:
+
+procedure_round_75.
+*** true .   /* contradiction within hypotheses. */
+
+
+
+For path(s) from assertion of line 147 to finish:
+
+procedure_round_76.
+H1:    mk__chain_pair(left := mk__chain(h0 := cla, h1 := clb, h2 := clc, h3 := 
+          cld, h4 := cle), right := mk__chain(h0 := cra, h1 := crb, h2 := crc, 
+          h3 := crd, h4 := cre)) = steps(mk__chain_pair(left := mk__chain(h0 := 
+          ca~, h1 := cb~, h2 := cc~, h3 := cd~, h4 := ce~), right := mk__chain(
+          h0 := ca~, h1 := cb~, h2 := cc~, h3 := cd~, h4 := ce~)), 80, x) .
+H2:    ca~ >= 0 .
+H3:    ca~ <= 4294967295 .
+H4:    cb~ >= 0 .
+H5:    cb~ <= 4294967295 .
+H6:    cc~ >= 0 .
+H7:    cc~ <= 4294967295 .
+H8:    cd~ >= 0 .
+H9:    cd~ <= 4294967295 .
+H10:   ce~ >= 0 .
+H11:   ce~ <= 4294967295 .
+H12:   for_all(i___1 : integer, 0 <= i___1 and i___1 <= 15 -> 0 <= element(x, [
+          i___1]) and element(x, [i___1]) <= 4294967295) .
+H13:   clc >= 0 .
+H14:   clc <= 4294967295 .
+H15:   crd >= 0 .
+H16:   crd <= 4294967295 .
+H17:   ((cb~ + clc) mod 4294967296 + crd) mod 4294967296 >= 0 .
+H18:   ((cb~ + clc) mod 4294967296 + crd) mod 4294967296 <= 4294967295 .
+H19:   cld >= 0 .
+H20:   cld <= 4294967295 .
+H21:   cre >= 0 .
+H22:   cre <= 4294967295 .
+H23:   ((cc~ + cld) mod 4294967296 + cre) mod 4294967296 >= 0 .
+H24:   ((cc~ + cld) mod 4294967296 + cre) mod 4294967296 <= 4294967295 .
+H25:   cle >= 0 .
+H26:   cle <= 4294967295 .
+H27:   cra >= 0 .
+H28:   cra <= 4294967295 .
+H29:   ((cd~ + cle) mod 4294967296 + cra) mod 4294967296 >= 0 .
+H30:   ((cd~ + cle) mod 4294967296 + cra) mod 4294967296 <= 4294967295 .
+H31:   cla >= 0 .
+H32:   cla <= 4294967295 .
+H33:   crb >= 0 .
+H34:   crb <= 4294967295 .
+H35:   ((ce~ + cla) mod 4294967296 + crb) mod 4294967296 >= 0 .
+H36:   ((ce~ + cla) mod 4294967296 + crb) mod 4294967296 <= 4294967295 .
+H37:   clb >= 0 .
+H38:   clb <= 4294967295 .
+H39:   crc >= 0 .
+H40:   crc <= 4294967295 .
+H41:   ((ca~ + clb) mod 4294967296 + crc) mod 4294967296 >= 0 .
+H42:   ((ca~ + clb) mod 4294967296 + crc) mod 4294967296 <= 4294967295 .
+H43:   integer__size >= 0 .
+H44:   interfaces__unsigned_32__size >= 0 .
+H45:   wordops__word__size >= 0 .
+H46:   wordops__rotate_amount__size >= 0 .
+H47:   word__size >= 0 .
+H48:   chain__size >= 0 .
+H49:   block_index__size >= 0 .
+H50:   block_index__base__first <= block_index__base__last .
+H51:   round_index__size >= 0 .
+H52:   round_index__base__first <= round_index__base__last .
+H53:   chain_pair__size >= 0 .
+H54:   rotate_amount__size >= 0 .
+H55:   block_index__base__first <= 0 .
+H56:   block_index__base__last >= 15 .
+H57:   round_index__base__first <= 0 .
+H58:   round_index__base__last >= 79 .
+       ->
+C1:    mk__chain(h0 := ((cb~ + clc) mod 4294967296 + crd) mod 4294967296, h1 := 
+          ((cc~ + cld) mod 4294967296 + cre) mod 4294967296, h2 := ((cd~ + cle) 
+          mod 4294967296 + cra) mod 4294967296, h3 := ((ce~ + cla) mod 
+          4294967296 + crb) mod 4294967296, h4 := ((ca~ + clb) mod 4294967296 + 
+          crc) mod 4294967296) = round_spec(mk__chain(h0 := ca~, h1 := cb~, h2 
+          := cc~, h3 := cd~, h4 := ce~), x) .
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_l.fdl	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,37 @@
+           {*******************************************************}
+                               {FDL Declarations}
+    {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039}
+             {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.}
+           {*******************************************************}
+
+
+                        {DATE : 29-NOV-2010 14:30:19.83}
+
+                              {function RMD.S_L}
+
+
+title function s_l;
+
+  function round__(real) : integer;
+  type round_index = integer;
+  type rotate_definition = array [integer] of integer;
+  const s_values : rotate_definition = pending;
+  const rotate_amount__base__first : integer = pending; 
+  const rotate_amount__base__last : integer = pending; 
+  const round_index__base__first : integer = pending; 
+  const round_index__base__last : integer = pending; 
+  const integer__base__first : integer = pending; 
+  const integer__base__last : integer = pending; 
+  const rotate_amount__first : integer = pending; 
+  const rotate_amount__last : integer = pending; 
+  const rotate_amount__size : integer = pending; 
+  const round_index__first : integer = pending; 
+  const round_index__last : integer = pending; 
+  const round_index__size : integer = pending; 
+  const integer__first : integer = pending; 
+  const integer__last : integer = pending; 
+  const integer__size : integer = pending; 
+  var j : integer;
+  function s_l_spec(integer) : integer;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_l.rls	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,81 @@
+           /*********************************************************/
+                           /*Proof Rule Declarations*/
+    /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/
+             /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/
+           /*********************************************************/
+
+
+                        /*DATE : 29-NOV-2010 14:30:19.83*/
+
+                              /*function RMD.S_L*/
+
+
+rule_family s_l_rules:
+     X      requires [X:any] &
+     X <= Y requires [X:ire, Y:ire] &
+     X >= Y requires [X:ire, Y:ire].
+
+s_l_rules(1): rotate_amount__first <= element(s_values, [I]) may_be_deduced_from [0 <= I, I <= 79].
+s_l_rules(2): element(s_values, [I]) <= rotate_amount__last may_be_deduced_from [0 <= I, I <= 79].
+s_l_rules(3): s_values may_be_replaced_by 
+           mk__rotate_definition([round_index__first] := 11, [
+           round_index__first + 1] := 14, [round_index__first + 2] := 
+           15, [round_index__first + 3] := 12, [round_index__first + 4] := 
+           5, [round_index__first + 5] := 8, [round_index__first + 6] := 
+           7, [round_index__first + 7] := 9, [round_index__first + 8] := 
+           11, [round_index__first + 9] := 13, [round_index__first + 10] := 
+           14, [round_index__first + 11] := 15, [
+           round_index__first + 12] := 6, [round_index__first + 13] := 
+           7, [round_index__first + 14] := 9, [round_index__first + 15] := 
+           8, [round_index__first + 16] := 7, [round_index__first + 17] := 
+           6, [round_index__first + 18] := 8, [round_index__first + 19] := 
+           13, [round_index__first + 20] := 11, [
+           round_index__first + 21] := 9, [round_index__first + 22] := 
+           7, [round_index__first + 23] := 15, [round_index__first + 24] := 
+           7, [round_index__first + 25] := 12, [round_index__first + 26] := 
+           15, [round_index__first + 27] := 9, [round_index__first + 28] := 
+           11, [round_index__first + 29] := 7, [round_index__first + 30] := 
+           13, [round_index__first + 31] := 12, [
+           round_index__first + 32] := 11, [round_index__first + 33] := 
+           13, [round_index__first + 34] := 6, [round_index__first + 35] := 
+           7, [round_index__first + 36] := 14, [round_index__first + 37] := 
+           9, [round_index__first + 38] := 13, [round_index__first + 39] := 
+           15, [round_index__first + 40] := 14, [
+           round_index__first + 41] := 8, [round_index__first + 42] := 
+           13, [round_index__first + 43] := 6, [round_index__first + 44] := 
+           5, [round_index__first + 45] := 12, [round_index__first + 46] := 
+           7, [round_index__first + 47] := 5, [round_index__first + 48] := 
+           11, [round_index__first + 49] := 12, [
+           round_index__first + 50] := 14, [round_index__first + 51] := 
+           15, [round_index__first + 52] := 14, [
+           round_index__first + 53] := 15, [round_index__first + 54] := 
+           9, [round_index__first + 55] := 8, [round_index__first + 56] := 
+           9, [round_index__first + 57] := 14, [round_index__first + 58] := 
+           5, [round_index__first + 59] := 6, [round_index__first + 60] := 
+           8, [round_index__first + 61] := 6, [round_index__first + 62] := 
+           5, [round_index__first + 63] := 12, [round_index__first + 64] := 
+           9, [round_index__first + 65] := 15, [round_index__first + 66] := 
+           5, [round_index__first + 67] := 11, [round_index__first + 68] := 
+           6, [round_index__first + 69] := 8, [round_index__first + 70] := 
+           13, [round_index__first + 71] := 12, [
+           round_index__first + 72] := 5, [round_index__first + 73] := 
+           12, [round_index__first + 74] := 13, [
+           round_index__first + 75] := 14, [round_index__first + 76] := 
+           11, [round_index__first + 77] := 8, [round_index__first + 78] := 
+           5, [round_index__first + 79] := 6).
+s_l_rules(4): integer__size >= 0 may_be_deduced.
+s_l_rules(5): integer__first may_be_replaced_by -2147483648.
+s_l_rules(6): integer__last may_be_replaced_by 2147483647.
+s_l_rules(7): integer__base__first may_be_replaced_by -2147483648.
+s_l_rules(8): integer__base__last may_be_replaced_by 2147483647.
+s_l_rules(9): round_index__size >= 0 may_be_deduced.
+s_l_rules(10): round_index__first may_be_replaced_by 0.
+s_l_rules(11): round_index__last may_be_replaced_by 79.
+s_l_rules(12): round_index__base__first <= round_index__base__last may_be_deduced.
+s_l_rules(13): round_index__base__first <= round_index__first may_be_deduced.
+s_l_rules(14): round_index__base__last >= round_index__last may_be_deduced.
+s_l_rules(15): rotate_amount__size >= 0 may_be_deduced.
+s_l_rules(16): rotate_amount__first may_be_replaced_by 0.
+s_l_rules(17): rotate_amount__last may_be_replaced_by 15.
+s_l_rules(18): rotate_amount__base__first may_be_replaced_by -2147483648.
+s_l_rules(19): rotate_amount__base__last may_be_replaced_by 2147483647.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_l.siv	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,79 @@
+*****************************************************************************
+                       Semantic Analysis of SPARK Text
+    Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+             Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+*****************************************************************************
+
+
+CREATED 29-NOV-2010, 14:30:19  SIMPLIFIED 29-NOV-2010, 14:30:29
+
+SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+
+function RMD.S_L
+
+
+
+
+For path(s) from start to run-time check associated with statement of line 87:
+
+function_s_l_1.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to finish:
+
+function_s_l_2.
+H1:    j >= 0 .
+H2:    j <= 79 .
+H3:    integer__size >= 0 .
+H4:    round_index__size >= 0 .
+H5:    round_index__base__first <= round_index__base__last .
+H6:    rotate_amount__size >= 0 .
+H7:    round_index__base__first <= 0 .
+H8:    round_index__base__last >= 79 .
+       ->
+C1:    element(mk__rotate_definition([0] := 11, [1] := 14, [2] := 15, [3] := 
+          12, [4] := 5, [5] := 8, [6] := 7, [7] := 9, [8] := 11, [9] := 13, [10]
+           := 14, [11] := 15, [12] := 6, [13] := 7, [14] := 9, [15] := 8, [16] 
+          := 7, [17] := 6, [18] := 8, [19] := 13, [20] := 11, [21] := 9, [22] 
+          := 7, [23] := 15, [24] := 7, [25] := 12, [26] := 15, [27] := 9, [28] 
+          := 11, [29] := 7, [30] := 13, [31] := 12, [32] := 11, [33] := 13, [34]
+           := 6, [35] := 7, [36] := 14, [37] := 9, [38] := 13, [39] := 15, [40] 
+          := 14, [41] := 8, [42] := 13, [43] := 6, [44] := 5, [45] := 12, [46] 
+          := 7, [47] := 5, [48] := 11, [49] := 12, [50] := 14, [51] := 15, [52] 
+          := 14, [53] := 15, [54] := 9, [55] := 8, [56] := 9, [57] := 14, [58] 
+          := 5, [59] := 6, [60] := 8, [61] := 6, [62] := 5, [63] := 12, [64] := 
+          9, [65] := 15, [66] := 5, [67] := 11, [68] := 6, [69] := 8, [70] := 
+          13, [71] := 12, [72] := 5, [73] := 12, [74] := 13, [75] := 14, [76] 
+          := 11, [77] := 8, [78] := 5, [79] := 6), [j]) = s_l_spec(j) .
+C2:    element(mk__rotate_definition([0] := 11, [1] := 14, [2] := 15, [3] := 
+          12, [4] := 5, [5] := 8, [6] := 7, [7] := 9, [8] := 11, [9] := 13, [10]
+           := 14, [11] := 15, [12] := 6, [13] := 7, [14] := 9, [15] := 8, [16] 
+          := 7, [17] := 6, [18] := 8, [19] := 13, [20] := 11, [21] := 9, [22] 
+          := 7, [23] := 15, [24] := 7, [25] := 12, [26] := 15, [27] := 9, [28] 
+          := 11, [29] := 7, [30] := 13, [31] := 12, [32] := 11, [33] := 13, [34]
+           := 6, [35] := 7, [36] := 14, [37] := 9, [38] := 13, [39] := 15, [40] 
+          := 14, [41] := 8, [42] := 13, [43] := 6, [44] := 5, [45] := 12, [46] 
+          := 7, [47] := 5, [48] := 11, [49] := 12, [50] := 14, [51] := 15, [52] 
+          := 14, [53] := 15, [54] := 9, [55] := 8, [56] := 9, [57] := 14, [58] 
+          := 5, [59] := 6, [60] := 8, [61] := 6, [62] := 5, [63] := 12, [64] := 
+          9, [65] := 15, [66] := 5, [67] := 11, [68] := 6, [69] := 8, [70] := 
+          13, [71] := 12, [72] := 5, [73] := 12, [74] := 13, [75] := 14, [76] 
+          := 11, [77] := 8, [78] := 5, [79] := 6), [j]) >= 0 .
+C3:    element(mk__rotate_definition([0] := 11, [1] := 14, [2] := 15, [3] := 
+          12, [4] := 5, [5] := 8, [6] := 7, [7] := 9, [8] := 11, [9] := 13, [10]
+           := 14, [11] := 15, [12] := 6, [13] := 7, [14] := 9, [15] := 8, [16] 
+          := 7, [17] := 6, [18] := 8, [19] := 13, [20] := 11, [21] := 9, [22] 
+          := 7, [23] := 15, [24] := 7, [25] := 12, [26] := 15, [27] := 9, [28] 
+          := 11, [29] := 7, [30] := 13, [31] := 12, [32] := 11, [33] := 13, [34]
+           := 6, [35] := 7, [36] := 14, [37] := 9, [38] := 13, [39] := 15, [40] 
+          := 14, [41] := 8, [42] := 13, [43] := 6, [44] := 5, [45] := 12, [46] 
+          := 7, [47] := 5, [48] := 11, [49] := 12, [50] := 14, [51] := 15, [52] 
+          := 14, [53] := 15, [54] := 9, [55] := 8, [56] := 9, [57] := 14, [58] 
+          := 5, [59] := 6, [60] := 8, [61] := 6, [62] := 5, [63] := 12, [64] := 
+          9, [65] := 15, [66] := 5, [67] := 11, [68] := 6, [69] := 8, [70] := 
+          13, [71] := 12, [72] := 5, [73] := 12, [74] := 13, [75] := 14, [76] 
+          := 11, [77] := 8, [78] := 5, [79] := 6), [j]) <= 15 .
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_r.fdl	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,37 @@
+           {*******************************************************}
+                               {FDL Declarations}
+    {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039}
+             {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.}
+           {*******************************************************}
+
+
+                        {DATE : 29-NOV-2010 14:30:19.84}
+
+                              {function RMD.S_R}
+
+
+title function s_r;
+
+  function round__(real) : integer;
+  type round_index = integer;
+  type rotate_definition = array [integer] of integer;
+  const s_values : rotate_definition = pending;
+  const rotate_amount__base__first : integer = pending; 
+  const rotate_amount__base__last : integer = pending; 
+  const round_index__base__first : integer = pending; 
+  const round_index__base__last : integer = pending; 
+  const integer__base__first : integer = pending; 
+  const integer__base__last : integer = pending; 
+  const rotate_amount__first : integer = pending; 
+  const rotate_amount__last : integer = pending; 
+  const rotate_amount__size : integer = pending; 
+  const round_index__first : integer = pending; 
+  const round_index__last : integer = pending; 
+  const round_index__size : integer = pending; 
+  const integer__first : integer = pending; 
+  const integer__last : integer = pending; 
+  const integer__size : integer = pending; 
+  var j : integer;
+  function s_r_spec(integer) : integer;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_r.rls	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,81 @@
+           /*********************************************************/
+                           /*Proof Rule Declarations*/
+    /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/
+             /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/
+           /*********************************************************/
+
+
+                        /*DATE : 29-NOV-2010 14:30:19.84*/
+
+                              /*function RMD.S_R*/
+
+
+rule_family s_r_rules:
+     X      requires [X:any] &
+     X <= Y requires [X:ire, Y:ire] &
+     X >= Y requires [X:ire, Y:ire].
+
+s_r_rules(1): rotate_amount__first <= element(s_values, [I]) may_be_deduced_from [0 <= I, I <= 79].
+s_r_rules(2): element(s_values, [I]) <= rotate_amount__last may_be_deduced_from [0 <= I, I <= 79].
+s_r_rules(3): s_values may_be_replaced_by 
+           mk__rotate_definition([round_index__first] := 8, [
+           round_index__first + 1] := 9, [round_index__first + 2] := 9, [
+           round_index__first + 3] := 11, [round_index__first + 4] := 
+           13, [round_index__first + 5] := 15, [round_index__first + 6] := 
+           15, [round_index__first + 7] := 5, [round_index__first + 8] := 
+           7, [round_index__first + 9] := 7, [round_index__first + 10] := 
+           8, [round_index__first + 11] := 11, [round_index__first + 12] := 
+           14, [round_index__first + 13] := 14, [
+           round_index__first + 14] := 12, [round_index__first + 15] := 
+           6, [round_index__first + 16] := 9, [round_index__first + 17] := 
+           13, [round_index__first + 18] := 15, [
+           round_index__first + 19] := 7, [round_index__first + 20] := 
+           12, [round_index__first + 21] := 8, [round_index__first + 22] := 
+           9, [round_index__first + 23] := 11, [round_index__first + 24] := 
+           7, [round_index__first + 25] := 7, [round_index__first + 26] := 
+           12, [round_index__first + 27] := 7, [round_index__first + 28] := 
+           6, [round_index__first + 29] := 15, [round_index__first + 30] := 
+           13, [round_index__first + 31] := 11, [
+           round_index__first + 32] := 9, [round_index__first + 33] := 
+           7, [round_index__first + 34] := 15, [round_index__first + 35] := 
+           11, [round_index__first + 36] := 8, [round_index__first + 37] := 
+           6, [round_index__first + 38] := 6, [round_index__first + 39] := 
+           14, [round_index__first + 40] := 12, [
+           round_index__first + 41] := 13, [round_index__first + 42] := 
+           5, [round_index__first + 43] := 14, [round_index__first + 44] := 
+           13, [round_index__first + 45] := 13, [
+           round_index__first + 46] := 7, [round_index__first + 47] := 
+           5, [round_index__first + 48] := 15, [round_index__first + 49] := 
+           5, [round_index__first + 50] := 8, [round_index__first + 51] := 
+           11, [round_index__first + 52] := 14, [
+           round_index__first + 53] := 14, [round_index__first + 54] := 
+           6, [round_index__first + 55] := 14, [round_index__first + 56] := 
+           6, [round_index__first + 57] := 9, [round_index__first + 58] := 
+           12, [round_index__first + 59] := 9, [round_index__first + 60] := 
+           12, [round_index__first + 61] := 5, [round_index__first + 62] := 
+           15, [round_index__first + 63] := 8, [round_index__first + 64] := 
+           8, [round_index__first + 65] := 5, [round_index__first + 66] := 
+           12, [round_index__first + 67] := 9, [round_index__first + 68] := 
+           12, [round_index__first + 69] := 5, [round_index__first + 70] := 
+           14, [round_index__first + 71] := 6, [round_index__first + 72] := 
+           8, [round_index__first + 73] := 13, [round_index__first + 74] := 
+           6, [round_index__first + 75] := 5, [round_index__first + 76] := 
+           15, [round_index__first + 77] := 13, [
+           round_index__first + 78] := 11, [round_index__first + 79] := 
+           11).
+s_r_rules(4): integer__size >= 0 may_be_deduced.
+s_r_rules(5): integer__first may_be_replaced_by -2147483648.
+s_r_rules(6): integer__last may_be_replaced_by 2147483647.
+s_r_rules(7): integer__base__first may_be_replaced_by -2147483648.
+s_r_rules(8): integer__base__last may_be_replaced_by 2147483647.
+s_r_rules(9): round_index__size >= 0 may_be_deduced.
+s_r_rules(10): round_index__first may_be_replaced_by 0.
+s_r_rules(11): round_index__last may_be_replaced_by 79.
+s_r_rules(12): round_index__base__first <= round_index__base__last may_be_deduced.
+s_r_rules(13): round_index__base__first <= round_index__first may_be_deduced.
+s_r_rules(14): round_index__base__last >= round_index__last may_be_deduced.
+s_r_rules(15): rotate_amount__size >= 0 may_be_deduced.
+s_r_rules(16): rotate_amount__first may_be_replaced_by 0.
+s_r_rules(17): rotate_amount__last may_be_replaced_by 15.
+s_r_rules(18): rotate_amount__base__first may_be_replaced_by -2147483648.
+s_r_rules(19): rotate_amount__base__last may_be_replaced_by 2147483647.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/rmd/s_r.siv	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,79 @@
+*****************************************************************************
+                       Semantic Analysis of SPARK Text
+    Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+             Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+*****************************************************************************
+
+
+CREATED 29-NOV-2010, 14:30:19  SIMPLIFIED 29-NOV-2010, 14:30:30
+
+SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+
+function RMD.S_R
+
+
+
+
+For path(s) from start to run-time check associated with statement of line 101:
+
+function_s_r_1.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to finish:
+
+function_s_r_2.
+H1:    j >= 0 .
+H2:    j <= 79 .
+H3:    integer__size >= 0 .
+H4:    round_index__size >= 0 .
+H5:    round_index__base__first <= round_index__base__last .
+H6:    rotate_amount__size >= 0 .
+H7:    round_index__base__first <= 0 .
+H8:    round_index__base__last >= 79 .
+       ->
+C1:    element(mk__rotate_definition([0] := 8, [1] := 9, [2] := 9, [3] := 11, [
+          4] := 13, [5] := 15, [6] := 15, [7] := 5, [8] := 7, [9] := 7, [10] := 
+          8, [11] := 11, [12] := 14, [13] := 14, [14] := 12, [15] := 6, [16] := 
+          9, [17] := 13, [18] := 15, [19] := 7, [20] := 12, [21] := 8, [22] := 
+          9, [23] := 11, [24] := 7, [25] := 7, [26] := 12, [27] := 7, [28] := 
+          6, [29] := 15, [30] := 13, [31] := 11, [32] := 9, [33] := 7, [34] := 
+          15, [35] := 11, [36] := 8, [37] := 6, [38] := 6, [39] := 14, [40] := 
+          12, [41] := 13, [42] := 5, [43] := 14, [44] := 13, [45] := 13, [46] 
+          := 7, [47] := 5, [48] := 15, [49] := 5, [50] := 8, [51] := 11, [52] 
+          := 14, [53] := 14, [54] := 6, [55] := 14, [56] := 6, [57] := 9, [58] 
+          := 12, [59] := 9, [60] := 12, [61] := 5, [62] := 15, [63] := 8, [64] 
+          := 8, [65] := 5, [66] := 12, [67] := 9, [68] := 12, [69] := 5, [70] 
+          := 14, [71] := 6, [72] := 8, [73] := 13, [74] := 6, [75] := 5, [76] 
+          := 15, [77] := 13, [78] := 11, [79] := 11), [j]) = s_r_spec(j) .
+C2:    element(mk__rotate_definition([0] := 8, [1] := 9, [2] := 9, [3] := 11, [
+          4] := 13, [5] := 15, [6] := 15, [7] := 5, [8] := 7, [9] := 7, [10] := 
+          8, [11] := 11, [12] := 14, [13] := 14, [14] := 12, [15] := 6, [16] := 
+          9, [17] := 13, [18] := 15, [19] := 7, [20] := 12, [21] := 8, [22] := 
+          9, [23] := 11, [24] := 7, [25] := 7, [26] := 12, [27] := 7, [28] := 
+          6, [29] := 15, [30] := 13, [31] := 11, [32] := 9, [33] := 7, [34] := 
+          15, [35] := 11, [36] := 8, [37] := 6, [38] := 6, [39] := 14, [40] := 
+          12, [41] := 13, [42] := 5, [43] := 14, [44] := 13, [45] := 13, [46] 
+          := 7, [47] := 5, [48] := 15, [49] := 5, [50] := 8, [51] := 11, [52] 
+          := 14, [53] := 14, [54] := 6, [55] := 14, [56] := 6, [57] := 9, [58] 
+          := 12, [59] := 9, [60] := 12, [61] := 5, [62] := 15, [63] := 8, [64] 
+          := 8, [65] := 5, [66] := 12, [67] := 9, [68] := 12, [69] := 5, [70] 
+          := 14, [71] := 6, [72] := 8, [73] := 13, [74] := 6, [75] := 5, [76] 
+          := 15, [77] := 13, [78] := 11, [79] := 11), [j]) >= 0 .
+C3:    element(mk__rotate_definition([0] := 8, [1] := 9, [2] := 9, [3] := 11, [
+          4] := 13, [5] := 15, [6] := 15, [7] := 5, [8] := 7, [9] := 7, [10] := 
+          8, [11] := 11, [12] := 14, [13] := 14, [14] := 12, [15] := 6, [16] := 
+          9, [17] := 13, [18] := 15, [19] := 7, [20] := 12, [21] := 8, [22] := 
+          9, [23] := 11, [24] := 7, [25] := 7, [26] := 12, [27] := 7, [28] := 
+          6, [29] := 15, [30] := 13, [31] := 11, [32] := 9, [33] := 7, [34] := 
+          15, [35] := 11, [36] := 8, [37] := 6, [38] := 6, [39] := 14, [40] := 
+          12, [41] := 13, [42] := 5, [43] := 14, [44] := 13, [45] := 13, [46] 
+          := 7, [47] := 5, [48] := 15, [49] := 5, [50] := 8, [51] := 11, [52] 
+          := 14, [53] := 14, [54] := 6, [55] := 14, [56] := 6, [57] := 9, [58] 
+          := 12, [59] := 9, [60] := 12, [61] := 5, [62] := 15, [63] := 8, [64] 
+          := 8, [65] := 5, [66] := 12, [67] := 9, [68] := 12, [69] := 5, [70] 
+          := 14, [71] := 6, [72] := 8, [73] := 13, [74] := 6, [75] := 5, [76] 
+          := 15, [77] := 13, [78] := 11, [79] := 11), [j]) <= 15 .
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/shadow/interfaces.ads	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,5 @@
+package Interfaces is
+
+   type Unsigned_32 is mod 2 ** 32;
+
+end Interfaces;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/wordops.adb	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,9 @@
+package body Wordops is
+
+   function Rotate(I : Rotate_Amount; W : Word) return Word
+   is
+   begin
+      return Interfaces.Rotate_Left (W, I);
+   end Rotate;
+
+end Wordops;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/RIPEMD-160/wordops.ads	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,19 @@
+
+with Interfaces;
+--# inherit Interfaces;
+
+package WordOps is
+
+   subtype Word is Interfaces.Unsigned_32;
+
+   subtype Rotate_Amount is Integer range 0..15;
+
+   --# function rotate_left(I : Rotate_Amount; W : Word) return Word;
+
+   function Rotate(I : Rotate_Amount; W : Word) return Word;
+   --# return rotate_left(I, W);
+   --# accept W, 3, "Expecting this warning";
+   pragma Inline (Rotate);
+   --# end accept;
+
+end Wordops;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/ROOT.ML	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,16 @@
+use_thys
+  ["Gcd/Greatest_Common_Divisor",
+
+   "Liseq/Longest_Increasing_Subsequence",
+
+   "RIPEMD-160/F",
+   "RIPEMD-160/Hash",
+   "RIPEMD-160/K_L",
+   "RIPEMD-160/K_R",
+   "RIPEMD-160/R_L",
+   "RIPEMD-160/Round",
+   "RIPEMD-160/R_R",
+   "RIPEMD-160/S_L",
+   "RIPEMD-160/S_R",
+
+   "Sqrt/Sqrt"];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Sqrt/Sqrt.adb	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,16 @@
+package body Sqrt is
+
+   function Isqrt(N: Natural) return Natural
+   is
+     R: Natural;
+   begin
+      R := 0;
+      loop
+         --# assert R * R <= N;
+         exit when N - R * R < 2 * R + 1;
+         R := R + 1;
+      end loop;
+      return R;
+   end Isqrt;
+
+end Sqrt;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Sqrt/Sqrt.ads	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,6 @@
+package Sqrt is
+
+   function Isqrt(N: Natural) return Natural;
+   --# return R => R * R <= N and (R + 1) * (R + 1) > N;
+
+end Sqrt;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Sqrt/Sqrt.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,28 @@
+(*  Title:      HOL/SPARK/Examples/Sqrt/Sqrt.thy
+    Author:     Stefan Berghofer
+    Copyright:  secunet Security Networks AG
+*)
+
+theory Sqrt
+imports SPARK
+begin
+
+spark_open "sqrt/isqrt.siv"
+
+spark_vc function_isqrt_4
+proof -
+  from `0 \<le> r` have "(r = 0 \<or> r = 1 \<or> r = 2) \<or> 2 < r" by auto
+  then show "2 * r \<le> 2147483646"
+  proof
+    assume "2 < r"
+    then have "0 < r" by simp
+    with `2 < r` have "2 * r < r * r" by (rule mult_strict_right_mono)
+    with `r * r \<le> n` and `n \<le> 2147483647` show ?thesis
+      by simp
+  qed auto
+  then show "2 * r \<le> 2147483647" by simp
+qed
+
+spark_end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Sqrt/sqrt/isqrt.fdl	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,29 @@
+           {*******************************************************}
+                               {FDL Declarations}
+    {Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039}
+             {Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.}
+           {*******************************************************}
+
+
+                        {DATE : 29-NOV-2010 14:30:17.95}
+
+                            {function Sqrt.Isqrt}
+
+
+title function isqrt;
+
+  function round__(real) : integer;
+  const natural__base__first : integer = pending; 
+  const natural__base__last : integer = pending; 
+  const integer__base__first : integer = pending; 
+  const integer__base__last : integer = pending; 
+  const natural__first : integer = pending; 
+  const natural__last : integer = pending; 
+  const natural__size : integer = pending; 
+  const integer__first : integer = pending; 
+  const integer__last : integer = pending; 
+  const integer__size : integer = pending; 
+  var n : integer;
+  var r : integer;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Sqrt/sqrt/isqrt.rls	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,27 @@
+           /*********************************************************/
+                           /*Proof Rule Declarations*/
+    /*Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039*/
+             /*Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.*/
+           /*********************************************************/
+
+
+                        /*DATE : 29-NOV-2010 14:30:17.95*/
+
+                            /*function Sqrt.Isqrt*/
+
+
+rule_family isqrt_rules:
+     X      requires [X:any] &
+     X <= Y requires [X:ire, Y:ire] &
+     X >= Y requires [X:ire, Y:ire].
+
+isqrt_rules(1): integer__size >= 0 may_be_deduced.
+isqrt_rules(2): integer__first may_be_replaced_by -2147483648.
+isqrt_rules(3): integer__last may_be_replaced_by 2147483647.
+isqrt_rules(4): integer__base__first may_be_replaced_by -2147483648.
+isqrt_rules(5): integer__base__last may_be_replaced_by 2147483647.
+isqrt_rules(6): natural__size >= 0 may_be_deduced.
+isqrt_rules(7): natural__first may_be_replaced_by 0.
+isqrt_rules(8): natural__last may_be_replaced_by 2147483647.
+isqrt_rules(9): natural__base__first may_be_replaced_by -2147483648.
+isqrt_rules(10): natural__base__last may_be_replaced_by 2147483647.
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Examples/Sqrt/sqrt/isqrt.siv	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,64 @@
+*****************************************************************************
+                       Semantic Analysis of SPARK Text
+    Examiner Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+             Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+*****************************************************************************
+
+
+CREATED 29-NOV-2010, 14:30:17  SIMPLIFIED 29-NOV-2010, 14:30:18
+
+SPARK Simplifier Pro Edition, Version 9.1.0, Build Date 20101119, Build 19039
+Copyright (C) 2010 Altran Praxis Limited, Bath, U.K.
+
+function Sqrt.Isqrt
+
+
+
+
+For path(s) from start to run-time check associated with statement of line 7:
+
+function_isqrt_1.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from start to assertion of line 9:
+
+function_isqrt_2.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 9 to assertion of line 9:
+
+function_isqrt_3.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 9 to run-time check associated with 
+          statement of line 10:
+
+function_isqrt_4.
+H1:    r * r <= n .
+H2:    n >= 0 .
+H3:    n <= 2147483647 .
+H4:    r >= 0 .
+H5:    r <= 2147483647 .
+H6:    integer__size >= 0 .
+H7:    natural__size >= 0 .
+       ->
+C1:    2 * r <= 2147483646 .
+C2:    2 * r <= 2147483647 .
+
+
+For path(s) from assertion of line 9 to run-time check associated with 
+          statement of line 11:
+
+function_isqrt_5.
+*** true .          /* all conclusions proved */
+
+
+For path(s) from assertion of line 9 to finish:
+
+function_isqrt_6.
+*** true .          /* all conclusions proved */
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/ROOT.ML	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,1 @@
+use_thys ["SPARK"];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/SPARK.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,278 @@
+(*  Title:      HOL/SPARK/SPARK.thy
+    Author:     Stefan Berghofer
+    Copyright:  secunet Security Networks AG
+
+Declaration of proof functions for SPARK/Ada verification environment.
+*)
+
+theory SPARK
+imports SPARK_Setup
+begin
+
+text {* Bitwise logical operators *}
+
+spark_proof_functions
+  bit__and (integer, integer) : integer = "op AND"
+  bit__or (integer, integer) : integer = "op OR"
+  bit__xor (integer, integer) : integer = "op XOR"
+
+lemma AND_lower [simp]:
+  fixes x :: int and y :: int
+  assumes "0 \<le> x"
+  shows "0 \<le> x AND y"
+  using assms
+proof (induct x arbitrary: y rule: bin_induct)
+  case (3 bin bit)
+  show ?case
+  proof (cases y rule: bin_exhaust)
+    case (1 bin' bit')
+    from 3 have "0 \<le> bin"
+      by (simp add: Bit_def bitval_def split add: bit.split_asm)
+    then have "0 \<le> bin AND bin'" by (rule 3)
+    with 1 show ?thesis
+      by simp (simp add: Bit_def bitval_def split add: bit.split)
+  qed
+next
+  case 2
+  then show ?case by (simp only: Min_def)
+qed simp
+
+lemma OR_lower [simp]:
+  fixes x :: int and y :: int
+  assumes "0 \<le> x" "0 \<le> y"
+  shows "0 \<le> x OR y"
+  using assms
+proof (induct x arbitrary: y rule: bin_induct)
+  case (3 bin bit)
+  show ?case
+  proof (cases y rule: bin_exhaust)
+    case (1 bin' bit')
+    from 3 have "0 \<le> bin"
+      by (simp add: Bit_def bitval_def split add: bit.split_asm)
+    moreover from 1 3 have "0 \<le> bin'"
+      by (simp add: Bit_def bitval_def split add: bit.split_asm)
+    ultimately have "0 \<le> bin OR bin'" by (rule 3)
+    with 1 show ?thesis
+      by simp (simp add: Bit_def bitval_def split add: bit.split)
+  qed
+qed simp_all
+
+lemma XOR_lower [simp]:
+  fixes x :: int and y :: int
+  assumes "0 \<le> x" "0 \<le> y"
+  shows "0 \<le> x XOR y"
+  using assms
+proof (induct x arbitrary: y rule: bin_induct)
+  case (3 bin bit)
+  show ?case
+  proof (cases y rule: bin_exhaust)
+    case (1 bin' bit')
+    from 3 have "0 \<le> bin"
+      by (simp add: Bit_def bitval_def split add: bit.split_asm)
+    moreover from 1 3 have "0 \<le> bin'"
+      by (simp add: Bit_def bitval_def split add: bit.split_asm)
+    ultimately have "0 \<le> bin XOR bin'" by (rule 3)
+    with 1 show ?thesis
+      by simp (simp add: Bit_def bitval_def split add: bit.split)
+  qed
+next
+  case 2
+  then show ?case by (simp only: Min_def)
+qed simp
+
+lemma AND_upper1 [simp]:
+  fixes x :: int and y :: int
+  assumes "0 \<le> x"
+  shows "x AND y \<le> x"
+  using assms
+proof (induct x arbitrary: y rule: bin_induct)
+  case (3 bin bit)
+  show ?case
+  proof (cases y rule: bin_exhaust)
+    case (1 bin' bit')
+    from 3 have "0 \<le> bin"
+      by (simp add: Bit_def bitval_def split add: bit.split_asm)
+    then have "bin AND bin' \<le> bin" by (rule 3)
+    with 1 show ?thesis
+      by simp (simp add: Bit_def bitval_def split add: bit.split)
+  qed
+next
+  case 2
+  then show ?case by (simp only: Min_def)
+qed simp
+
+lemmas AND_upper1' [simp] = order_trans [OF AND_upper1]
+lemmas AND_upper1'' [simp] = order_le_less_trans [OF AND_upper1]
+
+lemma AND_upper2 [simp]:
+  fixes x :: int and y :: int
+  assumes "0 \<le> y"
+  shows "x AND y \<le> y"
+  using assms
+proof (induct y arbitrary: x rule: bin_induct)
+  case (3 bin bit)
+  show ?case
+  proof (cases x rule: bin_exhaust)
+    case (1 bin' bit')
+    from 3 have "0 \<le> bin"
+      by (simp add: Bit_def bitval_def split add: bit.split_asm)
+    then have "bin' AND bin \<le> bin" by (rule 3)
+    with 1 show ?thesis
+      by simp (simp add: Bit_def bitval_def split add: bit.split)
+  qed
+next
+  case 2
+  then show ?case by (simp only: Min_def)
+qed simp
+
+lemmas AND_upper2' [simp] = order_trans [OF AND_upper2]
+lemmas AND_upper2'' [simp] = order_le_less_trans [OF AND_upper2]
+
+lemma OR_upper:
+  fixes x :: int and y :: int
+  assumes "0 \<le> x" "x < 2 ^ n" "y < 2 ^ n"
+  shows "x OR y < 2 ^ n"
+  using assms
+proof (induct x arbitrary: y n rule: bin_induct)
+  case (3 bin bit)
+  show ?case
+  proof (cases y rule: bin_exhaust)
+    case (1 bin' bit')
+    show ?thesis
+    proof (cases n)
+      case 0
+      with 3 have "bin BIT bit = 0" by simp
+      then have "bin = 0" "bit = 0"
+        by (auto simp add: Bit_def bitval_def split add: bit.split_asm) arith
+      then show ?thesis using 0 1 `y < 2 ^ n`
+        by simp (simp add: Bit0_def int_or_Pls [unfolded Pls_def])
+    next
+      case (Suc m)
+      from 3 have "0 \<le> bin"
+        by (simp add: Bit_def bitval_def split add: bit.split_asm)
+      moreover from 3 Suc have "bin < 2 ^ m"
+        by (simp add: Bit_def bitval_def split add: bit.split_asm)
+      moreover from 1 3 Suc have "bin' < 2 ^ m"
+        by (simp add: Bit_def bitval_def split add: bit.split_asm)
+      ultimately have "bin OR bin' < 2 ^ m" by (rule 3)
+      with 1 Suc show ?thesis
+        by simp (simp add: Bit_def bitval_def split add: bit.split)
+    qed
+  qed
+qed simp_all
+
+lemmas [simp] =
+  OR_upper [of _ 8, simplified zle_diff1_eq [symmetric], simplified]
+  OR_upper [of _ 8, simplified]
+  OR_upper [of _ 16, simplified zle_diff1_eq [symmetric], simplified]
+  OR_upper [of _ 16, simplified]
+  OR_upper [of _ 32, simplified zle_diff1_eq [symmetric], simplified]
+  OR_upper [of _ 32, simplified]
+  OR_upper [of _ 64, simplified zle_diff1_eq [symmetric], simplified]
+  OR_upper [of _ 64, simplified]
+
+lemma XOR_upper:
+  fixes x :: int and y :: int
+  assumes "0 \<le> x" "x < 2 ^ n" "y < 2 ^ n"
+  shows "x XOR y < 2 ^ n"
+  using assms
+proof (induct x arbitrary: y n rule: bin_induct)
+  case (3 bin bit)
+  show ?case
+  proof (cases y rule: bin_exhaust)
+    case (1 bin' bit')
+    show ?thesis
+    proof (cases n)
+      case 0
+      with 3 have "bin BIT bit = 0" by simp
+      then have "bin = 0" "bit = 0"
+        by (auto simp add: Bit_def bitval_def split add: bit.split_asm) arith
+      then show ?thesis using 0 1 `y < 2 ^ n`
+        by simp (simp add: Bit0_def int_xor_Pls [unfolded Pls_def])
+    next
+      case (Suc m)
+      from 3 have "0 \<le> bin"
+        by (simp add: Bit_def bitval_def split add: bit.split_asm)
+      moreover from 3 Suc have "bin < 2 ^ m"
+        by (simp add: Bit_def bitval_def split add: bit.split_asm)
+      moreover from 1 3 Suc have "bin' < 2 ^ m"
+        by (simp add: Bit_def bitval_def split add: bit.split_asm)
+      ultimately have "bin XOR bin' < 2 ^ m" by (rule 3)
+      with 1 Suc show ?thesis
+        by simp (simp add: Bit_def bitval_def split add: bit.split)
+    qed
+  qed
+next
+  case 2
+  then show ?case by (simp only: Min_def)
+qed simp
+
+lemmas [simp] =
+  XOR_upper [of _ 8, simplified zle_diff1_eq [symmetric], simplified]
+  XOR_upper [of _ 8, simplified]
+  XOR_upper [of _ 16, simplified zle_diff1_eq [symmetric], simplified]
+  XOR_upper [of _ 16, simplified]
+  XOR_upper [of _ 32, simplified zle_diff1_eq [symmetric], simplified]
+  XOR_upper [of _ 32, simplified]
+  XOR_upper [of _ 64, simplified zle_diff1_eq [symmetric], simplified]
+  XOR_upper [of _ 64, simplified]
+
+lemma bit_not_spark_eq:
+  "NOT (word_of_int x :: ('a::len0) word) =
+  word_of_int (2 ^ len_of TYPE('a) - 1 - x)"
+proof -
+  have "word_of_int x + NOT (word_of_int x) =
+    word_of_int x + (word_of_int (2 ^ len_of TYPE('a) - 1 - x)::'a word)"
+    by (simp only: bwsimps bin_add_not Min_def)
+      (simp add: word_of_int_hom_syms word_of_int_2p_len)
+  then show ?thesis by (rule add_left_imp_eq)
+qed
+
+lemmas [simp] =
+  bit_not_spark_eq [where 'a=8, simplified]
+  bit_not_spark_eq [where 'a=16, simplified]
+  bit_not_spark_eq [where 'a=32, simplified]
+  bit_not_spark_eq [where 'a=64, simplified]
+
+lemma power_BIT: "2 ^ (Suc n) - 1 = (2 ^ n - 1) BIT 1"
+  unfolding Bit_B1
+  by (induct n) simp_all
+
+lemma mod_BIT:
+  "bin BIT bit mod 2 ^ Suc n = (bin mod 2 ^ n) BIT bit"
+proof -
+  have "bin mod 2 ^ n < 2 ^ n" by simp
+  then have "bin mod 2 ^ n \<le> 2 ^ n - 1" by simp
+  then have "2 * (bin mod 2 ^ n) \<le> 2 * (2 ^ n - 1)"
+    by (rule mult_left_mono) simp
+  then have "2 * (bin mod 2 ^ n) + 1 < 2 * 2 ^ n" by simp
+  then show ?thesis
+    by (auto simp add: Bit_def bitval_def mod_mult_mult1 mod_add_left_eq [of "2 * bin"]
+      mod_pos_pos_trivial split add: bit.split)
+qed
+
+lemma AND_mod:
+  fixes x :: int
+  shows "x AND 2 ^ n - 1 = x mod 2 ^ n"
+proof (induct x arbitrary: n rule: bin_induct)
+  case 1
+  then show ?case
+    by simp (simp add: Pls_def)
+next
+  case 2
+  then show ?case
+    by (simp, simp only: Min_def, simp add: m1mod2k)
+next
+  case (3 bin bit)
+  show ?case
+  proof (cases n)
+    case 0
+    then show ?thesis by (simp add: int_and_extra_simps [unfolded Pls_def])
+  next
+    case (Suc m)
+    with 3 show ?thesis
+      by (simp only: power_BIT mod_BIT int_and_Bits) simp
+  qed
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/SPARK_Setup.thy	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,212 @@
+(*  Title:      HOL/SPARK/SPARK_Setup.thy
+    Author:     Stefan Berghofer
+    Copyright:  secunet Security Networks AG
+
+Setup for SPARK/Ada verification environment.
+*)
+
+theory SPARK_Setup
+imports Word
+uses
+  "Tools/fdl_lexer.ML"
+  "Tools/fdl_parser.ML"
+  ("Tools/spark_vcs.ML")
+  ("Tools/spark_commands.ML")
+begin
+
+text {*
+SPARK versions of div and mod, see section 4.4.1.1 of SPARK Proof Manual
+*}
+
+definition sdiv :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "sdiv" 70) where
+  "a sdiv b =
+     (if 0 \<le> a then
+        if 0 \<le> b then a div b
+        else - (a div - b)
+      else
+        if 0 \<le> b then - (- a div b)
+        else - a div - b)"
+
+definition smod :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "smod" 70) where
+  "a smod b = a - ((a sdiv b) * b)"
+
+lemma sdiv_minus_dividend: "- a sdiv b = - (a sdiv b)"
+  by (simp add: sdiv_def)
+
+lemma sdiv_minus_divisor: "a sdiv - b = - (a sdiv b)"
+  by (simp add: sdiv_def)
+
+lemma smod_minus_dividend: "- a smod b = - (a smod b)"
+  by (simp add: smod_def sdiv_minus_dividend)
+
+lemma smod_minus_divisor: "a smod - b = a smod b"
+  by (simp add: smod_def sdiv_minus_divisor)
+
+text {*
+Correspondence between HOL's and SPARK's versions of div and mod
+*}
+
+lemma sdiv_pos_pos: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> a sdiv b = a div b"
+  by (simp add: sdiv_def)
+
+lemma sdiv_pos_neg: "0 \<le> a \<Longrightarrow> b < 0 \<Longrightarrow> a sdiv b = - (a div - b)"
+  by (simp add: sdiv_def)
+
+lemma sdiv_neg_pos: "a < 0 \<Longrightarrow> 0 \<le> b \<Longrightarrow> a sdiv b = - (- a div b)"
+  by (simp add: sdiv_def)
+
+lemma sdiv_neg_neg: "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> a sdiv b = - a div - b"
+  by (simp add: sdiv_def)
+
+lemma smod_pos_pos: "0 \<le> a \<Longrightarrow> 0 \<le> b \<Longrightarrow> a smod b = a mod b"
+  by (simp add: smod_def sdiv_pos_pos zmod_zdiv_equality')
+
+lemma smod_pos_neg: "0 \<le> a \<Longrightarrow> b < 0 \<Longrightarrow> a smod b = a mod - b"
+  by (simp add: smod_def sdiv_pos_neg zmod_zdiv_equality')
+
+lemma smod_neg_pos: "a < 0 \<Longrightarrow> 0 \<le> b \<Longrightarrow> a smod b = - (- a mod b)"
+  by (simp add: smod_def sdiv_neg_pos zmod_zdiv_equality')
+
+lemma smod_neg_neg: "a < 0 \<Longrightarrow> b < 0 \<Longrightarrow> a smod b = - (- a mod - b)"
+  by (simp add: smod_def sdiv_neg_neg zmod_zdiv_equality')
+
+
+text {*
+Updating a function at a set of points. Useful for building arrays.
+*}
+
+definition fun_upds :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'b" where
+  "fun_upds f xs y z = (if z \<in> xs then y else f z)"
+
+syntax
+  "_updsbind" :: "['a, 'a] => updbind"             ("(2_ [:=]/ _)")
+
+translations
+  "f(xs[:=]y)" == "CONST fun_upds f xs y"
+
+lemma fun_upds_in [simp]: "z \<in> xs \<Longrightarrow> (f(xs [:=] y)) z = y"
+  by (simp add: fun_upds_def)
+
+lemma fun_upds_notin [simp]: "z \<notin> xs \<Longrightarrow> (f(xs [:=] y)) z = f z"
+  by (simp add: fun_upds_def)
+
+lemma upds_singleton [simp]: "f({x} [:=] y) = f(x := y)"
+  by (simp add: fun_eq_iff)
+
+
+text {* Enumeration types *}
+
+class enum = ord + finite +
+  fixes pos :: "'a \<Rightarrow> int"
+  assumes range_pos: "range pos = {0..<int (card (UNIV::'a set))}"
+  and less_pos: "(x < y) = (pos x < pos y)"
+  and less_eq_pos: "(x \<le> y) = (pos x \<le> pos y)"
+begin
+
+definition "val = inv pos"
+
+definition "succ x = val (pos x + 1)"
+
+definition "pred x = val (pos x - 1)"
+
+lemma inj_pos: "inj pos"
+  using finite_UNIV
+  by (rule eq_card_imp_inj_on) (simp add: range_pos)
+
+lemma val_pos: "val (pos x) = x"
+  unfolding val_def using inj_pos
+  by (rule inv_f_f)
+
+lemma pos_val: "z \<in> range pos \<Longrightarrow> pos (val z) = z"
+  unfolding val_def
+  by (rule f_inv_into_f)
+
+subclass linorder
+proof
+  fix x::'a and y show "(x < y) = (x \<le> y \<and> \<not> y \<le> x)"
+    by (simp add: less_pos less_eq_pos less_le_not_le)
+next
+  fix x::'a show "x \<le> x" by (simp add: less_eq_pos)
+next
+  fix x::'a and y z assume "x \<le> y" and "y \<le> z"
+  then show "x \<le> z" by (simp add: less_eq_pos)
+next
+  fix x::'a and y assume "x \<le> y" and "y \<le> x"
+  with inj_pos show "x = y"
+    by (auto dest: injD simp add: less_eq_pos)
+next
+  fix x::'a and y show "x \<le> y \<or> y \<le> x"
+    by (simp add: less_eq_pos linear)
+qed
+
+definition "first_el = val 0"
+
+definition "last_el = val (int (card (UNIV::'a set)) - 1)"
+
+lemma first_el_smallest: "first_el \<le> x"
+proof -
+  have "pos x \<in> range pos" by (rule rangeI)
+  then have "pos (val 0) \<le> pos x"
+    by (simp add: range_pos pos_val)
+  then show ?thesis by (simp add: first_el_def less_eq_pos)
+qed
+
+lemma last_el_greatest: "x \<le> last_el"
+proof -
+  have "pos x \<in> range pos" by (rule rangeI)
+  then have "pos x \<le> pos (val (int (card (UNIV::'a set)) - 1))"
+    by (simp add: range_pos pos_val)
+  then show ?thesis by (simp add: last_el_def less_eq_pos)
+qed
+
+lemma pos_succ:
+  assumes "x \<noteq> last_el"
+  shows "pos (succ x) = pos x + 1"
+proof -
+  have "x \<le> last_el" by (rule last_el_greatest)
+  with assms have "x < last_el" by simp
+  then have "pos x < pos last_el"
+    by (simp add: less_pos)
+  with rangeI [of pos x]
+  have "pos x + 1 \<in> range pos"
+    by (simp add: range_pos last_el_def pos_val)
+  then show ?thesis
+    by (simp add: succ_def pos_val)
+qed
+
+lemma pos_pred:
+  assumes "x \<noteq> first_el"
+  shows "pos (pred x) = pos x - 1"
+proof -
+  have "first_el \<le> x" by (rule first_el_smallest)
+  with assms have "first_el < x" by simp
+  then have "pos first_el < pos x"
+    by (simp add: less_pos)
+  with rangeI [of pos x]
+  have "pos x - 1 \<in> range pos"
+    by (simp add: range_pos first_el_def pos_val)
+  then show ?thesis
+    by (simp add: pred_def pos_val)
+qed
+
+lemma succ_val: "x \<in> range pos \<Longrightarrow> succ (val x) = val (x + 1)"
+  by (simp add: succ_def pos_val)
+
+lemma pred_val: "x \<in> range pos \<Longrightarrow> pred (val x) = val (x - 1)"
+  by (simp add: pred_def pos_val)
+
+end
+
+lemma interval_expand:
+  "x < y \<Longrightarrow> (z::int) \<in> {x..<y} = (z = x \<or> z \<in> {x+1..<y})"
+  by auto
+
+
+text {* Load the package *}
+
+use "Tools/spark_vcs.ML"
+use "Tools/spark_commands.ML"
+
+setup SPARK_Commands.setup
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Tools/fdl_lexer.ML	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,298 @@
+(*  Title:      HOL/SPARK/Tools/fdl_lexer.ML
+    Author:     Stefan Berghofer
+    Copyright:  secunet Security Networks AG
+
+Lexical analyzer for fdl files.
+*)
+
+signature FDL_LEXER =
+sig
+  type T
+  type chars
+  type banner
+  type date
+  type time
+  datatype kind = Keyword | Ident | Long_Ident | Traceability | Number | Comment | EOF
+  val tokenize: (chars -> 'a * chars) -> (chars -> T * chars) ->
+    Position.T -> string -> 'a * T list
+  val position_of: T -> Position.T
+  val pos_of: T -> string
+  val is_eof: T -> bool
+  val stopper: T Scan.stopper
+  val kind_of: T -> kind
+  val content_of: T -> string
+  val unparse: T -> string
+  val is_proper: T -> bool
+  val is_digit: string -> bool
+  val c_comment: chars -> T * chars
+  val curly_comment: chars -> T * chars
+  val percent_comment: chars -> T * chars
+  val vcg_header: chars -> (banner * (date * time) option) * chars
+  val siv_header: chars ->
+    (banner * (date * time) * (date * time) * (string * string)) * chars
+end;
+
+structure Fdl_Lexer: FDL_LEXER =
+struct
+
+(** tokens **)
+
+datatype kind = Keyword | Ident | Long_Ident | Traceability | Number | Comment | EOF;
+
+datatype T = Token of kind * string * Position.T;
+
+fun make_token k xs = Token (k, implode (map fst xs),
+  case xs of [] => Position.none | (_, p) :: _ => p);
+
+fun kind_of (Token (k, _, _)) = k;
+
+fun is_proper (Token (Comment, _, _)) = false
+  | is_proper _ = true;
+
+fun content_of (Token (_, s, _)) = s;
+
+fun unparse (Token (Traceability, s, _)) = "For " ^ s ^ ":"
+  | unparse (Token (_, s, _)) = s;
+
+fun position_of (Token (_, _, pos)) = pos;
+
+val pos_of = Position.str_of o position_of;
+
+fun is_eof (Token (EOF, _, _)) = true
+  | is_eof _ = false;
+
+fun mk_eof pos = Token (EOF, "", pos);
+val eof = mk_eof Position.none;
+
+val stopper =
+  Scan.stopper (fn [] => eof | toks => mk_eof (position_of (List.last toks))) is_eof;
+
+fun leq_token (Token (_, s, _), Token (_, s', _)) = size s <= size s';
+
+
+(** split up a string into a list of characters (with positions) **)
+
+type chars = (string * Position.T) list;
+
+(* a variant of Position.advance (see Pure/General/position.ML) *)
+fun advance x pos =
+  let
+    fun incr i = i+1;
+    fun change_prop s f props = case Properties.get_int props s of
+        NONE => props
+      | SOME i => Properties.put (s, Int.toString (f i)) props;
+  in
+    pos |>
+    Position.properties_of |>
+    change_prop "offset" incr |>
+    (case x of
+       "\n" =>
+         change_prop "line" incr #>
+         change_prop "column" (K 1)
+     | _ => change_prop "column" incr) |>
+    Position.of_properties
+  end;
+
+fun is_char_eof ("", _) = true
+  | is_char_eof _ = false;
+
+val char_stopper = Scan.stopper (K ("", Position.none)) is_char_eof;
+
+fun symbol (x : string, _ : Position.T) = x;
+
+fun explode_pos s pos = fst (fold_map (fn x => fn pos =>
+  ((x, pos), advance x pos)) (raw_explode s) pos);
+
+
+(** scanners **)
+
+val any = Scan.one (not o Scan.is_stopper char_stopper);
+
+fun prfx [] = Scan.succeed []
+  | prfx (x :: xs) = Scan.one (equal x o symbol) ::: prfx xs;
+
+val $$$ = prfx o raw_explode;
+
+val lexicon = Scan.make_lexicon (map raw_explode
+  ["rule_family",
+   "title",
+   "For",
+   ":",
+   "[",
+   "]",
+   "(",
+   ")",
+   ",",
+   "&",
+   ";",
+   "=",
+   ".",
+   "..",
+   "requires",
+   "may_be_replaced_by",
+   "may_be_deduced",
+   "may_be_deduced_from",
+   "are_interchangeable",
+   "if",
+   "end",
+   "function",
+   "procedure",
+   "type",
+   "var",
+   "const",
+   "array",
+   "record",
+   ":=",
+   "of",
+   "**",
+   "*",
+   "/",
+   "div",
+   "mod",
+   "+",
+   "-",
+   "<>",
+   "<",
+   ">",
+   "<=",
+   ">=",
+   "<->",
+   "->",
+   "not",
+   "and",
+   "or",
+   "for_some",
+   "for_all",
+   "***",
+   "!!!",
+   "element",
+   "update",
+   "pending"]);
+
+fun keyword s = Scan.literal lexicon :|--
+  (fn xs => if map symbol xs = raw_explode s then Scan.succeed xs else Scan.fail);
+
+fun is_digit x = "0" <= x andalso x <= "9";
+fun is_alpha x = "a" <= x andalso x <= "z" orelse "A" <= x andalso x <= "Z";
+val is_underscore = equal "_";
+val is_tilde = equal "~";
+val is_newline = equal "\n";
+val is_tab = equal "\t";
+val is_space = equal " ";
+val is_whitespace = is_space orf is_tab orf is_newline;
+val is_whitespace' = is_space orf is_tab;
+
+val number = Scan.many1 (is_digit o symbol);
+
+val identifier =
+  Scan.one (is_alpha o symbol) :::
+  Scan.many
+    ((is_alpha orf is_digit orf is_underscore) o symbol) @@@
+   Scan.optional (Scan.one (is_tilde o symbol) >> single) [];
+
+val long_identifier =
+  identifier @@@ (Scan.repeat1 ($$$ "." @@@ identifier) >> flat);
+
+val whitespace = Scan.many (is_whitespace o symbol);
+val whitespace' = Scan.many (is_whitespace' o symbol);
+val newline = Scan.one (is_newline o symbol);
+
+fun beginning n cs =
+  let
+    val drop_blanks = #1 o take_suffix is_whitespace;
+    val all_cs = drop_blanks cs;
+    val dots = if length all_cs > n then " ..." else "";
+  in
+    (drop_blanks (take n all_cs)
+      |> map (fn c => if is_whitespace c then " " else c)
+      |> implode) ^ dots
+  end;
+
+fun !!! text scan =
+  let
+    fun get_pos [] = " (past end-of-text!)"
+      | get_pos ((_, pos) :: _) = Position.str_of pos;
+
+    fun err (syms, msg) =
+      text ^ get_pos syms ^ " at " ^ beginning 10 (map symbol syms) ^
+      (case msg of NONE => "" | SOME s => "\n" ^ s);
+  in Scan.!! err scan end;
+
+val any_line' =
+  Scan.many (not o (Scan.is_stopper char_stopper orf (is_newline o symbol)));
+
+val any_line = whitespace' |-- any_line' --|
+  newline >> (implode o map symbol);
+
+fun gen_comment a b = $$$ a |-- !!! "missing end of comment"
+  (Scan.repeat (Scan.unless ($$$ b) any) --| $$$ b) >> make_token Comment;
+
+val c_comment = gen_comment "/*" "*/";
+val curly_comment = gen_comment "{" "}";
+
+val percent_comment = $$$ "%" |-- any_line' >> make_token Comment;
+
+fun repeatn 0 _ = Scan.succeed []
+  | repeatn n p = Scan.one p ::: repeatn (n-1) p;
+
+
+(** header of *.vcg and *.siv files (see simplifier/load_provenance.pro) **)
+
+type banner = string * string * string;
+type date = string * string * string;
+type time = string * string * string * string option;
+
+val asterisks = Scan.repeat1 (Scan.one (equal "*" o symbol));
+
+fun alphan n = repeatn n (is_alpha o symbol) >> (implode o map symbol);
+fun digitn n = repeatn n (is_digit o symbol) >> (implode o map symbol);
+
+val time =
+  digitn 2 --| $$$ ":" -- digitn 2 --| $$$ ":" -- digitn 2 --
+  Scan.option ($$$ "." |-- digitn 2) >>
+    (fn (((hr, mi), s), ms) => (hr, mi, s, ms));
+
+val date =
+  digitn 2 --| $$$ "-" -- alphan 3 --| $$$ "-" -- digitn 4 >>
+    (fn ((d, m), y) => (d, m, y));
+
+val banner = 
+  whitespace' |-- asterisks --| whitespace' --| newline :|-- (fn xs =>
+    (any_line -- any_line -- any_line >>
+       (fn ((l1, l2), l3) => (l1, l2, l3))) --|
+    whitespace' --| prfx (map symbol xs) --| whitespace' --| newline);
+
+val vcg_header = banner -- Scan.option (whitespace |--
+  $$$ "DATE :" |-- whitespace |-- date --| whitespace --|
+  Scan.option ($$$ "TIME :" -- whitespace) -- time);
+
+val siv_header = banner --| whitespace --
+  ($$$ "CREATED" |-- whitespace |-- (date --| $$$ "," --| whitespace -- time)) --|
+  whitespace --
+  ($$$ "SIMPLIFIED" |-- whitespace |-- (date --| $$$ "," --| whitespace -- time)) --|
+  newline --| newline -- (any_line -- any_line) >>
+    (fn (((b, c), s), ls) => (b, c, s, ls));
+
+
+(** the main tokenizer **)
+
+fun scan header comment =
+  !!! "bad header" header --| whitespace --
+  Scan.repeat (Scan.unless (Scan.one is_char_eof)
+    (!!! "bad input"
+       (   comment
+        || (keyword "For" -- whitespace) |--
+              Scan.repeat1 (Scan.unless (keyword ":") any) --|
+              keyword ":" >> make_token Traceability
+        || Scan.max leq_token
+             (Scan.literal lexicon >> make_token Keyword)
+             (   long_identifier >> make_token Long_Ident
+              || identifier >> make_token Ident)
+        || number >> make_token Number) --|
+     whitespace));
+
+fun tokenize header comment pos s =
+  fst (Scan.finite char_stopper
+    (Scan.error (scan header comment)) (explode_pos s pos));
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Tools/fdl_parser.ML	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,384 @@
+(*  Title:      HOL/SPARK/Tools/fdl_parser.ML
+    Author:     Stefan Berghofer
+    Copyright:  secunet Security Networks AG
+
+Parser for fdl files.
+*)
+
+signature FDL_PARSER =
+sig
+  datatype expr =
+      Ident of string
+    | Number of int
+    | Quantifier of string * string list * string * expr
+    | Funct of string * expr list
+    | Element of expr * expr list
+    | Update of expr * expr list * expr
+    | Record of string * (string * expr) list
+    | Array of string * expr option *
+        ((expr * expr option) list list * expr) list;
+
+  datatype fdl_type =
+      Basic_Type of string
+    | Enum_Type of string list
+    | Array_Type of string list * string
+    | Record_Type of (string list * string) list
+    | Pending_Type;
+
+  datatype fdl_rule =
+      Inference_Rule of expr list * expr
+    | Substitution_Rule of expr list * expr * expr;
+
+  type rules =
+    ((string * int) * fdl_rule) list *
+    (string * (expr * (string * string) list) list) list;
+
+  type vcs = (string * (string *
+    ((string * expr) list * (string * expr) list)) list) list;
+
+  type 'a tab = 'a Symtab.table * (string * 'a) list;
+
+  val lookup: 'a tab -> string -> 'a option;
+  val update: string * 'a -> 'a tab -> 'a tab;
+  val items: 'a tab -> (string * 'a) list;
+
+  type decls =
+    {types: fdl_type tab,
+     vars: string tab,
+     consts: string tab,
+     funs: (string list * string) tab};
+
+  val parse_vcs: (Fdl_Lexer.chars -> 'a * Fdl_Lexer.chars) -> Position.T ->
+    string -> 'a * ((string * string) * vcs);
+
+  val parse_declarations:  Position.T -> string -> (string * string) * decls;
+
+  val parse_rules: Position.T -> string -> rules;
+end;
+
+structure Fdl_Parser: FDL_PARSER =
+struct
+
+(** error handling **)
+
+fun beginning n cs =
+  let val dots = if length cs > n then " ..." else "";
+  in
+    space_implode " " (take n cs) ^ dots
+  end;
+
+fun !!! scan =
+  let
+    fun get_pos [] = " (past end-of-file!)"
+      | get_pos (tok :: _) = Fdl_Lexer.pos_of tok;
+
+    fun err (syms, msg) =
+      "Syntax error" ^ get_pos syms ^ " at " ^
+      beginning 10 (map Fdl_Lexer.unparse syms) ^
+      (case msg of NONE => "" | SOME s => "\n" ^ s ^ " expected");
+  in Scan.!! err scan end;
+
+fun parse_all p =
+  Scan.repeat (Scan.unless (Scan.one Fdl_Lexer.is_eof) (!!! p));
+
+
+(** parsers **)
+
+fun group s p = p || Scan.fail_with (K s);
+
+fun $$$ s = group s
+  (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Keyword andalso
+     Fdl_Lexer.content_of t = s) >> K s);
+
+val identifier = group "identifier"
+  (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Ident) >>
+     Fdl_Lexer.content_of);
+
+val long_identifier = group "long identifier"
+  (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Long_Ident) >>
+     Fdl_Lexer.content_of);
+
+fun the_identifier s = group s
+  (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Ident andalso
+     Fdl_Lexer.content_of t = s) >> K s);
+
+fun prfx_identifier g s = group g
+  (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Ident andalso
+     can (unprefix s) (Fdl_Lexer.content_of t)) >>
+     (unprefix s o Fdl_Lexer.content_of));
+
+val mk_identifier = prfx_identifier "identifier \"mk__...\"" "mk__";
+val hyp_identifier = prfx_identifier "hypothesis identifer" "H";
+val concl_identifier = prfx_identifier "conclusion identifier" "C";
+
+val number = group "number"
+  (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Number) >>
+     (the o Int.fromString o Fdl_Lexer.content_of));
+
+val traceability = group "traceability information"
+  (Scan.one (fn t => Fdl_Lexer.kind_of t = Fdl_Lexer.Traceability) >>
+     Fdl_Lexer.content_of);
+
+fun enum1 sep scan = scan ::: Scan.repeat ($$$ sep |-- !!! scan);
+fun enum sep scan = enum1 sep scan || Scan.succeed [];
+
+fun list1 scan = enum1 "," scan;
+fun list scan = enum "," scan;
+
+
+(* expressions, see section 4.4 of SPARK Proof Manual *)
+
+datatype expr =
+    Ident of string
+  | Number of int
+  | Quantifier of string * string list * string * expr
+  | Funct of string * expr list
+  | Element of expr * expr list
+  | Update of expr * expr list * expr
+  | Record of string * (string * expr) list
+  | Array of string * expr option *
+      ((expr * expr option) list list * expr) list;
+
+fun unop (f, x) = Funct (f, [x]);
+
+fun binop p q = p :|-- (fn x => Scan.optional
+  (q -- !!! p >> (fn (f, y) => Funct (f, [x, y]))) x);
+
+(* left-associative *)
+fun binops opp argp =
+  argp -- Scan.repeat (opp -- !!! argp) >> (fn (x, fys) =>
+    fold (fn (f, y) => fn x => Funct (f, [x, y])) fys x);
+
+(* right-associative *)
+fun binops' f p = enum1 f p >> foldr1 (fn (x, y) => Funct (f, [x, y]));
+
+val multiplying_operator = $$$ "*" || $$$ "/" || $$$ "div" || $$$ "mod";
+
+val adding_operator = $$$ "+" || $$$ "-";
+
+val relational_operator =
+     $$$ "=" || $$$ "<>"
+  || $$$ "<" || $$$ ">"
+  || $$$ "<="|| $$$ ">=";
+
+val quantification_kind = $$$ "for_all" || $$$ "for_some";
+
+val quantification_generator =
+  list1 identifier --| $$$ ":" -- identifier;
+
+fun expression xs = group "expression"
+  (binop disjunction ($$$ "->" || $$$ "<->")) xs
+
+and disjunction xs = binops' "or" conjunction xs
+
+and conjunction xs = binops' "and" negation xs
+
+and negation xs =
+  (   $$$ "not" -- !!! relation >> unop
+   || relation) xs
+
+and relation xs = binop sum relational_operator xs
+
+and sum xs = binops adding_operator term xs
+
+and term xs = binops multiplying_operator factor xs
+
+and factor xs =
+  (   $$$ "+" |-- !!! primary
+   || $$$ "-" -- !!! primary >> unop
+   || binop primary ($$$ "**")) xs
+
+and primary xs = group "primary"
+  (   number >> Number
+   || $$$ "(" |-- !!! (expression --| $$$ ")")
+   || quantified_expression
+   || function_designator
+   || identifier >> Ident) xs
+
+and quantified_expression xs = (quantification_kind --
+  !!! ($$$ "(" |-- quantification_generator --|  $$$ "," --
+    expression --| $$$ ")") >> (fn (q, ((xs, T), e)) =>
+      Quantifier (q, xs, T, e))) xs
+
+and function_designator xs =
+  (   mk_identifier --| $$$ "(" :|--
+        (fn s => record_args s || array_args s) --| $$$ ")"
+   || $$$ "element" |-- !!! ($$$ "(" |-- expression --| $$$ "," --| $$$ "[" --
+        list1 expression --| $$$ "]" --| $$$ ")") >> Element
+   || $$$ "update" |-- !!! ($$$ "(" |-- expression --| $$$ "," --| $$$ "[" --
+        list1 expression --| $$$ "]" --| $$$ "," -- expression --| $$$ ")") >>
+          (fn ((A, xs), x) => Update (A, xs, x))
+   || identifier --| $$$ "(" -- !!! (list1 expression --| $$$ ")") >> Funct) xs
+
+and record_args s xs =
+  (list1 (identifier --| $$$ ":=" -- !!! expression) >> (pair s #> Record)) xs
+
+and array_args s xs =
+  (   expression -- Scan.optional ($$$ "," |-- !!! array_associations) [] >>
+        (fn (default, assocs) => Array (s, SOME default, assocs))
+   || array_associations >> (fn assocs => Array (s, NONE, assocs))) xs
+
+and array_associations xs =
+  (list1 (enum1 "&" ($$$ "[" |--
+     !!! (list1 (expression -- Scan.option ($$$ ".." |-- !!! expression)) --|
+       $$$ "]")) --| $$$ ":=" -- expression)) xs;
+
+
+(* verification conditions *)
+
+type vcs = (string * (string *
+  ((string * expr) list * (string * expr) list)) list) list;
+
+val vc =
+  identifier --| $$$ "." -- !!!
+    (   $$$ "***" |-- !!! (the_identifier "true" --| $$$ ".") >>
+          (Ident #> pair "1" #> single #> pair [])
+     || $$$ "!!!" |-- !!! (the_identifier "false" --| $$$ ".") >>
+          (Ident #> pair "1" #> single #> pair [])
+     || Scan.repeat1 (hyp_identifier --
+          !!! ($$$ ":" |-- expression --| $$$ ".")) --| $$$ "->" --
+        Scan.repeat1 (concl_identifier --
+          !!! ($$$ ":" |-- expression --| $$$ ".")));
+
+val subprogram_kind = $$$ "function" || $$$ "procedure";
+
+val vcs =
+  subprogram_kind -- (long_identifier || identifier) --
+  parse_all (traceability -- !!! (Scan.repeat1 vc));
+
+fun parse_vcs header pos s =
+  s |>
+  Fdl_Lexer.tokenize header Fdl_Lexer.c_comment pos ||>
+  filter Fdl_Lexer.is_proper ||>
+  Scan.finite Fdl_Lexer.stopper (Scan.error (!!! vcs)) ||>
+  fst;
+
+
+(* fdl declarations, see section 4.3 of SPARK Proof Manual *)
+
+datatype fdl_type =
+    Basic_Type of string
+  | Enum_Type of string list
+  | Array_Type of string list * string
+  | Record_Type of (string list * string) list
+  | Pending_Type;
+
+(* also store items in a list to preserve order *)
+type 'a tab = 'a Symtab.table * (string * 'a) list;
+
+fun lookup ((tab, _) : 'a tab) = Symtab.lookup tab;
+fun update decl (tab, items) = (Symtab.update_new decl tab, decl :: items);
+fun items ((_, items) : 'a tab) = rev items;
+
+type decls =
+  {types: fdl_type tab,
+   vars: string tab,
+   consts: string tab,
+   funs: (string list * string) tab};
+
+val empty_decls : decls =
+  {types = (Symtab.empty, []), vars = (Symtab.empty, []),
+   consts = (Symtab.empty, []), funs = (Symtab.empty, [])};
+
+fun add_type_decl decl {types, vars, consts, funs} =
+  {types = update decl types,
+   vars = vars, consts = consts, funs = funs}
+  handle Symtab.DUP s => error ("Duplicate type " ^ s);
+
+fun add_var_decl (vs, ty) {types, vars, consts, funs} =
+  {types = types,
+   vars = fold (update o rpair ty) vs vars,
+   consts = consts, funs = funs}
+  handle Symtab.DUP s => error ("Duplicate variable " ^ s);
+
+fun add_const_decl decl {types, vars, consts, funs} =
+  {types = types, vars = vars,
+   consts = update decl consts,
+   funs = funs}
+  handle Symtab.DUP s => error ("Duplicate constant " ^ s);
+
+fun add_fun_decl decl {types, vars, consts, funs} =
+  {types = types, vars = vars, consts = consts,
+   funs = update decl funs}
+  handle Symtab.DUP s => error ("Duplicate function " ^ s);
+
+val type_decl = $$$ "type" |-- !!! (identifier --| $$$ "=" --
+  (   identifier >> Basic_Type
+   || $$$ "(" |-- !!! (list1 identifier --| $$$ ")") >> Enum_Type
+   || $$$ "array" |-- !!! ($$$ "[" |-- list1 identifier --| $$$ "]" --|
+        $$$ "of" -- identifier) >> Array_Type
+   || $$$ "record" |-- !!! (enum1 ";"
+        (list1 identifier -- !!! ($$$ ":" |-- identifier)) --|
+           $$$ "end") >> Record_Type
+   || $$$ "pending" >> K Pending_Type)) >> add_type_decl;
+
+val const_decl = $$$ "const" |-- !!! (identifier --| $$$ ":" -- identifier --|
+  $$$ "=" --| $$$ "pending") >> add_const_decl;
+
+val var_decl = $$$ "var" |-- !!! (list1 identifier --| $$$ ":" -- identifier) >>
+  add_var_decl;
+
+val fun_decl = $$$ "function" |-- !!! (identifier --
+  (Scan.optional ($$$ "(" |-- !!! (list1 identifier --| $$$ ")")) [] --|
+   $$$ ":" -- identifier)) >> add_fun_decl;
+
+val declarations =
+  $$$ "title" |-- subprogram_kind -- identifier --| $$$ ";" --
+  (Scan.repeat ((type_decl || const_decl || var_decl || fun_decl) --|
+     !!! ($$$ ";")) >> (fn ds => apply ds empty_decls)) --|
+  $$$ "end" --| $$$ ";"
+
+fun parse_declarations pos s =
+  s |>
+  Fdl_Lexer.tokenize (Scan.succeed ()) Fdl_Lexer.curly_comment pos |>
+  snd |> filter Fdl_Lexer.is_proper |>
+  Scan.finite Fdl_Lexer.stopper (Scan.error (!!! declarations)) |>
+  fst;
+
+
+(* rules, see section 5 of SPADE Proof Checker Rules Manual *)
+
+datatype fdl_rule =
+    Inference_Rule of expr list * expr
+  | Substitution_Rule of expr list * expr * expr;
+
+type rules =
+  ((string * int) * fdl_rule) list *
+  (string * (expr * (string * string) list) list) list;
+
+val condition_list = $$$ "[" |-- list expression --| $$$ "]";
+val if_condition_list = $$$ "if" |-- !!! condition_list;
+
+val rule =
+  identifier -- !!! ($$$ "(" |-- number --| $$$ ")" --| $$$ ":" --
+  (expression :|-- (fn e =>
+        $$$ "may_be_deduced" >> K (Inference_Rule ([], e))
+     || $$$ "may_be_deduced_from" |--
+          !!! condition_list >> (Inference_Rule o rpair e)
+     || $$$ "may_be_replaced_by" |-- !!! (expression --
+          Scan.optional if_condition_list []) >> (fn (e', cs) =>
+            Substitution_Rule (cs, e, e'))
+     || $$$ "&" |-- !!! (expression --| $$$ "are_interchangeable" --
+          Scan.optional if_condition_list []) >> (fn (e', cs) =>
+            Substitution_Rule (cs, e, e')))) --| $$$ ".") >>
+    (fn (id, (n, rl)) => ((id, n), rl));
+
+val rule_family = 
+  $$$ "rule_family" |-- identifier --| $$$ ":" --
+  enum1 "&" (expression -- !!! ($$$ "requires" |-- $$$ "[" |--
+    list (identifier -- !!! ($$$ ":" |-- identifier)) --| $$$ "]")) --|
+  $$$ ".";
+
+val rules =
+  parse_all (rule >> (apfst o cons) || rule_family >> (apsnd o cons)) >>
+  (fn rls => apply (rev rls) ([], []));
+
+fun parse_rules pos s =
+  s |>
+  Fdl_Lexer.tokenize (Scan.succeed ())
+    (Fdl_Lexer.c_comment || Fdl_Lexer.percent_comment) pos |>
+  snd |> filter Fdl_Lexer.is_proper |>
+  Scan.finite Fdl_Lexer.stopper (Scan.error (!!! rules)) |>
+  fst;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Tools/spark_commands.ML	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,155 @@
+(*  Title:      HOL/SPARK/Tools/spark_commands.ML
+    Author:     Stefan Berghofer
+    Copyright:  secunet Security Networks AG
+
+Isar commands for handling SPARK/Ada verification conditions.
+*)
+
+signature SPARK_COMMANDS =
+sig
+  val setup: theory -> theory
+end
+
+structure SPARK_Commands: SPARK_COMMANDS =
+struct
+
+fun read f path = f (Position.file (Path.implode path)) (File.read path);
+
+fun spark_open vc_name thy =
+  let
+    val (vc_path, _) = Thy_Load.check_file
+      [Thy_Load.master_directory thy] (Path.explode vc_name);
+    val (base, header) = (case Path.split_ext vc_path of
+        (base, "vcg") => (base, Fdl_Lexer.vcg_header >> K ())
+      | (base, "siv") => (base, Fdl_Lexer.siv_header >> K ())
+      | _ => error "File name must end with .vcg or .siv");
+    val fdl_path = Path.ext "fdl" base;
+    val rls_path = Path.ext "rls" base;
+  in
+    SPARK_VCs.set_vcs
+      (snd (read Fdl_Parser.parse_declarations fdl_path))
+      (read Fdl_Parser.parse_rules rls_path)
+      (snd (snd (read (Fdl_Parser.parse_vcs header) vc_path)))
+      base thy
+  end;
+
+fun add_proof_fun_cmd pf thy =
+  let val ctxt = ProofContext.init_global thy
+  in SPARK_VCs.add_proof_fun
+    (fn optT => Syntax.parse_term ctxt #>
+       the_default I (Option.map Type.constraint optT) #>
+       Syntax.check_term ctxt) pf thy
+  end;
+
+fun get_vc thy vc_name =
+  (case SPARK_VCs.lookup_vc thy vc_name of
+    SOME (ctxt, (_, proved, ctxt', stmt)) =>
+      if proved then
+        error ("The verification condition " ^
+          quote vc_name ^ " has already been proved.")
+      else (ctxt @ [ctxt'], stmt)
+  | NONE => error ("There is no verification condition " ^
+      quote vc_name ^ "."));
+
+fun prove_vc vc_name lthy =
+  let
+    val thy = ProofContext.theory_of lthy;
+    val (ctxt, stmt) = get_vc thy vc_name
+  in
+    Specification.theorem Thm.theoremK NONE
+      (K (Local_Theory.background_theory (SPARK_VCs.mark_proved vc_name)))
+      (Binding.name vc_name, []) ctxt stmt true lthy
+  end;
+
+fun string_of_status false = "(unproved)"
+  | string_of_status true = "(proved)";
+
+fun chunks ps = Pretty.blk (0,
+  flat (separate [Pretty.fbrk, Pretty.fbrk] (map single ps)));
+
+fun show_status (p, f) = Toplevel.no_timing o Toplevel.keep (fn state =>
+  let
+    val thy = Toplevel.theory_of state;
+
+    val (context, defs, vcs) = SPARK_VCs.get_vcs thy;
+
+    val vcs' = AList.coalesce (op =) (map_filter
+      (fn (name, (trace, status, ctxt, stmt)) =>
+         if p status then
+           SOME (trace, (name, status, ctxt, stmt))
+         else NONE) vcs);
+
+    val ctxt = state |>
+      Toplevel.theory_of |>
+      ProofContext.init_global |>
+      Context.proof_map (fold Element.init context)
+  in
+    (writeln "Context:\n";
+     Pretty.chunks (maps (Element.pretty_ctxt ctxt) context) |>
+     Pretty.writeln;
+
+     writeln "\nDefinitions:\n";
+     Pretty.chunks (map (fn (bdg, th) => Pretty.block
+       [Pretty.str (Binding.str_of bdg ^ ":"),
+        Pretty.brk 1,
+        Display.pretty_thm ctxt th])
+          defs) |>
+     Pretty.writeln;
+
+     writeln "\nVerification conditions:\n";
+     chunks (maps (fn (trace, vcs'') =>
+       Pretty.str trace ::
+       map (fn (name, status, context', stmt) =>
+         Pretty.big_list (name ^ " " ^ f status)
+           (Element.pretty_ctxt ctxt context' @
+            Element.pretty_stmt ctxt stmt)) vcs'') vcs') |>
+     Pretty.writeln)
+  end);
+
+val _ =
+  Outer_Syntax.command "spark_open"
+    "Open a new SPARK environment and load a SPARK-generated .vcg or .siv file."
+    Keyword.thy_decl
+    (Parse.name >> (Toplevel.theory o spark_open));
+
+val pfun_type = Scan.option
+  (Args.parens (Parse.list1 Parse.name) --| Args.colon -- Parse.name);
+
+val _ =
+  Outer_Syntax.command "spark_proof_functions"
+    "Associate SPARK proof functions with terms."
+    Keyword.thy_decl
+    (Scan.repeat1 (Parse.name -- (pfun_type --| Args.$$$ "=" -- Parse.term)) >>
+       (Toplevel.theory o fold add_proof_fun_cmd));
+
+val _ =
+  Outer_Syntax.command "spark_vc"
+    "Enter into proof mode for a specific verification condition."
+    Keyword.thy_goal
+    (Parse.name >> (fn name =>
+      (Toplevel.print o Toplevel.local_theory_to_proof NONE (prove_vc name))));
+
+val _ =
+  Outer_Syntax.improper_command "spark_status"
+    "Show the name and state of all loaded verification conditions."
+    Keyword.diag
+    (Scan.optional
+       (Args.parens
+          (   Args.$$$ "proved" >> K (I, K "")
+           || Args.$$$ "unproved" >> K (not, K "")))
+       (K true, string_of_status) >> show_status);
+
+val _ =
+  Outer_Syntax.command "spark_end"
+    "Close the current SPARK environment."
+    Keyword.thy_decl
+    (Scan.succeed (Toplevel.theory SPARK_VCs.close));
+
+val setup = Theory.at_end (fn thy =>
+  let
+    val _ = SPARK_VCs.is_closed thy
+      orelse error ("Found the end of the theory, " ^ 
+        "but the last SPARK environment is still open.")
+  in NONE end);
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SPARK/Tools/spark_vcs.ML	Sat Jan 15 12:35:29 2011 +0100
@@ -0,0 +1,870 @@
+(*  Title:      HOL/SPARK/Tools/spark_vcs.ML
+    Author:     Stefan Berghofer
+    Copyright:  secunet Security Networks AG
+
+Store for verification conditions generated by SPARK/Ada.
+*)
+
+signature SPARK_VCS =
+sig
+  val set_vcs: Fdl_Parser.decls -> Fdl_Parser.rules -> Fdl_Parser.vcs ->
+    Path.T -> theory -> theory
+  val add_proof_fun: (typ option -> 'a -> term) ->
+    string * ((string list * string) option * 'a) ->
+    theory -> theory
+  val lookup_vc: theory -> string -> (Element.context_i list *
+    (string * bool * Element.context_i * Element.statement_i)) option
+  val get_vcs: theory ->
+    Element.context_i list * (binding * thm) list *
+    (string * (string * bool * Element.context_i * Element.statement_i)) list
+  val mark_proved: string -> theory -> theory
+  val close: theory -> theory
+  val is_closed: theory -> bool
+end;
+
+structure SPARK_VCs: SPARK_VCS =
+struct
+
+open Fdl_Parser;
+
+
+(** utilities **)
+
+fun mk_unop s t =
+  let val T = fastype_of t
+  in Const (s, T --> T) $ t end;
+
+fun mk_times (t, u) =
+  let
+    val setT = fastype_of t;
+    val T = HOLogic.dest_setT setT;
+    val U = HOLogic.dest_setT (fastype_of u)
+  in
+    Const (@{const_name Sigma}, setT --> (T --> HOLogic.mk_setT U) -->
+      HOLogic.mk_setT (HOLogic.mk_prodT (T, U))) $ t $ Abs ("", T, u)
+  end;
+
+fun mk_type _ "integer" = HOLogic.intT
+  | mk_type _ "boolean" = HOLogic.boolT
+  | mk_type thy ty = Syntax.check_typ (ProofContext.init_global thy)
+      (Type (Sign.full_name thy (Binding.name ty), []));
+
+val booleanN = "boolean";
+val integerN = "integer";
+
+fun mk_qual_name thy s s' =
+  Sign.full_name thy (Binding.qualify true s (Binding.name s'));
+
+fun define_overloaded (def_name, eq) lthy =
+  let
+    val ((c, _), rhs) = eq |> Syntax.check_term lthy |>
+      Logic.dest_equals |>> dest_Free;
+    val ((_, (_, thm)), lthy') = Local_Theory.define
+      ((Binding.name c, NoSyn), ((Binding.name def_name, []), rhs)) lthy
+    val ctxt_thy = ProofContext.init_global (ProofContext.theory_of lthy');
+    val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm
+  in (thm', lthy') end;
+
+fun strip_underscores s =
+  strip_underscores (unsuffix "_" s) handle Fail _ => s;
+
+fun strip_tilde s =
+  unsuffix "~" s ^ "_init" handle Fail _ => s;
+
+val mangle_name = strip_underscores #> strip_tilde;
+
+fun mk_variables thy xs ty (tab, ctxt) =
+  let
+    val T = mk_type thy ty;
+    val (ys, ctxt') = Name.variants (map mangle_name xs) ctxt;
+    val zs = map (Free o rpair T) ys;
+  in (zs, (fold (Symtab.update o apsnd (rpair ty)) (xs ~~ zs) tab, ctxt')) end;
+
+
+(** generate properties of enumeration types **)
+
+fun add_enum_type tyname els (tab, ctxt) thy =
+  let
+    val tyb = Binding.name tyname;
+    val tyname' = Sign.full_name thy tyb;
+    val T = Type (tyname', []);
+    val case_name = mk_qual_name thy tyname (tyname ^ "_case");
+    val cs = map (fn s => Const (mk_qual_name thy tyname s, T)) els;
+    val k = length els;
+    val p = Const (@{const_name pos}, T --> HOLogic.intT);
+    val v = Const (@{const_name val}, HOLogic.intT --> T);
+    val card = Const (@{const_name card},
+      HOLogic.mk_setT T --> HOLogic.natT) $ HOLogic.mk_UNIV T;
+
+    fun mk_binrel_def s f = Logic.mk_equals
+      (Const (s, T --> T --> HOLogic.boolT),
+       Abs ("x", T, Abs ("y", T,
+         Const (s, HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $
+           (f $ Bound 1) $ (f $ Bound 0))));
+
+    val (((def1, def2), def3), lthy) = thy |>
+      Datatype.add_datatype {strict = true, quiet = true} [tyname]
+        [([], tyb, NoSyn,
+          map (fn s => (Binding.name s, [], NoSyn)) els)] |> snd |>
+
+      Class.instantiation ([tyname'], [], @{sort enum}) |>
+
+      define_overloaded ("pos_" ^ tyname ^ "_def", Logic.mk_equals
+        (p,
+         list_comb (Const (case_name, replicate k HOLogic.intT @
+             [T] ---> HOLogic.intT),
+           map (HOLogic.mk_number HOLogic.intT) (0 upto k - 1)))) ||>>
+
+      define_overloaded ("less_eq_" ^ tyname ^ "_def",
+        mk_binrel_def @{const_name less_eq} p) ||>>
+      define_overloaded ("less_" ^ tyname ^ "_def",
+        mk_binrel_def @{const_name less} p);
+
+    val UNIV_eq = Goal.prove lthy [] []
+      (HOLogic.mk_Trueprop (HOLogic.mk_eq
+         (HOLogic.mk_UNIV T, HOLogic.mk_set T cs)))
+      (fn _ =>
+         rtac @{thm subset_antisym} 1 THEN
+         rtac @{thm subsetI} 1 THEN
+         Datatype_Aux.exh_tac (K (#exhaust (Datatype_Data.the_info
+           (ProofContext.theory_of lthy) tyname'))) 1 THEN
+         ALLGOALS (asm_full_simp_tac (simpset_of lthy)));
+
+    val finite_UNIV = Goal.prove lthy [] []
+      (HOLogic.mk_Trueprop (Const (@{const_name finite},
+         HOLogic.mk_setT T --> HOLogic.boolT) $ HOLogic.mk_UNIV T))
+      (fn _ => simp_tac (simpset_of lthy addsimps [UNIV_eq]) 1);
+
+    val card_UNIV = Goal.prove lthy [] []
+      (HOLogic.mk_Trueprop (HOLogic.mk_eq
+         (card, HOLogic.mk_number HOLogic.natT k)))
+      (fn _ => simp_tac (simpset_of lthy addsimps [UNIV_eq]) 1);
+
+    val range_pos = Goal.prove lthy [] []
+      (HOLogic.mk_Trueprop (HOLogic.mk_eq
+         (Const (@{const_name image}, (T --> HOLogic.intT) -->
+            HOLogic.mk_setT T --> HOLogic.mk_setT HOLogic.intT) $
+              p $ HOLogic.mk_UNIV T,
+          Const (@{const_name atLeastLessThan}, HOLogic.intT -->
+            HOLogic.intT --> HOLogic.mk_setT HOLogic.intT) $
+              HOLogic.mk_number HOLogic.intT 0 $
+              (@{term int} $ card))))
+      (fn _ =>
+         simp_tac (simpset_of lthy addsimps [card_UNIV]) 1 THEN
+         simp_tac (simpset_of lthy addsimps [UNIV_eq, def1]) 1 THEN
+         rtac @{thm subset_antisym} 1 THEN
+         simp_tac (simpset_of lthy) 1 THEN
+         rtac @{thm subsetI} 1 THEN
+         asm_full_simp_tac (simpset_of lthy addsimps @{thms interval_expand}
+           delsimps @{thms atLeastLessThan_iff}) 1);
+
+    val lthy' =
+      Class.prove_instantiation_instance (fn _ =>
+        Class.intro_classes_tac [] THEN
+        rtac finite_UNIV 1 THEN
+        rtac range_pos 1 THEN
+        simp_tac (HOL_basic_ss addsimps [def3]) 1 THEN
+        simp_tac (HOL_basic_ss addsimps [def2]) 1) lthy;
+
+    val (pos_eqs, val_eqs) = split_list (map_index (fn (i, c) =>
+      let
+        val n = HOLogic.mk_number HOLogic.intT i;
+        val th = Goal.prove lthy' [] []
+          (HOLogic.mk_Trueprop (HOLogic.mk_eq (p $ c, n)))
+          (fn _ => simp_tac (simpset_of lthy' addsimps [def1]) 1);
+        val th' = Goal.prove lthy' [] []
+          (HOLogic.mk_Trueprop (HOLogic.mk_eq (v $ n, c)))
+          (fn _ =>
+             rtac (@{thm inj_pos} RS @{thm injD}) 1 THEN
+             simp_tac (simpset_of lthy' addsimps
+               [@{thm pos_val}, range_pos, card_UNIV, th]) 1)
+      in (th, th') end) cs);
+
+    val first_el = Goal.prove lthy' [] []
+      (HOLogic.mk_Trueprop (HOLogic.mk_eq
+         (Const (@{const_name first_el}, T), hd cs)))
+      (fn _ => simp_tac (simpset_of lthy' addsimps
+         [@{thm first_el_def}, hd val_eqs]) 1);
+
+    val last_el = Goal.prove lthy' [] []
+      (HOLogic.mk_Trueprop (HOLogic.mk_eq
+         (Const (@{const_name last_el}, T), List.last cs)))
+      (fn _ => simp_tac (simpset_of lthy' addsimps
+         [@{thm last_el_def}, List.last val_eqs, card_UNIV]) 1);
+
+    val simp_att = [Attrib.internal (K Simplifier.simp_add)]
+
+  in
+    ((fold (Symtab.update_new o apsnd (rpair tyname)) (els ~~ cs) tab,
+      fold Name.declare els ctxt),
+     lthy' |>
+     Local_Theory.note
+       ((Binding.name (tyname ^ "_card_UNIV"), simp_att), [card_UNIV]) ||>>
+     Local_Theory.note
+       ((Binding.name (tyname ^ "_pos"), simp_att), pos_eqs) ||>>
+     Local_Theory.note
+       ((Binding.name (tyname ^ "_val"), simp_att), val_eqs) ||>>
+     Local_Theory.note
+       ((Binding.name (tyname ^ "_first_el"), simp_att), [first_el]) ||>>
+     Local_Theory.note
+       ((Binding.name (tyname ^ "_last_el"), simp_att), [last_el]) |> snd |>
+     Local_Theory.exit_global)
+  end;
+
+
+fun add_type_def (s, Basic_Type ty) (ids, thy) =
+      (ids,
+       Typedecl.abbrev_global (Binding.name s, [], NoSyn)
+         (mk_type thy ty) thy |> snd)
+
+  | add_type_def (s, Enum_Type els) (ids, thy) = add_enum_type s els ids thy
+
+  | add_type_def (s, Array_Type (argtys, resty)) (ids, thy) =
+      (ids,
+       Typedecl.abbrev_global (Binding.name s, [], NoSyn)
+         (foldr1 HOLogic.mk_prodT (map (mk_type thy) argtys) -->
+            mk_type thy resty) thy |> snd)
+
+  | add_type_def (s, Record_Type fldtys) (ids, thy) =
+      (ids,
+       Record.add_record true ([], Binding.name s) NONE
+         (maps (fn (flds, ty) =>
+            let val T = mk_type thy ty
+            in map (fn fld => (Binding.name fld, T, NoSyn)) flds
+            end) fldtys) thy)
+
+  | add_type_def (s, Pending_Type) (ids, thy) =
+      (ids, Typedecl.typedecl_global (Binding.name s, [], NoSyn) thy |> snd);
+
+
+fun term_of_expr thy types funs pfuns =
+  let
+    fun tm_of vs (Funct ("->", [e, e'])) =
+          (HOLogic.mk_imp (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
+
+      | tm_of vs (Funct ("<->", [e, e'])) =
+          (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
+
+      | tm_of vs (Funct ("or", [e, e'])) =
+          (HOLogic.mk_disj (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
+
+      | tm_of vs (Funct ("and", [e, e'])) =
+          (HOLogic.mk_conj (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
+
+      | tm_of vs (Funct ("not", [e])) =
+          (HOLogic.mk_not (fst (tm_of vs e)), booleanN)
+
+      | tm_of vs (Funct ("=", [e, e'])) =
+          (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
+
+      | tm_of vs (Funct ("<>", [e, e'])) = (HOLogic.mk_not
+          (HOLogic.mk_eq (fst (tm_of vs e), fst (tm_of vs e'))), booleanN)
+
+      | tm_of vs (Funct ("<", [e, e'])) = (HOLogic.mk_binrel @{const_name less}
+          (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
+
+      | tm_of vs (Funct (">", [e, e'])) = (HOLogic.mk_binrel @{const_name less}
+          (fst (tm_of vs e'), fst (tm_of vs e)), booleanN)
+
+      | tm_of vs (Funct ("<=", [e, e'])) = (HOLogic.mk_binrel @{const_name less_eq}
+          (fst (tm_of vs e), fst (tm_of vs e')), booleanN)
+
+      | tm_of vs (Funct (">=", [e, e'])) = (HOLogic.mk_binrel @{const_name less_eq}
+          (fst (tm_of vs e'), fst (tm_of vs e)), booleanN)
+
+      | tm_of vs (Funct ("+", [e, e'])) = (HOLogic.mk_binop @{const_name plus}
+          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
+
+      | tm_of vs (Funct ("-", [e, e'])) = (HOLogic.mk_binop @{const_name minus}
+          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
+
+      | tm_of vs (Funct ("*", [e, e'])) = (HOLogic.mk_binop @{const_name times}
+          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
+
+      | tm_of vs (Funct ("/", [e, e'])) = (HOLogic.mk_binop @{const_name divide}
+          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
+
+      | tm_of vs (Funct ("div", [e, e'])) = (HOLogic.mk_binop @{const_name sdiv}
+          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
+
+      | tm_of vs (Funct ("mod", [e, e'])) = (HOLogic.mk_binop @{const_name smod}
+          (fst (tm_of vs e), fst (tm_of vs e')), integerN)
+
+      | tm_of vs (Funct ("-", [e])) =
+          (mk_unop @{const_name uminus} (fst (tm_of vs e)), integerN)
+
+      | tm_of vs (Funct ("**", [e, e'])) =
+          (Const (@{const_name power}, HOLogic.intT --> HOLogic.natT -->
+             HOLogic.intT) $ fst (tm_of vs e) $
+               (@{const nat} $ fst (tm_of vs e')), integerN)
+
+      | tm_of (tab, _) (Ident s) =
+          (case Symtab.lookup tab s of
+             SOME t_ty => t_ty
+           | NONE => error ("Undeclared identifier " ^ s))
+
+      | tm_of _ (Number i) = (HOLogic.mk_number HOLogic.intT i, integerN)
+
+      | tm_of vs (Quantifier (s, xs, ty, e)) =
+          let
+            val (ys, vs') = mk_variables thy xs ty vs;
+            val q = (case s of
+                "for_all" => HOLogic.mk_all
+              | "for_some" => HOLogic.mk_exists)
+          in
+            (fold_rev (fn Free (x, T) => fn t => q (x, T, t))
+               ys (fst (tm_of vs' e)),
+             booleanN)
+          end
+
+      | tm_of vs (Funct (s, es)) =
+
+          (* record field selection *)
+          (case try (unprefix "fld_") s of
+             SOME fname => (case es of
+               [e] =>
+                 let val (t, rcdty) = tm_of vs e
+                 in case lookup types rcdty of
+                     SOME (Record_Type fldtys) =>
+                       (case get_first (fn (flds, fldty) =>
+                            if member (op =) flds fname then SOME fldty
+                            else NONE) fldtys of
+                          SOME fldty =>
+                            (Const (mk_qual_name thy rcdty fname,
+                               mk_type thy rcdty --> mk_type thy fldty) $ t,
+                             fldty)
+                        | NONE => error ("Record " ^ rcdty ^
+                            " has no field named " ^ fname))
+                   | _ => error (rcdty ^ " is not a record type")
+                 end
+             | _ => error ("Function " ^ s ^ " expects one argument"))
+           | NONE =>
+
+          (* record field update *)
+          (case try (unprefix "upf_") s of
+             SOME fname => (case es of
+               [e, e'] =>
+                 let
+                   val (t, rcdty) = tm_of vs e;
+                   val rT = mk_type thy rcdty;
+                   val (u, fldty) = tm_of vs e';
+                   val fT = mk_type thy fldty
+                 in case lookup types rcdty of
+                     SOME (Record_Type fldtys) =>
+                       (case get_first (fn (flds, fldty) =>
+                            if member (op =) flds fname then SOME fldty
+                            else NONE) fldtys of
+                          SOME fldty' =>
+                            if fldty = fldty' then
+                              (Const (mk_qual_name thy rcdty (fname ^ "_update"),
+                                 (fT --> fT) --> rT --> rT) $
+                                   Abs ("x", fT, u) $ t,
+                               rcdty)
+                            else error ("Type " ^ fldty ^
+                              " does not match type " ^ fldty' ^ " of field " ^
+                              fname)
+                        | NONE => error ("Record " ^ rcdty ^
+                            " has no field named " ^ fname))
+                   | _ => error (rcdty ^ " is not a record type")
+                 end
+             | _ => error ("Function " ^ s ^ " expects two arguments"))
+           | NONE =>
+
+          (* enumeration type to integer *)
+          (case try (unsuffix "__pos") s of
+             SOME tyname => (case es of
+               [e] => (Const (@{const_name pos},
+                 mk_type thy tyname --> HOLogic.intT) $ fst (tm_of vs e), integerN)
+             | _ => error ("Function " ^ s ^ " expects one argument"))
+           | NONE =>
+
+          (* integer to enumeration type *)
+          (case try (unsuffix "__val") s of
+             SOME tyname => (case es of
+               [e] => (Const (@{const_name val},
+                 HOLogic.intT --> mk_type thy tyname) $ fst (tm_of vs e), tyname)
+             | _ => error ("Function " ^ s ^ " expects one argument"))
+           | NONE =>
+
+          (* successor / predecessor of enumeration type element *)
+          if s = "succ" orelse s = "pred" then (case es of
+              [e] =>
+                let
+                  val (t, tyname) = tm_of vs e;
+                  val T = mk_type thy tyname
+                in (Const
+                  (if s = "succ" then @{const_name succ}
+                   else @{const_name pred}, T --> T) $ t, tyname)
+                end
+            | _ => error ("Function " ^ s ^ " expects one argument"))
+
+          (* user-defined proof function *)
+          else
+            (case Symtab.lookup pfuns s of
+               SOME (SOME (_, resty), t) =>
+                 (list_comb (t, map (fst o tm_of vs) es), resty)
+             | _ => error ("Undeclared proof function " ^ s))))))
+
+      | tm_of vs (Element (e, es)) =
+          let val (t, ty) = tm_of vs e
+          in case lookup types ty of
+              SOME (Array_Type (_, elty)) =>
+                (t $ foldr1 HOLogic.mk_prod (map (fst o tm_of vs) es), elty)
+            | _ => error (ty ^ " is not an array type")
+          end
+
+      | tm_of vs (Update (e, es, e')) =
+          let val (t, ty) = tm_of vs e
+          in case lookup types ty of
+              SOME (Array_Type (idxtys, elty)) =>
+                let
+                  val T = foldr1 HOLogic.mk_prodT (map (mk_type thy) idxtys);
+                  val U = mk_type thy elty;
+                  val fT = T --> U
+                in
+                  (Const (@{const_name fun_upd}, fT --> T --> U --> fT) $
+                     t $ foldr1 HOLogic.mk_prod (map (fst o tm_of vs) es) $
+                       fst (tm_of vs e'),
+                   ty)
+                end
+            | _ => error (ty ^ " is not an array type")
+          end
+
+      | tm_of vs (Record (s, flds)) =
+          (case lookup types s of
+             SOME (Record_Type fldtys) =>
+               let
+                 val flds' = map (apsnd (tm_of vs)) flds;
+                 val fnames = maps fst fldtys;
+                 val fnames' = map fst flds;
+                 val (fvals, ftys) = split_list (map (fn s' =>
+                   case AList.lookup (op =) flds' s' of
+                     SOME fval_ty => fval_ty
+                   | NONE => error ("Field " ^ s' ^ " missing in record " ^ s))
+                       fnames);
+                 val _ = (case subtract (op =) fnames fnames' of
+                     [] => ()
+                   | xs => error ("Extra field(s) " ^ commas xs ^
+                       " in record " ^ s));
+                 val _ = (case duplicates (op =) fnames' of
+                     [] => ()
+                   | xs => error ("Duplicate field(s) " ^ commas xs ^
+                       " in record " ^ s))
+               in
+                 (list_comb
+                    (Const (mk_qual_name thy s (s ^ "_ext"),
+                       map (mk_type thy) ftys @ [HOLogic.unitT] --->
+                         mk_type thy s),
+                     fvals @ [HOLogic.unit]),
+                  s)
+               end
+           | _ => error (s ^ " is not a record type"))
+
+      | tm_of vs (Array (s, default, assocs)) =
+          (case lookup types s of
+             SOME (Array_Type (idxtys, elty)) =>
+               let
+                 val Ts = map (mk_type thy) idxtys;
+                 val T = foldr1 HOLogic.mk_prodT Ts;
+                 val U = mk_type thy elty;
+                 fun mk_idx' T (e, NONE) = HOLogic.mk_set T [fst (tm_of vs e)]
+                   | mk_idx' T (e, SOME e') = Const (@{const_name atLeastAtMost},
+                       T --> T --> HOLogic.mk_setT T) $
+                         fst (tm_of vs e) $ fst (tm_of vs e');
+                 fun mk_idx idx =
+                   if length Ts <> length idx then
+                     error ("Arity mismatch in construction of array " ^ s)
+                   else foldr1 mk_times (map2 mk_idx' Ts idx);
+                 fun mk_upd (idxs, e) t =
+                   if length idxs = 1 andalso forall (is_none o snd) (hd idxs)
+                   then
+                     Const (@{const_name fun_upd}, (T --> U) -->
+                         T --> U --> T --> U) $ t $
+                       foldl1 HOLogic.mk_prod
+                         (map (fst o tm_of vs o fst) (hd idxs)) $
+                       fst (tm_of vs e)
+                   else
+                     Const (@{const_name fun_upds}, (T --> U) -->
+                         HOLogic.mk_setT T --> U --> T --> U) $ t $
+                       foldl1 (HOLogic.mk_binop @{const_name sup})
+                         (map mk_idx idxs) $
+                       fst (tm_of vs e)
+               in
+                 (fold mk_upd assocs
+                    (case default of
+                       SOME e => Abs ("x", T, fst (tm_of vs e))
+                     | NONE => Const (@{const_name undefined}, T --> U)),
+                  s)
+               end
+           | _ => error (s ^ " is not an array type"))
+
+  in tm_of end;
+
+
+fun term_of_rule thy types funs pfuns ids rule =
+  let val tm_of = fst o term_of_expr thy types funs pfuns ids
+  in case rule of
+      Inference_Rule (es, e) => Logic.list_implies
+        (map (HOLogic.mk_Trueprop o tm_of) es, HOLogic.mk_Trueprop (tm_of e))
+    | Substitution_Rule (es, e, e') => Logic.list_implies
+        (map (HOLogic.mk_Trueprop o tm_of) es,
+         HOLogic.mk_Trueprop (HOLogic.mk_eq (tm_of e, tm_of e')))
+  end;
+
+
+val builtin = Symtab.make (map (rpair ())
+  ["->", "<->", "or", "and", "not", "=", "<>", "<", ">", "<=", ">=",
+   "+", "-", "*", "/", "div", "mod", "**"]);
+
+fun complex_expr (Number _) = false
+  | complex_expr (Ident _) = false 
+  | complex_expr (Funct (s, es)) =
+      not (Symtab.defined builtin s) orelse exists complex_expr es
+  | complex_expr (Quantifier (_, _, _, e)) = complex_expr e
+  | complex_expr _ = true;
+
+fun complex_rule (Inference_Rule (es, e)) =
+      complex_expr e orelse exists complex_expr es
+  | complex_rule (Substitution_Rule (es, e, e')) =
+      complex_expr e orelse complex_expr e' orelse
+      exists complex_expr es;
+
+val is_pfun =
+  Symtab.defined builtin orf
+  can (unprefix "fld_") orf can (unprefix "upf_") orf
+  can (unsuffix "__pos") orf can (unsuffix "__val") orf
+  equal "succ" orf equal "pred";
+
+fun fold_opt f = the_default I o Option.map f;
+fun fold_pair f g (x, y) = f x #> g y;
+
+fun fold_expr f g (Funct (s, es)) = f s #> fold (fold_expr f g) es
+  | fold_expr f g (Ident s) = g s
+  | fold_expr f g (Number _) = I
+  | fold_expr f g (Quantifier (_, _, _, e)) = fold_expr f g e
+  | fold_expr f g (Element (e, es)) =
+      fold_expr f g e #> fold (fold_expr f g) es
+  | fold_expr f g (Update (e, es, e')) =
+      fold_expr f g e #> fold (fold_expr f g) es #> fold_expr f g e'
+  | fold_expr f g (Record (_, flds)) = fold (fold_expr f g o snd) flds
+  | fold_expr f g (Array (_, default, assocs)) =
+      fold_opt (fold_expr f g) default #>
+      fold (fold_pair
+        (fold (fold (fold_pair
+          (fold_expr f g) (fold_opt (fold_expr f g)))))
+        (fold_expr f g)) assocs;
+
+val add_expr_pfuns = fold_expr
+  (fn s => if is_pfun s then I else insert (op =) s) (K I);
+
+val add_expr_idents = fold_expr (K I) (insert (op =));
+
+fun pfun_type thy (argtys, resty) =
+  map (mk_type thy) argtys ---> mk_type thy resty;
+
+fun check_pfun_type thy s t optty1 optty2 =
+  let
+    val T = fastype_of t;
+    fun check ty =
+      let val U = pfun_type thy ty
+      in
+        T = U orelse
+        error ("Type\n" ^
+          Syntax.string_of_typ_global thy T ^
+          "\nof function " ^
+          Syntax.string_of_term_global thy t ^
+          " associated with proof function " ^ s ^
+          "\ndoes not match declared type\n" ^
+          Syntax.string_of_typ_global thy U)
+      end
+  in (Option.map check optty1; Option.map check optty2; ()) end;
+
+fun upd_option x y = if is_some x then x else y;
+
+fun check_pfuns_types thy funs =
+  Symtab.map (fn s => fn (optty, t) =>
+   let val optty' = lookup funs s
+   in
+     (check_pfun_type thy s t optty optty';
+      (NONE |> upd_option optty |> upd_option optty', t))
+   end);
+
+
+(** the VC store **)
+
+fun err_unfinished () = error "An unfinished SPARK environment is still open."
+
+fun err_vcs names = error (Pretty.string_of
+  (Pretty.big_list "The following verification conditions have not been proved:"
+    (map Pretty.str names)))
+
+val strip_number = pairself implode o take_suffix Fdl_Lexer.is_digit o raw_explode;
+
+val name_ord = prod_ord string_ord (option_ord int_ord) o
+  pairself (strip_number ##> Int.fromString);
+
+structure VCtab = Table(type key = string val ord = name_ord);
+
+structure VCs = Theory_Data
+(
+  type T =
+    {pfuns: ((string list * string) option * term) Symtab.table,
+     env:
+       {ctxt: Element.context_i list,
+        defs: (binding * thm) list,
+        types: fdl_type tab,
+        funs: (string list * string) tab,
+        ids: (term * string) Symtab.table * Name.context,
+        proving: bool,
+        vcs: (string * bool *
+          (string * expr) list * (string * expr) list) VCtab.table,
+        path: Path.T} option}
+  val empty : T = {pfuns = Symtab.empty, env = NONE}
+  val extend = I
+  fun merge ({pfuns = pfuns1, env = NONE}, {pfuns = pfuns2, env = NONE}) =
+        {pfuns = Symtab.merge (eq_pair (op =) (op aconv)) (pfuns1, pfuns2),
+         env = NONE}
+    | merge _ = err_unfinished ()
+)
+
+fun set_env (env as {funs, ...}) thy = VCs.map (fn
+    {pfuns, env = NONE} =>
+      {pfuns = check_pfuns_types thy funs pfuns, env = SOME env}
+  | _ => err_unfinished ()) thy;
+
+fun mk_pat s = (case Int.fromString s of
+    SOME i => [HOLogic.mk_Trueprop (Var (("C", i), HOLogic.boolT))]
+  | NONE => error ("Bad conclusion identifier: C" ^ s));
+
+fun mk_vc thy types funs pfuns ids (tr, proved, ps, cs) =
+  let val prop_of =
+    HOLogic.mk_Trueprop o fst o term_of_expr thy types funs pfuns ids
+  in
+    (tr, proved,
+     Element.Assumes (map (fn (s', e) =>
+       ((Binding.name ("H" ^ s'), []), [(prop_of e, [])])) ps),
+     Element.Shows (map (fn (s', e) =>
+       (Attrib.empty_binding, [(prop_of e, mk_pat s')])) cs))
+  end;
+
+fun fold_vcs f vcs =
+  VCtab.fold (fn (_, (_, _, ps, cs)) => fold f ps #> fold f cs) vcs;
+
+fun pfuns_of_vcs pfuns vcs =
+  fold_vcs (add_expr_pfuns o snd) vcs [] |>
+  filter_out (Symtab.defined pfuns);
+
+fun declare_missing_pfuns thy funs pfuns vcs (tab, ctxt) =
+  let
+    val (fs, (tys, Ts)) =
+      pfuns_of_vcs pfuns vcs |>
+      map_filter (fn s => lookup funs s |>
+        Option.map (fn ty => (s, (SOME ty, pfun_type thy ty)))) |>
+      split_list ||> split_list;
+    val (fs', ctxt') = Name.variants fs ctxt
+  in
+    (fold Symtab.update_new (fs ~~ (tys ~~ map Free (fs' ~~ Ts))) pfuns,
+     Element.Fixes (map2 (fn s => fn T =>
+       (Binding.name s, SOME T, NoSyn)) fs' Ts),
+     (tab, ctxt'))
+  end;
+
+fun add_proof_fun prep (s, (optty, raw_t)) thy =
+  VCs.map (fn
+      {env = SOME {proving = true, ...}, ...} => err_unfinished ()
+    | {pfuns, env} =>
+        let
+          val optty' = (case env of
+              SOME {funs, ...} => lookup funs s
+            | NONE => NONE);
+          val optty'' = NONE |> upd_option optty |> upd_option optty';
+          val t = prep (Option.map (pfun_type thy) optty'') raw_t
+        in
+          (check_pfun_type thy s t optty optty';
+           if is_some optty'' orelse is_none env then
+             {pfuns = Symtab.update_new (s, (optty'', t)) pfuns,
+              env = env}
+               handle Symtab.DUP _ => error ("Proof function " ^ s ^
+                 " already associated with function")
+           else error ("Undeclared proof function " ^ s))
+        end) thy;
+
+val is_closed = is_none o #env o VCs.get;
+
+fun lookup_vc thy name =
+  (case VCs.get thy of
+    {env = SOME {vcs, types, funs, ids, ctxt, ...}, pfuns} =>
+      (case VCtab.lookup vcs name of
+         SOME vc =>           
+           let val (pfuns', ctxt', ids') =
+             declare_missing_pfuns thy funs pfuns vcs ids
+           in SOME (ctxt @ [ctxt'], mk_vc thy types funs pfuns' ids' vc) end
+       | NONE => NONE)
+  | _ => NONE);
+
+fun get_vcs thy = (case VCs.get thy of
+    {env = SOME {vcs, types, funs, ids, ctxt, defs, ...}, pfuns} =>
+      let val (pfuns', ctxt', ids') =
+        declare_missing_pfuns thy funs pfuns vcs ids
+      in
+        (ctxt @ [ctxt'], defs,
+         VCtab.dest vcs |>
+         map (apsnd (mk_vc thy types funs pfuns' ids')))
+      end
+  | _ => ([], [], []));
+
+fun mark_proved name = VCs.map (fn
+    {pfuns, env = SOME {ctxt, defs, types, funs, ids, vcs, path, ...}} =>
+      {pfuns = pfuns,
+       env = SOME {ctxt = ctxt, defs = defs,
+         types = types, funs = funs, ids = ids,
+         proving = true,
+         vcs = VCtab.map_entry name (fn (trace, _, ps, cs) =>
+           (trace, true, ps, cs)) vcs,
+         path = path}}
+  | x => x);
+
+fun close thy = VCs.map (fn
+    {pfuns, env = SOME {vcs, path, ...}} =>
+      (case VCtab.fold_rev (fn (s, (_, p, _, _)) =>
+           (if p then apfst else apsnd) (cons s)) vcs ([], []) of
+         (proved, []) =>
+           (File.write (Path.ext "prv" path)
+              (concat (map (fn s => snd (strip_number s) ^
+                 " -- proved by " ^ Distribution.version ^ "\n") proved));
+            {pfuns = pfuns, env = NONE})
+       | (_, unproved) => err_vcs unproved)
+  | x => x) thy;
+
+
+(** set up verification conditions **)
+
+fun partition_opt f =
+  let
+    fun part ys zs [] = (rev ys, rev zs)
+      | part ys zs (x :: xs) = (case f x of
+            SOME y => part (y :: ys) zs xs
+          | NONE => part ys (x :: zs) xs)
+  in part [] [] end;
+
+fun dest_def (id, (Substitution_Rule ([], Ident s, rhs))) = SOME (id, (s, rhs))
+  | dest_def _ = NONE;
+
+fun mk_rulename (s, i) = Binding.name (s ^ string_of_int i);
+
+fun add_const (s, ty) ((tab, ctxt), thy) =
+  let
+    val T = mk_type thy ty;
+    val b = Binding.name s;
+    val c = Const (Sign.full_name thy b, T)
+  in
+    (c,
+     ((Symtab.update (s, (c, ty)) tab, Name.declare s ctxt),
+      Sign.add_consts_i [(b, T, NoSyn)] thy))
+  end;
+
+fun add_def types funs pfuns consts (id, (s, e)) (ids as (tab, ctxt), thy) =
+  (case lookup consts s of
+     SOME ty =>
+       let
+         val (t, ty') = term_of_expr thy types funs pfuns ids e;
+         val _ = ty = ty' orelse
+           error ("Declared type " ^ ty ^ " of " ^ s ^
+             "\ndoes not match type " ^ ty' ^ " in definition");
+         val id' = mk_rulename id;
+         val lthy = Named_Target.theory_init thy;
+         val ((t', (_, th)), lthy') = Specification.definition
+           (NONE, ((id', []), HOLogic.mk_Trueprop (HOLogic.mk_eq
+             (Free (s, mk_type thy ty), t)))) lthy;
+         val phi = ProofContext.export_morphism lthy' lthy
+       in
+         ((id', Morphism.thm phi th),
+          ((Symtab.update (s, (Morphism.term phi t', ty)) tab,
+            Name.declare s ctxt),
+           Local_Theory.exit_global lthy'))
+       end
+   | NONE => error ("Undeclared constant " ^ s));
+
+fun add_var (s, ty) (ids, thy) =
+  let val ([Free p], ids') = mk_variables thy [s] ty ids
+  in (p, (ids', thy)) end;
+
+fun add_init_vars vcs (ids_thy as ((tab, _), _)) =
+  fold_map add_var
+    (map_filter
+       (fn s => case try (unsuffix "~") s of
+          SOME s' => (case Symtab.lookup tab s' of
+            SOME (_, ty) => SOME (s, ty)
+          | NONE => error ("Undeclared identifier " ^ s'))
+        | NONE => NONE)
+       (fold_vcs (add_expr_idents o snd) vcs []))
+    ids_thy;
+
+fun is_trivial_vc ([], [(_, Ident "true")]) = true
+  | is_trivial_vc _ = false;
+
+fun rulenames rules = commas
+  (map (fn ((s, i), _) => s ^ "(" ^ string_of_int i ^ ")") rules);
+
+(* sort definitions according to their dependency *)
+fun sort_defs _ _ [] sdefs = rev sdefs
+  | sort_defs pfuns consts defs sdefs =
+      (case find_first (fn (_, (_, e)) =>
+         forall (Symtab.defined pfuns) (add_expr_pfuns e []) andalso
+         forall (fn id =>
+           member (fn (s, (_, (s', _))) => s = s') sdefs id orelse
+           member (fn (s, (s', _)) => s = s') consts id)
+             (add_expr_idents e [])) defs of
+         SOME d => sort_defs pfuns consts
+           (remove (op =) d defs) (d :: sdefs)
+       | NONE => error ("Bad definitions: " ^ rulenames defs));
+
+fun set_vcs ({types, vars, consts, funs} : decls) (rules, _) vcs path thy =
+  let
+    val {pfuns, ...} = VCs.get thy;
+    val (defs', rules') = partition_opt dest_def rules;
+    val consts' =
+      subtract (fn ((_, (s, _)), (s', _)) => s = s') defs' (items consts);
+    val defs = sort_defs pfuns consts' defs' [];
+    (* ignore all complex rules in rls files *)
+    val (rules'', other_rules) =
+      List.partition (complex_rule o snd) rules';
+    val _ = if null rules'' then ()
+      else warning ("Ignoring rules: " ^ rulenames rules'');
+
+    val vcs' = VCtab.make (maps (fn (tr, vcs) =>
+      map (fn (s, (ps, cs)) => (s, (tr, false, ps, cs)))
+        (filter_out (is_trivial_vc o snd) vcs)) vcs);
+
+    val _ = (case filter_out (is_some o lookup funs)
+        (pfuns_of_vcs pfuns vcs') of
+        [] => ()
+      | fs => error ("Undeclared proof function(s) " ^ commas fs));
+
+    val (((defs', vars''), ivars), (ids, thy')) =
+      ((Symtab.empty |>
+        Symtab.update ("false", (HOLogic.false_const, booleanN)) |>
+        Symtab.update ("true", (HOLogic.true_const, booleanN)),
+        Name.context), thy) |>
+      fold add_type_def (items types) |>
+      fold (snd oo add_const) consts' |>
+      fold_map (add_def types funs pfuns consts) defs ||>>
+      fold_map add_var (items vars) ||>>
+      add_init_vars vcs';
+
+    val ctxt =
+      [Element.Fixes (map (fn (s, T) =>
+         (Binding.name s, SOME T, NoSyn)) (vars'' @ ivars)),
+       Element.Assumes (map (fn (id, rl) =>
+         ((mk_rulename id, []),
+          [(term_of_rule thy' types funs pfuns ids rl, [])]))
+           other_rules),
+       Element.Notes (Thm.definitionK,
+         [((Binding.name "defns", []), map (rpair [] o single o snd) defs')])]
+          
+  in
+    set_env {ctxt = ctxt, defs = defs', types = types, funs = funs,
+      ids = ids, proving = false, vcs = vcs', path = path} thy'
+  end;
+
+end;