initial commit for Isabelle 2013-1
authorChristian Urban <christian dot urban at kcl dot ac dot uk>
Thu, 06 Mar 2014 13:28:38 +0000
changeset 0 1378b654acde
child 1 ed280ad05133
initial commit for Isabelle 2013-1
ROOT
Separation_Algebra/Sep_Tactics.thy
Separation_Algebra/Separation_Algebra.thy
Separation_Algebra/sep_tactics.ML
thys/Data_slot.thy
thys/Hoare_abc.thy
thys/Hoare_gen.thy
thys/Hoare_tm.thy
thys/My_block.thy
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/ROOT	Thu Mar 06 13:28:38 2014 +0000
@@ -0,0 +1,11 @@
+session "Hoare_gen" in "thys" = "HOL" +
+  theories 
+     Hoare_gen
+
+session "Hoare_tm" in "thys" = "Hoare_gen" +
+  theories 
+     Hoare_tm
+
+session "Hoare_abc" in "thys" = "Hoare_tm" +
+  theories
+     Hoare_abc
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Separation_Algebra/Sep_Tactics.thy	Thu Mar 06 13:28:38 2014 +0000
@@ -0,0 +1,127 @@
+(* Authors: Gerwin Klein and Rafal Kolanski, 2012
+   Maintainers: Gerwin Klein <kleing at cse.unsw.edu.au>
+                Rafal Kolanski <rafal.kolanski at nicta.com.au>
+*)
+
+header "Separation Logic Tactics"
+
+theory Sep_Tactics
+imports Separation_Algebra
+begin
+
+ML_file "sep_tactics.ML"
+
+text {* A number of proof methods to assist with reasoning about separation logic. *}
+
+
+section {* Selection (move-to-front) tactics *}
+
+method_setup sep_select = {*
+  Scan.lift Parse.int >> (fn n => fn ctxt => SIMPLE_METHOD' (sep_select_tac ctxt n))
+*} "Select nth separation conjunct in conclusion"
+
+method_setup sep_select_asm = {*
+  Scan.lift Parse.int >> (fn n => fn ctxt => SIMPLE_METHOD' (sep_select_asm_tac ctxt n))
+*} "Select nth separation conjunct in assumptions"
+
+
+section {* Substitution *}
+
+method_setup "sep_subst" = {*
+  Scan.lift (Args.mode "asm" -- Scan.optional (Args.parens (Scan.repeat Parse.nat)) [0]) --
+    Attrib.thms >> (fn ((asm, occs), thms) => fn ctxt =>
+      SIMPLE_METHOD' ((if asm then sep_subst_asm_tac else sep_subst_tac) ctxt occs thms))
+*}
+"single-step substitution after solving one separation logic assumption"
+
+
+section {* Forward Reasoning *}
+
+method_setup "sep_drule" = {*
+  Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD' (sep_dtac ctxt thms))
+*} "drule after separating conjunction reordering"
+
+method_setup "sep_frule" = {*
+  Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD' (sep_ftac ctxt thms))
+*} "frule after separating conjunction reordering"
+
+
+section {* Backward Reasoning *}
+
+method_setup "sep_rule" = {*
+  Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD' (sep_rtac ctxt thms))
+*} "applies rule after separating conjunction reordering"
+
+
+section {* Cancellation of Common Conjuncts via Elimination Rules *}
+
+ML {*
+  structure SepCancel_Rules = Named_Thms (
+    val name = @{binding "sep_cancel"};
+    val description = "sep_cancel rules";
+  );
+*}
+setup SepCancel_Rules.setup
+
+text {*
+  The basic @{text sep_cancel_tac} is minimal. It only eliminates
+  erule-derivable conjuncts between an assumption and the conclusion.
+
+  To have a more useful tactic, we augment it with more logic, to proceed as
+  follows:
+  \begin{itemize}
+  \item try discharge the goal first using @{text tac}
+  \item if that fails, invoke @{text sep_cancel_tac}
+  \item if @{text sep_cancel_tac} succeeds
+    \begin{itemize}
+    \item try to finish off with tac (but ok if that fails)
+    \item try to finish off with @{term sep_true} (but ok if that fails)
+    \end{itemize}
+  \end{itemize}
+  *}
+
+ML {*
+  fun sep_cancel_smart_tac ctxt tac =
+    let fun TRY' tac = tac ORELSE' (K all_tac)
+    in
+      tac
+      ORELSE' (sep_cancel_tac ctxt tac
+               THEN' TRY' tac
+               THEN' TRY' (rtac @{thm TrueI}))
+      ORELSE' (etac @{thm sep_conj_sep_emptyE}
+               THEN' sep_cancel_tac ctxt tac
+               THEN' TRY' tac
+               THEN' TRY' (rtac @{thm TrueI}))
+    end;
+
+  fun sep_cancel_smart_tac_rules ctxt etacs =
+      sep_cancel_smart_tac ctxt (FIRST' ([atac] @ etacs));
+
+  val sep_cancel_syntax = Method.sections [
+    Args.add -- Args.colon >> K (I, SepCancel_Rules.add)] >> K ();
+*}
+
+method_setup sep_cancel = {*
+  sep_cancel_syntax >> (fn _ => fn ctxt =>
+    let
+      val etacs = map etac (SepCancel_Rules.get ctxt);
+    in
+      SIMPLE_METHOD' (sep_cancel_smart_tac_rules ctxt etacs)
+    end)
+*} "Separating conjunction conjunct cancellation"
+
+text {*
+  As above, but use blast with a depth limit to figure out where cancellation
+  can be done. *}
+
+method_setup sep_cancel_blast = {*
+  sep_cancel_syntax >> (fn _ => fn ctxt =>
+    let
+      val rules = SepCancel_Rules.get ctxt;
+      val tac = Blast.depth_tac (ctxt addIs rules) 10;
+    in
+      SIMPLE_METHOD' (sep_cancel_smart_tac ctxt tac)
+    end)
+*} "Separating conjunction conjunct cancellation using blast"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Separation_Algebra/Separation_Algebra.thy	Thu Mar 06 13:28:38 2014 +0000
@@ -0,0 +1,816 @@
+(* Authors: Gerwin Klein and Rafal Kolanski, 2012
+   Maintainers: Gerwin Klein <kleing at cse.unsw.edu.au>
+                Rafal Kolanski <rafal.kolanski at nicta.com.au>
+*)
+
+header "Abstract Separation Algebra"
+
+theory Separation_Algebra
+imports Main
+begin
+
+
+text {* This theory is the main abstract separation algebra development *}
+
+
+section {* Input syntax for lifting boolean predicates to separation predicates *}
+
+abbreviation (input)
+  pred_and :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" (infixr "and" 35) where
+  "a and b \<equiv> \<lambda>s. a s \<and> b s"
+
+abbreviation (input)
+  pred_or :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" (infixr "or" 30) where
+  "a or b \<equiv> \<lambda>s. a s \<or> b s"
+
+abbreviation (input)
+  pred_not :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" ("not _" [40] 40) where
+  "not a \<equiv> \<lambda>s. \<not>a s"
+
+abbreviation (input)
+  pred_imp :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" (infixr "imp" 25) where
+  "a imp b \<equiv> \<lambda>s. a s \<longrightarrow> b s"
+
+abbreviation (input)
+  pred_K :: "'b \<Rightarrow> 'a \<Rightarrow> 'b" ("\<langle>_\<rangle>") where
+  "\<langle>f\<rangle> \<equiv> \<lambda>s. f"
+
+(* abbreviation *)
+  definition pred_ex :: "('b \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" (binder "EXS " 10) where
+  "EXS x. P x \<equiv> \<lambda>s. \<exists>x. P x s"
+
+abbreviation (input)
+  pred_all :: "('b \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" (binder "ALLS " 10) where
+  "ALLS x. P x \<equiv> \<lambda>s. \<forall>x. P x s"
+
+
+section {* Associative/Commutative Monoid Basis of Separation Algebras *}
+
+class pre_sep_algebra = zero + plus +
+  fixes sep_disj :: "'a => 'a => bool" (infix "##" 60)
+
+  assumes sep_disj_zero [simp]: "x ## 0"
+  assumes sep_disj_commuteI: "x ## y \<Longrightarrow> y ## x"
+
+  assumes sep_add_zero [simp]: "x + 0 = x"
+  assumes sep_add_commute: "x ## y \<Longrightarrow> x + y = y + x"
+
+  assumes sep_add_assoc:
+    "\<lbrakk> x ## y; y ## z; x ## z \<rbrakk> \<Longrightarrow> (x + y) + z = x + (y + z)"
+begin
+
+lemma sep_disj_commute: "x ## y = y ## x"
+  by (blast intro: sep_disj_commuteI)
+
+lemma sep_add_left_commute:
+  assumes a: "a ## b" "b ## c" "a ## c"
+  shows "b + (a + c) = a + (b + c)" (is "?lhs = ?rhs")
+proof -
+  have "?lhs = b + a + c" using a
+    by (simp add: sep_add_assoc[symmetric] sep_disj_commute)
+  also have "... = a + b + c" using a
+    by (simp add: sep_add_commute sep_disj_commute)
+  also have "... = ?rhs" using a
+    by (simp add: sep_add_assoc sep_disj_commute)
+  finally show ?thesis .
+qed
+
+lemmas sep_add_ac = sep_add_assoc sep_add_commute sep_add_left_commute
+                    sep_disj_commute (* nearly always necessary *)
+
+end
+
+
+section {* Separation Algebra as Defined by Calcagno et al. *}
+
+class sep_algebra = pre_sep_algebra +
+  assumes sep_disj_addD1: "\<lbrakk> x ## y + z; y ## z \<rbrakk> \<Longrightarrow> x ## y"
+  assumes sep_disj_addI1: "\<lbrakk> x ## y + z; y ## z \<rbrakk> \<Longrightarrow> x + y ##  z"
+begin
+
+subsection {* Basic Construct Definitions and Abbreviations *}
+
+definition
+  sep_conj :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool)" (infixr "**" 35)
+  where
+  "P ** Q \<equiv> \<lambda>h. \<exists>x y. x ## y \<and> h = x + y \<and> P x \<and> Q y"
+
+notation
+  sep_conj (infixr "\<and>*" 35)
+
+definition
+  sep_empty :: "'a \<Rightarrow> bool" ("\<box>") where
+  "\<box> \<equiv> \<lambda>h. h = 0"
+
+definition
+  sep_impl :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool)" (infixr "\<longrightarrow>*" 25)
+  where
+  "P \<longrightarrow>* Q \<equiv> \<lambda>h. \<forall>h'. h ## h' \<and> P h' \<longrightarrow> Q (h + h')"
+
+definition
+  sep_substate :: "'a => 'a => bool" (infix "\<preceq>" 60) where
+  "x \<preceq> y \<equiv> \<exists>z. x ## z \<and> x + z = y"
+
+(* We want these to be abbreviations not definitions, because basic True and
+   False will occur by simplification in sep_conj terms *)
+abbreviation
+  "sep_true \<equiv> \<langle>True\<rangle>"
+
+abbreviation
+  "sep_false \<equiv> \<langle>False\<rangle>"
+
+definition
+  sep_list_conj :: "('a \<Rightarrow> bool) list \<Rightarrow> ('a \<Rightarrow> bool)"  ("\<And>* _" [60] 90) where
+  "sep_list_conj Ps \<equiv> foldl (op **) \<box> Ps"
+
+
+subsection {* Disjunction/Addition Properties *}
+
+lemma disjoint_zero_sym [simp]: "0 ## x"
+  by (simp add: sep_disj_commute)
+
+lemma sep_add_zero_sym [simp]: "0 + x = x"
+  by (simp add: sep_add_commute)
+
+lemma sep_disj_addD2: "\<lbrakk> x ## y + z; y ## z \<rbrakk> \<Longrightarrow> x ## z"
+  by (metis sep_disj_addD1 sep_add_ac)
+
+lemma sep_disj_addD: "\<lbrakk> x ## y + z; y ## z \<rbrakk> \<Longrightarrow> x ## y \<and> x ## z"
+  by (metis sep_disj_addD1 sep_disj_addD2)
+
+lemma sep_add_disjD: "\<lbrakk> x + y ## z; x ## y \<rbrakk> \<Longrightarrow> x ## z \<and> y ## z"
+  by (metis sep_disj_addD sep_disj_commuteI)
+
+lemma sep_disj_addI2:
+  "\<lbrakk> x ## y + z; y ## z \<rbrakk> \<Longrightarrow> x + z ## y"
+  by (metis sep_add_ac sep_disj_addI1)
+
+lemma sep_add_disjI1:
+  "\<lbrakk> x + y ## z; x ## y \<rbrakk> \<Longrightarrow> x + z ## y"
+  by (metis sep_add_ac sep_add_disjD sep_disj_addI2)
+
+lemma sep_add_disjI2:
+  "\<lbrakk> x + y ## z; x ## y \<rbrakk> \<Longrightarrow> z + y ## x"
+  by (metis sep_add_ac sep_add_disjD sep_disj_addI2)
+
+lemma sep_disj_addI3:
+   "x + y ## z \<Longrightarrow> x ## y \<Longrightarrow> x ## y + z"
+   by (metis sep_add_ac sep_add_disjD sep_add_disjI2)
+
+lemma sep_disj_add:
+  "\<lbrakk> y ## z; x ## y \<rbrakk> \<Longrightarrow> x ## y + z = x + y ## z"
+  by (metis sep_disj_addI1 sep_disj_addI3)
+
+
+subsection {* Substate Properties *}
+
+lemma sep_substate_disj_add:
+  "x ## y \<Longrightarrow> x \<preceq> x + y"
+  unfolding sep_substate_def by blast
+
+lemma sep_substate_disj_add':
+  "x ## y \<Longrightarrow> x \<preceq> y + x"
+  by (simp add: sep_add_ac sep_substate_disj_add)
+
+
+subsection {* Separating Conjunction Properties *}
+
+lemma sep_conjD:
+  "(P \<and>* Q) h \<Longrightarrow> \<exists>x y. x ## y \<and> h = x + y \<and> P x \<and> Q y"
+  by (simp add: sep_conj_def)
+
+lemma sep_conjE:
+  "\<lbrakk> (P ** Q) h; \<And>x y. \<lbrakk> P x; Q y; x ## y; h = x + y \<rbrakk> \<Longrightarrow> X \<rbrakk> \<Longrightarrow> X"
+  by (auto simp: sep_conj_def)
+
+lemma sep_conjI:
+  "\<lbrakk> P x; Q y; x ## y; h = x + y \<rbrakk> \<Longrightarrow> (P ** Q) h"
+  by (auto simp: sep_conj_def)
+
+lemma sep_conj_commuteI:
+  "(P ** Q) h \<Longrightarrow> (Q ** P) h"
+  by (auto intro!: sep_conjI elim!: sep_conjE simp: sep_add_ac)
+
+lemma sep_conj_commute:
+  "(P ** Q) = (Q ** P)"
+  by (rule ext) (auto intro: sep_conj_commuteI)
+
+lemma sep_conj_assoc:
+  "((P ** Q) ** R) = (P ** Q ** R)" (is "?lhs = ?rhs")
+proof (rule ext, rule iffI)
+  fix h
+  assume a: "?lhs h"
+  then obtain x y z where "P x" and "Q y" and "R z"
+                      and "x ## y" and "x ## z" and "y ## z" and "x + y ## z"
+                      and "h = x + y + z"
+    by (auto dest!: sep_conjD dest: sep_add_disjD)
+  moreover
+  then have "x ## y + z"
+    by (simp add: sep_disj_add)
+  ultimately
+  show "?rhs h"
+    by (auto simp: sep_add_ac intro!: sep_conjI)
+next
+  fix h
+  assume a: "?rhs h"
+  then obtain x y z where "P x" and "Q y" and "R z"
+                      and "x ## y" and "x ## z" and "y ## z" and "x ## y + z"
+                      and "h = x + y + z"
+    thm sep_disj_addD
+    by (fastforce elim!: sep_conjE simp: sep_add_ac dest: sep_disj_addD)
+  thus "?lhs h"
+    by (metis sep_conj_def sep_disj_addI1)
+qed
+
+lemma sep_conj_impl:
+  "\<lbrakk> (P ** Q) h; \<And>h. P h \<Longrightarrow> P' h; \<And>h. Q h \<Longrightarrow> Q' h \<rbrakk> \<Longrightarrow> (P' ** Q') h"
+  by (erule sep_conjE, auto intro!: sep_conjI)
+
+lemma sep_conj_impl1:
+  assumes P: "\<And>h. P h \<Longrightarrow> I h"
+  shows "(P ** R) h \<Longrightarrow> (I ** R) h"
+  by (auto intro: sep_conj_impl P)
+
+lemma sep_globalise:
+  "\<lbrakk> (P ** R) h; (\<And>h. P h \<Longrightarrow> Q h) \<rbrakk> \<Longrightarrow> (Q ** R) h"
+  by (fast elim: sep_conj_impl)
+
+lemma sep_conj_trivial_strip2:
+  "Q = R \<Longrightarrow> (Q ** P) = (R ** P)" by simp
+
+lemma disjoint_subheaps_exist:
+  "\<exists>x y. x ## y \<and> h = x + y"
+  by (rule_tac x=0 in exI, auto)
+
+lemma sep_conj_left_commute: (* for permutative rewriting *)
+  "(P ** (Q ** R)) = (Q ** (P ** R))" (is "?x = ?y")
+proof -
+  have "?x = ((Q ** R) ** P)" by (simp add: sep_conj_commute)
+  also have "\<dots> = (Q ** (R ** P))" by (subst sep_conj_assoc, simp)
+  finally show ?thesis by (simp add: sep_conj_commute)
+qed
+
+lemmas sep_conj_ac = sep_conj_commute sep_conj_assoc sep_conj_left_commute
+
+lemma ab_semigroup_mult_sep_conj: "class.ab_semigroup_mult op **"
+  by (unfold_locales)
+     (auto simp: sep_conj_ac)
+
+lemma sep_empty_zero [simp,intro!]: "\<box> 0"
+  by (simp add: sep_empty_def)
+
+
+subsection {* Properties of @{text sep_true} and @{text sep_false} *}
+
+lemma sep_conj_sep_true:
+  "P h \<Longrightarrow> (P ** sep_true) h"
+  by (simp add: sep_conjI[where y=0])
+
+lemma sep_conj_sep_true':
+  "P h \<Longrightarrow> (sep_true ** P) h"
+  by (simp add: sep_conjI[where x=0])
+
+lemma sep_conj_true [simp]:
+  "(sep_true ** sep_true) = sep_true"
+  unfolding sep_conj_def
+  by (auto intro!: ext intro: disjoint_subheaps_exist)
+
+lemma sep_conj_false_right [simp]:
+  "(P ** sep_false) = sep_false"
+  by (force elim: sep_conjE intro!: ext)
+
+lemma sep_conj_false_left [simp]:
+  "(sep_false ** P) = sep_false"
+  by (subst sep_conj_commute) (rule sep_conj_false_right)
+
+
+
+subsection {* Properties of zero (@{const sep_empty}) *}
+
+lemma sep_conj_empty [simp]:
+  "(P ** \<box>) = P"
+  by (simp add: sep_conj_def sep_empty_def)
+
+lemma sep_conj_empty'[simp]:
+  "(\<box> ** P) = P"
+  by (subst sep_conj_commute, rule sep_conj_empty)
+
+lemma sep_conj_sep_emptyI:
+  "P h \<Longrightarrow> (P ** \<box>) h"
+  by simp
+
+lemma sep_conj_sep_emptyE:
+  "\<lbrakk> P s; (P ** \<box>) s \<Longrightarrow> (Q ** R) s \<rbrakk> \<Longrightarrow> (Q ** R) s"
+  by simp
+
+lemma monoid_add: "class.monoid_add (op **) \<box>"
+  by (unfold_locales) (auto simp: sep_conj_ac)
+
+lemma comm_monoid_add: "class.comm_monoid_add op ** \<box>"
+  by (unfold_locales) (auto simp: sep_conj_ac)
+
+
+subsection {* Properties of top (@{text sep_true}) *}
+
+lemma sep_conj_true_P [simp]:
+  "(sep_true ** (sep_true ** P)) = (sep_true ** P)"
+  by (simp add: sep_conj_assoc[symmetric])
+
+lemma sep_conj_disj:
+  "((P or Q) ** R) = ((P ** R) or (Q ** R))"
+  by (auto simp: sep_conj_def intro!: ext)
+
+lemma sep_conj_sep_true_left:
+  "(P ** Q) h \<Longrightarrow> (sep_true ** Q) h"
+  by (erule sep_conj_impl, simp+)
+
+lemma sep_conj_sep_true_right:
+  "(P ** Q) h \<Longrightarrow> (P ** sep_true) h"
+  by (subst (asm) sep_conj_commute, drule sep_conj_sep_true_left,
+      simp add: sep_conj_ac)
+
+
+subsection {* Separating Conjunction with Quantifiers *}
+
+lemma sep_conj_conj:
+  "((P and Q) ** R) h \<Longrightarrow> ((P ** R) and (Q ** R)) h"
+  by (force intro: sep_conjI elim!: sep_conjE)
+
+lemma sep_conj_exists1:
+  "((EXS x. P x) ** Q) = (EXS x. (P x ** Q))"
+  by (unfold pred_ex_def, force intro!: ext intro: sep_conjI elim: sep_conjE)
+
+lemma sep_conj_exists2:
+  "(P ** (EXS x. Q x)) = (EXS x. P ** Q x)"
+  by (unfold pred_ex_def, force intro!: sep_conjI ext elim!: sep_conjE)
+
+lemmas sep_conj_exists = sep_conj_exists1 sep_conj_exists2
+
+lemma sep_conj_spec:
+  "((ALLS x. P x) ** Q) h \<Longrightarrow> (P x ** Q) h"
+  by (force intro: sep_conjI elim: sep_conjE)
+
+
+subsection {* Properties of Separating Implication *}
+
+lemma sep_implI:
+  assumes a: "\<And>h'. \<lbrakk> h ## h'; P h' \<rbrakk> \<Longrightarrow> Q (h + h')"
+  shows "(P \<longrightarrow>* Q) h"
+  unfolding sep_impl_def by (auto elim: a)
+
+lemma sep_implD:
+  "(x \<longrightarrow>* y) h \<Longrightarrow> \<forall>h'. h ## h' \<and> x h' \<longrightarrow> y (h + h')"
+  by (force simp: sep_impl_def)
+
+lemma sep_implE:
+  "(x \<longrightarrow>* y) h \<Longrightarrow> (\<forall>h'. h ## h' \<and> x h' \<longrightarrow> y (h + h') \<Longrightarrow> Q) \<Longrightarrow> Q"
+  by (auto dest: sep_implD)
+
+lemma sep_impl_sep_true [simp]:
+  "(P \<longrightarrow>* sep_true) = sep_true"
+  by (force intro!: sep_implI ext)
+
+lemma sep_impl_sep_false [simp]:
+  "(sep_false \<longrightarrow>* P) = sep_true"
+  by (force intro!: sep_implI ext)
+
+lemma sep_impl_sep_true_P:
+  "(sep_true \<longrightarrow>* P) h \<Longrightarrow> P h"
+  by (clarsimp dest!: sep_implD elim!: allE[where x=0])
+
+lemma sep_impl_sep_true_false [simp]:
+  "(sep_true \<longrightarrow>* sep_false) = sep_false"
+  by (force intro!: ext dest: sep_impl_sep_true_P)
+
+lemma sep_conj_sep_impl:
+  "\<lbrakk> P h; \<And>h. (P ** Q) h \<Longrightarrow> R h \<rbrakk> \<Longrightarrow> (Q \<longrightarrow>* R) h"
+proof (rule sep_implI)
+  fix h' h
+  assume "P h" and "h ## h'" and "Q h'"
+  hence "(P ** Q) (h + h')" by (force intro: sep_conjI)
+  moreover assume "\<And>h. (P ** Q) h \<Longrightarrow> R h"
+  ultimately show "R (h + h')" by simp
+qed
+
+lemma sep_conj_sep_impl2:
+  "\<lbrakk> (P ** Q) h; \<And>h. P h \<Longrightarrow> (Q \<longrightarrow>* R) h \<rbrakk> \<Longrightarrow> R h"
+  by (force dest: sep_implD elim: sep_conjE)
+
+lemma sep_conj_sep_impl_sep_conj2:
+  "(P ** R) h \<Longrightarrow> (P ** (Q \<longrightarrow>* (Q ** R))) h"
+  by (erule (1) sep_conj_impl, erule sep_conj_sep_impl, simp add: sep_conj_ac)
+
+
+subsection {* Pure assertions *}
+
+definition
+  pure :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where
+  "pure P \<equiv> \<forall>h h'. P h = P h'"
+
+lemma pure_sep_true:
+  "pure sep_true"
+  by (simp add: pure_def)
+
+lemma pure_sep_false:
+  "pure sep_true"
+  by (simp add: pure_def)
+
+lemma pure_split:
+  "pure P = (P = sep_true \<or> P = sep_false)"
+  by (force simp: pure_def intro!: ext)
+
+lemma pure_sep_conj:
+  "\<lbrakk> pure P; pure Q \<rbrakk> \<Longrightarrow> pure (P \<and>* Q)"
+  by (force simp: pure_split)
+
+lemma pure_sep_impl:
+  "\<lbrakk> pure P; pure Q \<rbrakk> \<Longrightarrow> pure (P \<longrightarrow>* Q)"
+  by (force simp: pure_split)
+
+lemma pure_conj_sep_conj:
+  "\<lbrakk> (P and Q) h; pure P \<or> pure Q \<rbrakk> \<Longrightarrow> (P \<and>* Q) h"
+  by (metis pure_def sep_add_zero sep_conjI sep_conj_commute sep_disj_zero)
+
+lemma pure_sep_conj_conj:
+  "\<lbrakk> (P \<and>* Q) h; pure P; pure Q \<rbrakk> \<Longrightarrow> (P and Q) h"
+  by (force simp: pure_split)
+
+lemma pure_conj_sep_conj_assoc:
+  "pure P \<Longrightarrow> ((P and Q) \<and>* R) = (P and (Q \<and>* R))"
+  by (auto simp: pure_split)
+
+lemma pure_sep_impl_impl:
+  "\<lbrakk> (P \<longrightarrow>* Q) h; pure P \<rbrakk> \<Longrightarrow> P h \<longrightarrow> Q h"
+  by (force simp: pure_split dest: sep_impl_sep_true_P)
+
+lemma pure_impl_sep_impl:
+  "\<lbrakk> P h \<longrightarrow> Q h; pure P; pure Q \<rbrakk> \<Longrightarrow> (P \<longrightarrow>* Q) h"
+  by (force simp: pure_split)
+
+lemma pure_conj_right: "(Q \<and>* (\<langle>P'\<rangle> and Q')) = (\<langle>P'\<rangle> and (Q \<and>* Q'))"
+  by (rule ext, rule, rule, clarsimp elim!: sep_conjE)
+     (erule sep_conj_impl, auto)
+
+lemma pure_conj_right': "(Q \<and>* (P' and \<langle>Q'\<rangle>)) = (\<langle>Q'\<rangle> and (Q \<and>* P'))"
+  by (simp add: conj_comms pure_conj_right)
+
+lemma pure_conj_left: "((\<langle>P'\<rangle> and Q') \<and>* Q) = (\<langle>P'\<rangle> and (Q' \<and>* Q))"
+  by (simp add: pure_conj_right sep_conj_ac)
+
+lemma pure_conj_left': "((P' and \<langle>Q'\<rangle>) \<and>* Q) = (\<langle>Q'\<rangle> and (P' \<and>* Q))"
+  by (subst conj_comms, subst pure_conj_left, simp)
+
+lemmas pure_conj = pure_conj_right pure_conj_right' pure_conj_left
+    pure_conj_left'
+
+declare pure_conj[simp add]
+
+
+subsection {* Intuitionistic assertions *}
+
+definition intuitionistic :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where
+  "intuitionistic P \<equiv> \<forall>h h'. P h \<and> h \<preceq> h' \<longrightarrow> P h'"
+
+lemma intuitionisticI:
+  "(\<And>h h'. \<lbrakk> P h; h \<preceq> h' \<rbrakk> \<Longrightarrow> P h') \<Longrightarrow> intuitionistic P"
+  by (unfold intuitionistic_def, fast)
+
+lemma intuitionisticD:
+  "\<lbrakk> intuitionistic P; P h; h \<preceq> h' \<rbrakk> \<Longrightarrow> P h'"
+  by (unfold intuitionistic_def, fast)
+
+lemma pure_intuitionistic:
+  "pure P \<Longrightarrow> intuitionistic P"
+  by (clarsimp simp: intuitionistic_def pure_def, fast)
+
+lemma intuitionistic_conj:
+  "\<lbrakk> intuitionistic P; intuitionistic Q \<rbrakk> \<Longrightarrow> intuitionistic (P and Q)"
+  by (force intro: intuitionisticI dest: intuitionisticD)
+
+lemma intuitionistic_disj:
+  "\<lbrakk> intuitionistic P; intuitionistic Q \<rbrakk> \<Longrightarrow> intuitionistic (P or Q)"
+  by (force intro: intuitionisticI dest: intuitionisticD)
+
+lemma intuitionistic_forall:
+  "(\<And>x. intuitionistic (P x)) \<Longrightarrow> intuitionistic (ALLS x. P x)"
+  by (force intro: intuitionisticI dest: intuitionisticD)
+
+lemma intuitionistic_exists:
+  "(\<And>x. intuitionistic (P x)) \<Longrightarrow> intuitionistic (EXS x. P x)"
+  by (unfold pred_ex_def, force intro: intuitionisticI dest: intuitionisticD)
+
+lemma intuitionistic_sep_conj_sep_true:
+  "intuitionistic (sep_true \<and>* P)"
+proof (rule intuitionisticI)
+  fix h h' r
+  assume a: "(sep_true \<and>* P) h"
+  then obtain x y where P: "P y" and h: "h = x + y" and xyd: "x ## y"
+    by - (drule sep_conjD, clarsimp)
+  moreover assume a2: "h \<preceq> h'"
+  then obtain z where h': "h' = h + z" and hzd: "h ## z"
+    by (clarsimp simp: sep_substate_def)
+
+  moreover have "(P \<and>* sep_true) (y + (x + z))"
+    using P h hzd xyd
+    by (metis sep_add_disjI1 sep_disj_commute sep_conjI)
+  ultimately show "(sep_true \<and>* P) h'" using hzd
+    by (auto simp: sep_conj_commute sep_add_ac dest!: sep_disj_addD)
+qed
+
+lemma intuitionistic_sep_impl_sep_true:
+  "intuitionistic (sep_true \<longrightarrow>* P)"
+proof (rule intuitionisticI)
+  fix h h'
+  assume imp: "(sep_true \<longrightarrow>* P) h" and hh': "h \<preceq> h'"
+
+  from hh' obtain z where h': "h' = h + z" and hzd: "h ## z"
+    by (clarsimp simp: sep_substate_def)
+  show "(sep_true \<longrightarrow>* P) h'" using imp h' hzd
+    apply (clarsimp dest!: sep_implD)
+    apply (metis sep_add_assoc sep_add_disjD sep_disj_addI3 sep_implI)
+    done
+qed
+
+lemma intuitionistic_sep_conj:
+  assumes ip: "intuitionistic (P::('a \<Rightarrow> bool))"
+  shows "intuitionistic (P \<and>* Q)"
+proof (rule intuitionisticI)
+  fix h h'
+  assume sc: "(P \<and>* Q) h" and hh': "h \<preceq> h'"
+
+  from hh' obtain z where h': "h' = h + z" and hzd: "h ## z"
+    by (clarsimp simp: sep_substate_def)
+
+  from sc obtain x y where px: "P x" and qy: "Q y"
+                       and h: "h = x + y" and xyd: "x ## y"
+    by (clarsimp simp: sep_conj_def)
+
+  have "x ## z" using hzd h xyd
+    by (metis sep_add_disjD)
+
+  with ip px have "P (x + z)"
+    by (fastforce elim: intuitionisticD sep_substate_disj_add)
+
+  thus "(P \<and>* Q) h'" using h' h hzd qy xyd
+    by (metis (full_types) sep_add_commute sep_add_disjD sep_add_disjI2
+              sep_add_left_commute sep_conjI)
+qed
+
+lemma intuitionistic_sep_impl:
+  assumes iq: "intuitionistic Q"
+  shows "intuitionistic (P \<longrightarrow>* Q)"
+proof (rule intuitionisticI)
+  fix h h'
+  assume imp: "(P \<longrightarrow>* Q) h" and hh': "h \<preceq> h'"
+
+  from hh' obtain z where h': "h' = h + z" and hzd: "h ## z"
+    by (clarsimp simp: sep_substate_def)
+
+  {
+    fix x
+    assume px: "P x" and hzx: "h + z ## x"
+
+    have "h + x \<preceq> h + x + z" using hzx hzd
+    by (metis sep_add_disjI1 sep_substate_def)
+
+    with imp hzd iq px hzx
+    have "Q (h + z + x)"
+    by (metis intuitionisticD sep_add_assoc sep_add_ac sep_add_disjD sep_implE)
+  }
+
+  with imp h' hzd iq show "(P \<longrightarrow>* Q) h'"
+    by (fastforce intro: sep_implI)
+qed
+
+lemma strongest_intuitionistic:
+  "\<not> (\<exists>Q. (\<forall>h. (Q h \<longrightarrow> (P \<and>* sep_true) h)) \<and> intuitionistic Q \<and>
+      Q \<noteq> (P \<and>* sep_true) \<and> (\<forall>h. P h \<longrightarrow> Q h))"
+  by (fastforce intro!: ext sep_substate_disj_add
+                dest!: sep_conjD intuitionisticD)
+
+lemma weakest_intuitionistic:
+  "\<not> (\<exists>Q. (\<forall>h. ((sep_true \<longrightarrow>* P) h \<longrightarrow> Q h)) \<and> intuitionistic Q \<and>
+      Q \<noteq> (sep_true \<longrightarrow>* P) \<and> (\<forall>h. Q h \<longrightarrow> P h))"
+  apply (clarsimp intro!: ext)
+  apply (rule iffI)
+   apply (rule sep_implI)
+   apply (drule_tac h="x" and h'="x + h'" in intuitionisticD)
+     apply (clarsimp simp: sep_add_ac sep_substate_disj_add)+
+  done
+
+lemma intuitionistic_sep_conj_sep_true_P:
+  "\<lbrakk> (P \<and>* sep_true) s; intuitionistic P \<rbrakk> \<Longrightarrow> P s"
+  by (force dest: intuitionisticD elim: sep_conjE sep_substate_disj_add)
+
+lemma intuitionistic_sep_conj_sep_true_simp:
+  "intuitionistic P \<Longrightarrow> (P \<and>* sep_true) = P"
+  by (fast intro!: sep_conj_sep_true ext
+           elim: intuitionistic_sep_conj_sep_true_P)
+
+lemma intuitionistic_sep_impl_sep_true_P:
+  "\<lbrakk> P h; intuitionistic P \<rbrakk> \<Longrightarrow> (sep_true \<longrightarrow>* P) h"
+  by (force intro!: sep_implI dest: intuitionisticD
+            intro: sep_substate_disj_add)
+
+lemma intuitionistic_sep_impl_sep_true_simp:
+  "intuitionistic P \<Longrightarrow> (sep_true \<longrightarrow>* P) = P"
+  by (fast intro!: ext
+           elim: sep_impl_sep_true_P intuitionistic_sep_impl_sep_true_P)
+
+
+subsection {* Strictly exact assertions *}
+
+definition strictly_exact :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where
+  "strictly_exact P \<equiv> \<forall>h h'. P h \<and> P h' \<longrightarrow> h = h'"
+
+lemma strictly_exactD:
+  "\<lbrakk> strictly_exact P; P h; P h' \<rbrakk> \<Longrightarrow> h = h'"
+  by (unfold strictly_exact_def, fast)
+
+lemma strictly_exactI:
+  "(\<And>h h'. \<lbrakk> P h; P h' \<rbrakk> \<Longrightarrow> h = h') \<Longrightarrow> strictly_exact P"
+  by (unfold strictly_exact_def, fast)
+
+lemma strictly_exact_sep_conj:
+  "\<lbrakk> strictly_exact P; strictly_exact Q \<rbrakk> \<Longrightarrow> strictly_exact (P \<and>* Q)"
+  apply (rule strictly_exactI)
+  apply (erule sep_conjE)+
+  apply (drule_tac h="x" and h'="xa" in strictly_exactD, assumption+)
+  apply (drule_tac h="y" and h'="ya" in strictly_exactD, assumption+)
+  apply clarsimp
+  done
+
+lemma strictly_exact_conj_impl:
+  "\<lbrakk> (Q \<and>* sep_true) h; P h; strictly_exact Q \<rbrakk> \<Longrightarrow> (Q \<and>* (Q \<longrightarrow>* P)) h"
+  by (force intro: sep_conjI sep_implI dest: strictly_exactD elim!: sep_conjE
+            simp: sep_add_commute sep_add_assoc)
+
+end
+
+interpretation sep: ab_semigroup_mult "op **"
+  by (rule ab_semigroup_mult_sep_conj)
+
+interpretation sep: comm_monoid_add "op **" \<box>
+  by (rule comm_monoid_add)
+
+
+section {* Separation Algebra with Stronger, but More Intuitive Disjunction Axiom *}
+
+class stronger_sep_algebra = pre_sep_algebra +
+  assumes sep_add_disj_eq [simp]: "y ## z \<Longrightarrow> x ## y + z = (x ## y \<and> x ## z)"
+begin
+
+lemma sep_disj_add_eq [simp]: "x ## y \<Longrightarrow> x + y ## z = (x ## z \<and> y ## z)"
+  by (metis sep_add_disj_eq sep_disj_commute)
+
+subclass sep_algebra by default auto
+
+end
+
+
+section {* Folding separating conjunction over lists of predicates *}
+
+lemma sep_list_conj_Nil [simp]: "\<And>* [] = \<box>"
+  by (simp add: sep_list_conj_def)
+
+(* apparently these two are rarely used and had to be removed from List.thy *)
+lemma (in semigroup_add) foldl_assoc:
+shows "foldl op+ (x+y) zs = x + (foldl op+ y zs)"
+by (induct zs arbitrary: y) (simp_all add:add_assoc)
+
+lemma (in monoid_add) foldl_absorb0:
+shows "x + (foldl op+ 0 zs) = foldl op+ x zs"
+by (induct zs) (simp_all add:foldl_assoc)
+
+lemma sep_list_conj_Cons [simp]: "\<And>* (x#xs) = (x ** \<And>* xs)"
+  by (simp add: sep_list_conj_def sep.foldl_absorb0)
+
+lemma sep_list_conj_append [simp]: "\<And>* (xs @ ys) = (\<And>* xs ** \<And>* ys)"
+  by (simp add: sep_list_conj_def sep.foldl_absorb0)
+
+lemma (in comm_monoid_add) foldl_map_filter:
+  "foldl op + 0 (map f (filter P xs)) +
+     foldl op + 0 (map f (filter (not P) xs))
+   = foldl op + 0 (map f xs)"
+proof (induct xs)
+  case Nil thus ?case by clarsimp
+next
+  case (Cons x xs)
+  hence IH: "foldl op + 0 (map f xs) =
+               foldl op + 0 (map f (filter P xs)) +
+               foldl op + 0 (map f [x\<leftarrow>xs . \<not> P x])"
+               by (simp only: eq_commute)
+
+  have foldl_Cons':
+    "\<And>x xs. foldl op + 0 (x # xs) = x + (foldl op + 0 xs)"
+    by (simp, subst foldl_absorb0[symmetric], rule refl)
+
+  { assume "P x"
+    hence ?case by (auto simp del: foldl_Cons simp add: foldl_Cons' IH add_ac)
+  } moreover {
+    assume "\<not> P x"
+    hence ?case by (auto simp del: foldl_Cons simp add: foldl_Cons' IH add_ac)
+  }
+  ultimately show ?case by blast
+qed
+
+
+section {* Separation Algebra with a Cancellative Monoid (for completeness) *}
+
+text {*
+  Separation algebra with a cancellative monoid. The results of being a precise
+  assertion (distributivity over separating conjunction) require this.
+  although we never actually use this property in our developments, we keep
+  it here for completeness.
+  *}
+class cancellative_sep_algebra = sep_algebra +
+  assumes sep_add_cancelD: "\<lbrakk> x + z = y + z ; x ## z ; y ## z \<rbrakk> \<Longrightarrow> x = y"
+begin
+
+definition
+  (* In any heap, there exists at most one subheap for which P holds *)
+  precise :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where
+  "precise P = (\<forall>h hp hp'. hp \<preceq> h \<and> P hp \<and> hp' \<preceq> h \<and> P hp' \<longrightarrow> hp = hp')"
+
+lemma "precise (op = s)"
+  by (metis (full_types) precise_def)
+
+lemma sep_add_cancel:
+  "x ## z \<Longrightarrow> y ## z \<Longrightarrow> (x + z = y + z) = (x = y)"
+  by (metis sep_add_cancelD)
+
+lemma precise_distribute:
+  "precise P = (\<forall>Q R. ((Q and R) \<and>* P) = ((Q \<and>* P) and (R \<and>* P)))"
+proof (rule iffI)
+  assume pp: "precise P"
+  {
+    fix Q R
+    fix h hp hp' s
+
+    { assume a: "((Q and R) \<and>* P) s"
+      hence "((Q \<and>* P) and (R \<and>* P)) s"
+        by (fastforce dest!: sep_conjD elim: sep_conjI)
+    }
+    moreover
+    { assume qs: "(Q \<and>* P) s" and qr: "(R \<and>* P) s"
+
+      from qs obtain x y where sxy: "s = x + y" and xy: "x ## y"
+                           and x: "Q x" and y: "P y"
+        by (fastforce dest!: sep_conjD)
+      from qr obtain x' y' where sxy': "s = x' + y'" and xy': "x' ## y'"
+                           and x': "R x'" and y': "P y'"
+        by (fastforce dest!: sep_conjD)
+
+      from sxy have ys: "y \<preceq> x + y" using xy
+        by (fastforce simp: sep_substate_disj_add' sep_disj_commute)
+      from sxy' have ys': "y' \<preceq> x' + y'" using xy'
+        by (fastforce simp: sep_substate_disj_add' sep_disj_commute)
+
+      from pp have yy: "y = y'" using sxy sxy' xy xy' y y' ys ys'
+        by (fastforce simp: precise_def)
+
+      hence "x = x'" using sxy sxy' xy xy'
+        by (fastforce dest!: sep_add_cancelD)
+
+      hence "((Q and R) \<and>* P) s" using sxy x x' yy y' xy'
+        by (fastforce intro: sep_conjI)
+    }
+    ultimately
+    have "((Q and R) \<and>* P) s = ((Q \<and>* P) and (R \<and>* P)) s" using pp by blast
+  }
+  thus "\<forall>Q R. ((Q and R) \<and>* P) = ((Q \<and>* P) and (R \<and>* P))" by (blast intro!: ext)
+
+next
+  assume a: "\<forall>Q R. ((Q and R) \<and>* P) = ((Q \<and>* P) and (R \<and>* P))"
+  thus "precise P"
+  proof (clarsimp simp: precise_def)
+    fix h hp hp' Q R
+    assume hp: "hp \<preceq> h" and hp': "hp' \<preceq> h" and php: "P hp" and php': "P hp'"
+
+    obtain z where hhp: "h = hp + z" and hpz: "hp ## z" using hp
+      by (clarsimp simp: sep_substate_def)
+    obtain z' where hhp': "h = hp' + z'" and hpz': "hp' ## z'" using hp'
+      by (clarsimp simp: sep_substate_def)
+
+    have h_eq: "z' + hp' = z + hp" using hhp hhp' hpz hpz'
+      by (fastforce simp: sep_add_ac)
+
+    from hhp hhp' a hpz hpz' h_eq
+    have "\<forall>Q R. ((Q and R) \<and>* P) (z + hp) = ((Q \<and>* P) and (R \<and>* P)) (z' + hp')"
+      by (fastforce simp: h_eq sep_add_ac sep_conj_commute)
+
+    hence "((op = z and op = z') \<and>* P) (z + hp) =
+           ((op = z \<and>* P) and (op = z' \<and>* P)) (z' + hp')" by blast
+
+    thus  "hp = hp'" using php php' hpz hpz' h_eq
+      by (fastforce dest!: iffD2 cong: conj_cong
+                    simp: sep_add_ac sep_add_cancel sep_conj_def)
+  qed
+qed
+
+lemma strictly_precise: "strictly_exact P \<Longrightarrow> precise P"
+  by (metis precise_def strictly_exactD)
+
+end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Separation_Algebra/sep_tactics.ML	Thu Mar 06 13:28:38 2014 +0000
@@ -0,0 +1,375 @@
+(* Title: Tactics for abstract separation algebras
+   Authors: Gerwin Klein and Rafal Kolanski, 2012
+   Maintainers: Gerwin Klein <kleing at cse.unsw.edu.au>
+                Rafal Kolanski <rafal.kolanski at nicta.com.au>
+*)
+
+(* Separating Conjunction (and Top, AKA sep_true) {{{
+
+  This defines all the constants and theorems necessary for the conjunct
+  selection and cancelling tactic, as well as utility functions.
+*)
+
+structure SepConj =
+struct
+
+val sep_conj_term = @{term sep_conj};
+val sep_conj_str = "**";
+val sep_conj_ac = @{thms sep_conj_ac};
+val sep_conj_impl = @{thm sep_conj_impl}
+
+fun is_sep_conj_const (Const (@{const_name sep_conj}, _)) = true
+  | is_sep_conj_const _ = false;
+
+fun is_sep_conj_term
+      (Const t $ _ $ _ $ _) = is_sep_conj_const (Const t)
+  | is_sep_conj_term _ = false;
+
+fun is_sep_conj_prop
+      (Const Trueprop $ t) = is_sep_conj_term t
+  | is_sep_conj_prop _ = false;
+
+fun break_sep_conj (Const (@{const_name sep_conj},_) $ t1 $ t2 $ _) =
+  [t1] @ (break_sep_conj t2)
+  | break_sep_conj (Const (@{const_name sep_conj},_) $ t1 $ t2) =
+  [t1] @ (break_sep_conj t2)
+  (* dig through eta exanded terms: *)
+  | break_sep_conj (Abs (_, _, t $ Bound 0)) = break_sep_conj t
+  | break_sep_conj t = [t];
+
+fun is_sep_true_term (Abs (_, _, Const (@{const_name True}, _))) = true
+  | is_sep_true_term _ = false;
+
+end;
+
+(* }}} *)
+
+(* Convenience functions for lists {{{ *)
+structure ListExtra =
+struct
+
+fun init l = List.take (l, List.length l - 1);
+
+fun range _ 0 = []
+  | range from howmany = from :: (range (from+1) (howmany-1));
+
+(* move nth element in list to the front *)
+fun nth_to_front i xs =
+      (nth xs i) :: (List.take (xs, i)) @ (List.drop (xs,i+1));
+
+(* assign all members of list an index in the list *)
+fun index_list xs = ListPair.zip (range 0 (length xs), xs);
+
+end; (* }}} *)
+
+(* Function application terms {{{ *)
+(* Dealing with function applications of the type
+     Const/Free(name,type) $ arg1 $ arg2 $ ... $ last_arg *)
+structure FunApp =
+struct
+
+(* apply a transformation to the args in a function application term *)
+fun app_args_op f t = strip_comb t |> apsnd f |> list_comb;
+
+(* remove last argument *)
+fun app_del_last_arg t = app_args_op ListExtra.init t;
+
+(* apply a function term to a Free with given name *)
+fun fun_app_free t free_name = t $ Free (free_name, type_of t |> domain_type);
+
+(* fold two-argument function over a list of arg names using fun_app_free *)
+fun fun_app_foldr f [a,b] = fun_app_free (fun_app_free f a) b
+  | fun_app_foldr f (x::xs) = (fun_app_free f x) $ (fun_app_foldr f xs)
+  | fun_app_foldr _ _ = raise Fail "fun_app_foldr";
+
+end; (* }}} *)
+
+(* Selecting Conjuncts in Premise or Conclusion {{{ *)
+
+(* Constructs a rearrangement lemma of the kind:
+   (A ** B ** C) s ==> (C ** A ** B) s
+   When cjt_select = 2 (0-based index of C) and
+   cjt_select = 3 (number of conjuncts to use), conclusion = true
+   "conclusion" specifies whether the rearrangement occurs in conclusion
+   (for dtac) or the premise (for rtac) of the rule.
+*)
+fun mk_sep_select_rule ctxt conclusion (cjt_count, cjt_select) =
+let
+  fun variants nctxt names = fold_map Name.variant names nctxt;
+
+  val (state, nctxt0) = Name.variant "s" (Variable.names_of ctxt);
+  fun sep_conj_prop cjts =
+        FunApp.fun_app_free
+          (FunApp.fun_app_foldr SepConj.sep_conj_term cjts) state
+        |> HOLogic.mk_Trueprop;
+
+  (* concatenate string and string of an int *)
+  fun conc_str_int str int = str ^ Int.toString int;
+
+  (* make the conjunct names *)
+  val (cjts, _) = ListExtra.range 1 cjt_count
+                  |> map (conc_str_int "a") |> variants nctxt0;
+
+  (* make normal-order separation conjunction terms *)
+  val orig = sep_conj_prop cjts;
+
+  (* make reordered separation conjunction terms *)
+  val reordered = sep_conj_prop (ListExtra.nth_to_front cjt_select cjts);
+
+  val goal = Logic.mk_implies
+               (if conclusion then (orig, reordered) else (reordered, orig));
+
+  (* simp add: sep_conj_ac *)
+  val sep_conj_ac_tac = Simplifier.asm_full_simp_tac
+                          (put_simpset HOL_basic_ss ctxt addsimps SepConj.sep_conj_ac);
+
+in
+  (* XXX: normally you'd want to keep track of what variables we want to make
+     schematic and which ones are bound, but we don't use fixed names for
+     the rules we make, so we use Drule.export_without_context to schematise
+     all. *)
+  Goal.prove ctxt [] [] goal (fn _ => sep_conj_ac_tac 1)
+  |> Drule.export_without_context
+end;
+
+(* }}} *)
+
+local
+  (* Common tactic functionality {{{ *)
+
+  (* given two terms of some 'a to bool, can you prove
+     \<And>s. t1 s \<Longrightarrow> t2 s
+     using the supplied proof method?
+     NOTE: t1 and t2 MUST have a function type with one argument,
+     or TYPE dest_Type is raised
+     NOTE: if asm or concl is sep_true, returns false
+  *)
+  fun can_prove ctxt tac asm concl =
+    let
+      fun variant name = Name.variant name (Variable.names_of ctxt) |> fst;
+      val arg_name = variant "s";
+      val left = FunApp.fun_app_free asm arg_name |> HOLogic.mk_Trueprop;
+      val right = FunApp.fun_app_free concl arg_name |> HOLogic.mk_Trueprop;
+      val goal = Logic.mk_implies (left, right);
+    in
+      if (SepConj.is_sep_true_term asm) orelse (SepConj.is_sep_true_term concl)
+      then false
+      else (Goal.prove ctxt [] [] goal (fn _ => tac 1); true)
+            handle ERROR _ => false
+    end;
+
+  (* apply function in list until it returns SOME *)
+  fun find_some (x::xs) f =
+    (case f x of SOME v => SOME v
+               | NONE => find_some xs f)
+    | find_some [] _ = NONE;
+
+  (* Given indices into the separating conjunctions in the assumption and
+     conclusion, rewrite them so that the targeted conjuncts are at the
+     front, then remove them. *)
+  fun eliminate_target_tac ctxt tac
+                           ((prem_total,prem_idx), (concl_total,concl_idx)) =
+    let
+      val asm_select = mk_sep_select_rule ctxt true (prem_total,prem_idx);
+      val concl_select = mk_sep_select_rule ctxt false
+                           (concl_total,concl_idx);
+    in
+      dtac asm_select THEN' rtac concl_select
+      THEN' etac SepConj.sep_conj_impl THEN' tac
+    end;
+
+  fun find_target ctxt tac cprem cconcl =
+    let
+      val prem_cjts = cprem |> term_of |> SepConj.break_sep_conj;
+      val concl_cjts = cconcl |> term_of |> SepConj.break_sep_conj;
+
+      val iprems = ListExtra.index_list prem_cjts;
+      val iconcls = ListExtra.index_list concl_cjts;
+
+      fun can_prove' (pi,p) (ci,c) =
+            if can_prove ctxt tac p c then SOME (pi, ci) else NONE;
+
+      val target = find_some iconcls
+                     (fn c => find_some iprems (fn p => can_prove' p c));
+    in
+      case target
+        of SOME (pi,ci) => SOME ((List.length prem_cjts, pi),
+                                 (List.length concl_cjts, ci))
+         | NONE => NONE
+    end;
+
+  fun strip_cprop ct =
+    let val thy = theory_of_cterm ct
+    in ct |> term_of |> HOLogic.dest_Trueprop |> cterm_of thy
+    end;
+
+  (* }}} *)
+in
+
+  (* Tactic: Select nth conjunct in assumption {{{ *)
+  local
+    fun sep_select_asm_tac' ctxt n (ct, i) =
+      let
+        (* digging out prems *)
+        val ((_, ct'), _) = Variable.focus_cterm ct ctxt;
+        val prems = Drule.strip_imp_prems ct';
+
+        fun prem_ok ct = SepConj.is_sep_conj_prop (term_of ct)
+
+        fun mk_tac prem =
+            let
+              val prem = HOLogic.dest_Trueprop (term_of prem)
+              val p = length (SepConj.break_sep_conj prem)
+              val th = mk_sep_select_rule ctxt true (p,n)
+                  handle Subscript => error "Conjunct index out of range"
+             in
+               dtac th i
+             end;
+      in
+        if length prems = 0
+        then error ("No assumption of form: (_ " ^ SepConj.sep_conj_str ^
+                    " _) _")
+        else
+          (* backtrack until first premise that does something *)
+          map mk_tac (filter prem_ok prems) |> FIRST
+      end;
+    in
+      fun sep_select_asm_tac ctxt n = CSUBGOAL (sep_select_asm_tac' ctxt (n-1))
+    end; (* }}} *)
+
+  (* Tactic: Select nth conjunct in conclusion {{{ *)
+  local
+    fun sep_select_tac' ctxt n (ct, i) =
+      let
+        (* digging out conclusions *)
+        val ((_, ct'), _) = Variable.focus_cterm ct ctxt;
+        val concl = ct' |> Drule.strip_imp_concl |> term_of;
+        val p = concl |> HOLogic.dest_Trueprop |> SepConj.break_sep_conj
+                |> length;
+        val th = mk_sep_select_rule ctxt false (p,n)
+                 handle Subscript => error "Conjunct index out of range"
+      in
+        if not (SepConj.is_sep_conj_prop concl)
+        then error ("Goal not of form: (_ " ^ SepConj.sep_conj_str ^ " _) _")
+        else rtac th i
+      end;
+  in
+    fun sep_select_tac ctxt n = CSUBGOAL (sep_select_tac' ctxt (n-1))
+  end; (* }}} *)
+
+  (* Tactic: for all reorderings of the premises try apply tac {{{ *)
+    local
+      fun sep_assm_tac' ctxt tac (ct,i) =
+        let
+          (* digging out prems *)
+          val ((_, ct'), _) = Variable.focus_cterm ct ctxt;
+          val prems = Drule.strip_imp_prems ct';
+
+          fun prem_ok ct = SepConj.is_sep_conj_prop (term_of ct)
+
+          fun mk_tac prem =
+            let
+              val prem = HOLogic.dest_Trueprop (term_of prem)
+              val p = length (SepConj.break_sep_conj prem)
+              fun ord n = mk_sep_select_rule ctxt true (p,n)
+              val ord_thms = map ord (0 upto p-1)
+            in
+                (dresolve_tac ord_thms THEN' tac) i
+            end;
+        in
+          (* backtrack until first premise that does something *)
+          map mk_tac (filter prem_ok prems) |> FIRST
+        end;
+    in
+      fun sep_assm_tac ctxt tac = CSUBGOAL (sep_assm_tac' ctxt tac)
+    end; (* }}} *)
+
+  (* Tactic: for all reorderings of the conclusion, try apply tac {{{ *)
+  local
+    fun sep_concl_tac' ctxt tac (ct, i) =
+      let
+        (* digging out conclusion *)
+        val ((_, ct'), _) = Variable.focus_cterm ct ctxt;
+        val concl = ct' |> Drule.strip_imp_concl |> term_of;
+        val p = concl |> HOLogic.dest_Trueprop |> SepConj.break_sep_conj
+                |> length;
+        fun ord n = mk_sep_select_rule ctxt false (p,n);
+        val ord_thms = map ord (0 upto p-1);
+      in
+        if not (SepConj.is_sep_conj_prop concl)
+        then (tracing ("Goal not of form: (_ " ^ SepConj.sep_conj_str ^
+                      " _) _");
+              no_tac)
+        else (resolve_tac ord_thms THEN' tac) i
+      end;
+  in
+    fun sep_concl_tac ctxt tac = CSUBGOAL (sep_concl_tac' ctxt tac)
+  end; (* }}} *)
+
+  (* Tactic: Cancel conjuncts of assumption and conclusion via tac {{{ *)
+  local
+    fun sep_cancel_tac' ctxt tac (ct, i) =
+      let
+        (* digging out prems and conclusions *)
+        val ((vars, ct'), ctxt') = Variable.focus_cterm ct ctxt;
+        val concl = Drule.strip_imp_concl ct';
+        val prems = Drule.strip_imp_prems ct';
+
+        fun prem_ok ct =
+          let
+            (* name of state in sep conj (should be Free after focus) *)
+            fun state_get (_ $ _ $ _ $ s) = s
+              | state_get t = raise Fail "prem_ok: state_get";
+            val state_get_ct = state_get o HOLogic.dest_Trueprop o term_of;
+
+            val concl_state = concl |> state_get_ct;
+            (* states considered equal if they alpha-convert *)
+            fun state_ok ct = (state_get_ct ct) aconv concl_state;
+          in
+            SepConj.is_sep_conj_prop (term_of ct) andalso state_ok ct
+          end;
+
+        fun mk_tac prem =
+              case find_target ctxt tac (prem |> strip_cprop)
+                                        (strip_cprop concl)
+                of SOME target => eliminate_target_tac ctxt tac target i
+                 | NONE => no_tac;
+      in
+        if (not (concl |> term_of |> SepConj.is_sep_conj_prop))
+        then (tracing ("Goal not of form: (_ " ^ SepConj.sep_conj_str ^
+                       " _) _");
+              no_tac)
+        else if (length prems = 0)
+        then (tracing ("No assumption of form: (_ " ^ SepConj.sep_conj_str ^
+                       " _) _");
+              no_tac)
+        else
+          (* backtrack until first premise that does something *)
+          map mk_tac (filter prem_ok prems) |> FIRST
+      end;
+  in
+    fun sep_cancel_tac ctxt tac = CSUBGOAL (sep_cancel_tac' ctxt tac)
+  end;
+  (* }}} *)
+
+  (* Derived Tactics *)
+
+  fun sep_atac ctxt = sep_assm_tac ctxt atac;
+
+  (* Substitution *)
+  fun sep_subst_tac ctxt occs thms =
+        EqSubst.eqsubst_tac ctxt occs thms THEN' sep_atac ctxt;
+  fun sep_subst_asm_tac ctxt occs thms =
+        EqSubst.eqsubst_asm_tac ctxt occs thms THEN' sep_atac ctxt;
+
+  (* Forward reasoning *)
+  fun sep_dtac ctxt thms = sep_assm_tac ctxt (dresolve_tac thms)
+  fun sep_ftac ctxt thms = sep_assm_tac ctxt (forward_tac thms)
+
+  (* Backward reasoning *)
+  fun sep_rtac ctxt thms = sep_concl_tac ctxt (resolve_tac thms)
+
+end;
+
+(* vim: set foldmethod=marker sw=2 sts=2 et: *)
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/Data_slot.thy	Thu Mar 06 13:28:38 2014 +0000
@@ -0,0 +1,25 @@
+theory Data_slot
+imports Main
+begin
+
+ML {*
+  fun ML_exec prog = ML_Compiler.eval true Position.none
+                        (ML_Lex.tokenize prog)
+
+  fun prog_data_slot name tp empty  = 
+  "structure " ^ name ^ " = Theory_Data ( " ^
+  "type T = "^ tp ^" " ^
+  "val empty = " ^ empty ^ " ; "^
+  "val extend = I; "^
+  "fun merge (x, y) = x) "^
+  "fun "^ name ^ "_store v =  Context.>>(Context.map_theory (" ^ name ^".map (fn x => v))) "^
+  "fun "^ name ^ "_get () = Context.>>>(Context.map_theory_result (fn th => "^
+  "                                     ("^ name ^".get th, th))) "
+
+  fun data_slot name tp empty  = ML_exec (prog_data_slot name tp empty)
+  val empty_term = @{term "x"}
+  val empty_cterm = @{cterm "x"}
+  val empty_thm = @{thm ext}
+*}
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/Hoare_abc.thy	Thu Mar 06 13:28:38 2014 +0000
@@ -0,0 +1,1217 @@
+header {* 
+ {\em Abacus} defined as macros of TM
+ *}
+
+theory Hoare_abc
+imports Hoare_tm Finite_Set
+begin
+
+
+text {*
+  {\em Abacus} instructions
+*}
+
+(*
+text {* The following Abacus instructions will be replaced by TM macros. *}
+datatype abc_inst =
+  -- {* @{text "Inc n"} increments the memory cell (or register) 
+         with address @{text "n"} by one.
+     *}
+     Inc nat
+  -- {*
+     @{text "Dec n label"} decrements the memory cell with address @{text "n"} by one. 
+      If cell @{text "n"} is already zero, no decrements happens and the executio jumps to
+      the instruction labeled by @{text "label"}.
+     *}
+   | Dec nat nat
+  -- {*
+  @{text "Goto label"} unconditionally jumps to the instruction labeled by @{text "label"}.
+  *}
+   | Goto nat
+
+*)
+
+datatype aresource = 
+    M nat nat
+  (* | C nat abc_inst *) (* C resource is not needed because there is no Abacus code any more *)  
+  | At nat
+  | Faults nat
+
+
+section {* An interpretation from Abacus to Turing Machine *}
+
+fun recse_map :: "nat list \<Rightarrow> aresource \<Rightarrow> tassert" where
+  "recse_map ks (M a v) = <(a < length ks \<and> ks!a = v \<or> a \<ge> length ks \<and> v = 0)>" |
+  "recse_map ks (At l) = st l" |
+  "recse_map ks (Faults n) = sg {TFaults n}"
+
+definition "IA ars = (EXS ks i. ps 2 \<and>* zero 0 \<and>* zero 1 \<and>* (reps 2 i ks) \<and>* 
+                                     fam_conj {i<..} zero \<and>* 
+                                fam_conj ars (recse_map ks))"
+
+
+section {* A virtually defined Abacus *}
+
+text {* The following Abacus instructions are to be defined as TM macros *}
+
+definition "pc l = sg {At l}"
+
+definition "mm a v =sg ({M a v})"
+
+type_synonym assert = "aresource set \<Rightarrow> bool"
+
+lemma tm_hoare_inc1:
+  assumes h: "a < length ks \<and> ks ! a = v \<or> length ks \<le> a \<and> v = 0"
+  shows "
+    \<lbrace>st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u ia ks \<and>* fam_conj {ia<..} zero\<rbrace> 
+    i :[Inc a ]: j
+    \<lbrace>st j \<and>*
+     ps u \<and>*
+     zero (u - 2) \<and>*
+     zero (u - 1) \<and>*
+     reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1) ((list_ext a ks)[a := Suc v]) \<and>*
+     fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1<..} zero\<rbrace>"
+  using h
+proof
+  assume hh: "a < length ks \<and> ks ! a = v"
+  hence "a < length ks" by simp 
+  from list_ext_lt [OF this] tm_hoare_inc00[OF hh]
+  show ?thesis by simp
+next
+  assume "length ks \<le> a \<and> v = 0"
+  from tm_hoare_inc01[OF this]
+  show ?thesis by simp
+qed
+
+lemma tm_hoare_inc2: 
+  assumes "mm a v sr"
+  shows "
+    \<lbrace> (fam_conj sr (recse_map ks) \<and>*
+       st i \<and>* ps 2 \<and>* zero 0 \<and>* zero 1 \<and>* reps 2 ia ks \<and>* fam_conj {ia<..} zero) \<rbrace> 
+          i:[ (Inc a) ]:j   
+    \<lbrace> (fam_conj {M a (Suc v)} (recse_map (list_ext a ks[a := Suc v])) \<and>*
+           st j \<and>*
+           ps 2 \<and>*
+           zero 0 \<and>*
+           zero 1 \<and>*
+           reps 2 (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1) ((list_ext a ks)[a := Suc v]) \<and>*
+           fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1<..} zero)\<rbrace>"
+proof -
+  from `mm a v sr` have eq_sr: "sr = {M a v}" by (auto simp:mm_def sg_def)
+  from tm_hoare_inc1[where u = 2]
+  have "a < length ks \<and> ks ! a = v \<or> length ks \<le> a \<and> v = 0 \<Longrightarrow>
+        \<lbrace>st i \<and>* ps 2 \<and>* zero 0 \<and>* zero 1 \<and>* reps 2 ia ks \<and>* fam_conj {ia<..} zero\<rbrace> 
+             i :[Inc a ]: j
+        \<lbrace>(st j \<and>*
+           ps 2 \<and>*
+           zero 0 \<and>*
+           zero 1 \<and>*
+           reps 2 (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1) ((list_ext a ks)[a := Suc v]) \<and>*
+           fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1<..} zero)\<rbrace>" by simp
+  thus ?thesis
+    apply (unfold eq_sr)
+    apply (simp add:fam_conj_simps list_ext_get_upd list_ext_len)
+    by (rule tm.pre_condI, blast)
+qed
+
+locale IA_disjoint = 
+  fixes s s' s1 cnf
+  assumes h_IA: "IA (s + s') s1"
+  and h_disj: "s ## s'"
+  and h_conf: "s1 \<subseteq> trset_of cnf"
+begin
+
+lemma at_disj1: 
+  assumes at_in: "At i \<in> s"
+  shows "At j \<notin> s'"
+proof
+  from h_IA[unfolded IA_def]
+  obtain ks idx
+    where "((ps 2 \<and>* zero 0 \<and>* zero 1 \<and>* 
+              reps 2 idx ks \<and>* fam_conj {idx<..} zero) \<and>* 
+              fam_conj (s + s') (recse_map ks)) s1" (is "(?P \<and>* ?Q) s1")
+    by (auto elim!:EXS_elim simp:sep_conj_ac)
+  then obtain ss1 ss2 where "s1 = ss1 + ss2" "ss1 ## ss2" "?P ss1" "?Q ss2"
+    by (auto elim:sep_conjE)
+  from `?Q ss2`[unfolded fam_conj_disj_simp[OF h_disj]]
+  obtain tt1 tt2
+    where "ss2 = tt1 + tt2" "tt1 ## tt2" 
+            "(fam_conj s (recse_map ks)) tt1"
+            "(fam_conj s' (recse_map ks)) tt2"
+    by (auto elim:sep_conjE)
+  assume "At j \<in> s'"
+  from `(fam_conj s' (recse_map ks)) tt2` [unfolded fam_conj_elm_simp[OF this]]
+       `ss2 = tt1 + tt2` `s1 = ss1 + ss2` h_conf
+  have "TAt j \<in> trset_of cnf"
+    by (auto elim!:sep_conjE simp: st_def sg_def tpn_set_def set_ins_def)
+  moreover from `(fam_conj s (recse_map ks)) tt1` [unfolded fam_conj_elm_simp[OF at_in]]
+                `ss2 = tt1 + tt2` `s1 = ss1 + ss2` h_conf
+  have "TAt i \<in> trset_of cnf"
+    by (auto elim!:sep_conjE simp: st_def sg_def tpn_set_def set_ins_def)
+  ultimately have "i = j"
+    by (cases cnf, simp add:trset_of.simps tpn_set_def)
+  from at_in `At j \<in> s'` h_disj
+  show False 
+    by (unfold `i = j`, auto simp:set_ins_def)
+qed
+
+lemma at_disj2: "At i \<in> s' \<Longrightarrow> At j \<notin> s"
+  by (metis at_disj1)
+
+lemma m_disj1: 
+  assumes m_in: "M a v \<in> s"
+  shows "M a v' \<notin> s'"
+proof
+  from h_IA[unfolded IA_def]
+  obtain ks idx
+    where "((ps 2 \<and>* zero 0 \<and>* zero 1 \<and>* 
+              reps 2 idx ks \<and>* fam_conj {idx<..} zero) \<and>* 
+              fam_conj (s + s') (recse_map ks)) s1" (is "(?P \<and>* ?Q) s1")
+    by (auto elim!:EXS_elim simp:sep_conj_ac)
+  then obtain ss1 ss2 where "s1 = ss1 + ss2" "ss1 ## ss2" "?P ss1" "?Q ss2"
+    by (auto elim:sep_conjE)
+  from `?Q ss2`[unfolded fam_conj_disj_simp[OF h_disj]]
+  obtain tt1 tt2
+    where "ss2 = tt1 + tt2" "tt1 ## tt2" 
+            "(fam_conj s (recse_map ks)) tt1"
+            "(fam_conj s' (recse_map ks)) tt2"
+    by (auto elim:sep_conjE)
+  assume "M a v' \<in> s'"
+  from `(fam_conj s' (recse_map ks)) tt2` [unfolded fam_conj_elm_simp[OF this]
+        recse_map.simps]
+  have "(a < length ks \<and> ks ! a = v' \<or> length ks \<le> a \<and> v' = 0)"
+    by (auto simp:pasrt_def)
+  moreover from `(fam_conj s (recse_map ks)) tt1` [unfolded fam_conj_elm_simp[OF m_in]
+                         recse_map.simps]
+  have "a < length ks \<and> ks ! a = v \<or> length ks \<le> a \<and> v = 0"
+    by (auto simp:pasrt_def)
+  moreover note m_in `M a v' \<in> s'` h_disj
+  ultimately show False
+    by (auto simp:set_ins_def)
+qed
+
+lemma m_disj2: "M a v \<in> s' \<Longrightarrow> M a v' \<notin> s"
+  by (metis m_disj1)
+
+end
+
+lemma EXS_elim1: 
+  assumes "((EXS x. P(x)) \<and>* r) s"
+  obtains x where "(P(x) \<and>* r) s"
+  by (metis EXS_elim assms sep_conj_exists1)
+
+lemma hoare_inc[step]: "IA. \<lbrace> pc i ** mm a v \<rbrace> 
+                        i:[ (Inc a) ]:j   
+                      \<lbrace> pc j ** mm a (Suc v)\<rbrace>"
+                      (is "IA. \<lbrace> pc i ** ?P \<rbrace> 
+                        i:[ ?code ?e ]:j   
+                      \<lbrace> pc ?e ** ?Q\<rbrace>")
+proof(induct rule:tm.IHoareI)
+  case (IPre s' s r cnf)
+  let ?cnf = "(trset_of cnf)"
+  from IPre
+  have h: "(pc i \<and>* ?P) s" "(s ## s')" "(IA (s + s') \<and>* i :[ ?code(?e) ]: j \<and>* r) ?cnf"
+    by (metis condD)+
+  from h(1) obtain sr where
+      eq_s: "s = {At i} \<union>  sr" "{At i} ##  sr" "?P sr"
+    by (auto dest!:sep_conjD simp:set_ins_def pc_def sg_def)
+  hence "At i \<in> s" by auto
+  from h(3) obtain s1 s2 s3
+    where hh: "?cnf = s1 + s2 + s3"
+              "s1 ## s2 \<and> s2 ## s3 \<and> s1 ## s3" 
+              "IA (s + s') s1" 
+              "(i :[ ?code ?e ]: j) s2"
+              "r s3"
+    apply (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac)
+    by (metis sep_add_commute sep_disj_addD1 sep_disj_addD2 sep_disj_commuteI)
+  interpret ia_disj: IA_disjoint s s' s1 cnf
+  proof
+    from `IA (s + s') s1` show "IA (s + s') s1" .
+  next
+    from `s ## s'` show "s ## s'" .
+  next
+    from hh(1) show "s1 \<subseteq> ?cnf" by (auto simp:set_ins_def)
+  qed
+  from hh(1) have s1_belong: "s1 \<subseteq> ?cnf" by (auto simp:set_ins_def)
+  from hh(3)
+  have "(EXS ks ia.
+         ps 2 \<and>*
+         zero 0 \<and>*
+         zero 1 \<and>*
+         reps 2 ia ks \<and>*
+         fam_conj {ia<..} zero \<and>*
+         (st i \<and>* fam_conj sr (recse_map ks)) \<and>* fam_conj s' (recse_map ks))
+        s1"
+    apply (unfold IA_def fam_conj_disj_simp[OF h(2)])
+    apply (unfold eq_s)
+    by (insert eq_s(2), simp add: fam_conj_disj_simp[OF eq_s(2)] fam_conj_simps set_ins_def)
+  then obtain ks ia
+    where "((fam_conj sr (recse_map ks) \<and>* st i \<and>* ps 2 \<and>* zero 0 \<and>* zero 1 \<and>*
+             reps 2 ia ks \<and>* fam_conj {ia<..} zero) \<and>* fam_conj s' (recse_map ks)) s1" 
+    (is "(?PP \<and>* ?QQ) s1")
+    by (unfold pred_ex_def, auto simp:sep_conj_ac)
+  then obtain ss1 ss2 where pres:
+        "s1 = ss1 + ss2" "ss1 ## ss2"
+        "?PP ss1"
+        "?QQ ss2"
+    by (auto elim!:sep_conjE intro!:sep_conjI)
+  from ia_disj.at_disj1 [OF `At i \<in> s`]
+  have at_fresh_s': "At ?e \<notin>  s'" .
+  have at_fresh_sr: "At ?e \<notin> sr"
+  proof
+    assume at_in: "At ?e \<in> sr"
+    from h(3)[unfolded IA_def fam_conj_disj_simp[OF h(2)]]
+    have "TAt ?e \<in> trset_of cnf"
+      apply (elim EXS_elim1)
+      apply (unfold eq_s fam_conj_disj_simp[OF eq_s(2), unfolded set_ins_def]
+             fam_conj_elm_simp[OF at_in])
+      apply (erule_tac sep_conjE, unfold set_ins_def)+
+      by (auto simp:sep_conj_ac fam_conj_simps st_def sg_def tpn_set_def)
+    moreover from pres(1, 3) hh(1) have "TAt i \<in> trset_of cnf" 
+      apply(erule_tac sep_conjE)
+      apply(erule_tac sep_conjE)
+      by (auto simp:st_def tpc_set_def sg_def set_ins_def)
+    ultimately have "i = ?e"
+      by (cases cnf, auto simp:tpn_set_def trset_of.simps)
+    from eq_s[unfolded this] at_in
+    show "False" by (auto simp:set_ins_def)
+  qed
+  from pres(3) and hh(2, 4, 5) pres(2, 4)
+  have pres1: "(?PP \<and>* i :[ ?code(?e)]: j \<and>*  (r \<and>* (fam_conj s' (recse_map ks))))
+         (trset_of cnf)"
+    apply (unfold hh(1) pres(1))
+    apply (rule sep_conjI[where x=ss1 and y = "ss2 + s2 + s3"], assumption+)
+    apply (rule sep_conjI[where x="s2" and y = "ss2 + s3"], assumption+)
+    apply (rule sep_conjI[where x="s3" and y = ss2], assumption+)
+    by (auto simp:set_ins_def)
+  (*****************************************************************************)
+  let ?ks_f = "\<lambda> sr ks. list_ext a ks[a := Suc v]"
+  let ?elm_f = "\<lambda> sr. {M a (Suc v)}"
+  let ?idx_f = "\<lambda> sr ks ia. (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1)"
+  (*----------------------------------------------------------------------------*)
+  (******************************************************************************)
+  from tm_hoare_inc2 [OF eq_s(3), unfolded tm.Hoare_gen_def, rule_format, OF pres1]
+  obtain k where "((fam_conj (?elm_f sr) (recse_map (?ks_f sr ks)) \<and>*
+        st ?e \<and>*
+        ps 2 \<and>*
+        zero 0 \<and>* zero 1 \<and>* reps 2 (?idx_f sr ks ia) (?ks_f sr ks) \<and>* 
+        fam_conj {?idx_f sr ks ia<..} zero) \<and>*
+       i :[ ?code ?e ]: j \<and>* r \<and>* fam_conj s' (recse_map ks))
+       (trset_of (sep_exec.run tstep (Suc k) cnf))" by blast
+  (*----------------------------------------------------------------------------*)
+  moreover have "(pc ?e \<and>* ?Q) ({At ?e} \<union> ?elm_f sr)"
+  proof -
+    have "pc ?e {At ?e}" by (simp add:pc_def sg_def)
+  (******************************************************************************)
+    moreover have "?Q (?elm_f sr)" 
+      by (simp add:mm_def sg_def)
+  (*----------------------------------------------------------------------------*)
+    moreover 
+  (******************************************************************************)
+    have "{At ?e} ## ?elm_f sr" by (simp add:set_ins_def)
+  (*----------------------------------------------------------------------------*)
+    ultimately show ?thesis by (auto intro!:sep_conjI simp:set_ins_def)
+  qed
+  moreover 
+  (******************************************************************************)
+  from ia_disj.m_disj1 `?P sr` `s = {At i} \<union> sr` 
+  have "?elm_f sr ## s'" by (auto simp:set_ins_def mm_def sg_def)
+  (*----------------------------------------------------------------------------*)
+  with at_fresh_s' 
+  have fresh_atm: "{At ?e} \<union> ?elm_f sr ## s'" by (auto simp:set_ins_def)
+  moreover have eq_map: "\<And> elm. elm \<in> s' \<Longrightarrow> 
+                           (recse_map ks elm) = (recse_map (?ks_f sr ks) elm)"
+  proof -
+    fix elm
+    assume elm_in: "elm \<in> s'"
+    show "(recse_map ks elm) = (recse_map (?ks_f sr ks) elm)"
+    proof(cases elm)
+      (*******************************************************************)
+      case (M a' v')
+      from eq_s have "M a v \<in> s" by (auto simp:set_ins_def mm_def sg_def)
+      with elm_in ia_disj.m_disj1[OF this] M
+      have "a \<noteq> a'" by auto
+      thus ?thesis
+        apply (auto simp:M recse_map.simps pasrt_def list_ext_len)
+        apply (case_tac "a < length ks", auto simp:list_ext_len_eq list_ext_lt)
+        apply (case_tac "a' < Suc a", auto simp:list_ext_len_eq list_ext_lt)
+        by (metis (full_types) bot_nat_def 
+              leI le_trans less_Suc_eq_le list_ext_lt_get list_ext_tail set_zero_def)
+      (*-----------------------------------------------------------------*)
+    qed auto
+  qed
+  ultimately show ?case 
+    apply (rule_tac x = k in exI, rule_tac x = "{At ?e} \<union> ?elm_f sr" in exI)
+    apply (unfold IA_def, intro condI, assumption+)
+    apply (rule_tac x = "?ks_f sr ks" in pred_exI)
+    apply (rule_tac x = "?idx_f sr ks ia" in pred_exI)
+    apply (unfold fam_conj_disj_simp[OF fresh_atm])
+    apply (auto simp:sep_conj_ac fam_conj_simps)
+    (***************************************************************************)
+    (* apply (unfold fam_conj_insert_simp [OF h_fresh1[OF at_fresh_sr]], simp) *)
+    (*-------------------------------------------------------------------------*)
+    apply (sep_cancel)+
+    by (simp add: fam_conj_ext_eq[where I = "s'", OF eq_map])
+qed
+
+lemma tm_hoare_dec_fail:
+  assumes "mm a 0 sr"
+  shows 
+  "\<lbrace> fam_conj sr (recse_map ks) \<and>*
+    st i \<and>* ps 2 \<and>* zero 0 \<and>* zero 1 \<and>* reps 2 ia ks \<and>* fam_conj {ia<..} zero \<rbrace> 
+       i:[ (Dec a e) ]:j 
+   \<lbrace> fam_conj {M a 0} (recse_map (list_ext a ks[a := 0])) \<and>*
+     st e \<and>*
+     ps 2 \<and>*
+     zero 0 \<and>*
+     zero 1 \<and>*
+     reps 2 (ia + int (reps_len (list_ext a ks)) - int (reps_len ks)) (list_ext a ks[a := 0]) \<and>*
+     fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) <..} zero\<rbrace>"
+proof - 
+  from `mm a 0 sr` have eq_sr: "sr = {M a 0}" by (auto simp:mm_def sg_def)
+  { assume h: "a < length ks \<and> ks ! a = 0 \<or> length ks \<le> a"
+    from tm_hoare_dec_fail1[where u = 2, OF this]
+    have "\<lbrace>st i \<and>* ps 2 \<and>* zero 0 \<and>* zero 1 \<and>* reps 2 ia ks \<and>* fam_conj {ia<..} zero\<rbrace>  
+                  i :[ Dec a e ]: j
+          \<lbrace>st e \<and>* ps 2 \<and>* zero 0 \<and>* zero 1 \<and>*
+            reps 2 (ia + int (reps_len (list_ext a ks)) - int (reps_len ks)) (list_ext a ks[a := 0]) \<and>*
+            fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks)<..} zero\<rbrace>"
+    by (simp)
+  }
+  thus ?thesis
+    apply (unfold eq_sr)
+    apply (simp add:fam_conj_simps list_ext_get_upd list_ext_len)
+    by (rule tm.pre_condI, blast)
+qed
+
+lemma hoare_dec_fail: "IA. \<lbrace> pc i ** mm a 0 \<rbrace>
+                              i:[ (Dec a e) ]:j   
+                           \<lbrace> pc e ** mm a 0 \<rbrace>"
+                      (is "IA. \<lbrace> pc i ** ?P \<rbrace> 
+                                i:[ ?code ?e]:j   
+                               \<lbrace> pc ?e ** ?Q\<rbrace>")
+proof(induct rule:tm.IHoareI)
+  case (IPre s' s r cnf)
+  let ?cnf = "(trset_of cnf)"
+  from IPre
+  have h: "(pc i \<and>* ?P) s" "(s ## s')" "(IA (s + s') \<and>* i :[ ?code(?e) ]: j \<and>* r) ?cnf"
+    by (metis condD)+
+  from h(1) obtain sr where
+      eq_s: "s = {At i} \<union>  sr" "{At i} ##  sr" "?P sr"
+    by (auto dest!:sep_conjD simp:set_ins_def pc_def sg_def)
+  hence "At i \<in> s" by auto
+  from h(3) obtain s1 s2 s3
+    where hh: "?cnf = s1 + s2 + s3"
+              "s1 ## s2 \<and> s2 ## s3 \<and> s1 ## s3" 
+              "IA (s + s') s1" 
+              "(i :[ ?code ?e ]: j) s2"
+              "r s3"
+    apply (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac)
+    by (metis sep_add_commute sep_disj_addD1 sep_disj_addD2 sep_disj_commuteI)
+  interpret ia_disj: IA_disjoint s s' s1 cnf
+  proof
+    from `IA (s + s') s1` show "IA (s + s') s1" .
+  next
+    from `s ## s'` show "s ## s'" .
+  next
+    from hh(1) show "s1 \<subseteq> ?cnf" by (auto simp:set_ins_def)
+  qed
+  from hh(1) have s1_belong: "s1 \<subseteq> ?cnf" by (auto simp:set_ins_def)
+  from hh(3)
+  have "(EXS ks ia.
+         ps 2 \<and>*
+         zero 0 \<and>*
+         zero 1 \<and>*
+         reps 2 ia ks \<and>*
+         fam_conj {ia<..} zero \<and>*
+         (st i \<and>* fam_conj sr (recse_map ks)) \<and>* fam_conj s' (recse_map ks))
+        s1"
+    apply (unfold IA_def fam_conj_disj_simp[OF h(2)])
+    apply (unfold eq_s)
+    by (insert eq_s(2), simp add: fam_conj_disj_simp[OF eq_s(2)] fam_conj_simps set_ins_def)
+  then obtain ks ia
+    where "((fam_conj sr (recse_map ks) \<and>* st i \<and>* ps 2 \<and>* zero 0 \<and>* zero 1 \<and>*
+             reps 2 ia ks \<and>* fam_conj {ia<..} zero) \<and>* fam_conj s' (recse_map ks)) s1" 
+    (is "(?PP \<and>* ?QQ) s1")
+    by (unfold pred_ex_def, auto simp:sep_conj_ac)
+  then obtain ss1 ss2 where pres:
+        "s1 = ss1 + ss2" "ss1 ## ss2"
+        "?PP ss1"
+        "?QQ ss2"
+    by (auto elim!:sep_conjE intro!:sep_conjI)
+  from ia_disj.at_disj1 [OF `At i \<in> s`]
+  have at_fresh_s': "At ?e \<notin>  s'" .
+  have at_fresh_sr: "At ?e \<notin> sr"
+  proof
+    assume at_in: "At ?e \<in> sr"
+    from h(3)[unfolded IA_def fam_conj_disj_simp[OF h(2)]]
+    have "TAt ?e \<in> trset_of cnf"
+      apply (elim EXS_elim1)
+      apply (unfold eq_s fam_conj_disj_simp[OF eq_s(2), unfolded set_ins_def]
+             fam_conj_elm_simp[OF at_in])
+      apply (erule_tac sep_conjE, unfold set_ins_def)+
+      by (auto simp:sep_conj_ac fam_conj_simps st_def sg_def tpn_set_def)
+    moreover from pres(1, 3) hh(1) have "TAt i \<in> trset_of cnf" 
+      apply(erule_tac sep_conjE)
+      apply(erule_tac sep_conjE)
+      by (auto simp:st_def tpc_set_def sg_def set_ins_def)
+    ultimately have "i = ?e"
+      by (cases cnf, auto simp:tpn_set_def trset_of.simps)
+    from eq_s[unfolded this] at_in
+    show "False" by (auto simp:set_ins_def)
+  qed
+  from pres(3) and hh(2, 4, 5) pres(2, 4)
+  have pres1: "(?PP \<and>* i :[ ?code(?e)]: j \<and>*  (r \<and>* (fam_conj s' (recse_map ks))))
+         (trset_of cnf)"
+    apply (unfold hh(1) pres(1))
+    apply (rule sep_conjI[where x=ss1 and y = "ss2 + s2 + s3"], assumption+)
+    apply (rule sep_conjI[where x="s2" and y = "ss2 + s3"], assumption+)
+    apply (rule sep_conjI[where x="s3" and y = ss2], assumption+)
+    by (auto simp:set_ins_def)
+  (*****************************************************************************)
+  let ?ks_f = "\<lambda> sr ks. list_ext a ks[a:=0]"
+  let ?elm_f = "\<lambda> sr. {M a 0}"
+  let ?idx_f = "\<lambda> sr ks ia. (ia + int (reps_len (list_ext a ks)) - int (reps_len ks))"
+  (*----------------------------------------------------------------------------*)
+  (******************************************************************************)
+  from tm_hoare_dec_fail[OF `?P sr`, unfolded tm.Hoare_gen_def, rule_format, OF pres1]
+  obtain k where "((fam_conj (?elm_f sr) (recse_map (?ks_f sr ks)) \<and>*
+        st ?e \<and>*
+        ps 2 \<and>*
+        zero 0 \<and>* zero 1 \<and>* reps 2 (?idx_f sr ks ia) (?ks_f sr ks) \<and>* 
+        fam_conj {?idx_f sr ks ia<..} zero) \<and>*
+       i :[ ?code ?e ]: j \<and>* r \<and>* fam_conj s' (recse_map ks))
+       (trset_of (sep_exec.run tstep (Suc k) cnf))" by blast
+  (*----------------------------------------------------------------------------*)
+  moreover have "(pc ?e \<and>* ?Q) ({At ?e} \<union> ?elm_f sr)"
+  proof -
+    have "pc ?e {At ?e}" by (simp add:pc_def sg_def)
+  (******************************************************************************)
+    moreover have "?Q (?elm_f sr)" 
+      by (simp add:mm_def sg_def)
+  (*----------------------------------------------------------------------------*)
+    moreover 
+  (******************************************************************************)
+    have "{At ?e} ## ?elm_f sr" by (simp add:set_ins_def)
+  (*----------------------------------------------------------------------------*)
+    ultimately show ?thesis by (auto intro!:sep_conjI simp:set_ins_def)
+  qed
+  moreover 
+  (******************************************************************************)
+  from ia_disj.m_disj1 `?P sr` `s = {At i} \<union> sr` 
+  have "?elm_f sr ## s'" by (auto simp:set_ins_def mm_def sg_def)
+  (*----------------------------------------------------------------------------*)
+  with at_fresh_s' 
+  have fresh_atm: "{At ?e} \<union> ?elm_f sr ## s'" by (auto simp:set_ins_def)
+  moreover have eq_map: "\<And> elm. elm \<in> s' \<Longrightarrow> 
+                           (recse_map ks elm) = (recse_map (?ks_f sr ks) elm)"
+  proof -
+    fix elm
+    assume elm_in: "elm \<in> s'"
+    show "(recse_map ks elm) = (recse_map (?ks_f sr ks) elm)"
+    proof(cases elm)
+      (*******************************************************************)
+      case (M a' v')
+      from eq_s have "M a 0 \<in> s" by (auto simp:set_ins_def mm_def sg_def)
+      with elm_in ia_disj.m_disj1[OF this] M
+      have "a \<noteq> a'" by auto
+      thus ?thesis
+        apply (auto simp:M recse_map.simps pasrt_def list_ext_len)
+        apply (case_tac "a < length ks", auto simp:list_ext_len_eq list_ext_lt)
+        apply (case_tac "a' < Suc a", auto simp:list_ext_len_eq list_ext_lt)
+        by (metis (full_types) bot_nat_def 
+              leI le_trans less_Suc_eq_le list_ext_lt_get list_ext_tail set_zero_def)
+      (*-----------------------------------------------------------------*)
+    qed auto
+  qed
+  ultimately show ?case 
+    apply (rule_tac x = k in exI, rule_tac x = "{At ?e} \<union> ?elm_f sr" in exI)
+    apply (unfold IA_def, intro condI, assumption+)
+    apply (rule_tac x = "?ks_f sr ks" in pred_exI)
+    apply (rule_tac x = "?idx_f sr ks ia" in pred_exI)
+    apply (unfold fam_conj_disj_simp[OF fresh_atm])
+    apply (auto simp:sep_conj_ac fam_conj_simps)
+    (***************************************************************************)
+    (* apply (unfold fam_conj_insert_simp [OF h_fresh1[OF at_fresh_sr]], simp) *)
+    (*-------------------------------------------------------------------------*)
+    apply (sep_cancel)+
+    by (simp add: fam_conj_ext_eq[where I = "s'", OF eq_map])
+qed
+
+lemma hoare_dec_fail_gen[step]: 
+  assumes "v = 0"
+  shows 
+  "IA. \<lbrace> pc i ** mm a v \<rbrace>
+       i:[ (Dec a e) ]:j   
+       \<lbrace> pc e ** mm a v \<rbrace>"
+  by (unfold assms, rule hoare_dec_fail)
+
+
+lemma tm_hoare_dec_suc2:
+  assumes "mm a (Suc v) sr"
+  shows "\<lbrace> fam_conj sr (recse_map ks) \<and>*
+           st i \<and>* ps 2 \<and>* zero 0 \<and>* zero 1 \<and>* reps 2 ia ks \<and>* fam_conj {ia<..} zero \<rbrace>
+             i:[(Dec a e)]:j
+         \<lbrace> fam_conj {M a v} (recse_map (list_ext a ks[a := v])) \<and>*
+           st j \<and>*
+           ps 2 \<and>*
+           zero 0 \<and>*
+           zero 1 \<and>*
+           reps 2 (ia  - 1) (list_ext a ks[a := v]) \<and>*
+           fam_conj {ia  - 1<..} zero\<rbrace>"
+proof -
+  from `mm a (Suc v) sr` have eq_sr: "sr = {M a (Suc v)}" by (auto simp:mm_def sg_def)
+  thus ?thesis
+    apply (unfold eq_sr)
+    apply (simp add:fam_conj_simps list_ext_get_upd list_ext_len)
+    apply (rule tm.pre_condI)
+    by (drule tm_hoare_dec_suc1[where u = "2"], simp)
+qed
+
+lemma hoare_dec_suc2: 
+  "IA. \<lbrace>(pc i \<and>* mm a (Suc v))\<rbrace>  
+          i :[ Dec a e ]: j 
+       \<lbrace>pc j \<and>* mm a v\<rbrace>"
+       (is "IA. \<lbrace> pc i ** ?P \<rbrace> 
+                   i:[ ?code ?e]:j   
+                \<lbrace> pc ?e ** ?Q\<rbrace>")
+proof(induct rule:tm.IHoareI)
+  case (IPre s' s r cnf)
+  let ?cnf = "(trset_of cnf)"
+  from IPre
+  have h: "(pc i \<and>* ?P) s" "(s ## s')" "(IA (s + s') \<and>* i :[ ?code(?e) ]: j \<and>* r) ?cnf"
+    by (metis condD)+
+  from h(1) obtain sr where
+      eq_s: "s = {At i} \<union>  sr" "{At i} ##  sr" "?P sr"
+    by (auto dest!:sep_conjD simp:set_ins_def pc_def sg_def)
+  hence "At i \<in> s" by auto
+  from h(3) obtain s1 s2 s3
+    where hh: "?cnf = s1 + s2 + s3"
+              "s1 ## s2 \<and> s2 ## s3 \<and> s1 ## s3" 
+              "IA (s + s') s1" 
+              "(i :[ ?code ?e ]: j) s2"
+              "r s3"
+    apply (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac)
+    by (metis sep_add_commute sep_disj_addD1 sep_disj_addD2 sep_disj_commuteI)
+  interpret ia_disj: IA_disjoint s s' s1 cnf
+  proof
+    from `IA (s + s') s1` show "IA (s + s') s1" .
+  next
+    from `s ## s'` show "s ## s'" .
+  next
+    from hh(1) show "s1 \<subseteq> ?cnf" by (auto simp:set_ins_def)
+  qed
+  from hh(1) have s1_belong: "s1 \<subseteq> ?cnf" by (auto simp:set_ins_def)
+  from hh(3)
+  have "(EXS ks ia.
+         ps 2 \<and>*
+         zero 0 \<and>*
+         zero 1 \<and>*
+         reps 2 ia ks \<and>*
+         fam_conj {ia<..} zero \<and>*
+         (st i \<and>* fam_conj sr (recse_map ks)) \<and>* fam_conj s' (recse_map ks))
+        s1"
+    apply (unfold IA_def fam_conj_disj_simp[OF h(2)])
+    apply (unfold eq_s)
+    by (insert eq_s(2), simp add: fam_conj_disj_simp[OF eq_s(2)] fam_conj_simps set_ins_def)
+  then obtain ks ia
+    where "((fam_conj sr (recse_map ks) \<and>* st i \<and>* ps 2 \<and>* zero 0 \<and>* zero 1 \<and>*
+             reps 2 ia ks \<and>* fam_conj {ia<..} zero) \<and>* fam_conj s' (recse_map ks)) s1" 
+    (is "(?PP \<and>* ?QQ) s1")
+    by (unfold pred_ex_def, auto simp:sep_conj_ac)
+  then obtain ss1 ss2 where pres:
+        "s1 = ss1 + ss2" "ss1 ## ss2"
+        "?PP ss1"
+        "?QQ ss2"
+    by (auto elim!:sep_conjE intro!:sep_conjI)
+  from ia_disj.at_disj1 [OF `At i \<in> s`]
+  have at_fresh_s': "At ?e \<notin>  s'" .
+  have at_fresh_sr: "At ?e \<notin> sr"
+  proof
+    assume at_in: "At ?e \<in> sr"
+    from h(3)[unfolded IA_def fam_conj_disj_simp[OF h(2)]]
+    have "TAt ?e \<in> trset_of cnf"
+      apply (elim EXS_elim1)
+      apply (unfold eq_s fam_conj_disj_simp[OF eq_s(2), unfolded set_ins_def]
+             fam_conj_elm_simp[OF at_in])
+      apply (erule_tac sep_conjE, unfold set_ins_def)+
+      by (auto simp:sep_conj_ac fam_conj_simps st_def sg_def tpn_set_def)
+    moreover from pres(1, 3) hh(1) have "TAt i \<in> trset_of cnf" 
+      apply(erule_tac sep_conjE)
+      apply(erule_tac sep_conjE)
+      by (auto simp:st_def tpc_set_def sg_def set_ins_def)
+    ultimately have "i = ?e"
+      by (cases cnf, auto simp:tpn_set_def trset_of.simps)
+    from eq_s[unfolded this] at_in
+    show "False" by (auto simp:set_ins_def)
+  qed
+  from pres(3) and hh(2, 4, 5) pres(2, 4)
+  have pres1: "(?PP \<and>* i :[ ?code(?e)]: j \<and>*  (r \<and>* (fam_conj s' (recse_map ks))))
+         (trset_of cnf)"
+    apply (unfold hh(1) pres(1))
+    apply (rule sep_conjI[where x=ss1 and y = "ss2 + s2 + s3"], assumption+)
+    apply (rule sep_conjI[where x="s2" and y = "ss2 + s3"], assumption+)
+    apply (rule sep_conjI[where x="s3" and y = ss2], assumption+)
+    by (auto simp:set_ins_def)
+  (*****************************************************************************)
+  let ?ks_f = "\<lambda> sr ks. list_ext a ks[a:=v]"
+  let ?elm_f = "\<lambda> sr. {M a v}"
+  let ?idx_f = "\<lambda> sr ks ia. ia  - 1"
+  (*----------------------------------------------------------------------------*)
+  (******************************************************************************)
+  from tm_hoare_dec_suc2[OF `?P sr`, unfolded tm.Hoare_gen_def, rule_format, OF pres1]
+  obtain k where "((fam_conj (?elm_f sr) (recse_map (?ks_f sr ks)) \<and>*
+        st ?e \<and>*
+        ps 2 \<and>*
+        zero 0 \<and>* zero 1 \<and>* reps 2 (?idx_f sr ks ia) (?ks_f sr ks) \<and>* 
+        fam_conj {?idx_f sr ks ia<..} zero) \<and>*
+       i :[ ?code ?e ]: j \<and>* r \<and>* fam_conj s' (recse_map ks))
+       (trset_of (sep_exec.run tstep (Suc k) cnf))" by blast
+  (*----------------------------------------------------------------------------*)
+  moreover have "(pc ?e \<and>* ?Q) ({At ?e} \<union> ?elm_f sr)"
+  proof -
+    have "pc ?e {At ?e}" by (simp add:pc_def sg_def)
+  (******************************************************************************)
+    moreover have "?Q (?elm_f sr)" 
+      by (simp add:mm_def sg_def)
+  (*----------------------------------------------------------------------------*)
+    moreover 
+  (******************************************************************************)
+    have "{At ?e} ## ?elm_f sr" by (simp add:set_ins_def)
+  (*----------------------------------------------------------------------------*)
+    ultimately show ?thesis by (auto intro!:sep_conjI simp:set_ins_def)
+  qed
+  moreover 
+  (******************************************************************************)
+  from ia_disj.m_disj1 `?P sr` `s = {At i} \<union> sr` 
+  have "?elm_f sr ## s'" by (auto simp:set_ins_def mm_def sg_def)
+  (*----------------------------------------------------------------------------*)
+  with at_fresh_s' 
+  have fresh_atm: "{At ?e} \<union> ?elm_f sr ## s'" by (auto simp:set_ins_def)
+  moreover have eq_map: "\<And> elm. elm \<in> s' \<Longrightarrow> 
+                           (recse_map ks elm) = (recse_map (?ks_f sr ks) elm)"
+  proof -
+    fix elm
+    assume elm_in: "elm \<in> s'"
+    show "(recse_map ks elm) = (recse_map (?ks_f sr ks) elm)"
+    proof(cases elm)
+      (*******************************************************************)
+      case (M a' v')
+      from eq_s have "M a (Suc v) \<in> s" by (auto simp:set_ins_def mm_def sg_def)
+      with elm_in ia_disj.m_disj1[OF this] M
+      have "a \<noteq> a'" by auto
+      thus ?thesis
+        apply (auto simp:M recse_map.simps pasrt_def list_ext_len)
+        apply (case_tac "a < length ks", auto simp:list_ext_len_eq list_ext_lt)
+        apply (case_tac "a' < Suc a", auto simp:list_ext_len_eq list_ext_lt)
+        by (metis (full_types) bot_nat_def 
+              leI le_trans less_Suc_eq_le list_ext_lt_get list_ext_tail set_zero_def)
+      (*-----------------------------------------------------------------*)
+    qed auto
+  qed
+  ultimately show ?case 
+    apply (rule_tac x = k in exI, rule_tac x = "{At ?e} \<union> ?elm_f sr" in exI)
+    apply (unfold IA_def, intro condI, assumption+)
+    apply (rule_tac x = "?ks_f sr ks" in pred_exI)
+    apply (rule_tac x = "?idx_f sr ks ia" in pred_exI)
+    apply (unfold fam_conj_disj_simp[OF fresh_atm])
+    apply (auto simp:sep_conj_ac fam_conj_simps)
+    (***************************************************************************)
+    (* apply (unfold fam_conj_insert_simp [OF h_fresh1[OF at_fresh_sr]], simp) *)
+    (*-------------------------------------------------------------------------*)
+    apply (sep_cancel)+
+    by (simp add: fam_conj_ext_eq[where I = "s'", OF eq_map])
+qed
+
+lemma hoare_dec_suc2_gen[step]: 
+  assumes "v > 0"
+  shows 
+  "IA. \<lbrace>pc i \<and>* mm a v\<rbrace>  
+          i :[ Dec a e ]: j 
+       \<lbrace>pc j \<and>* mm a (v - 1)\<rbrace>"
+proof -
+  from assms obtain v' where "v = Suc v'"
+    by (metis gr_implies_not0 nat.exhaust)
+  show ?thesis
+    apply (unfold `v = Suc v'`, simp)
+    by (rule hoare_dec_suc2)
+qed
+
+definition "Goto e = jmp e"
+
+lemma hoare_jmp_reps2:
+      "\<lbrace> st i \<and>* ps u \<and>* reps u v ks \<and>* tm (v + 1) x \<rbrace>
+                 i:[(jmp e)]:j
+       \<lbrace> st e \<and>* ps u \<and>* reps u v ks \<and>* tm (v + 1) x \<rbrace>"
+proof(cases "ks")
+  case Nil
+  thus ?thesis
+    by (simp add:sep_conj_cond reps.simps, intro tm.pre_condI, simp, hsteps)
+next
+  case (Cons k ks')
+  thus ?thesis
+  proof(cases "ks' = []")
+    case True with Cons
+    show ?thesis
+      apply(simp add:sep_conj_cond reps.simps, intro tm.pre_condI, simp add:ones_simps)
+      by (hgoto hoare_jmp[where p = u])
+  next
+    case False
+    show ?thesis
+      apply (unfold `ks = k#ks'` reps_simp3[OF False], simp add:ones_simps)
+      by (hgoto hoare_jmp[where p = u])
+  qed
+qed
+
+lemma tm_hoare_goto_pre: (* ccc *)
+  assumes "(<True>) sr"
+  shows "\<lbrace> fam_conj sr (recse_map ks) \<and>*
+           st i \<and>* ps 2 \<and>* zero 0 \<and>* zero 1 \<and>* reps 2 ia ks \<and>* fam_conj {ia<..} zero \<rbrace>
+             i:[(Goto e)]:j
+         \<lbrace> fam_conj {} (recse_map ks) \<and>*
+           st e \<and>* ps 2 \<and>* zero 0 \<and>* zero 1 \<and>* reps 2 ia ks \<and>* fam_conj {ia<..} zero \<rbrace>"
+  apply (unfold Goto_def)
+  apply (subst (1 2) fam_conj_interv_simp)
+  apply (unfold zero_def)
+  apply (hstep hoare_jmp_reps2)
+  apply (simp add:sep_conj_ac)
+  my_block
+    from assms have "sr = {}"
+      by (simp add:pasrt_def set_ins_def)
+  my_block_end
+  by (unfold this, sep_cancel+)
+  
+lemma hoare_goto_pre: 
+  "IA. \<lbrace> pc i \<and>* <True> \<rbrace> 
+        i:[ (Goto e) ]:j   
+       \<lbrace> pc e \<and>* <True> \<rbrace>"
+  (is "IA. \<lbrace> pc i ** ?P \<rbrace> 
+              i:[ ?code ?e]:j   
+           \<lbrace> pc ?e ** ?Q\<rbrace>")
+proof(induct rule:tm.IHoareI)
+  case (IPre s' s r cnf)
+  let ?cnf = "(trset_of cnf)"
+  from IPre
+  have h: "(pc i \<and>* ?P) s" "(s ## s')" "(IA (s + s') \<and>* i :[ ?code(?e) ]: j \<and>* r) ?cnf"
+    by (metis condD)+
+  from h(1) obtain sr where
+      eq_s: "s = {At i} \<union>  sr" "{At i} ##  sr" "?P sr"
+    by (auto dest!:sep_conjD simp:set_ins_def pc_def sg_def pasrt_def)
+  hence "At i \<in> s" by auto
+  from h(3) obtain s1 s2 s3
+    where hh: "?cnf = s1 + s2 + s3"
+              "s1 ## s2 \<and> s2 ## s3 \<and> s1 ## s3" 
+              "IA (s + s') s1" 
+              "(i :[ ?code ?e ]: j) s2"
+              "r s3"
+    apply (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac)
+    by (metis sep_add_commute sep_disj_addD1 sep_disj_addD2 sep_disj_commuteI)
+  interpret ia_disj: IA_disjoint s s' s1 cnf
+  proof
+    from `IA (s + s') s1` show "IA (s + s') s1" .
+  next
+    from `s ## s'` show "s ## s'" .
+  next
+    from hh(1) show "s1 \<subseteq> ?cnf" by (auto simp:set_ins_def)
+  qed
+  from hh(1) have s1_belong: "s1 \<subseteq> ?cnf" by (auto simp:set_ins_def)
+  from hh(3)
+  have "(EXS ks ia.
+         ps 2 \<and>*
+         zero 0 \<and>*
+         zero 1 \<and>*
+         reps 2 ia ks \<and>*
+         fam_conj {ia<..} zero \<and>*
+         (st i \<and>* fam_conj sr (recse_map ks)) \<and>* fam_conj s' (recse_map ks))
+        s1"
+    apply (unfold IA_def fam_conj_disj_simp[OF h(2)])
+    apply (unfold eq_s)
+    by (insert eq_s(2), simp add: fam_conj_disj_simp[OF eq_s(2)] fam_conj_simps set_ins_def)
+  then obtain ks ia
+    where "((fam_conj sr (recse_map ks) \<and>* st i \<and>* ps 2 \<and>* zero 0 \<and>* zero 1 \<and>*
+             reps 2 ia ks \<and>* fam_conj {ia<..} zero) \<and>* fam_conj s' (recse_map ks)) s1" 
+    (is "(?PP \<and>* ?QQ) s1")
+    by (unfold pred_ex_def, auto simp:sep_conj_ac)
+  then obtain ss1 ss2 where pres:
+        "s1 = ss1 + ss2" "ss1 ## ss2"
+        "?PP ss1"
+        "?QQ ss2"
+    by (auto elim!:sep_conjE intro!:sep_conjI)
+  from ia_disj.at_disj1 [OF `At i \<in> s`]
+  have at_fresh_s': "At ?e \<notin>  s'" .
+  have at_fresh_sr: "At ?e \<notin> sr"
+  proof
+    assume at_in: "At ?e \<in> sr"
+    from h(3)[unfolded IA_def fam_conj_disj_simp[OF h(2)]]
+    have "TAt ?e \<in> trset_of cnf"
+      apply (elim EXS_elim1)
+      apply (unfold eq_s fam_conj_disj_simp[OF eq_s(2), unfolded set_ins_def]
+             fam_conj_elm_simp[OF at_in])
+      apply (erule_tac sep_conjE, unfold set_ins_def)+
+      by (auto simp:sep_conj_ac fam_conj_simps st_def sg_def tpn_set_def)
+    moreover from pres(1, 3) hh(1) have "TAt i \<in> trset_of cnf" 
+      apply(erule_tac sep_conjE)
+      apply(erule_tac sep_conjE)
+      by (auto simp:st_def tpc_set_def sg_def set_ins_def)
+    ultimately have "i = ?e"
+      by (cases cnf, auto simp:tpn_set_def trset_of.simps)
+    from eq_s[unfolded this] at_in
+    show "False" by (auto simp:set_ins_def)
+  qed
+  from pres(3) and hh(2, 4, 5) pres(2, 4)
+  have pres1: "(?PP \<and>* i :[ ?code(?e)]: j \<and>*  (r \<and>* (fam_conj s' (recse_map ks))))
+         (trset_of cnf)"
+    apply (unfold hh(1) pres(1))
+    apply (rule sep_conjI[where x=ss1 and y = "ss2 + s2 + s3"], assumption+)
+    apply (rule sep_conjI[where x="s2" and y = "ss2 + s3"], assumption+)
+    apply (rule sep_conjI[where x="s3" and y = ss2], assumption+)
+    by (auto simp:set_ins_def)
+  (*****************************************************************************)
+  let ?ks_f = "\<lambda> sr ks. ks"
+  let ?elm_f = "\<lambda> sr. {}"
+  let ?idx_f = "\<lambda> sr ks ia. ia"
+  (*----------------------------------------------------------------------------*)
+  (******************************************************************************)
+  from tm_hoare_goto_pre[OF `?P sr`, unfolded tm.Hoare_gen_def, rule_format, OF pres1]
+  obtain k where "((fam_conj (?elm_f sr) (recse_map (?ks_f sr ks)) \<and>*
+        st ?e \<and>*
+        ps 2 \<and>*
+        zero 0 \<and>* zero 1 \<and>* reps 2 (?idx_f sr ks ia) (?ks_f sr ks) \<and>* 
+        fam_conj {?idx_f sr ks ia<..} zero) \<and>*
+       i :[ ?code ?e ]: j \<and>* r \<and>* fam_conj s' (recse_map ks))
+       (trset_of (sep_exec.run tstep (Suc k) cnf))" by blast
+  (*----------------------------------------------------------------------------*)
+  moreover have "(pc ?e \<and>* ?Q) ({At ?e} \<union> ?elm_f sr)"
+  proof -
+    have "pc ?e {At ?e}" by (simp add:pc_def sg_def)
+  (******************************************************************************)
+    moreover have "?Q (?elm_f sr)" 
+     by (simp add:pasrt_def set_ins_def)
+  (*----------------------------------------------------------------------------*)
+    moreover 
+  (******************************************************************************)
+    have "{At ?e} ## ?elm_f sr" by (simp add:set_ins_def)
+  (*----------------------------------------------------------------------------*)
+    ultimately show ?thesis by (auto intro!:sep_conjI simp:set_ins_def)
+  qed
+  moreover 
+  (******************************************************************************)
+  from ia_disj.m_disj1 `?P sr` `s = {At i} \<union> sr` 
+  have "?elm_f sr ## s'" by (auto simp:set_ins_def mm_def sg_def)
+  (*----------------------------------------------------------------------------*)
+  with at_fresh_s' 
+  have fresh_atm: "{At ?e} \<union> ?elm_f sr ## s'" by (auto simp:set_ins_def)
+  moreover have eq_map: "\<And> elm. elm \<in> s' \<Longrightarrow> 
+                           (recse_map ks elm) = (recse_map (?ks_f sr ks) elm)"
+  proof -
+    fix elm
+    assume elm_in: "elm \<in> s'"
+    show "(recse_map ks elm) = (recse_map (?ks_f sr ks) elm)" 
+      by simp
+  qed
+  ultimately show ?case 
+    apply (rule_tac x = k in exI, rule_tac x = "{At ?e} \<union> ?elm_f sr" in exI)
+    apply (unfold IA_def, intro condI, assumption+)
+    apply (rule_tac x = "?ks_f sr ks" in pred_exI)
+    apply (rule_tac x = "?idx_f sr ks ia" in pred_exI)
+    apply (unfold fam_conj_disj_simp[OF fresh_atm])
+    by (auto simp:sep_conj_ac fam_conj_simps)
+qed
+
+lemma hoare_goto[step]: "IA. \<lbrace> pc i \<rbrace> 
+                          i:[ (Goto e) ]:j   
+                       \<lbrace> pc e \<rbrace>"
+proof(rule tm.I_hoare_adjust [OF hoare_goto_pre])
+  fix s assume "pc i s" thus "(pc i \<and>* <True>) s"
+    by (metis cond_true_eq2)
+next
+  fix s assume "(pc e \<and>* <True>) s" thus "pc e s"
+    by (metis cond_true_eq2)
+qed
+
+lemma I_hoare_sequence: 
+  assumes h1: "\<And> i j. I. \<lbrace>pc i ** p\<rbrace> i:[c1]:j \<lbrace>pc j ** q\<rbrace>"
+  and h2: "\<And> j k. I. \<lbrace>pc j ** q\<rbrace> j:[c2]:k \<lbrace>pc k ** r\<rbrace>"
+  shows "I. \<lbrace>pc i ** p\<rbrace> i:[(c1 ; c2)]:k \<lbrace>pc k ** r\<rbrace>"
+proof(unfold tassemble_to.simps, intro tm.I_code_exI)
+  fix j'
+  show "I.\<lbrace>pc i \<and>* p\<rbrace>  i :[ c1 ]: j' \<and>* j' :[ c2 ]: k \<lbrace>pc k \<and>* r\<rbrace>"
+  proof(rule tm.I_sequencing)
+    from tm.I_code_extension[OF h1 [of i j'], of" j' :[ c2 ]: k"]
+    show "I.\<lbrace>pc i \<and>* p\<rbrace>  i :[ c1 ]: j' \<and>* j' :[ c2 ]: k \<lbrace>pc j' \<and>* q\<rbrace>" .
+  next
+    from tm.I_code_extension[OF h2 [of j' k], of" i :[ c1 ]: j'"]
+    show "I.\<lbrace>pc j' \<and>* q\<rbrace>  i :[ c1 ]: j' \<and>* j' :[ c2 ]: k \<lbrace>pc k \<and>* r\<rbrace>"
+      by (auto simp:sep_conj_ac)
+  qed
+qed
+
+lemma I_hoare_seq1:
+  assumes h1: "\<And>j'. I. \<lbrace>pc i ** p\<rbrace> i:[c1]:j' \<lbrace>pc j' ** q\<rbrace>"
+  and h2: "\<And>j' . I. \<lbrace>pc j' ** q\<rbrace> j':[c2]:k \<lbrace>pc k' ** r\<rbrace>"
+  shows "I. \<lbrace>pc i ** p\<rbrace> i:[(c1 ; c2)]:k \<lbrace>pc k' ** r\<rbrace>"
+proof(unfold tassemble_to.simps, intro tm.I_code_exI)
+  fix j'
+  show "I.\<lbrace>pc i \<and>* p\<rbrace>  i :[ c1 ]: j' \<and>* j' :[ c2 ]: k \<lbrace>pc k' \<and>* r\<rbrace>"
+  proof(rule tm.I_sequencing)
+    from tm.I_code_extension[OF h1 [of j'], of "j' :[ c2 ]: k "]
+    show "I.\<lbrace>pc i \<and>* p\<rbrace>  i :[ c1 ]: j' \<and>* j' :[ c2 ]: k \<lbrace>pc j' \<and>* q\<rbrace>" .
+  next
+    from tm.I_code_extension[OF h2 [of j'], of" i :[ c1 ]: j'"]
+    show "I.\<lbrace>pc j' \<and>* q\<rbrace>  i :[ c1 ]: j' \<and>* j' :[ c2 ]: k \<lbrace>pc k' \<and>* r\<rbrace>"
+      by (auto simp:sep_conj_ac)
+  qed
+qed
+
+lemma t_hoare_local1: 
+  "(\<And>l. \<lbrace>p\<rbrace>  i :[ c l ]: j \<lbrace>q\<rbrace>) \<Longrightarrow>
+    \<lbrace>p\<rbrace> i:[TLocal c]:j \<lbrace>q\<rbrace>"
+by (unfold tassemble_to.simps, rule tm.code_exI, auto)
+
+lemma I_hoare_local:
+  assumes h: "(\<And>l. I.\<lbrace>pc i \<and>* p\<rbrace>  i :[ c l ]: j \<lbrace>pc k \<and>* q\<rbrace>)"
+  shows "I. \<lbrace>pc i ** p\<rbrace> i:[TLocal c]:j \<lbrace>pc k ** q\<rbrace>"
+proof(unfold tassemble_to.simps, rule tm.I_code_exI)
+  fix l
+  from h[of l]
+  show " I.\<lbrace>pc i \<and>* p\<rbrace>  i :[ c l ]: j \<lbrace>pc k \<and>* q\<rbrace>" .
+qed
+
+lemma t_hoare_label1: 
+      "(\<And>l. l = i \<Longrightarrow> \<lbrace>p\<rbrace>  l :[ c l ]: j \<lbrace>q\<rbrace>) \<Longrightarrow>
+             \<lbrace>p \<rbrace> 
+               i:[(TLabel l; c l)]:j
+             \<lbrace>q\<rbrace>"
+by (unfold tassemble_to.simps, intro tm.code_exI tm.code_condI, clarify, auto)
+
+lemma I_hoare_label: 
+  assumes h:"\<And>l. l = i \<Longrightarrow> I. \<lbrace>pc l \<and>* p\<rbrace>  l :[ c l ]: j \<lbrace>pc k \<and>* q\<rbrace>"
+  shows "I. \<lbrace>pc i \<and>* p \<rbrace> 
+               i:[(TLabel l; c l)]:j
+             \<lbrace>pc k \<and>* q\<rbrace>"
+proof(unfold tm.IHoare_def, default)
+  fix s'
+  show " \<lbrace>EXS s. <(pc i \<and>* p) s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>  i :[ (TLabel l ; c l) ]: j
+         \<lbrace>EXS s. <(pc k \<and>* q) s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>"
+  proof(rule t_hoare_label1)
+    fix l assume "l = i"
+    from h[OF this, unfolded tm.IHoare_def]
+    show "\<lbrace>EXS s. <(pc i \<and>* p) s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>  l :[ c l ]: j
+          \<lbrace>EXS s. <(pc k \<and>* q) s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>" 
+      by (simp add:`l = i`)
+  qed
+qed
+
+lemma I_hoare_label_last: 
+  assumes h1: "t_last_cmd c = Some (TLabel l)"
+  and h2: "l = j \<Longrightarrow> I. \<lbrace>p\<rbrace> i:[t_blast_cmd c]:j \<lbrace>q\<rbrace>"
+  shows "I. \<lbrace>p\<rbrace> i:[c]:j \<lbrace>q\<rbrace>"
+proof(unfold tm.IHoare_def, default)
+  fix s'
+  show "\<lbrace>EXS s. <p s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>  i :[ c ]: j 
+        \<lbrace>EXS s. <q s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>"
+  proof(rule t_hoare_label_last[OF h1])
+    assume "l = j"
+    from h2[OF this, unfolded tm.IHoare_def]
+    show "\<lbrace>EXS s. <p s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>  i :[ t_blast_cmd c ]: j
+          \<lbrace>EXS s. <q s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>"
+      by fast
+  qed
+qed
+
+lemma I_hoare_seq2:
+ assumes h: "\<And>j'. I. \<lbrace>pc i ** p\<rbrace> i:[c1]:j' \<lbrace>pc k' \<and>* r\<rbrace>"
+ shows "I. \<lbrace>pc i ** p\<rbrace> i:[(c1 ; c2)]:j \<lbrace>pc k' ** r\<rbrace>"
+  apply (unfold tassemble_to.simps, intro tm.I_code_exI)
+  apply (unfold tm.IHoare_def, default)
+  apply (rule tm.code_extension)
+  by (rule h[unfolded tm.IHoare_def, rule_format])
+
+lemma IA_pre_stren:
+  assumes h1: "IA. \<lbrace>p\<rbrace> c \<lbrace>q\<rbrace>"
+  and h2:  "\<And>s. r s \<Longrightarrow> p s"
+  shows "IA. \<lbrace>r\<rbrace> c \<lbrace>q\<rbrace>"
+  by (rule tm.I_pre_stren[OF assms], simp)
+
+lemma IA_post_weaken: 
+  assumes h1: "IA. \<lbrace>p\<rbrace> c \<lbrace>q\<rbrace>"
+    and h2: "\<And> s. q s \<Longrightarrow> r s"
+  shows "IA. \<lbrace>p\<rbrace> c \<lbrace>r\<rbrace>"
+  by (rule tm.I_post_weaken[OF assms], simp)
+
+section {* Making triple processor for IA *}
+
+ML {* (* Functions specific to Hoare triple: IA {P} c {Q} *)
+  fun get_pre ctxt t = 
+    let val pat = term_of @{cpat "IA. \<lbrace>?P\<rbrace> ?c \<lbrace>?Q\<rbrace>"} 
+        val env = match ctxt pat t 
+    in inst env (term_of @{cpat "?P::aresource set \<Rightarrow> bool"}) end
+
+  fun can_process ctxt t = ((get_pre ctxt t; true) handle _ => false)
+
+  fun get_post ctxt t = 
+    let val pat = term_of @{cpat "IA. \<lbrace>?P\<rbrace> ?c \<lbrace>?Q\<rbrace>"} 
+        val env = match ctxt pat t 
+    in inst env (term_of @{cpat "?Q::aresource set \<Rightarrow> bool"}) end;
+
+  fun get_mid ctxt t = 
+    let val pat = term_of @{cpat "IA. \<lbrace>?P\<rbrace> ?c \<lbrace>?Q\<rbrace>"} 
+        val env = match ctxt pat t 
+    in inst env (term_of @{cpat "?c::tresource set \<Rightarrow> bool"}) end;
+
+  fun is_pc_term (Const (@{const_name pc}, _) $ _) = true
+    | is_pc_term _ = false
+
+  fun mk_pc_term x =
+     Const (@{const_name pc}, @{typ "nat \<Rightarrow> aresource set \<Rightarrow> bool"}) $ Free (x, @{typ "nat"})
+
+  val sconj_term = term_of @{cterm "sep_conj::assert \<Rightarrow> assert \<Rightarrow> assert"}
+
+  val abc_triple = {binding = @{binding "abc_triple"}, 
+                   can_process = can_process,
+                   get_pre = get_pre,
+                   get_mid = get_mid,
+                   get_post = get_post,
+                   is_pc_term = is_pc_term,
+                   mk_pc_term = mk_pc_term,
+                   sconj_term = sconj_term,
+                   sep_conj_ac_tac = sep_conj_ac_tac,
+                   hoare_seq1 = @{thm I_hoare_seq1},
+                   hoare_seq2 = @{thm I_hoare_seq2},
+                   pre_stren = @{thm IA_pre_stren},
+                   post_weaken = @{thm IA_post_weaken},
+                   frame_rule = @{thm tm.I_frame_rule}
+                  }:HoareTriple
+
+  val _ = (HoareTriples_get ()) |> (fn orig => HoareTriples_store (abc_triple::orig))
+*}
+
+section {* Example proofs *}
+
+definition "clear a = (TL start exit. TLabel start; Dec a exit; Goto start; TLabel exit)"
+
+lemma hoare_clear[step]: 
+  "IA. \<lbrace>pc i ** mm a v\<rbrace> 
+         i:[clear a]:j 
+       \<lbrace>pc j ** mm a 0\<rbrace>"
+proof(unfold clear_def, intro I_hoare_local I_hoare_label, simp,
+      rule I_hoare_label_last, simp+, prune)
+  show "IA.\<lbrace>pc i \<and>* mm a v\<rbrace>  i :[ (Dec a j ; Goto i) ]: j \<lbrace>pc j \<and>* mm a 0\<rbrace>"
+  proof(induct v)
+    case 0
+    show ?case
+      by hgoto
+  next
+    case (Suc v)
+    show ?case
+      apply (rule_tac Q = "pc i \<and>* mm a v" in tm.I_sequencing)
+      by hsteps
+  qed
+qed
+
+definition "dup a b c = 
+       (TL start exit. TLabel start; Dec a exit; Inc b; Inc c; Goto start; TLabel exit)"
+
+lemma hoare_dup[step]: 
+        "IA. \<lbrace>pc i ** mm a va ** mm b vb ** mm c vc \<rbrace> 
+                  i:[dup a b c]:j 
+             \<lbrace>pc j ** mm a 0 ** mm b (va + vb) ** mm c (va + vc)\<rbrace>"
+proof(unfold dup_def, intro I_hoare_local I_hoare_label, clarsimp,
+     rule I_hoare_label_last, simp+, prune)
+  show "IA. \<lbrace>pc i \<and>* mm a va \<and>* mm b vb \<and>* mm c vc\<rbrace>  
+               i :[ (Dec a j ; Inc b ; Inc c ; Goto i) ]: j
+            \<lbrace>pc j \<and>* mm a 0 \<and>* mm b (va + vb) \<and>* mm c (va + vc)\<rbrace>"
+  proof(induct va arbitrary: vb vc)
+    case (0 vb vc)
+    show ?case
+      by hgoto
+  next
+    case (Suc va vb vc)
+    show ?case
+      apply (rule_tac Q = "pc i \<and>* mm a va \<and>* mm b (Suc vb) \<and>* mm c (Suc vc)" in tm.I_sequencing)
+      by (hsteps Suc)
+  qed
+qed
+
+definition "clear_add a b = 
+               (TL start exit. TLabel start; Dec a exit; Inc b; Goto start; TLabel exit)"
+
+lemma hoare_clear_add[step]: 
+                 "IA. \<lbrace>pc i ** mm a va ** mm b vb \<rbrace> 
+                          i:[clear_add a b]:j 
+                      \<lbrace>pc j ** mm a 0 ** mm b (va + vb)\<rbrace>"
+proof(unfold clear_add_def, intro I_hoare_local I_hoare_label, clarsimp,
+      rule I_hoare_label_last, simp+, prune)
+  show "IA. \<lbrace>pc i \<and>* mm a va \<and>* mm b vb\<rbrace>  
+               i :[ (Dec a j ; Inc b ; Goto i) ]: j 
+            \<lbrace>pc j \<and>* mm a 0 \<and>* mm b (va + vb)\<rbrace>"
+  proof(induct va arbitrary: vb)
+    case 0
+    show ?case
+      by hgoto
+  next
+    case (Suc va vb)
+    show ?case
+      apply (rule_tac Q = "pc i \<and>* mm a va \<and>* mm b (Suc vb)" in tm.I_sequencing)
+      by (hsteps Suc)
+  qed
+qed
+
+definition "copy_to a b c = clear b; clear c; dup a b c; clear_add c a"
+
+lemma hoare_copy_to[step]: 
+    "IA. \<lbrace>pc i ** mm a va ** mm b vb ** mm c vc \<rbrace> 
+               i:[copy_to a b c]:j 
+         \<lbrace>pc j ** mm a va ** mm b va ** mm c 0\<rbrace>"
+  by (unfold copy_to_def, hsteps)
+
+definition "preserve_add a b c = clear c; dup a b c; clear_add c a"
+
+lemma hoare_preserve_add[step]: 
+    "IA. \<lbrace>pc i ** mm a va ** mm b vb ** mm c vc \<rbrace> 
+              i:[preserve_add a b c]:j 
+         \<lbrace>pc j ** mm a va ** mm b (va + vb) ** mm c 0\<rbrace>"
+  by (unfold preserve_add_def, hsteps)
+
+definition "mult a b c t1 t2 = 
+                clear c;
+                copy_to a t2 t1; 
+                (TL start exit. 
+                    TLabel start;
+                    Dec a exit;
+                    preserve_add b c t1;
+                    Goto start;
+                    TLabel exit
+                );
+                clear_add t2 a"
+
+lemma hoare_mult[step]: 
+    "IA. \<lbrace>pc i ** mm a va ** mm b vb ** mm c vc ** mm t1 vt1 ** mm t2 vt2 \<rbrace> 
+                         i:[mult a b c t1 t2]:j 
+         \<lbrace>pc j ** mm a va ** mm b vb ** mm c (va * vb) ** mm t1 0 ** mm t2 0 \<rbrace>"
+  apply (unfold mult_def, hsteps)
+  apply (rule_tac q = "mm a 0 \<and>* mm b vb \<and>* mm c (va * vb) \<and>* mm t1 0 \<and>* mm t2 va" in I_hoare_seq1)
+  apply (intro I_hoare_local I_hoare_label, clarify,
+      rule I_hoare_label_last, simp+, clarify, prune)
+  my_block
+  fix i j vc
+    have "IA. \<lbrace>pc i \<and>* mm a va \<and>* mm t1 0 \<and>* mm c vc \<and>* mm b vb\<rbrace> 
+               i :[ (Dec a j ; preserve_add b c t1 ; Goto i) ]: j
+              \<lbrace>pc j \<and>* mm a 0 \<and>* mm b vb \<and>* mm c (va * vb + vc) \<and>* mm t1 0 \<rbrace>"
+    proof(induct va arbitrary:vc)
+      case (0 vc)
+      show ?case
+        by hgoto
+    next
+      case (Suc va vc)
+      show ?case
+        apply (rule_tac Q = "pc i \<and>* mm a va \<and>* mm t1 0 \<and>* mm c (vb + vc) \<and>* mm b vb"
+                in tm.I_sequencing)
+        apply (hsteps Suc)
+        by (sep_cancel+, simp, smt)
+    qed
+  my_block_end
+  by (hsteps this)
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/Hoare_gen.thy	Thu Mar 06 13:28:38 2014 +0000
@@ -0,0 +1,602 @@
+header {* 
+  Generic Separation Logic
+*}
+
+theory Hoare_gen
+imports Main  
+  (*"../progtut/Tactical" *)
+  "../Separation_Algebra/Sep_Tactics"
+begin
+
+ML_file "../Separation_Algebra/sep_tactics.ML"
+
+definition pasrt :: "bool \<Rightarrow> (('a::sep_algebra) \<Rightarrow> bool)" ("<_>" [72] 71)
+where "pasrt b = (\<lambda> s . s = 0 \<and> b)"
+
+lemma sep_conj_cond1: "(p \<and>* <cond> \<and>* q) = (<cond> \<and>* p \<and>* q)"
+  by(simp add: sep_conj_ac)
+
+lemma sep_conj_cond2: "(p \<and>* <cond>) = (<cond> \<and>* p)"
+  by(simp add: sep_conj_ac)
+
+lemma sep_conj_cond3: "((<cond> \<and>* p) \<and>* r) = (<cond> \<and>* p \<and>* r)"
+  by (metis sep.mult_assoc)
+
+lemmas sep_conj_cond = sep_conj_cond1 sep_conj_cond2 sep_conj_cond3
+
+lemma cond_true_eq[simp]: "<True> = \<box>"
+  by(unfold sep_empty_def pasrt_def, auto)
+
+lemma cond_true_eq1: "(<True> \<and>* p) = p"
+  by(simp)
+
+lemma false_simp [simp]: "<False> = sep_false" (* move *)
+  by (simp add:pasrt_def)
+
+lemma cond_true_eq2: "(p \<and>* <True>) = p"
+  by simp
+
+lemma condD: "(<b> ** r) s \<Longrightarrow> b \<and> r s" 
+by (unfold sep_conj_def pasrt_def, auto)
+
+locale sep_exec = 
+   fixes step :: "'conf \<Rightarrow> 'conf"
+    and  recse:: "'conf \<Rightarrow> 'a::sep_algebra"
+begin 
+
+definition "run n = step ^^ n"
+
+lemma run_add: "run (n1 + n2) s = run n1 (run n2 s)"
+  apply (unfold run_def)
+  by (metis funpow_add o_apply)
+
+definition
+  Hoare_gen :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool)  \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" 
+                  ("(\<lbrace>(1_)\<rbrace> / (_)/ \<lbrace>(1_)\<rbrace>)" 50)
+where
+  "\<lbrace> p \<rbrace> c \<lbrace> q \<rbrace> \<equiv> 
+      (\<forall> s r. (p**c**r) (recse s) \<longrightarrow> (\<exists> k. ((q ** c ** r) (recse (run (Suc k) s)))))"
+
+lemma HoareI [case_names Pre]: 
+  assumes h: "\<And> r s. (p**c**r) (recse s) \<Longrightarrow> (\<exists> k. ((q ** c ** r) (recse (run (Suc k) s))))"
+  shows "\<lbrace> p \<rbrace> c \<lbrace> q \<rbrace>"
+  using h
+  by (unfold Hoare_gen_def, auto)
+
+lemma frame_rule: 
+  assumes h: "\<lbrace> p \<rbrace> c \<lbrace> q \<rbrace>"
+  shows "\<lbrace> p ** r \<rbrace> c \<lbrace> q ** r \<rbrace>"
+proof(induct rule: HoareI)
+  case (Pre r' s')
+  hence "(p \<and>* c \<and>* r \<and>* r') (recse s')" by (auto simp:sep_conj_ac)
+  from h[unfolded Hoare_gen_def, rule_format, OF this]
+  show ?case
+    by (metis sep_conj_assoc sep_conj_left_commute)
+qed
+
+lemma sequencing: 
+  assumes h1: "\<lbrace>p\<rbrace> c \<lbrace>q\<rbrace>"
+  and h2: "\<lbrace>q\<rbrace> c \<lbrace>r\<rbrace>"
+  shows "\<lbrace>p\<rbrace> c \<lbrace>r\<rbrace>"
+proof(induct rule:HoareI)
+  case (Pre r' s')
+  from h1[unfolded Hoare_gen_def, rule_format, OF Pre]
+  obtain k1 where "(q \<and>* c \<and>* r') (recse (run (Suc k1) s'))" by auto
+  from h2[unfolded Hoare_gen_def, rule_format, OF this]
+  obtain k2 where "(r \<and>* c \<and>* r') (recse (run (Suc k2) (run (Suc k1) s')))" by auto
+  thus ?case
+    apply (rule_tac x = "Suc (k1 + k2)" in exI)
+    by (metis add_Suc_right nat_add_commute sep_exec.run_add)
+qed
+
+lemma pre_stren: 
+  assumes h1: "\<lbrace>p\<rbrace> c \<lbrace>q\<rbrace>"
+  and h2:  "\<And>s. r s \<Longrightarrow> p s"
+  shows "\<lbrace>r\<rbrace> c \<lbrace>q\<rbrace>"
+proof(induct rule:HoareI)
+  case (Pre r' s')
+  with h2
+  have "(p \<and>* c \<and>* r') (recse s')"
+    by (metis sep_conj_impl1)
+  from h1[unfolded Hoare_gen_def, rule_format, OF this]
+  show ?case .
+qed
+
+lemma post_weaken: 
+  assumes h1: "\<lbrace>p\<rbrace> c \<lbrace>q\<rbrace>"
+    and h2: "\<And> s. q s \<Longrightarrow> r s"
+  shows "\<lbrace>p\<rbrace> c \<lbrace>r\<rbrace>"
+proof(induct rule:HoareI)
+  case (Pre r' s')
+  from h1[unfolded Hoare_gen_def, rule_format, OF this]
+  obtain k where "(q \<and>* c \<and>* r') (recse (run (Suc k) s'))" by blast
+  with h2
+  show ?case
+    by (metis sep_conj_impl1)
+qed
+
+lemma hoare_adjust:
+  assumes h1: "\<lbrace>p1\<rbrace> c \<lbrace>q1\<rbrace>"
+  and h2: "\<And>s. p s \<Longrightarrow> p1 s"
+  and h3: "\<And>s. q1 s \<Longrightarrow> q s"
+  shows "\<lbrace>p\<rbrace> c \<lbrace>q\<rbrace>"
+  using h1 h2 h3 post_weaken pre_stren
+  by (metis)
+
+lemma code_exI: 
+  assumes h: "\<And> k. \<lbrace>p\<rbrace> c(k) \<lbrace>q\<rbrace>"
+  shows "\<lbrace>p\<rbrace> EXS k. c(k) \<lbrace>q\<rbrace>"
+proof(unfold pred_ex_def, induct rule:HoareI)
+  case (Pre r' s')
+  then obtain k where "(p \<and>* (\<lambda> s. c k s) \<and>* r') (recse s')"
+    by (auto elim!:sep_conjE intro!:sep_conjI)
+  from h[unfolded Hoare_gen_def, rule_format, OF this]
+  show ?case
+   by (auto elim!:sep_conjE intro!:sep_conjI)
+qed
+
+lemma code_extension: 
+  assumes h: "\<lbrace> p \<rbrace> c \<lbrace> q \<rbrace>"
+  shows "\<lbrace> p \<rbrace> c ** e \<lbrace> q \<rbrace>"
+proof(induct rule:HoareI)
+  case (Pre r' s')
+  hence "(p \<and>* c \<and>* e \<and>* r') (recse s')" by (auto simp:sep_conj_ac)
+  from h[unfolded Hoare_gen_def, rule_format, OF this]
+  show ?case
+    by (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac)
+qed
+
+lemma code_extension1: 
+  assumes h: "\<lbrace> p \<rbrace> c \<lbrace> q \<rbrace>"
+  shows "\<lbrace> p \<rbrace> e ** c \<lbrace> q \<rbrace>"
+  by (metis code_extension h sep.mult_commute)
+
+lemma composition: 
+  assumes h1: "\<lbrace>p\<rbrace> c1 \<lbrace>q\<rbrace>"
+    and h2: "\<lbrace>q\<rbrace> c2 \<lbrace>r\<rbrace>"
+  shows "\<lbrace>p\<rbrace> c1 ** c2 \<lbrace>r\<rbrace>"
+proof(induct rule:HoareI)
+  case (Pre r' s')
+  hence "(p \<and>* c1 \<and>* c2 \<and>* r') (recse s')" by (auto simp:sep_conj_ac)
+  from h1[unfolded Hoare_gen_def, rule_format, OF this]
+  obtain k1 where "(q \<and>* c2 \<and>* c1 \<and>* r') (recse (run (Suc k1) s'))" 
+    by (auto simp:sep_conj_ac)
+  from h2[unfolded Hoare_gen_def, rule_format, OF this, folded run_add]
+  show ?case
+    by (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac)
+qed
+
+
+definition
+  IHoare :: "('b::sep_algebra \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 
+                ('b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool)  \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> bool" 
+                  ("(1_).(\<lbrace>(1_)\<rbrace> / (_)/ \<lbrace>(1_)\<rbrace>)" 50)
+where
+  "I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace> = (\<forall>s'. \<lbrace> EXS s. <P s> \<and>* <(s ## s')> \<and>* I(s + s')\<rbrace> c 
+                         \<lbrace> EXS s. <Q s> \<and>* <(s ## s')> \<and>* I(s + s')\<rbrace>)"
+
+lemma IHoareI [case_names IPre]: 
+  assumes h: "\<And>s' s r cnf .  (<P s> \<and>* <(s ## s')> \<and>* I(s + s') \<and>* c \<and>* r) (recse cnf) \<Longrightarrow> 
+                   (\<exists>k t. (<Q t> \<and>* <(t ## s')>  \<and>* I(t + s') \<and>* c \<and>* r) (recse (run (Suc k) cnf)))"
+  shows "I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace>"
+  unfolding IHoare_def
+proof
+  fix s'
+  show " \<lbrace>EXS s. <P s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>  c
+         \<lbrace>EXS s. <Q s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>"
+  proof(unfold pred_ex_def, induct rule:HoareI)
+    case (Pre r s)
+    then obtain sa where "(<P sa> \<and>* <(sa ## s')> \<and>* I (sa + s') \<and>* c \<and>* r) (recse s)"
+      by (auto elim!:sep_conjE intro!:sep_conjI simp:sep_conj_ac)
+    hence " (\<exists>k t. (<Q t> \<and>* <(t##s')> \<and>* I(t + s') \<and>* c \<and>* r) (recse (run (Suc k) s)))" 
+      by (rule h)
+    then obtain k t where h2: "(<Q t> \<and>* <(t ## s')> \<and>* I(t + s') \<and>* c \<and>* r) (recse (run (Suc k) s))"
+      by auto
+    thus "\<exists>k. ((\<lambda>s. \<exists>sa. (<Q sa> \<and>* <(sa ## s')> \<and>* I (sa + s')) s) \<and>* c \<and>* r) (recse (run (Suc k) s))"
+      apply (rule_tac x = k in exI)
+      by (auto intro!:sep_conjI elim!:sep_conjE simp:sep_conj_ac)
+    qed
+  qed
+
+lemma I_frame_rule: 
+  assumes h: "I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace>"
+  shows "I. \<lbrace>P \<and>* R\<rbrace> c \<lbrace>Q \<and>* R\<rbrace>"
+proof(induct rule:IHoareI)
+  case (IPre s' s r cnf)
+  hence "((<(P \<and>* R) s> \<and>* <(s##s')> \<and>* I (s + s')) \<and>* c \<and>* r) (recse cnf)"
+    by (auto simp:sep_conj_ac)
+  then obtain s1 s2 
+  where h1: "((<P s1> \<and>* <((s1 + s2) ## s')> \<and>*I (s1 + s2 + s')) \<and>* c \<and>* r) (recse cnf)" 
+              "s1 ## s2" "s1 + s2 ## s'" "P s1" "R s2"
+    by (unfold pasrt_def, auto elim!:sep_conjE intro!:sep_conjI)
+  hence "((EXS s. <P s> \<and>* <(s ## s2 +s')> \<and>*I (s + (s2 + s'))) \<and>* c \<and>* r) (recse cnf)"
+    apply (sep_cancel, unfold pred_ex_def, auto intro!:sep_conjI sep_disj_addI3 elim!:sep_conjE)
+    apply (rule_tac x = s1 in exI, unfold pasrt_def,
+       auto intro!:sep_conjI sep_disj_addI3 elim!:sep_conjE simp:sep_conj_ac)
+    by (metis sep_add_assoc sep_add_disjD)
+  from h[unfolded IHoare_def Hoare_gen_def, rule_format, OF this]
+  obtain k s where
+     "((<Q s> \<and>* <(s ## s2 + s')> \<and>* I (s + (s2 + s'))) \<and>* c \<and>* r) (recse (run (Suc k) cnf))"
+    by (unfold pasrt_def pred_ex_def, auto elim!:sep_conjE intro!:sep_conjI)
+  thus ?case
+  proof(rule_tac x = k in exI, rule_tac x = "s + s2" in exI, sep_cancel+)
+    fix  h ha
+    assume hh: "(<Q s> \<and>* <(s ## s2 + s')> \<and>* I (s + (s2 + s'))) ha"
+    show " (<(Q \<and>* R) (s + s2)> \<and>* <(s + s2 ## s')> \<and>* I (s + s2 + s')) ha"
+    proof -
+      from hh have h0: "s ## s2 + s'"
+        by (metis pasrt_def sep_conjD)
+      with h1(2, 3)
+      have h2: "s + s2 ## s'" 
+        by (metis sep_add_disjD sep_disj_addI1)
+      moreover from h1(2, 3) h2 have h3: "(s + (s2 + s')) = (s + s2 + s')"
+        by (metis `s ## s2 + s'` sep_add_assoc sep_add_disjD sep_disj_addD1)
+      moreover from hh have "Q s" by (metis pasrt_def sep_conjD)
+      moreover from h0 h1(2) h1(3) have "s ## s2"
+        by (metis sep_add_disjD sep_disj_addD)
+      moreover note h1(5)
+      ultimately show ?thesis
+        by (smt h0 hh sep_conjI)
+    qed
+  qed
+qed
+
+lemma I_sequencing: 
+  assumes h1: "I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace>"
+  and h2: "I. \<lbrace>Q\<rbrace> c \<lbrace>R\<rbrace>"
+  shows "I. \<lbrace>P\<rbrace> c \<lbrace>R\<rbrace>"
+  using h1 h2 sequencing
+  by (smt IHoare_def)
+
+lemma I_pre_stren: 
+  assumes h1: "I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace>"
+  and h2:  "\<And>s. R s \<Longrightarrow> P s"
+  shows "I. \<lbrace>R\<rbrace> c \<lbrace>Q\<rbrace>"
+proof(unfold IHoare_def, default)
+  fix s'
+  show "\<lbrace>EXS s. <R s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>  c 
+       \<lbrace>EXS s. <Q s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>"
+  proof(rule pre_stren)
+    from h1[unfolded IHoare_def, rule_format, of s']
+    show "\<lbrace>EXS s. <P s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>  c 
+          \<lbrace>EXS s. <Q s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>" .
+  next
+    fix s
+    show "(EXS s. <R s> \<and>* <(s ## s')> \<and>* I (s + s')) s \<Longrightarrow> 
+            (EXS s. <P s> \<and>* <(s ## s')> \<and>* I (s + s')) s"
+      apply (unfold pred_ex_def, clarify)
+      apply (rule_tac x = sa in exI, sep_cancel+)
+      by (insert h2, auto simp:pasrt_def)
+  qed
+qed
+
+lemma I_post_weaken: 
+  assumes h1: "I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace>"
+    and h2: "\<And> s. Q s \<Longrightarrow> R s"
+  shows "I. \<lbrace>P\<rbrace> c \<lbrace>R\<rbrace>"
+proof(unfold IHoare_def, default)
+  fix s'
+  show "\<lbrace>EXS s. <P s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>  c 
+        \<lbrace>EXS s. <R s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>"
+  proof(rule post_weaken)
+    from h1[unfolded IHoare_def, rule_format, of s']
+    show "\<lbrace>EXS s. <P s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>  c 
+          \<lbrace>EXS s. <Q s> \<and>* <(s ## s')> \<and>* I (s + s')\<rbrace>" .
+  next
+    fix s
+    show "(EXS s. <Q s> \<and>* <(s ## s')> \<and>* I (s + s')) s \<Longrightarrow> 
+          (EXS s. <R s> \<and>* <(s ## s')> \<and>* I (s + s')) s"
+      apply (unfold pred_ex_def, clarify)
+      apply (rule_tac x = sa in exI, sep_cancel+)
+      by (insert h2, auto simp:pasrt_def)
+  qed
+qed
+
+lemma I_hoare_adjust:
+  assumes h1: "I. \<lbrace>P1\<rbrace> c \<lbrace>Q1\<rbrace>"
+  and h2: "\<And>s. P s \<Longrightarrow> P1 s"
+  and h3: "\<And>s. Q1 s \<Longrightarrow> Q s"
+  shows "I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace>"
+  using h1 h2 h3 I_post_weaken I_pre_stren
+  by (metis)
+
+lemma I_code_exI: 
+  assumes h: "\<And> k. I. \<lbrace>P\<rbrace> c(k) \<lbrace>Q\<rbrace>"
+  shows "I. \<lbrace>P\<rbrace> EXS k. c(k) \<lbrace>Q\<rbrace>"
+using h
+by (smt IHoare_def code_exI)
+
+lemma I_code_extension: 
+  assumes h: "I. \<lbrace> P \<rbrace> c \<lbrace> Q \<rbrace>"
+  shows "I. \<lbrace> P \<rbrace> c ** e \<lbrace> Q \<rbrace>"
+  using h
+  by (smt IHoare_def sep_exec.code_extension)
+
+lemma I_composition: 
+  assumes h1: "I. \<lbrace>P\<rbrace> c1 \<lbrace>Q\<rbrace>"
+    and h2: "I. \<lbrace>Q\<rbrace> c2 \<lbrace>R\<rbrace>"
+  shows "I. \<lbrace>P\<rbrace> c1 ** c2 \<lbrace>R\<rbrace>"
+  using h1 h2
+by (smt IHoare_def sep_exec.composition)
+
+lemma pre_condI: 
+  assumes h: "cond \<Longrightarrow> \<lbrace>p\<rbrace> c \<lbrace>q\<rbrace>" 
+  shows "\<lbrace><cond> \<and>* p\<rbrace> c \<lbrace>q\<rbrace>"
+proof(induct rule:HoareI)
+  case (Pre r s)
+  hence "cond" "(p \<and>* c \<and>* r) (recse s)"
+    apply (metis pasrt_def sep_conjD)
+    by (smt Pre.hyps pasrt_def sep_add_zero sep_conj_commute sep_conj_def)
+  from h[OF this(1), unfolded Hoare_gen_def, rule_format, OF this(2)]
+  show ?case .
+qed
+
+lemma I_pre_condI: 
+  assumes h: "cond \<Longrightarrow> I.\<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace>" 
+  shows "I.\<lbrace><cond> \<and>* P\<rbrace> c \<lbrace>Q\<rbrace>"
+  using h
+by (smt IHoareI condD cond_true_eq2 sep.mult_commute)
+
+lemma code_condI: 
+  assumes h: "b \<Longrightarrow> \<lbrace>p\<rbrace> c \<lbrace>q\<rbrace>"
+  shows "\<lbrace>p\<rbrace> <b>**c \<lbrace>q\<rbrace>"
+proof(induct rule: HoareI)
+  case (Pre r s)
+  hence h1: "b" "(p \<and>* c \<and>* r) (recse s)"
+    apply (metis condD sep_conjD sep_conj_assoc)
+    by (smt Pre.hyps condD sep_conj_impl)
+  from h[OF h1(1), unfolded Hoare_gen_def, rule_format, OF h1(2)]
+  and h1(1)
+  show ?case
+    by (metis (full_types) cond_true_eq1)
+qed
+
+lemma I_code_condI: 
+  assumes h: "b \<Longrightarrow> I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace>"
+  shows "I.\<lbrace>P\<rbrace> <b>**c \<lbrace>Q\<rbrace>"
+  using h
+by (smt IHoareI condD cond_true_eq2 sep.mult_commute sep_conj_cond1)
+
+lemma precond_exI: 
+  assumes h:"\<And>x. \<lbrace>p x\<rbrace> c \<lbrace>q\<rbrace>"
+  shows "\<lbrace>EXS x. p x\<rbrace> c \<lbrace>q\<rbrace>"
+proof(induct rule:HoareI)
+  case (Pre r s)
+  then obtain x where "(p x \<and>* c \<and>* r) (recse s)"
+    by (unfold pred_ex_def, auto elim!:sep_conjE intro!:sep_conjI)
+  from h[of x, unfolded Hoare_gen_def, rule_format, OF this]  
+  show ?case .
+qed
+
+lemma I_precond_exI: 
+  assumes h:"\<And>x. I. \<lbrace>P x\<rbrace> c \<lbrace>Q\<rbrace>"
+  shows "I.\<lbrace>EXS x. P x\<rbrace> c \<lbrace>Q\<rbrace>"
+proof(induct rule:IHoareI)
+  case (IPre s' s r cnf)
+  then obtain x
+    where "((EXS s. <P x s> \<and>* <(s ## s')> \<and>* I (s + s')) \<and>* c \<and>* r) (recse cnf)"
+    by ( auto elim!:sep_conjE intro!:sep_conjI simp:pred_ex_def pasrt_def sep_conj_ac)
+  from h[unfolded IHoare_def Hoare_gen_def, rule_format, OF this]
+  obtain k t 
+    where "((<Q t> \<and>* <(t ## s')> \<and>* I (t + s')) \<and>* c \<and>* r) (recse (run (Suc k) cnf))"
+    by (unfold pred_ex_def, auto elim!:sep_conjE intro!:sep_conjI)
+  thus ?case 
+    by (auto simp:sep_conj_ac)
+qed
+
+lemma hoare_sep_false: 
+     "\<lbrace>sep_false\<rbrace> c
+      \<lbrace>q\<rbrace>" 
+  by(unfold Hoare_gen_def, clarify, simp)
+
+lemma I_hoare_sep_false:
+  "I. \<lbrace>sep_false\<rbrace> c
+      \<lbrace>Q\<rbrace>"
+by (smt IHoareI condD)
+
+end
+
+instantiation set :: (type)sep_algebra
+begin
+
+definition set_zero_def: "0 = {}"
+
+definition plus_set_def: "s1 + s2 = s1 \<union> s2"
+
+definition sep_disj_set_def: "sep_disj s1 s2 = (s1 \<inter> s2 = {})"
+
+lemmas set_ins_def = sep_disj_set_def plus_set_def set_zero_def
+
+instance
+  apply(default)
+  apply(simp add:set_ins_def)
+  apply(simp add:sep_disj_set_def plus_set_def set_zero_def)
+  apply (metis inf_commute)
+  apply(simp add:sep_disj_set_def plus_set_def set_zero_def)
+  apply(simp add:sep_disj_set_def plus_set_def set_zero_def)
+  apply (metis sup_commute)
+  apply(simp add:sep_disj_set_def plus_set_def set_zero_def)
+  apply (metis (lifting) Un_assoc)
+  apply(simp add:sep_disj_set_def plus_set_def set_zero_def)
+  apply (metis (lifting) Int_Un_distrib Un_empty inf_sup_distrib1 sup_eq_bot_iff)
+  apply(simp add:sep_disj_set_def plus_set_def set_zero_def)
+  by (metis (lifting) Int_Un_distrib Int_Un_distrib2 sup_eq_bot_iff)
+end
+
+section {* A big operator of infinite separation conjunction *}
+
+definition "fam_conj I cpt s = (\<exists> p. (s = (\<Union> i \<in> I. p(i))) \<and>
+                                     (\<forall> i \<in> I. cpt i (p i)) \<and>
+                                     (\<forall> i \<in> I. \<forall> j \<in> I. i \<noteq> j \<longrightarrow> p(i) ## p(j)))"
+
+lemma fam_conj_zero_simp: "fam_conj {} cpt = <True>"
+proof
+  fix s
+  show "fam_conj {} cpt s = (<True>) s"
+  proof
+    assume "fam_conj {} cpt s"
+    then obtain p where "s = (\<Union> i \<in> {}. p(i))"
+      by (unfold fam_conj_def, auto)
+    hence "s = {}" by auto
+    thus "(<True>) s" by (metis pasrt_def set_zero_def) 
+  next
+    assume "(<True>) s"
+    hence eq_s: "s = {}" by (metis pasrt_def set_zero_def)
+    let ?p = "\<lambda> i. {}"
+    have "(s = (\<Union> i \<in> {}. ?p(i)))" by (unfold eq_s, auto)
+    moreover have "(\<forall> i \<in> {}. cpt i (?p i))" by auto
+    moreover have "(\<forall> i \<in> {}. \<forall> j \<in> {}. i \<noteq> j \<longrightarrow> ?p(i) ## ?p(j))" by auto
+    ultimately show "fam_conj {} cpt s"
+      by (unfold eq_s fam_conj_def, auto)
+  qed
+qed
+
+lemma fam_conj_disj_simp_pre:
+  assumes h1: "I = I1 + I2"
+  and h2: "I1 ## I2"
+  shows "fam_conj I cpt = (fam_conj I1 cpt \<and>* fam_conj I2 cpt)"
+proof
+  fix s
+  let ?fm = "fam_conj I cpt" and ?fm1 = "fam_conj I1 cpt" and ?fm2 = "fam_conj I2 cpt"
+  show "?fm s = (?fm1 \<and>* ?fm2) s"
+  proof
+    assume "?fm s"
+    then obtain p where pre:
+            "s = (\<Union> i \<in> I. p(i))"
+            "(\<forall> i \<in> I. cpt i (p i))"
+            "(\<forall> i \<in> I. \<forall> j \<in> I. i \<noteq> j \<longrightarrow> p(i) ## p(j))"
+      unfolding fam_conj_def by metis
+    from pre(1) h1 h2 have "s = (\<Union> i \<in> I1. p(i)) + (\<Union> i \<in> I2. p(i))"
+      by (auto simp:set_ins_def)
+    moreover from pre h1 have "?fm1 (\<Union> i \<in> I1. p(i))"
+      by (unfold fam_conj_def, rule_tac x = p in exI, auto simp:set_ins_def)
+    moreover from pre h1 have "?fm2 (\<Union> i \<in> I2. p(i))"
+      by (unfold fam_conj_def, rule_tac x = p in exI, auto simp:set_ins_def)
+    moreover have "(\<Union> i \<in> I1. p(i)) ## (\<Union> i \<in> I2. p(i))"
+    proof -
+      { fix x xa xb
+        assume h: "I1 \<inter> I2 = {}"
+                  "\<forall>i\<in>I1 \<union> I2. \<forall>j\<in>I1 \<union> I2. i \<noteq> j \<longrightarrow> p i \<inter> p j = {}"
+                  "xa \<in> I1" "x \<in> p xa" "xb \<in> I2" "x \<in> p xb"
+        have "p xa \<inter> p xb = {}"
+        proof(rule h(2)[rule_format])
+          from h(1, 3, 5) show "xa \<in> I1 \<union> I2" by auto
+        next
+          from h(1, 3, 5) show "xb \<in> I1 \<union> I2" by auto
+        next
+          from h(1, 3, 5) show "xa \<noteq> xb" by auto
+        qed with h(4, 6) have False by auto
+      } with h1 h2 pre(3) show ?thesis by (auto simp:set_ins_def) 
+    qed
+    ultimately show "(?fm1 \<and>* ?fm2) s" by (auto intro!:sep_conjI)
+  next
+    assume "(?fm1 \<and>* ?fm2) s"
+    then obtain s1 s2 where pre:
+      "s = s1 + s2" "s1 ## s2" "?fm1 s1" "?fm2 s2"
+      by (auto dest!:sep_conjD)
+    from pre(3) obtain p1 where pre1:
+            "s1 = (\<Union> i \<in> I1. p1(i))"
+            "(\<forall> i \<in> I1. cpt i (p1 i))"
+            "(\<forall> i \<in> I1. \<forall> j \<in> I1. i \<noteq> j \<longrightarrow> p1(i) ## p1(j))"
+       unfolding fam_conj_def by metis
+    from pre(4) obtain p2 where pre2:
+            "s2 = (\<Union> i \<in> I2. p2(i))"
+            "(\<forall> i \<in> I2. cpt i (p2 i))"
+            "(\<forall> i \<in> I2. \<forall> j \<in> I2. i \<noteq> j \<longrightarrow> p2(i) ## p2(j))"
+       unfolding fam_conj_def by metis
+     let ?p = "\<lambda> i. if i \<in> I1 then p1 i else p2 i"
+     from h2 pre(2)
+     have "s = (\<Union> i \<in> I. ?p(i))" 
+       apply (unfold h1 pre(1) pre1(1) pre2(1) set_ins_def, auto split:if_splits)
+       by (metis disjoint_iff_not_equal)
+     moreover from h2 pre1(2) pre2(2) have "(\<forall> i \<in> I. cpt i (?p i))" 
+       by (unfold h1 set_ins_def, auto split:if_splits)
+     moreover from pre1(1, 3) pre2(1, 3) pre(2) h2
+     have "(\<forall> i \<in> I. \<forall> j \<in> I. i \<noteq> j \<longrightarrow> ?p(i) ## ?p(j))" 
+       apply (unfold h1 set_ins_def, auto split:if_splits)
+       by (metis IntI empty_iff)
+     ultimately show "?fm s" by (unfold fam_conj_def, auto)
+  qed
+qed
+
+lemma fam_conj_disj_simp:
+  assumes h: "I1 ## I2"
+  shows "fam_conj (I1 + I2) cpt = (fam_conj I1 cpt \<and>* fam_conj I2 cpt)"
+proof -
+  from fam_conj_disj_simp_pre[OF _ h, of "I1 + I2"]
+  show ?thesis by simp
+qed
+
+lemma fam_conj_elm_simp:
+  assumes h: "i \<in> I"
+  shows "fam_conj I cpt = (cpt(i) \<and>* fam_conj (I - {i}) cpt)"
+proof
+  fix s
+  show "fam_conj I cpt s = (cpt i \<and>* fam_conj (I - {i}) cpt) s"
+  proof
+    assume pre: "fam_conj I cpt s"
+    show "(cpt i \<and>* fam_conj (I - {i}) cpt) s"
+    proof -
+      from pre obtain p where pres:
+            "s = (\<Union> i \<in> I. p(i))"
+            "(\<forall> i \<in> I. cpt i (p i))"
+            "(\<forall> i \<in> I. \<forall> j \<in> I. i \<noteq> j \<longrightarrow> p(i) ## p(j))"
+        unfolding fam_conj_def by metis
+      let ?s = "(\<Union>i\<in>(I - {i}). p i)"
+      from pres(3) h have disj: "(p i) ## ?s"
+        by (auto simp:set_ins_def)
+      moreover from pres(1) h have eq_s: "s = (p i) +  ?s"
+        by (auto simp:set_ins_def)
+      moreover from pres(2) h have "cpt i (p i)" by auto
+      moreover from pres have "(fam_conj (I - {i}) cpt) ?s"
+        by (unfold fam_conj_def, rule_tac x = p in exI, auto)
+      ultimately show ?thesis by (auto intro!:sep_conjI)
+    qed
+  next
+    let ?fam = "fam_conj (I - {i}) cpt"
+    assume "(cpt i \<and>* ?fam) s"
+    then obtain s1 s2 where pre:
+      "s = s1 + s2" "s1 ## s2" "cpt i s1" "?fam s2"
+      by (auto dest!:sep_conjD)
+    from pre(4) obtain p where pres:
+            "s2 = (\<Union> ia \<in> I - {i}. p(ia))"
+            "(\<forall> ia \<in> I - {i}. cpt ia (p ia))"
+            "(\<forall> ia \<in> I - {i}. \<forall> j \<in> I - {i}. ia \<noteq> j \<longrightarrow> p(ia) ## p(j))"
+      unfolding fam_conj_def by metis
+    let ?p = "p(i:=s1)"
+    from h pres(3) pre(2)[unfolded pres(1)] 
+    have " \<forall>ia\<in>I. \<forall>j\<in>I. ia \<noteq> j \<longrightarrow> ?p ia ## ?p j" by  (auto simp:set_ins_def)
+    moreover from pres(1) pre(1) h pre(2)
+    have "(s = (\<Union> i \<in> I. ?p(i)))" by (auto simp:set_ins_def split:if_splits)
+    moreover from pre(3) pres(2) h
+    have "(\<forall> i \<in> I. cpt i (?p i))" by (auto simp:set_ins_def split:if_splits)
+    ultimately show "fam_conj I cpt s"
+      by (unfold fam_conj_def, auto)
+  qed
+qed
+
+lemma fam_conj_insert_simp:
+  assumes h:"i \<notin> I"
+  shows "fam_conj (insert i I) cpt = (cpt(i) \<and>* fam_conj I cpt)"
+proof -
+  have "i \<in> insert i I" by auto
+  from fam_conj_elm_simp[OF this] and h
+  show ?thesis by auto
+qed
+
+lemmas fam_conj_simps = fam_conj_insert_simp fam_conj_zero_simp
+
+lemma "fam_conj {0, 2, 3} cpt = (cpt 2 \<and>* cpt (0::int) \<and>* cpt 3)"
+  by (simp add:fam_conj_simps sep_conj_ac)
+
+lemma fam_conj_ext_eq:
+  assumes h: "\<And> i . i \<in> I \<Longrightarrow> f i = g i"
+  shows "fam_conj I f = fam_conj I g"
+proof
+  fix s
+  show "fam_conj I f s = fam_conj I g s"
+   by (unfold fam_conj_def, auto simp:h)
+qed
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/Hoare_tm.thy	Thu Mar 06 13:28:38 2014 +0000
@@ -0,0 +1,5718 @@
+header {* 
+  Separation logic for TM
+*}
+
+theory Hoare_tm
+imports Hoare_gen My_block Data_slot
+begin
+
+
+ML {*
+fun pretty_terms ctxt trms =
+  Pretty.block (Pretty.commas (map (Syntax.pretty_term ctxt) trms))
+*}
+
+lemma int_add_C :"x + (y::int) = y + x"
+  by simp
+
+lemma int_add_A : "(x + y) + z = x + (y + (z::int))"
+  by simp
+
+lemma int_add_LC: "x + (y + (z::int)) = y + (x + z)"
+  by simp
+
+lemmas int_add_ac = int_add_A int_add_C int_add_LC
+
+method_setup prune = {* Scan.succeed (SIMPLE_METHOD' o (K (K prune_params_tac))) *} 
+                       "pruning parameters"
+
+ML {*
+structure StepRules = Named_Thms
+  (val name = @{binding "step"}
+   val description = "Theorems for hoare rules for every step")
+*}
+
+ML {*
+structure FwdRules = Named_Thms
+  (val name = @{binding "fwd"}
+   val description = "Theorems for fwd derivation of seperation implication")
+*}
+
+setup {* StepRules.setup *}
+
+setup {* FwdRules.setup *}
+
+section {* Operational Semantics of TM *}
+
+datatype taction = W0 | W1 | L | R
+
+type_synonym tstate = nat
+
+type_synonym tm_inst = "(taction \<times> tstate) \<times> (taction \<times> tstate)"
+
+datatype Block = Oc | Bk
+
+type_synonym tconf = "nat \<times> (nat \<rightharpoonup> tm_inst) \<times> nat \<times> int \<times> (int \<rightharpoonup> Block)"
+
+fun next_tape :: "taction \<Rightarrow> (int \<times>  (int \<rightharpoonup> Block)) \<Rightarrow> (int \<times>  (int \<rightharpoonup> Block))"
+where "next_tape W0 (pos, m) = (pos, m(pos \<mapsto> Bk))" |
+      "next_tape W1 (pos, m) = (pos, m(pos \<mapsto> Oc))" |
+      "next_tape L  (pos, m) = (pos - 1, m)" |
+      "next_tape R  (pos, m) = (pos + 1, m)"
+
+fun tstep :: "tconf \<Rightarrow> tconf"
+  where "tstep (faults, prog, pc, pos, m) = 
+              (case (prog pc) of
+                  Some ((action0, pc0), (action1, pc1)) \<Rightarrow> 
+                     case (m pos) of
+                       Some Bk \<Rightarrow> (faults, prog, pc0, next_tape action0 (pos, m)) |
+                       Some Oc \<Rightarrow> (faults, prog, pc1, next_tape action1 (pos, m)) |
+                       None \<Rightarrow> (faults + 1, prog, pc, pos, m)
+                | None \<Rightarrow> (faults + 1, prog, pc, pos, m))"
+
+lemma tstep_splits: 
+  "(P (tstep s)) = ((\<forall> faults prog pc pos m action0 pc0 action1 pc1. 
+                          s = (faults, prog, pc, pos, m) \<longrightarrow> 
+                          prog pc = Some ((action0, pc0), (action1, pc1)) \<longrightarrow> 
+                          m pos = Some Bk \<longrightarrow> P (faults, prog, pc0, next_tape action0 (pos, m))) \<and>
+                    (\<forall> faults prog pc pos m action0 pc0 action1 pc1. 
+                          s = (faults, prog, pc, pos, m) \<longrightarrow> 
+                          prog pc = Some ((action0, pc0), (action1, pc1)) \<longrightarrow> 
+                          m pos = Some Oc \<longrightarrow> P (faults, prog, pc1, next_tape action1 (pos, m))) \<and>
+                    (\<forall> faults prog pc pos m action0 pc0 action1 pc1. 
+                          s = (faults, prog, pc, pos, m) \<longrightarrow> 
+                          prog pc = Some ((action0, pc0), (action1, pc1)) \<longrightarrow> 
+                          m pos = None \<longrightarrow> P (faults + 1, prog, pc, pos, m)) \<and>
+                    (\<forall> faults prog pc pos m . 
+                          s =  (faults, prog, pc, pos, m) \<longrightarrow>
+                          prog pc  = None \<longrightarrow> P (faults + 1, prog, pc, pos, m))
+                   )"
+  by (case_tac s, auto split:option.splits Block.splits)
+
+datatype tresource = 
+    TM int Block
+  | TC nat tm_inst
+  | TAt nat
+  | TPos int
+  | TFaults nat
+
+definition "tprog_set prog = {TC i inst | i inst. prog i = Some inst}"
+definition "tpc_set pc = {TAt pc}"
+definition "tmem_set m = {TM i n | i n. m (i) = Some n}"
+definition "tpos_set pos = {TPos pos}"
+definition "tfaults_set faults = {TFaults faults}"
+
+lemmas tpn_set_def = tprog_set_def tpc_set_def tmem_set_def tfaults_set_def tpos_set_def
+
+fun trset_of :: "tconf \<Rightarrow> tresource set"
+  where "trset_of (faults, prog, pc, pos, m) = 
+               tprog_set prog \<union> tpc_set pc \<union> tpos_set pos \<union> tmem_set m \<union> tfaults_set faults"
+
+interpretation tm: sep_exec tstep trset_of .
+
+section {* Assembly language and its program logic *}
+
+subsection {* The assembly language *}
+
+datatype tpg = 
+   TInstr tm_inst
+ | TLabel nat
+ | TSeq tpg tpg
+ | TLocal "(nat \<Rightarrow> tpg)"
+
+notation TLocal (binder "TL " 10)
+
+abbreviation tprog_instr :: "tm_inst \<Rightarrow> tpg" ("\<guillemotright> _" [61] 61)
+where "\<guillemotright> i \<equiv> TInstr i"
+
+abbreviation tprog_seq :: "tpg \<Rightarrow> tpg \<Rightarrow> tpg" (infixr ";" 52)
+where "c1 ; c2 \<equiv> TSeq c1 c2"
+
+definition "sg e = (\<lambda> s. s = e)"
+
+type_synonym tassert = "tresource set \<Rightarrow> bool"
+
+abbreviation t_hoare :: 
+  "tassert \<Rightarrow> tassert  \<Rightarrow> tassert \<Rightarrow> bool" ("(\<lbrace>(1_)\<rbrace> / (_)/ \<lbrace>(1_)\<rbrace>)" 50)
+  where "\<lbrace> p \<rbrace> c \<lbrace> q \<rbrace> == sep_exec.Hoare_gen tstep trset_of p c q"
+
+abbreviation it_hoare ::
+  "(('a::sep_algebra) \<Rightarrow> tresource set \<Rightarrow> bool) \<Rightarrow> 
+      ('a \<Rightarrow> bool) \<Rightarrow> (tresource set \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
+  ("(1_).(\<lbrace>(1_)\<rbrace> / (_)/ \<lbrace>(1_)\<rbrace>)" 50)
+where "I. \<lbrace>P\<rbrace> c \<lbrace>Q\<rbrace> == sep_exec.IHoare tstep trset_of I P c Q"
+
+(*
+primrec tpg_len :: "tpg \<Rightarrow> nat" where 
+  "tpg_len (TInstr ai) = 1" |
+  "tpg_len (TSeq p1 p2) = tpg_len p1 + tpg_len " |
+  "tpg_len (TLocal fp) = tpg_len (fp 0)" |
+  "tpg_len (TLabel l) = 0" *)
+
+primrec tassemble_to :: "tpg \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> tassert" 
+  where 
+  "tassemble_to (TInstr ai) i j = (sg ({TC i ai}) ** <(j = i + 1)>)" |
+  "tassemble_to (TSeq p1 p2) i j = (EXS j'. (tassemble_to p1 i j') ** (tassemble_to p2 j' j))" |
+  "tassemble_to (TLocal fp) i j  = (EXS l. (tassemble_to (fp l) i j))" | 
+  "tassemble_to (TLabel l) i j = <(i = j \<and> j = l)>"
+
+declare tassemble_to.simps [simp del]
+
+lemmas tasmp = tassemble_to.simps (2, 3, 4)
+
+abbreviation tasmb_to :: "nat \<Rightarrow> tpg \<Rightarrow> nat \<Rightarrow> tassert" ("_ :[ _ ]: _" [60, 60, 60] 60)
+  where "i :[ tpg ]: j \<equiv> tassemble_to tpg i j"
+
+lemma EXS_intro: 
+  assumes h: "(P x) s"
+  shows "((EXS x. P(x))) s"
+  by (smt h pred_ex_def sep_conj_impl)
+
+lemma EXS_elim: 
+  assumes "(EXS x. P x) s"
+  obtains x where "P x s"
+  by (metis assms pred_ex_def)
+
+lemma EXS_eq:
+  fixes x
+  assumes h: "Q (p x)" 
+  and h1: "\<And> y s. \<lbrakk>p y s\<rbrakk> \<Longrightarrow> y = x"
+  shows "p x = (EXS x. p x)"
+proof
+  fix s
+  show "p x s = (EXS x. p x) s"
+  proof
+    assume "p x s"
+    thus "(EXS x. p x) s" by (auto simp:pred_ex_def)
+  next
+    assume "(EXS x. p x) s"
+    thus "p x s"
+    proof(rule EXS_elim)
+      fix y
+      assume "p y s"
+      from this[unfolded h1[OF this]] show "p x s" .
+    qed
+  qed
+qed
+
+definition "tpg_fold tpgs = foldr TSeq (butlast tpgs) (last tpgs)"
+
+lemma tpg_fold_sg: "tpg_fold [tpg] = tpg"
+  by (simp add:tpg_fold_def)
+
+lemma tpg_fold_cons: 
+  assumes h: "tpgs \<noteq> []"
+  shows "tpg_fold (tpg#tpgs) = (tpg; (tpg_fold tpgs))"
+  using h
+proof(induct tpgs arbitrary:tpg)
+  case (Cons tpg1 tpgs1)
+  thus ?case
+  proof(cases "tpgs1 = []")
+    case True
+    thus ?thesis by (simp add:tpg_fold_def)
+  next
+    case False
+    show ?thesis
+    proof -
+      have eq_1: "butlast (tpg # tpg1 # tpgs1) = tpg # (butlast (tpg1 # tpgs1))"
+        by simp
+      from False have eq_2: "last (tpg # tpg1 # tpgs1) = last (tpg1 # tpgs1)"
+        by simp
+      have eq_3: "foldr (op ;) (tpg # butlast (tpg1 # tpgs1)) (last (tpg1 # tpgs1)) = 
+            (tpg ; (foldr (op ;) (butlast (tpg1 # tpgs1)) (last (tpg1 # tpgs1))))"
+        by simp
+      show ?thesis
+        apply (subst (1) tpg_fold_def, unfold eq_1 eq_2 eq_3)
+        by (fold tpg_fold_def, simp)
+    qed
+  qed
+qed auto
+
+lemmas tpg_fold_simps = tpg_fold_sg tpg_fold_cons
+
+lemma tpg_fold_app:
+  assumes h1: "tpgs1 \<noteq> []" 
+  and h2: "tpgs2 \<noteq> []"
+  shows "i:[(tpg_fold (tpgs1 @ tpgs2))]:j = i:[(tpg_fold (tpgs1); tpg_fold tpgs2)]:j"
+  using h1 h2
+proof(induct tpgs1 arbitrary: i j tpgs2)
+  case (Cons tpg1' tpgs1')
+  thus ?case (is "?L = ?R")
+  proof(cases "tpgs1' = []")
+    case False
+    from h2 have "(tpgs1' @ tpgs2) \<noteq> []"
+      by (metis Cons.prems(2) Nil_is_append_conv) 
+    have "?L = (i :[ tpg_fold (tpg1' # (tpgs1' @ tpgs2)) ]: j )" by simp
+    also have "\<dots> =  (i:[(tpg1'; (tpg_fold (tpgs1' @ tpgs2)))]:j )"
+      by (simp add:tpg_fold_cons[OF `(tpgs1' @ tpgs2) \<noteq> []`])
+    also have "\<dots> = ?R"
+    proof -
+      have "(EXS j'. i :[ tpg1' ]: j' \<and>* j' :[ tpg_fold (tpgs1' @ tpgs2) ]: j) =
+              (EXS j'. (EXS j'a. i :[ tpg1' ]: j'a \<and>* j'a :[ tpg_fold tpgs1' ]: j') \<and>* 
+                               j' :[ tpg_fold tpgs2 ]: j)"
+      proof(default+)
+        fix s
+        assume "(EXS j'. i :[ tpg1' ]: j' \<and>* j' :[ tpg_fold (tpgs1' @ tpgs2) ]: j) s"
+        thus "(EXS j'. (EXS j'a. i :[ tpg1' ]: j'a \<and>* j'a :[ tpg_fold tpgs1' ]: j') \<and>*
+                  j' :[ tpg_fold tpgs2 ]: j) s"
+        proof(elim EXS_elim)
+          fix j'
+          assume "(i :[ tpg1' ]: j' \<and>* j' :[ tpg_fold (tpgs1' @ tpgs2) ]: j) s"
+          from this[unfolded Cons(1)[OF False Cons(3)] tassemble_to.simps]
+          show "(EXS j'. (EXS j'a. i :[ tpg1' ]: j'a \<and>* j'a :[ tpg_fold tpgs1' ]: j') \<and>*
+                           j' :[ tpg_fold tpgs2 ]: j) s"
+          proof(elim sep_conjE EXS_elim)
+            fix x y j'a xa ya
+            assume h: "(i :[ tpg1' ]: j') x"
+                      "x ## y" "s = x + y"
+                      "(j' :[ tpg_fold tpgs1' ]: j'a) xa"
+                      "(j'a :[ tpg_fold tpgs2 ]: j) ya"
+                      " xa ## ya" "y = xa + ya"
+            show "(EXS j'. (EXS j'a. i :[ tpg1' ]: j'a \<and>* 
+                                j'a :[ tpg_fold tpgs1' ]: j') \<and>* j' :[ tpg_fold tpgs2 ]: j) s"
+               (is "(EXS j'. (?P j' \<and>* ?Q j')) s")
+            proof(rule EXS_intro[where x = "j'a"])
+              from `(j'a :[ tpg_fold tpgs2 ]: j) ya` have "(?Q j'a) ya" .
+              moreover have "(?P j'a) (x + xa)" 
+              proof(rule EXS_intro[where x = j'])
+                have "x + xa = x + xa" by simp
+                moreover from `x ## y` `y = xa + ya` `xa ## ya` 
+                have "x ## xa" by (metis sep_disj_addD) 
+                moreover note `(i :[ tpg1' ]: j') x` `(j' :[ tpg_fold tpgs1' ]: j'a) xa`
+                ultimately show "(i :[ tpg1' ]: j' \<and>* j' :[ tpg_fold tpgs1' ]: j'a) (x + xa)"
+                  by (auto intro!:sep_conjI)
+              qed
+              moreover from `x##y` `y = xa + ya` `xa ## ya` 
+              have "(x + xa) ## ya"
+                by (metis sep_disj_addI1 sep_disj_commuteI)
+              moreover from `s = x + y` `y = xa + ya`
+              have "s = (x + xa) + ya"
+                by (metis h(2) h(6) sep_add_assoc sep_disj_addD1 sep_disj_addD2) 
+              ultimately show "(?P j'a \<and>* ?Q j'a) s"
+                by (auto intro!:sep_conjI)
+            qed
+          qed
+        qed
+      next
+        fix s
+        assume "(EXS j'. (EXS j'a. i :[ tpg1' ]: j'a \<and>* j'a :[ tpg_fold tpgs1' ]: j') \<and>*
+                                    j' :[ tpg_fold tpgs2 ]: j) s"
+        thus "(EXS j'. i :[ tpg1' ]: j' \<and>* j' :[ tpg_fold (tpgs1' @ tpgs2) ]: j) s"
+        proof(elim EXS_elim sep_conjE)
+          fix j' x y j'a xa ya
+          assume h: "(j' :[ tpg_fold tpgs2 ]: j) y"
+                    "x ## y" "s = x + y" "(i :[ tpg1' ]: j'a) xa"
+                    "(j'a :[ tpg_fold tpgs1' ]: j') ya" "xa ## ya" "x = xa + ya"
+          show "(EXS j'. i :[ tpg1' ]: j' \<and>* j' :[ tpg_fold (tpgs1' @ tpgs2) ]: j) s"
+          proof(rule EXS_intro[where x = j'a])
+            from `(i :[ tpg1' ]: j'a) xa` have "(i :[ tpg1' ]: j'a) xa" .
+            moreover have "(j'a :[ tpg_fold (tpgs1' @ tpgs2) ]: j) (ya + y)"
+            proof(unfold Cons(1)[OF False Cons(3)] tassemble_to.simps)
+              show "(EXS j'. j'a :[ tpg_fold tpgs1' ]: j' \<and>* j' :[ tpg_fold tpgs2 ]: j) (ya + y)"
+              proof(rule EXS_intro[where x = "j'"])
+                from `x ## y` `x = xa + ya` `xa ## ya`
+                have "ya ## y" by (metis sep_add_disjD)
+                moreover have "ya + y = ya + y" by simp
+                moreover note `(j'a :[ tpg_fold tpgs1' ]: j') ya` 
+                               `(j' :[ tpg_fold tpgs2 ]: j) y`
+                ultimately show "(j'a :[ tpg_fold tpgs1' ]: j' \<and>* 
+                                 j' :[ tpg_fold tpgs2 ]: j) (ya + y)"
+                  by (auto intro!:sep_conjI)
+              qed
+            qed
+            moreover from `s = x + y` `x = xa + ya`
+            have "s = xa + (ya + y)"
+              by (metis h(2) h(6) sep_add_assoc sep_add_disjD)
+            moreover from `xa ## ya` `x ## y` `x = xa + ya`
+            have "xa ## (ya + y)" by (metis sep_disj_addI3) 
+            ultimately show "(i :[ tpg1' ]: j'a \<and>* j'a :[ tpg_fold (tpgs1' @ tpgs2) ]: j) s"
+              by (auto intro!:sep_conjI)
+          qed
+        qed
+      qed
+      thus ?thesis 
+        by (simp add:tassemble_to.simps, unfold tpg_fold_cons[OF False], 
+             unfold tassemble_to.simps, simp)
+    qed
+    finally show ?thesis . 
+  next
+    case True
+    thus ?thesis
+      by (simp add:tpg_fold_cons[OF Cons(3)] tpg_fold_sg)
+  qed 
+qed auto
+ 
+
+subsection {* Assertions and program logic for this assembly language *}
+
+definition "st l = sg (tpc_set l)"
+definition "ps p = sg (tpos_set p)" 
+definition "tm a v =sg ({TM a v})"
+definition "one i = tm i Oc"
+definition "zero i= tm i Bk"
+definition "any i = (EXS v. tm i v)"
+
+declare trset_of.simps[simp del]
+
+lemma stimes_sgD: "(sg x ** q) s \<Longrightarrow> q (s - x) \<and> x \<subseteq> s"
+  apply(erule_tac sep_conjE)
+  apply(unfold set_ins_def sg_def)
+  by (metis Diff_Int2 Diff_Int_distrib2 Diff_Un Diff_cancel 
+    Diff_empty Diff_idemp Diff_triv Int_Diff Un_Diff 
+    Un_Diff_cancel inf_commute inf_idem sup_bot_right sup_commute sup_ge2)
+
+lemma stD: "(st i ** r) (trset_of (ft, prog, i', pos, mem))
+       \<Longrightarrow> i' = i"
+proof -
+  assume "(st i ** r) (trset_of (ft, prog, i', pos, mem))"
+  from stimes_sgD [OF this[unfolded st_def], unfolded trset_of.simps]
+  have "tpc_set i \<subseteq> tprog_set prog \<union> tpc_set i' \<union> tpos_set pos \<union>  
+            tmem_set mem \<union> tfaults_set ft" by auto
+  thus ?thesis
+    by (unfold tpn_set_def, auto)
+qed
+
+lemma psD: "(ps p ** r) (trset_of (ft, prog, i', pos, mem))
+       \<Longrightarrow> pos = p"
+proof -
+  assume "(ps p ** r) (trset_of (ft, prog, i', pos, mem))"
+  from stimes_sgD [OF this[unfolded ps_def], unfolded trset_of.simps]
+  have "tpos_set p \<subseteq> tprog_set prog \<union> tpc_set i' \<union> 
+                       tpos_set pos \<union> tmem_set mem \<union> tfaults_set ft"
+    by simp
+  thus ?thesis
+    by (unfold tpn_set_def, auto)
+qed
+
+lemma codeD: "(st i ** sg {TC i inst} ** r) (trset_of (ft, prog, i', pos, mem))
+       \<Longrightarrow> prog i = Some inst"
+proof -
+  assume "(st i ** sg {TC i inst} ** r) (trset_of (ft, prog, i', pos, mem))"
+  thus ?thesis
+    apply(unfold sep_conj_def set_ins_def sg_def trset_of.simps tpn_set_def)
+    by auto
+qed
+
+lemma memD: "((tm a v) ** r) (trset_of (ft, prog, i, pos, mem))  \<Longrightarrow> mem a = Some v"
+proof -
+  assume "((tm a v) ** r) (trset_of (ft, prog, i, pos, mem))"
+  from stimes_sgD[OF this[unfolded trset_of.simps tpn_set_def tm_def]]
+  have "{TM a v} \<subseteq> {TC i inst |i inst. prog i = Some inst} \<union> {TAt i} \<union> 
+    {TPos pos} \<union> {TM i n |i n. mem i = Some n} \<union> {TFaults ft}" by simp
+  thus ?thesis by auto
+qed
+
+lemma t_hoare_seq: 
+  "\<lbrakk>\<And> i j. \<lbrace>st i ** p\<rbrace> i:[c1]:j \<lbrace>st j ** q\<rbrace>; 
+    \<And> j k. \<lbrace>st j ** q\<rbrace> j:[c2]:k \<lbrace>st k ** r\<rbrace>\<rbrakk> \<Longrightarrow>  \<lbrace>st i ** p\<rbrace> i:[(c1 ; c2)]:k \<lbrace>st k ** r\<rbrace>"
+proof -
+  assume h: "\<And> i j. \<lbrace>st i ** p\<rbrace> i:[c1]:j \<lbrace>st j ** q\<rbrace>" 
+            "\<And> j k. \<lbrace>st j ** q\<rbrace> j:[c2]:k \<lbrace>st k ** r\<rbrace>"
+  show "\<lbrace>st i ** p\<rbrace> i:[(c1 ; c2)]:k \<lbrace>st k ** r\<rbrace>"
+  proof(subst tassemble_to.simps, rule tm.code_exI)
+    fix j'
+    show "\<lbrace>st i \<and>* p\<rbrace>  i :[ c1 ]: j' \<and>* j' :[ c2 ]: k \<lbrace>st k \<and>* r\<rbrace>"
+    proof(rule tm.composition)
+      from h(1) show "\<lbrace>st i \<and>* p\<rbrace>  i :[ c1 ]: j' \<lbrace> st j' \<and>* q \<rbrace>" by auto
+    next
+      from h(2) show "\<lbrace>st j' \<and>* q\<rbrace>  j' :[ c2 ]: k \<lbrace>st k \<and>* r\<rbrace>" by auto
+    qed
+  qed
+qed
+
+lemma t_hoare_seq1:
+   "\<lbrakk>\<And>j'. \<lbrace>st i ** p\<rbrace> i:[c1]:j' \<lbrace>st j' ** q\<rbrace>;
+    \<And>j'. \<lbrace>st j' ** q\<rbrace> j':[c2]:k \<lbrace>st k' ** r\<rbrace>\<rbrakk> \<Longrightarrow>  
+           \<lbrace>st i ** p\<rbrace> i:[(c1 ; c2)]:k \<lbrace>st k' ** r\<rbrace>"
+proof -
+  assume h: "\<And>j'. \<lbrace>st i \<and>* p\<rbrace>  i :[ c1 ]: j' \<lbrace>st j' \<and>* q\<rbrace>" 
+            "\<And>j'. \<lbrace>st j' \<and>* q\<rbrace>  j' :[ c2 ]: k \<lbrace>st k' \<and>* r\<rbrace>"
+  show "\<lbrace>st i \<and>* p\<rbrace>  i :[ (c1 ; c2) ]: k \<lbrace>st k' \<and>* r\<rbrace>"
+  proof(subst tassemble_to.simps, rule tm.code_exI)
+    fix j'
+    show "\<lbrace>st i \<and>* p\<rbrace>  i :[ c1 ]: j' \<and>* j' :[ c2 ]: k \<lbrace>st k' \<and>* r\<rbrace>"
+    proof(rule tm.composition)
+      from h(1) show "\<lbrace>st i \<and>* p\<rbrace>  i :[ c1 ]: j' \<lbrace> st j' \<and>* q \<rbrace>" by auto
+    next
+      from h(2) show " \<lbrace>st j' \<and>* q\<rbrace>  j' :[ c2 ]: k \<lbrace>st k' \<and>* r\<rbrace>" by auto
+    qed
+  qed
+qed
+
+lemma t_hoare_seq2:
+ assumes h: "\<And>j. \<lbrace>st i ** p\<rbrace> i:[c1]:j \<lbrace>st k' \<and>* r\<rbrace>"
+ shows "\<lbrace>st i ** p\<rbrace> i:[(c1 ; c2)]:j \<lbrace>st k' ** r\<rbrace>"
+  apply (unfold tassemble_to.simps, rule tm.code_exI)
+  by (rule tm.code_extension, rule h)
+
+lemma t_hoare_local: 
+  assumes h: "(\<And>l. \<lbrace>st i \<and>* p\<rbrace>  i :[ c l ]: j \<lbrace>st k \<and>* q\<rbrace>)"
+  shows "\<lbrace>st i ** p\<rbrace> i:[TLocal c]:j \<lbrace>st k ** q\<rbrace>" using h
+  by (unfold tassemble_to.simps, intro tm.code_exI, simp)
+
+lemma t_hoare_label: 
+      "(\<And>l. l = i \<Longrightarrow> \<lbrace>st l \<and>* p\<rbrace>  l :[ c l ]: j \<lbrace>st k \<and>* q\<rbrace>) \<Longrightarrow>
+             \<lbrace>st i \<and>* p \<rbrace> 
+               i:[(TLabel l; c l)]:j
+             \<lbrace>st k \<and>* q\<rbrace>"
+by (unfold tassemble_to.simps, intro tm.code_exI tm.code_condI, clarify, auto)
+
+primrec t_split_cmd :: "tpg \<Rightarrow> (tpg \<times> tpg option)"
+  where "t_split_cmd (\<guillemotright>inst) = (\<guillemotright>inst, None)" |
+        "t_split_cmd (TLabel l) = (TLabel l, None)" |
+        "t_split_cmd (TSeq c1 c2) = (case (t_split_cmd c2) of
+                                      (c21, Some c22) \<Rightarrow> (TSeq c1 c21, Some c22) |
+                                      (c21, None) \<Rightarrow> (c1, Some c21))" |
+        "t_split_cmd (TLocal c) = (TLocal c, None)"
+
+definition "t_last_cmd tpg = (snd (t_split_cmd tpg))"
+
+declare t_last_cmd_def [simp]
+
+definition "t_blast_cmd tpg = (fst (t_split_cmd tpg))"
+
+declare t_blast_cmd_def [simp]
+
+lemma "t_last_cmd (c1; c2; (TLabel l)) = (Some (TLabel l))"
+  by simp
+
+lemma "t_blast_cmd (c1; c2; TLabel l) = (c1; c2)"
+  by simp
+
+lemma t_split_cmd_eq:
+  assumes "t_split_cmd c = (c1, Some c2)"
+  shows "(i:[c]:j) = (i:[(c1; c2)]:j)"
+  using assms
+proof(induct c arbitrary: c1 c2 i j)
+  case (TSeq cx cy)
+  from `t_split_cmd (cx ; cy) = (c1, Some c2)`
+  show ?case
+    apply (simp split:prod.splits option.splits)
+    apply (cases cy, auto split:prod.splits option.splits)
+    proof -
+      fix a
+      assume h: "t_split_cmd cy = (a, Some c2)"
+      show "i :[ (cx ; cy) ]: j = i :[ ((cx ; a) ; c2) ]: j"
+        apply (simp only: tassemble_to.simps(2) TSeq(2)[OF h] sep_conj_exists)
+        apply (simp add:sep_conj_ac)
+        by (simp add:pred_ex_def, default, auto)
+    qed
+qed auto
+
+lemma t_hoare_label_last_pre: 
+  fixes l
+  assumes h1: "t_split_cmd c = (c', Some (TLabel l))"
+  and h2: "l = j \<Longrightarrow> \<lbrace>p\<rbrace> i:[c']:j \<lbrace>q\<rbrace>"
+  shows "\<lbrace>p\<rbrace> i:[c]:j \<lbrace>q\<rbrace>"
+by (unfold t_split_cmd_eq[OF h1], 
+    simp only:tassemble_to.simps sep_conj_cond, 
+    intro tm.code_exI tm.code_condI, insert h2, auto)
+
+lemma t_hoare_label_last: 
+  fixes l
+  assumes h1: "t_last_cmd c = Some (TLabel l)"
+  and h2: "l = j \<Longrightarrow> \<lbrace>p\<rbrace> i:[t_blast_cmd c]:j \<lbrace>q\<rbrace>"
+  shows "\<lbrace>p\<rbrace> i:[c]:j \<lbrace>q\<rbrace>"
+proof -
+    have "t_split_cmd c = (t_blast_cmd c, t_last_cmd c)"
+      by simp
+  from t_hoare_label_last_pre[OF this[unfolded h1]] h2
+  show ?thesis by auto
+qed
+
+subsection {* Basic assertions for TM *}
+
+function ones :: "int \<Rightarrow> int \<Rightarrow> tassert" where
+  "ones i j = (if j < i then <(i = j + 1)> else
+                (one i) ** ones (i + 1) j)"
+by auto
+termination by (relation "measure(\<lambda> (i::int, j). nat (j - i + 1))") auto
+
+lemma ones_base_simp: "j < i \<Longrightarrow> ones i j = <(i = j + 1)>"
+  by simp
+
+lemma ones_step_simp: "\<not> j < i \<Longrightarrow> ones i j =  ((one i) ** ones (i + 1) j)"
+  by simp
+
+lemmas ones_simps = ones_base_simp ones_step_simp
+
+declare ones.simps [simp del] ones_simps [simp]
+
+lemma ones_induct [case_names Base Step]:
+  fixes P i j
+  assumes h1: "\<And> i j. j < i \<Longrightarrow> P i j (<(i = j + (1::int))>)"
+  and h2: "\<And> i j. \<lbrakk>\<not> j < i; P (i + 1) j (ones (i + 1) j)\<rbrakk> \<Longrightarrow> P i j ((one i) ** ones (i + 1) j)"
+  shows "P i j (ones i j)"
+proof(induct i j rule:ones.induct)
+  fix i j 
+  assume ih: "(\<not> j < i \<Longrightarrow> P (i + 1) j (ones (i + 1) j))"
+  show "P i j (ones i j)"
+  proof(cases "j < i")
+    case True
+    with h1 [OF True]
+    show ?thesis by auto
+  next
+    case False
+    from h2 [OF False] and ih[OF False]
+    have "P i j (one i \<and>* ones (i + 1) j)" by blast
+    with False show ?thesis by auto
+  qed
+qed
+
+function ones' ::  "int \<Rightarrow> int \<Rightarrow> tassert" where
+  "ones' i j = (if j < i then <(i = j + 1)> else
+                ones' i (j - 1) ** one j)"
+by auto
+termination by (relation "measure(\<lambda> (i::int, j). nat (j - i + 1))") auto
+
+lemma ones_rev: "\<not> j < i \<Longrightarrow> (ones i j) = ((ones i (j - 1)) ** one j)"
+proof(induct i j rule:ones_induct)
+  case Base
+  thus ?case by auto
+next
+  case (Step i j)
+  show ?case
+  proof(cases "j < i + 1")
+    case True
+    with Step show ?thesis
+      by simp
+  next
+    case False
+    with Step show ?thesis 
+      by (auto simp:sep_conj_ac)
+  qed
+qed
+
+lemma ones_rev_induct [case_names Base Step]:
+  fixes P i j
+  assumes h1: "\<And> i j. j < i \<Longrightarrow> P i j (<(i = j + (1::int))>)"
+  and h2: "\<And> i j. \<lbrakk>\<not> j < i; P i (j - 1) (ones i (j - 1))\<rbrakk> \<Longrightarrow> P i j ((ones i (j - 1)) ** one j)"
+  shows "P i j (ones i j)"
+proof(induct i j rule:ones'.induct)
+  fix i j 
+  assume ih: "(\<not> j < i \<Longrightarrow> P i (j - 1) (ones i (j - 1)))"
+  show "P i j (ones i j)"
+  proof(cases "j < i")
+    case True
+    with h1 [OF True]
+    show ?thesis by auto
+  next
+    case False
+    from h2 [OF False] and ih[OF False]
+    have "P i j (ones i (j - 1) \<and>* one j)" by blast
+    with ones_rev and False
+    show ?thesis
+      by simp
+  qed
+qed
+
+function zeros :: "int \<Rightarrow> int \<Rightarrow> tassert" where
+  "zeros i j = (if j < i then <(i = j + 1)> else
+                (zero i) ** zeros (i + 1) j)"
+by auto
+termination by (relation "measure(\<lambda> (i::int, j). nat (j - i + 1))") auto
+
+lemma zeros_base_simp: "j < i \<Longrightarrow> zeros i j = <(i = j + 1)>"
+  by simp
+
+lemma zeros_step_simp: "\<not> j < i \<Longrightarrow> zeros i j = ((zero i) ** zeros (i + 1) j)"
+  by simp
+
+declare zeros.simps [simp del]
+lemmas zeros_simps [simp] = zeros_base_simp zeros_step_simp
+
+lemma zeros_induct [case_names Base Step]:
+  fixes P i j
+  assumes h1: "\<And> i j. j < i \<Longrightarrow> P i j (<(i = j + (1::int))>)"
+  and h2: "\<And> i j. \<lbrakk>\<not> j < i; P (i + 1) j (zeros (i + 1) j)\<rbrakk> \<Longrightarrow> 
+                                   P i j ((zero i) ** zeros (i + 1) j)"
+  shows "P i j (zeros i j)"
+proof(induct i j rule:zeros.induct)
+  fix i j 
+  assume ih: "(\<not> j < i \<Longrightarrow> P (i + 1) j (zeros (i + 1) j))"
+  show "P i j (zeros i j)"
+  proof(cases "j < i")
+    case True
+    with h1 [OF True]
+    show ?thesis by auto
+  next
+    case False
+    from h2 [OF False] and ih[OF False]
+    have "P i j (zero i \<and>* zeros (i + 1) j)" by blast
+    with False show ?thesis by auto
+  qed
+qed
+
+lemma zeros_rev: "\<not> j < i \<Longrightarrow> (zeros i j) = ((zeros i (j - 1)) ** zero j)"
+proof(induct i j rule:zeros_induct)
+  case (Base i j)
+  thus ?case by auto
+next
+  case (Step i j)
+  show ?case
+  proof(cases "j < i + 1")
+    case True
+    with Step show ?thesis by auto
+  next
+    case False
+    with Step show ?thesis by (auto simp:sep_conj_ac)
+  qed
+qed
+
+lemma zeros_rev_induct [case_names Base Step]:
+  fixes P i j
+  assumes h1: "\<And> i j. j < i \<Longrightarrow> P i j (<(i = j + (1::int))>)"
+  and h2: "\<And> i j. \<lbrakk>\<not> j < i; P i (j - 1) (zeros i (j - 1))\<rbrakk> \<Longrightarrow> 
+                       P i j ((zeros i (j - 1)) ** zero j)"
+  shows "P i j (zeros i j)"
+proof(induct i j rule:ones'.induct)
+  fix i j 
+  assume ih: "(\<not> j < i \<Longrightarrow> P i (j - 1) (zeros i (j - 1)))"
+  show "P i j (zeros i j)"
+  proof(cases "j < i")
+    case True
+    with h1 [OF True]
+    show ?thesis by auto
+  next
+    case False
+    from h2 [OF False] and ih[OF False]
+    have "P i j (zeros i (j - 1) \<and>* zero j)" by blast
+    with zeros_rev and False
+    show ?thesis by simp
+  qed
+qed
+
+definition "rep i j k = ((ones i (i + (int k))) ** <(j = i + int k)>)"
+
+fun reps :: "int \<Rightarrow> int \<Rightarrow> nat list\<Rightarrow> tassert"
+  where
+  "reps i j [] = <(i = j + 1)>" |
+  "reps i j [k] = (ones i (i + int k) ** <(j = i + int k)>)" |
+  "reps i j (k # ks) = (ones i (i + int k) ** zero (i + int k + 1) ** reps (i + int k + 2) j ks)"
+
+lemma reps_simp3: "ks \<noteq> [] \<Longrightarrow> 
+  reps i j (k # ks) = (ones i (i + int k) ** zero (i + int k + 1) ** reps (i + int k + 2) j ks)"
+  by (cases ks, auto)
+
+definition "reps_sep_len ks = (if length ks \<le> 1 then 0 else (length ks) - 1)"
+
+definition "reps_ctnt_len ks = (\<Sum> k \<leftarrow> ks. (1 + k))"
+
+definition "reps_len ks = (reps_sep_len ks) + (reps_ctnt_len ks)"
+
+definition "splited xs ys zs = ((xs = ys @ zs) \<and> (ys \<noteq> []) \<and> (zs \<noteq> []))"
+
+lemma list_split: 
+  assumes h: "k # ks = ys @ zs"
+      and h1: "ys \<noteq> []"
+  shows "(ys = [k] \<and> zs = ks) \<or> (\<exists> ys'. ys' \<noteq> [] \<and> ys = k # ys' \<and> ks = ys' @ zs)"
+proof(cases ys)
+  case Nil
+  with h1 show ?thesis by auto
+next
+  case (Cons y' ys')
+  with h have "k#ks = y'#(ys' @ zs)" by simp
+  hence hh: "y' = k" "ks = ys' @ zs" by auto
+  show ?thesis
+  proof(cases "ys' = []")
+    case False
+    show ?thesis
+    proof(rule disjI2)
+      show " \<exists>ys'. ys' \<noteq> [] \<and> ys = k # ys' \<and> ks = ys' @ zs"
+      proof(rule exI[where x = ys'])
+        from False hh Cons show "ys' \<noteq> [] \<and> ys = k # ys' \<and> ks = ys' @ zs" by auto
+      qed
+    qed
+  next
+    case True
+    show ?thesis
+    proof(rule disjI1)
+      from hh True Cons
+      show "ys = [k] \<and> zs = ks" by auto
+    qed
+  qed
+qed
+
+lemma splited_cons[elim_format]: 
+  assumes h: "splited (k # ks) ys zs"
+  shows "(ys = [k] \<and> zs = ks) \<or> (\<exists> ys'. ys' \<noteq> [] \<and> ys = k # ys' \<and> splited ks ys' zs)"
+proof -
+  from h have "k # ks = ys @ zs" "ys \<noteq> [] " unfolding splited_def by auto
+  from list_split[OF this]
+  have "ys = [k] \<and> zs = ks \<or> (\<exists>ys'. ys' \<noteq> [] \<and> ys = k # ys' \<and> ks = ys' @ zs)" .
+  thus ?thesis
+  proof
+    assume " ys = [k] \<and> zs = ks" thus ?thesis by auto
+  next
+    assume "\<exists>ys'. ys' \<noteq> [] \<and> ys = k # ys' \<and> ks = ys' @ zs"
+    then obtain ys' where hh: "ys' \<noteq> []" "ys = k # ys'" "ks = ys' @ zs" by auto
+    show ?thesis
+    proof(rule disjI2)
+      show "\<exists>ys'. ys' \<noteq> [] \<and> ys = k # ys' \<and> splited ks ys' zs"
+      proof(rule exI[where x = ys'])
+        from h have "zs \<noteq> []" by (unfold splited_def, simp)
+        with hh(1) hh(3)
+        have "splited ks ys' zs"
+          by (unfold splited_def, auto)
+        with hh(1) hh(2) show "ys' \<noteq> [] \<and> ys = k # ys' \<and> splited ks ys' zs" by auto
+      qed
+    qed
+  qed
+qed
+
+lemma splited_cons_elim:
+  assumes h: "splited (k # ks) ys zs"
+  and h1: "\<lbrakk>ys = [k]; zs = ks\<rbrakk> \<Longrightarrow> P"
+  and h2: "\<And> ys'. \<lbrakk>ys' \<noteq> []; ys = k#ys'; splited ks ys' zs\<rbrakk> \<Longrightarrow> P"
+  shows P
+proof(rule splited_cons[OF h])
+  assume "ys = [k] \<and> zs = ks \<or> (\<exists>ys'. ys' \<noteq> [] \<and> ys = k # ys' \<and> splited ks ys' zs)"
+  thus P
+  proof
+    assume hh: "ys = [k] \<and> zs = ks"
+    show P
+    proof(rule h1)
+      from hh show "ys = [k]" by simp
+    next
+      from hh show "zs = ks" by simp
+    qed
+  next
+    assume "\<exists>ys'. ys' \<noteq> [] \<and> ys = k # ys' \<and> splited ks ys' zs"
+    then obtain ys' where hh: "ys' \<noteq> []" "ys = k # ys'"  "splited ks ys' zs" by auto
+    from h2[OF this]
+    show P .
+  qed
+qed
+
+lemma list_nth_expand:
+  "i < length xs \<Longrightarrow> xs = take i xs @ [xs!i] @ drop (Suc i) xs"
+  by (metis Cons_eq_appendI append_take_drop_id drop_Suc_conv_tl eq_Nil_appendI)
+
+lemma reps_len_nil: "reps_len [] = 0"
+   by (auto simp:reps_len_def reps_sep_len_def reps_ctnt_len_def)
+
+lemma reps_len_sg: "reps_len [k] = 1 + k"
+  by (auto simp:reps_len_def reps_sep_len_def reps_ctnt_len_def)
+
+lemma reps_len_cons: "ks \<noteq> [] \<Longrightarrow> (reps_len (k # ks)) = 2 + k + reps_len ks "
+proof(induct ks arbitrary:k)
+  case (Cons n ns)
+  thus ?case
+    by(cases "ns = []", 
+      auto simp:reps_len_def reps_sep_len_def reps_ctnt_len_def)
+qed auto
+
+lemma reps_len_correct:
+  assumes h: "(reps i j ks) s"
+  shows "j = i + int (reps_len ks) - 1"
+  using h
+proof(induct ks arbitrary:i j s)
+  case Nil
+  thus ?case
+    by (auto simp:reps_len_nil pasrt_def)
+next
+  case (Cons n ns)
+  thus ?case
+  proof(cases "ns = []")
+    case False
+    from Cons(2)[unfolded reps_simp3[OF False]]
+    obtain s' where "(reps (i + int n + 2) j ns) s'"
+      by (auto elim!:sep_conjE)
+    from Cons.hyps[OF this]
+    show ?thesis
+      by (unfold reps_len_cons[OF False], simp)
+  next
+    case True
+    with Cons(2)
+    show ?thesis
+      by (auto simp:reps_len_sg pasrt_def)
+  qed
+qed
+
+lemma reps_len_expand: 
+  shows "(EXS j. (reps i j ks)) = (reps i (i + int (reps_len ks) - 1) ks)"
+proof(default+)
+  fix s
+  assume "(EXS j. reps i j ks) s"
+  with reps_len_correct show "reps i (i + int (reps_len ks) - 1) ks s"
+    by (auto simp:pred_ex_def)
+next
+  fix s
+  assume "reps i (i + int (reps_len ks) - 1) ks s"
+  thus "(EXS j. reps i j ks) s"  by (auto simp:pred_ex_def)
+qed
+
+lemma reps_len_expand1: "(EXS j. (reps i j ks \<and>* r)) = (reps i (i + int (reps_len ks) - 1) ks \<and>* r)"
+by (metis reps_len_def reps_len_expand sep.mult_commute sep_conj_exists1)
+
+lemma reps_splited:
+  assumes h: "splited xs ys zs"
+  shows "reps i j xs = (reps i (i + int (reps_len ys) - 1) ys \<and>* 
+                        zero (i + int (reps_len ys)) \<and>* 
+                        reps (i + int (reps_len ys) + 1) j zs)"
+  using h
+proof(induct xs arbitrary: i j ys zs)
+  case Nil
+  thus ?case
+    by (unfold splited_def, auto)
+next
+  case (Cons k ks)
+  from Cons(2)
+  show ?case
+  proof(rule splited_cons_elim)
+    assume h: "ys = [k]" "zs = ks"
+    with Cons(2)
+    show ?thesis
+    proof(cases "ks = []")
+      case True
+      with Cons(2) have False
+        by (simp add:splited_def, cases ys, auto)
+      thus ?thesis by auto
+    next
+      case False
+      have ss: "i + int k + 1 = i + (1 + int k)"
+           "i + int k + 2 = 2 + (i + int k)" by auto
+      show ?thesis
+        by (unfold h(1), unfold reps_simp3[OF False] reps.simps(2) 
+            reps_len_sg, auto simp:sep_conj_ac,
+            unfold ss h(2), simp)
+    qed
+  next
+    fix ys'
+    assume h: "ys' \<noteq> []" "ys = k # ys'" "splited ks ys' zs"
+    hence nnks: "ks \<noteq> []"
+      by (unfold splited_def, auto)
+    have ss: "i + int k + 2 + int (reps_len ys') = 
+              i + (2 + (int k + int (reps_len ys')))" by auto
+    show ?thesis
+      by (simp add: reps_simp3[OF nnks] reps_simp3[OF h(1)] 
+                    Cons(1)[OF h(3)] h(2) 
+                    reps_len_cons[OF h(1)]
+                    sep_conj_ac, subst ss, simp)
+  qed
+qed
+
+subsection {* A theory of list extension *}
+
+definition "list_ext n xs = xs @ replicate ((Suc n) - length xs) 0"
+
+lemma list_ext_len_eq: "length (list_ext a xs) = length xs + (Suc a - length xs)"
+  by (metis length_append length_replicate list_ext_def)
+
+lemma list_ext_len: "a < length (list_ext a xs)"
+  by (unfold list_ext_len_eq, auto)
+
+lemma list_ext_lt: "a < length xs \<Longrightarrow> list_ext a xs = xs"
+  by (smt append_Nil2 list_ext_def replicate_0)
+
+lemma list_ext_lt_get: 
+  assumes h: "a' < length xs"
+  shows "(list_ext a xs)!a' = xs!a'"
+proof(cases "a < length xs")
+  case True
+  with h
+  show ?thesis by (auto simp:list_ext_lt)
+next
+  case False
+  with h show ?thesis
+    apply (unfold list_ext_def)
+    by (metis nth_append)
+qed
+
+lemma list_ext_get_upd: "((list_ext a xs)[a:=v])!a = v"
+  by (metis list_ext_len nth_list_update_eq)
+
+lemma nth_app: "length xs \<le> a \<Longrightarrow> (xs @ ys)!a = ys!(a - length xs)"
+  by (metis not_less nth_append)
+
+lemma pred_exI: 
+  assumes "(P(x) \<and>* r) s"
+  shows "((EXS x. P(x)) \<and>* r) s"
+  by (metis assms pred_ex_def sep_globalise)
+
+lemma list_ext_tail:
+  assumes h1: "length xs \<le> a"
+  and h2: "length xs \<le> a'"
+  and h3: "a' \<le> a"
+  shows "(list_ext a xs)!a' = 0"
+proof -
+  from h1 h2
+  have "a' - length xs < length (replicate (Suc a - length xs) 0)"
+    by (metis diff_less_mono h3 le_imp_less_Suc length_replicate)
+  moreover from h1 have "replicate (Suc a - length xs) 0 \<noteq> []"
+    by (smt empty_replicate)
+  ultimately have "(replicate (Suc a - length xs) 0)!(a' - length xs) = 0"
+    by (metis length_replicate nth_replicate)
+  moreover have "(xs @ (replicate (Suc a - length xs) 0))!a' = 
+            (replicate (Suc a - length xs) 0)!(a' - length xs)"
+    by (rule nth_app[OF h2])
+  ultimately show ?thesis
+    by (auto simp:list_ext_def)
+qed
+
+lemmas list_ext_simps = list_ext_lt_get list_ext_lt list_ext_len list_ext_len_eq
+
+lemma listsum_upd_suc:
+  "a < length ks \<Longrightarrow> listsum (map Suc (ks[a := Suc (ks ! a)]))= (Suc (listsum (map Suc ks)))"
+by (smt Ex_list_of_length Skolem_list_nth elem_le_listsum_nat 
+     length_induct length_list_update length_map length_splice 
+     list_eq_iff_nth_eq list_ext_get_upd list_ext_lt_get list_update_beyond 
+     list_update_id list_update_overwrite list_update_same_conv list_update_swap 
+     listsum_update_nat map_eq_imp_length_eq map_update neq_if_length_neq 
+     nth_equalityI nth_list_update nth_list_update_eq nth_list_update_neq nth_map reps_sep_len_def)
+
+lemma reps_len_suc:
+  assumes "a < length ks"
+  shows "reps_len (ks[a:=Suc(ks!a)]) = 1 + reps_len ks"
+proof(cases "length ks \<le> 1")
+  case True
+  with `a < length ks` 
+  obtain k where "ks = [k]" "a = 0"
+    by (smt length_0_conv length_Suc_conv)
+  thus ?thesis
+      apply (unfold `ks = [k]` `a = 0`)
+      by (unfold reps_len_def reps_sep_len_def reps_ctnt_len_def, auto)
+next
+  case False
+  have "Suc = (op + (Suc 0))"
+    by (default, auto)
+  with listsum_upd_suc[OF `a < length ks`] False
+  show ?thesis
+     by (unfold reps_len_def reps_sep_len_def reps_ctnt_len_def, simp)
+qed
+  
+lemma ks_suc_len:
+  assumes h1: "(reps i j ks) s"
+  and h2: "ks!a = v"
+  and h3: "a < length ks"
+  and h4: "(reps i j' (ks[a:=Suc v])) s'"
+  shows "j' = j + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1"
+proof -
+  from reps_len_correct[OF h1, unfolded list_ext_len_eq]
+       reps_len_correct[OF h4, unfolded list_ext_len_eq] 
+       reps_len_suc[OF `a < length ks`] h2 h3
+  show ?thesis
+    by (unfold list_ext_lt[OF `a < length ks`], auto)
+qed
+
+lemma ks_ext_len:
+  assumes h1: "(reps i j ks) s"
+  and h3: "length ks \<le> a"
+  and h4: "reps i j' (list_ext a ks) s'"
+  shows "j' = j + int (reps_len (list_ext a ks)) - int (reps_len ks)"
+proof -
+  from reps_len_correct[OF h1, unfolded  list_ext_len_eq]
+    and reps_len_correct[OF h4, unfolded list_ext_len_eq]
+  h3
+  show ?thesis by auto
+qed
+
+definition 
+  "reps' i j ks = 
+     (if ks = [] then (<(i = j + 1)>)  else (reps i (j - 1) ks \<and>* zero j))"
+
+lemma reps'_expand: 
+  assumes h: "ks \<noteq> []"
+  shows "(EXS j. reps' i j ks) = reps' i (i + int (reps_len ks)) ks"
+proof -
+  show ?thesis
+  proof(default+)
+    fix s
+    assume "(EXS j. reps' i j ks) s"
+    with h have "(EXS j. reps i (j - 1) ks \<and>* zero j) s" 
+      by (simp add:reps'_def)
+    hence "(reps i (i + int (reps_len ks) - 1) ks \<and>* zero (i + int (reps_len ks))) s"
+    proof(elim EXS_elim)
+      fix j
+      assume "(reps i (j - 1) ks \<and>* zero j) s"
+      then obtain s1 s2 where "s = s1 + s2" "s1 ## s2" "reps i (j - 1) ks s1" "zero j s2"
+        by (auto elim!:sep_conjE)
+      from reps_len_correct[OF this(3)]
+      have "j = i + int (reps_len ks)" by auto
+      with `reps i (j - 1) ks s1` have "reps i (i + int (reps_len ks) - 1) ks s1"
+        by simp
+      moreover from `j = i + int (reps_len ks)` and `zero j s2`
+      have "zero (i + int (reps_len ks)) s2" by auto
+      ultimately show "(reps i (i + int (reps_len ks) - 1) ks \<and>* zero (i + int (reps_len ks))) s"
+        using `s = s1 + s2` `s1 ## s2` by (auto intro!:sep_conjI)
+    qed
+    thus "reps' i (i + int (reps_len ks)) ks s"
+      by (simp add:h reps'_def)
+  next
+    fix s 
+    assume "reps' i (i + int (reps_len ks)) ks s"
+    thus "(EXS j. reps' i j ks) s"
+      by (auto intro!:EXS_intro)
+  qed
+qed
+
+lemma reps'_len_correct: 
+  assumes h: "(reps' i j ks) s"
+  and h1: "ks \<noteq> []"
+  shows "(j = i + int (reps_len ks))"
+proof -
+  from h1 have "reps' i j ks s = (reps i (j - 1) ks \<and>* zero j) s" by (simp add:reps'_def)
+  from h[unfolded this]
+  obtain s' where "reps i (j - 1) ks s'"
+    by (auto elim!:sep_conjE)
+  from reps_len_correct[OF this]
+  show ?thesis by simp
+qed
+
+lemma reps'_append:
+  "reps' i j (ks1 @ ks2) = (EXS m. (reps' i (m - 1) ks1 \<and>* reps' m j ks2))"
+proof(cases "ks1 = []")
+  case True
+  thus ?thesis
+    by (auto simp:reps'_def pred_ex_def pasrt_def set_ins_def sep_conj_def)
+next
+  case False
+  show ?thesis
+  proof(cases "ks2 = []")
+    case False
+    from `ks1 \<noteq> []` and `ks2 \<noteq> []` 
+    have "splited (ks1 @ ks2) ks1 ks2" by (auto simp:splited_def)
+    from reps_splited[OF this, of i "j - 1"]
+    have eq_1: "reps i (j - 1) (ks1 @ ks2) =
+           (reps i (i + int (reps_len ks1) - 1) ks1 \<and>*
+           zero (i + int (reps_len ks1)) \<and>* 
+           reps (i + int (reps_len ks1) + 1) (j - 1) ks2)" .
+    from `ks1 \<noteq> []`
+    have eq_2: "reps' i j (ks1 @ ks2) = (reps i (j - 1) (ks1 @ ks2) \<and>* zero j)"
+      by (unfold reps'_def, simp)
+    show ?thesis
+    proof(default+)
+      fix s
+      assume "reps' i j (ks1 @ ks2) s"
+      show "(EXS m. reps' i (m - 1) ks1 \<and>* reps' m j ks2) s"
+      proof(rule EXS_intro[where x = "i + int(reps_len ks1) + 1"])
+        from `reps' i j (ks1 @ ks2) s`[unfolded eq_2 eq_1]
+        and `ks1 \<noteq> []` `ks2 \<noteq> []`
+        show "(reps' i (i + int (reps_len ks1) + 1 - 1) ks1 \<and>* 
+                         reps' (i + int (reps_len ks1) + 1) j ks2) s"
+          by (unfold reps'_def, simp, sep_cancel+)
+      qed
+    next
+      fix s
+      assume "(EXS m. reps' i (m - 1) ks1 \<and>* reps' m j ks2) s"
+      thus "reps' i j (ks1 @ ks2) s"
+      proof(elim EXS_elim)
+        fix m
+        assume "(reps' i (m - 1) ks1 \<and>* reps' m j ks2) s"
+        then obtain s1 s2 where h: 
+          "s = s1 + s2" "s1 ## s2" "reps' i (m - 1) ks1 s1"
+          " reps' m j ks2 s2" by (auto elim!:sep_conjE)
+        from reps'_len_correct[OF this(3) `ks1 \<noteq> []`]
+        have eq_m: "m = i + int (reps_len ks1) + 1" by simp
+        have "((reps i (i + int (reps_len ks1) - 1) ks1 \<and>* zero (i + int (reps_len ks1))) \<and>* 
+               ((reps (i + int (reps_len ks1) + 1) (j - 1) ks2) \<and>* zero j)) s"
+          (is "(?P \<and>* ?Q) s") 
+        proof(rule sep_conjI)
+          from h(3) eq_m `ks1 \<noteq> []` show "?P s1"
+            by (simp add:reps'_def)
+        next
+          from h(4) eq_m `ks2 \<noteq> []` show "?Q s2"
+            by (simp add:reps'_def)
+        next
+          from h(2) show "s1 ## s2" .
+        next
+          from h(1) show "s = s1 + s2" .
+        qed
+        thus "reps' i j (ks1 @ ks2) s"
+          by (unfold eq_2 eq_1, auto simp:sep_conj_ac)
+      qed
+    qed
+  next
+    case True
+    have "-1 + j = j - 1" by auto
+    with `ks1 \<noteq> []` True
+    show ?thesis
+      apply (auto simp:reps'_def pred_ex_def pasrt_def)
+      apply (unfold `-1 + j = j - 1`)
+      by (fold sep_empty_def, simp only:sep_conj_empty)
+  qed
+qed
+
+lemma reps'_sg: 
+  "reps' i j [k] = 
+       (<(j = i + int k + 1)> \<and>* ones i (i + int k) \<and>* zero j)"
+  apply (unfold reps'_def, default, auto simp:sep_conj_ac)
+  by (sep_cancel+, simp add:pasrt_def)+
+
+
+section {* Basic macros for TM *}
+
+definition "write_zero = (TL exit. \<guillemotright>((W0, exit), (W0, exit)); TLabel exit)"
+
+lemma st_upd: 
+  assumes pre: "(st i' ** r) (trset_of (f, x, i, y, z))"
+  shows "(st i'' ** r) (trset_of (f, x,  i'', y, z))"
+proof -
+  from stimes_sgD[OF pre[unfolded st_def], unfolded trset_of.simps, unfolded stD[OF pre]]
+  have "r (tprog_set x \<union> tpc_set i' \<union> tpos_set y \<union> tmem_set z \<union> tfaults_set f - tpc_set i')"
+    by blast
+  moreover have 
+    "(tprog_set x \<union> tpc_set i' \<union> tpos_set y \<union> tmem_set z \<union> tfaults_set f - tpc_set i') =
+     (tprog_set x \<union> tpos_set y \<union> tmem_set z \<union> tfaults_set f)"
+    by (unfold tpn_set_def, auto)
+  ultimately have r_rest: "r (tprog_set x \<union> tpos_set y \<union> tmem_set z \<union> tfaults_set f)"
+    by auto
+  show ?thesis
+  proof(rule sep_conjI[where Q = r, OF _ r_rest])
+    show "st i'' (tpc_set i'')" 
+      by (unfold st_def sg_def, simp)
+  next
+    show "tpc_set i'' ## tprog_set x \<union> tpos_set y \<union> tmem_set z \<union> tfaults_set f"
+      by (unfold tpn_set_def sep_disj_set_def, auto)
+  next
+    show "trset_of (f, x, i'', y, z) =
+             tpc_set i'' + (tprog_set x \<union> tpos_set y \<union> tmem_set z \<union> tfaults_set f)"
+      by (unfold trset_of.simps plus_set_def, auto)
+  qed
+qed
+
+lemma pos_upd: 
+  assumes pre: "(ps i ** r) (trset_of (f, x, y, i', z))"
+  shows "(ps j ** r) (trset_of (f, x, y, j, z))"
+proof -
+  from stimes_sgD[OF pre[unfolded ps_def], unfolded trset_of.simps, unfolded psD[OF pre]]
+  have "r (tprog_set x \<union> tpc_set y \<union> tpos_set i \<union> tmem_set z \<union> 
+              tfaults_set f - tpos_set i)" (is "r ?xs")
+    by blast
+  moreover have 
+    "?xs = tprog_set x \<union> tpc_set y  \<union> tmem_set z \<union> tfaults_set f"
+    by (unfold tpn_set_def, auto)
+  ultimately have r_rest: "r \<dots>"
+    by auto
+  show ?thesis
+  proof(rule sep_conjI[where Q = r, OF _ r_rest])
+    show "ps j (tpos_set j)" 
+      by (unfold ps_def sg_def, simp)
+  next
+    show "tpos_set j ## tprog_set x \<union> tpc_set y \<union> tmem_set z \<union> tfaults_set f"
+      by (unfold tpn_set_def sep_disj_set_def, auto)
+  next
+    show "trset_of (f, x, y, j, z) = 
+             tpos_set j + (tprog_set x \<union> tpc_set y \<union> tmem_set z \<union> tfaults_set f)"
+      by (unfold trset_of.simps plus_set_def, auto)
+  qed
+qed
+
+lemma TM_in_simp: "({TM a v} \<subseteq> 
+                      tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> tmem_set mem \<union> tfaults_set f) = 
+                             ({TM a v} \<subseteq> tmem_set mem)"
+  by (unfold tpn_set_def, auto)
+
+lemma tmem_set_upd: 
+  "{TM a v} \<subseteq> tmem_set mem \<Longrightarrow> 
+        tmem_set (mem(a:=Some v')) = ((tmem_set mem) - {TM a v}) \<union> {TM a v'}"
+  by (unfold tpn_set_def, auto)
+
+lemma tmem_set_disj: "{TM a v} \<subseteq> tmem_set mem \<Longrightarrow> 
+                            {TM a v'} \<inter>  (tmem_set mem - {TM a v}) = {}"
+  by (unfold tpn_set_def, auto)
+
+lemma smem_upd: "((tm a v) ** r) (trset_of (f, x, y, z, mem))  \<Longrightarrow> 
+                    ((tm a v') ** r) (trset_of (f, x, y, z, mem(a := Some v')))"
+proof -
+  have eq_s: "(tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> tmem_set mem \<union> tfaults_set f - {TM a v}) =
+    (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f)"
+    by (unfold tpn_set_def, auto)
+  assume g: "(tm a v \<and>* r) (trset_of (f, x, y, z, mem))"
+  from this[unfolded trset_of.simps tm_def]
+  have h: "(sg {TM a v} \<and>* r) (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> tmem_set mem \<union> tfaults_set f)" .
+  hence h0: "r (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f)"
+    by(sep_drule stimes_sgD, clarify, unfold eq_s, auto)
+  from h TM_in_simp have "{TM a v} \<subseteq> tmem_set mem"
+    by(sep_drule stimes_sgD, auto)
+  from tmem_set_upd [OF this] tmem_set_disj[OF this]
+  have h2: "tmem_set (mem(a \<mapsto> v')) = {TM a v'} \<union> (tmem_set mem - {TM a v})" 
+           "{TM a v'} \<inter> (tmem_set mem - {TM a v}) = {}" by auto
+  show ?thesis
+  proof -
+    have "(tm a v' ** r) (tmem_set (mem(a \<mapsto> v')) \<union> tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> tfaults_set f)"
+    proof(rule sep_conjI)
+      show "tm a v' ({TM a v'})" by (unfold tm_def sg_def, simp)
+    next
+      from h0 show "r (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f)" .
+    next
+      show "{TM a v'} ## tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f"
+      proof -
+        from g have " mem a = Some v"
+          by(sep_frule memD, simp)
+        thus "?thesis"
+          by(unfold tpn_set_def set_ins_def, auto)
+      qed
+    next
+      show "tmem_set (mem(a \<mapsto> v')) \<union> tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> tfaults_set f =
+    {TM a v'} + (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f)"
+        by (unfold h2(1) set_ins_def eq_s, auto)
+    qed
+    thus ?thesis 
+      apply (unfold trset_of.simps)
+      by (metis sup_commute sup_left_commute)
+  qed
+qed
+
+lemma hoare_write_zero: 
+  "\<lbrace>st i ** ps p ** tm p v\<rbrace> 
+     i:[write_zero]:j
+   \<lbrace>st j ** ps p ** tm p Bk\<rbrace>"
+proof(unfold write_zero_def, intro t_hoare_local, rule t_hoare_label_last, simp, simp)
+  show "\<lbrace>st i \<and>* ps p \<and>* tm p v\<rbrace>  i :[ \<guillemotright> ((W0, j), W0, j) ]: j \<lbrace>st j \<and>* ps p \<and>* tm p Bk\<rbrace>"
+  proof(unfold tassemble_to.simps, simp only:sep_conj_cond,
+        intro tm.code_condI, simp)
+    assume eq_j: "j = Suc i"
+    show "\<lbrace>st i \<and>* ps p \<and>* tm p v\<rbrace>  sg {TC i ((W0, Suc i), W0, Suc i)} 
+          \<lbrace>st (Suc i) \<and>* ps p \<and>* tm p Bk\<rbrace>"
+    proof(fold eq_j, unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify)
+      fix ft prog cs i' mem r
+      assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v \<and>* sg {TC i ((W0, j), W0, j)})
+              (trset_of (ft, prog, cs, i', mem))"
+      from h have "prog i = Some ((W0, j), W0, j)"
+        apply(rule_tac r = "r \<and>* ps p \<and>* tm p v" in codeD)
+        by(simp add: sep_conj_ac)
+      from h and this have stp:
+        "tm.run 1 (ft, prog, cs, i', mem) = (ft, prog, j, i', mem(i'\<mapsto> Bk))" (is "?x = ?y")
+        apply(sep_frule psD)
+        apply(sep_frule stD)
+        apply(sep_frule memD, simp)
+        by(cases v, unfold tm.run_def, auto)
+      from h have "i' = p"
+        by(sep_drule psD, simp)
+      with h
+      have "(r \<and>* ps p \<and>* st j \<and>* tm p Bk \<and>* sg {TC i ((W0, j), W0, j)}) (trset_of ?x)"
+        apply(unfold stp)
+        apply(sep_drule pos_upd, sep_drule st_upd, sep_drule smem_upd)
+        apply(auto simp: sep_conj_ac)
+        done
+      thus "\<exists>k. (r \<and>* ps p \<and>* st j \<and>* tm p Bk \<and>* sg {TC i ((W0, j), W0, j)}) 
+             (trset_of (tm.run (Suc k) (ft, prog, cs, i', mem)))"
+        apply (rule_tac x = 0 in exI)
+        by auto
+    qed
+  qed
+qed
+
+lemma hoare_write_zero_gen[step]: 
+  assumes "p = q"
+  shows  "\<lbrace>st i ** ps p ** tm q v\<rbrace> 
+            i:[write_zero]:j
+          \<lbrace>st j ** ps p ** tm q Bk\<rbrace>"
+  by (unfold assms, rule hoare_write_zero)
+
+definition "write_one = (TL exit. \<guillemotright>((W1, exit), (W1, exit)); TLabel exit)"
+
+lemma hoare_write_one: 
+  "\<lbrace>st i ** ps p ** tm p v\<rbrace> 
+     i:[write_one]:j
+   \<lbrace>st j ** ps p ** tm p Oc\<rbrace>"
+proof(unfold write_one_def, intro t_hoare_local, rule t_hoare_label_last, simp+)
+  fix l
+  show "\<lbrace>st i \<and>* ps p \<and>* tm p v\<rbrace>  i :[ \<guillemotright> ((W1, j), W1, j) ]: j \<lbrace>st j \<and>* ps p \<and>* tm p Oc\<rbrace>"
+  proof(unfold tassemble_to.simps, simp only:sep_conj_cond,
+        rule_tac tm.code_condI, simp add: sep_conj_ac)
+    let ?j = "Suc i"
+    show "\<lbrace>ps p \<and>* st i \<and>* tm p v\<rbrace>  sg {TC i ((W1, ?j), W1, ?j)} 
+          \<lbrace>ps p \<and>* st ?j \<and>* tm p Oc\<rbrace>"
+    proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify)
+      fix ft prog cs i' mem r
+      assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v \<and>* sg {TC i ((W1, ?j), W1, ?j)})
+              (trset_of (ft, prog, cs, i', mem))"
+      from h have "prog i = Some ((W1, ?j), W1, ?j)"
+        apply(rule_tac r = "r \<and>* ps p \<and>* tm p v" in codeD)
+        by(simp add: sep_conj_ac)
+      from h and this have stp:
+        "tm.run 1 (ft, prog, cs, i', mem) = 
+                     (ft, prog, ?j, i', mem(i'\<mapsto> Oc))" (is "?x = ?y")
+        apply(sep_frule psD)
+        apply(sep_frule stD)
+        apply(sep_frule memD, simp)
+        by(cases v, unfold tm.run_def, auto)
+      from h have "i' = p"
+        by(sep_drule psD, simp)
+      with h
+      have "(r \<and>* ps p \<and>* st ?j \<and>* tm p Oc \<and>* sg {TC i ((W1, ?j), W1, ?j)}) (trset_of ?x)"
+        apply(unfold stp)
+        apply(sep_drule pos_upd, sep_drule st_upd, sep_drule smem_upd)
+        apply(auto simp: sep_conj_ac)
+        done
+      thus "\<exists>k. (r \<and>* ps p \<and>* st ?j \<and>* tm p Oc \<and>* sg {TC i ((W1, ?j), W1, ?j)}) 
+             (trset_of (tm.run (Suc k) (ft, prog, cs, i', mem)))"
+        apply (rule_tac x = 0 in exI)
+        by auto
+    qed
+  qed
+qed
+
+lemma hoare_write_one_gen[step]: 
+  assumes "p = q"
+  shows  "\<lbrace>st i ** ps p ** tm q v\<rbrace> 
+              i:[write_one]:j
+          \<lbrace>st j ** ps p ** tm q Oc\<rbrace>"
+  by (unfold assms, rule hoare_write_one)
+
+definition "move_left = (TL exit . \<guillemotright>((L, exit), (L, exit)); TLabel exit)"
+
+lemma hoare_move_left: 
+  "\<lbrace>st i ** ps p ** tm p v2\<rbrace> 
+     i:[move_left]:j
+   \<lbrace>st j ** ps (p - 1) **  tm p v2\<rbrace>"
+proof(unfold move_left_def, intro t_hoare_local, rule t_hoare_label_last, simp+)
+  fix l
+  show "\<lbrace>st i \<and>* ps p \<and>* tm p v2\<rbrace>  i :[ \<guillemotright> ((L, l), L, l) ]: l
+        \<lbrace>st l \<and>* ps (p - 1) \<and>* tm p v2\<rbrace>"
+  proof(unfold tassemble_to.simps, simp only:sep_conj_cond,
+      intro tm.code_condI, simp add: sep_conj_ac)
+    let ?j = "Suc i"
+    show "\<lbrace>ps p \<and>* st i \<and>* tm p v2\<rbrace>  sg {TC i ((L, ?j), L, ?j)} 
+          \<lbrace>st ?j \<and>* tm p v2 \<and>* ps (p - 1)\<rbrace>"
+    proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify)
+      fix ft prog cs i' mem r
+      assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v2 \<and>* sg {TC i ((L, ?j), L, ?j)}) 
+                       (trset_of (ft, prog, cs, i',  mem))"
+      from h have "prog i = Some ((L, ?j), L, ?j)"
+        apply(rule_tac r = "r \<and>* ps p \<and>* tm p v2" in codeD)
+        by(simp add: sep_conj_ac)
+      from h and this have stp:
+        "tm.run 1 (ft, prog, cs, i', mem) = (ft, prog, ?j, i' - 1, mem)" (is "?x = ?y")
+        apply(sep_frule psD)
+        apply(sep_frule stD)
+        apply(sep_frule memD, simp)
+        apply(unfold tm.run_def, case_tac v2, auto)
+        done
+      from h have "i' = p"
+        by(sep_drule psD, simp)
+      with h
+      have "(r \<and>* st ?j \<and>* tm p v2 \<and>* ps (p - 1) \<and>* sg {TC i ((L, ?j), L, ?j)}) 
+               (trset_of ?x)"
+        apply(unfold stp)
+        apply(sep_drule pos_upd, sep_drule st_upd, simp)
+      proof -
+        assume "(st ?j \<and>* ps (p - 1) \<and>* r \<and>* tm p v2 \<and>* sg {TC i ((L, ?j), L, ?j)}) 
+                   (trset_of (ft, prog, ?j, p - 1, mem))"
+        thus "(r \<and>* st ?j \<and>* tm p v2 \<and>* ps (p - 1) \<and>* sg {TC i ((L, ?j), L, ?j)}) 
+                    (trset_of (ft, prog, ?j, p - 1, mem))"
+          by(simp add: sep_conj_ac)
+      qed
+      thus "\<exists>k. (r \<and>* st ?j \<and>* tm p v2 \<and>* ps (p - 1) \<and>* sg {TC i ((L, ?j), L, ?j)}) 
+             (trset_of (tm.run (Suc k) (ft, prog, cs, i', mem)))"
+        apply (rule_tac x = 0 in exI)
+        by auto
+    qed
+  qed
+qed
+
+lemma hoare_move_left_gen[step]: 
+  assumes "p = q"
+  shows "\<lbrace>st i ** ps p ** tm q v2\<rbrace> 
+            i:[move_left]:j
+         \<lbrace>st j ** ps (p - 1) **  tm q v2\<rbrace>"
+  by (unfold assms, rule hoare_move_left)
+
+definition "move_right = (TL exit . \<guillemotright>((R, exit), (R, exit)); TLabel exit)"
+
+lemma hoare_move_right: 
+  "\<lbrace>st i ** ps p ** tm p v1 \<rbrace> 
+     i:[move_right]:j
+   \<lbrace>st j ** ps (p + 1) ** tm p v1 \<rbrace>"
+proof(unfold move_right_def, intro t_hoare_local, rule t_hoare_label_last, simp+)
+  fix l
+  show "\<lbrace>st i \<and>* ps p \<and>* tm p v1\<rbrace>  i :[ \<guillemotright> ((R, l), R, l) ]: l
+        \<lbrace>st l \<and>* ps (p + 1) \<and>* tm p v1\<rbrace>"
+  proof(unfold tassemble_to.simps, simp only:sep_conj_cond, 
+      intro tm.code_condI, simp add: sep_conj_ac)
+    let ?j = "Suc i"
+    show "\<lbrace>ps p \<and>* st i \<and>* tm p v1\<rbrace>  sg {TC i ((R, ?j), R, ?j)} 
+          \<lbrace>st ?j \<and>* tm p v1 \<and>* ps (p + 1)\<rbrace>"
+    proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify)
+      fix ft prog cs i' mem r
+      assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v1 \<and>* sg {TC i ((R, ?j), R, ?j)}) 
+                       (trset_of (ft, prog, cs, i',  mem))"
+      from h have "prog i = Some ((R, ?j), R, ?j)"
+        apply(rule_tac r = "r \<and>* ps p \<and>* tm p v1" in codeD)
+        by(simp add: sep_conj_ac)
+      from h and this have stp:
+        "tm.run 1 (ft, prog, cs, i', mem) = (ft, prog, ?j, i'+ 1, mem)" (is "?x = ?y")
+        apply(sep_frule psD)
+        apply(sep_frule stD)
+        apply(sep_frule memD, simp)
+        apply(unfold tm.run_def, case_tac v1, auto)
+        done
+      from h have "i' = p"
+        by(sep_drule psD, simp)
+      with h
+      have "(r \<and>* st ?j \<and>* tm p v1 \<and>* ps (p + 1) \<and>* 
+                sg {TC i ((R, ?j), R, ?j)}) (trset_of ?x)"
+        apply(unfold stp)
+        apply(sep_drule pos_upd, sep_drule st_upd, simp)
+      proof -
+        assume "(st ?j \<and>* ps (p + 1) \<and>* r \<and>* tm p v1 \<and>* sg {TC i ((R, ?j), R, ?j)}) 
+                   (trset_of (ft, prog, ?j, p + 1, mem))"
+        thus "(r \<and>* st ?j \<and>* tm p v1 \<and>* ps (p + 1) \<and>* sg {TC i ((R, ?j), R, ?j)}) 
+                    (trset_of (ft, prog, ?j, p + 1, mem))"
+          by (simp add: sep_conj_ac)
+      qed
+      thus "\<exists>k. (r \<and>* st ?j \<and>* tm p v1 \<and>* ps (p + 1) \<and>* sg {TC i ((R, ?j), R, ?j)}) 
+             (trset_of (tm.run (Suc k) (ft, prog, cs, i', mem)))"
+        apply (rule_tac x = 0 in exI)
+        by auto
+    qed
+  qed
+qed
+
+lemma hoare_move_right_gen[step]: 
+  assumes "p = q"
+  shows "\<lbrace>st i ** ps p ** tm q v1 \<rbrace> 
+           i:[move_right]:j
+         \<lbrace>st j ** ps (p + 1) ** tm q v1 \<rbrace>"
+  by (unfold assms, rule hoare_move_right)
+
+definition "if_one e = (TL exit . \<guillemotright>((W0, exit), (W1, e)); TLabel exit)"
+
+lemma hoare_if_one_true: 
+  "\<lbrace>st i ** ps p ** one p\<rbrace> 
+     i:[if_one e]:j
+   \<lbrace>st e ** ps p ** one p\<rbrace>"
+proof(unfold if_one_def, intro t_hoare_local, rule t_hoare_label_last, simp+)
+  fix l
+  show " \<lbrace>st i \<and>* ps p \<and>* one p\<rbrace>  i :[ \<guillemotright> ((W0, l), W1, e) ]: l \<lbrace>st e \<and>* ps p \<and>* one p\<rbrace>"
+  proof(unfold tassemble_to.simps, simp only:sep_conj_cond,
+        intro tm.code_condI, simp add: sep_conj_ac)
+    let ?j = "Suc i"
+    show "\<lbrace>one p \<and>* ps p \<and>* st i\<rbrace>  sg {TC i ((W0, ?j), W1, e)} \<lbrace>one p \<and>* ps p \<and>* st e\<rbrace>"
+    proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify)
+      fix ft prog cs pc mem r
+      assume h: "(r \<and>* one p \<and>* ps p \<and>* st i \<and>* sg {TC i ((W0, ?j), W1, e)}) 
+        (trset_of (ft, prog, cs, pc, mem))"
+      from h have k1: "prog i = Some ((W0, ?j), W1, e)"
+        apply(rule_tac r = "r \<and>* one p \<and>* ps p" in codeD)
+        by(simp add: sep_conj_ac)
+      from h have k2: "pc = p"
+        by(sep_drule psD, simp)
+      from h and this have k3: "mem pc = Some Oc"
+        apply(unfold one_def)
+        by(sep_drule memD, simp)
+      from h k1 k2 k3 have stp:
+        "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, e, pc, mem)" (is "?x = ?y")
+        apply(sep_drule stD)
+        by(unfold tm.run_def, auto)
+      from h k2 
+      have "(r \<and>* one p \<and>* ps p \<and>* st e \<and>* sg {TC i ((W0, ?j), W1, e)})  (trset_of ?x)"
+        apply(unfold stp)
+        by(sep_drule st_upd, simp add: sep_conj_ac)
+      thus "\<exists>k.(r \<and>* one p \<and>* ps p \<and>* st e \<and>* sg {TC i ((W0, ?j), W1, e)})
+             (trset_of (tm.run (Suc k) (ft, prog, cs, pc, mem)))"
+        apply (rule_tac x = 0 in exI)
+        by auto
+    qed
+  qed
+qed
+
+text {*
+  The following problematic lemma is not provable now 
+  lemma hoare_self: "\<lbrace>p\<rbrace> i :[ap]: j \<lbrace>p\<rbrace>" 
+  apply(simp add: tm.Hoare_gen_def, clarify)
+  apply(rule_tac x = 0 in exI, simp add: tm.run_def)
+  done
+*}
+
+lemma hoare_if_one_true_gen[step]: 
+  assumes "p = q"
+  shows
+  "\<lbrace>st i ** ps p ** one q\<rbrace> 
+     i:[if_one e]:j
+   \<lbrace>st e ** ps p ** one q\<rbrace>"
+  by (unfold assms, rule hoare_if_one_true)
+
+lemma hoare_if_one_true1: 
+  "\<lbrace>st i ** ps p ** one p\<rbrace> 
+     i:[(if_one e; c)]:j
+   \<lbrace>st e ** ps p ** one p\<rbrace>"
+proof(unfold tassemble_to.simps, rule tm.code_exI, 
+       simp add: sep_conj_ac tm.Hoare_gen_def, clarify)  
+  fix j' ft prog cs pos mem r
+  assume h: "(r \<and>* one p \<and>* ps p \<and>* st i \<and>* j' :[ c ]: j \<and>* i :[ if_one e ]: j') 
+    (trset_of (ft, prog, cs, pos, mem))"
+  from tm.frame_rule[OF hoare_if_one_true]
+  have "\<And> r. \<lbrace>st i \<and>* ps p \<and>* one p \<and>* r\<rbrace>  i :[ if_one e ]: j' \<lbrace>st e \<and>* ps p \<and>* one p \<and>* r\<rbrace>"
+    by(simp add: sep_conj_ac)
+  from this[unfolded tm.Hoare_gen_def tassemble_to.simps, rule_format, of "j' :[ c ]: j"] h
+  have "\<exists> k. (r \<and>* one p \<and>* ps p \<and>* st e \<and>* i :[ if_one e ]: j' \<and>* j' :[ c ]: j)
+    (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))"
+    by(auto simp: sep_conj_ac)
+  thus "\<exists>k. (r \<and>* one p \<and>* ps p \<and>* st e \<and>* j' :[ c ]: j \<and>* i :[ if_one e ]: j') 
+    (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))"
+    by(simp add: sep_conj_ac)
+qed
+
+lemma hoare_if_one_true1_gen[step]: 
+  assumes "p = q"
+  shows
+  "\<lbrace>st i ** ps p ** one q\<rbrace> 
+     i:[(if_one e; c)]:j
+   \<lbrace>st e ** ps p ** one q\<rbrace>"
+  by (unfold assms, rule hoare_if_one_true1)
+
+lemma hoare_if_one_false: 
+  "\<lbrace>st i ** ps p ** zero p\<rbrace> 
+       i:[if_one e]:j
+   \<lbrace>st j ** ps p ** zero p\<rbrace>"
+proof(unfold if_one_def, intro t_hoare_local, rule t_hoare_label_last, simp+)
+  show "\<lbrace>st i \<and>* ps p \<and>* zero p\<rbrace>  i :[ (\<guillemotright> ((W0, j), W1, e)) ]: j
+        \<lbrace>st j \<and>* ps p \<and>* zero p\<rbrace>"
+  proof(unfold tassemble_to.simps, simp only:sep_conj_cond,
+        intro tm.code_condI, simp add: sep_conj_ac)
+    let ?j = "Suc i"
+    show "\<lbrace>ps p \<and>* st i \<and>* zero p\<rbrace>  sg {TC i ((W0, ?j), W1, e)} \<lbrace>ps p \<and>*  zero p \<and>* st ?j \<rbrace>"
+    proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify)
+      fix ft prog cs pc mem r
+      assume h: "(r \<and>* ps p \<and>* st i \<and>* zero p \<and>* sg {TC i ((W0, ?j), W1, e)})
+        (trset_of (ft, prog, cs, pc, mem))"
+      from h have k1: "prog i = Some ((W0, ?j), W1, e)"
+        apply(rule_tac r = "r \<and>* zero p \<and>* ps p" in codeD)
+        by(simp add: sep_conj_ac)
+      from h have k2: "pc = p"
+        by(sep_drule psD, simp)
+      from h and this have k3: "mem pc = Some Bk"
+        apply(unfold zero_def)
+        by(sep_drule memD, simp)
+      from h k1 k2 k3 have stp:
+        "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, ?j, pc, mem)" (is "?x = ?y")
+        apply(sep_drule stD)
+        by(unfold tm.run_def, auto)
+      from h k2 
+      have "(r \<and>* zero p \<and>* ps p \<and>* st ?j \<and>* sg {TC i ((W0, ?j), W1, e)})  (trset_of ?x)"
+        apply (unfold stp)
+        by (sep_drule st_upd[where i''="?j"], auto simp:sep_conj_ac)
+      thus "\<exists>k. (r \<and>* ps p \<and>* zero p \<and>* st ?j \<and>*  sg {TC i ((W0, ?j), W1, e)})
+             (trset_of (tm.run (Suc k) (ft, prog, cs, pc, mem)))"
+        by(auto simp: sep_conj_ac)
+    qed
+  qed
+qed
+
+lemma hoare_if_one_false_gen[step]: 
+  assumes "p = q"
+  shows "\<lbrace>st i ** ps p ** zero q\<rbrace> 
+             i:[if_one e]:j
+         \<lbrace>st j ** ps p ** zero q\<rbrace>"
+  by (unfold assms, rule hoare_if_one_false)
+
+definition "if_zero e = (TL exit . \<guillemotright>((W0, e), (W1, exit)); TLabel exit)"
+
+lemma hoare_if_zero_true: 
+  "\<lbrace>st i ** ps p ** zero p\<rbrace> 
+     i:[if_zero e]:j
+   \<lbrace>st e ** ps p ** zero p\<rbrace>"
+proof(unfold if_zero_def, intro t_hoare_local, rule t_hoare_label_last, simp+)
+  fix l
+  show "\<lbrace>st i \<and>* ps p \<and>* zero p\<rbrace>  i :[ \<guillemotright> ((W0, e), W1, l) ]: l \<lbrace>st e \<and>* ps p \<and>* zero p\<rbrace>"
+  proof(unfold tassemble_to.simps, simp only:sep_conj_cond,
+        intro tm.code_condI, simp add: sep_conj_ac)
+    let ?j = "Suc i"
+    show "\<lbrace>ps p \<and>* st i \<and>* zero p\<rbrace>  sg {TC i ((W0, e), W1, ?j)} \<lbrace>ps p \<and>* st e \<and>* zero p\<rbrace>"
+    proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify)
+      fix ft prog cs pc mem r
+      assume h: "(r \<and>* ps p \<and>* st i \<and>* zero p \<and>* sg {TC i ((W0, e), W1, ?j)})
+        (trset_of (ft, prog, cs, pc, mem))"
+      from h have k1: "prog i = Some ((W0, e), W1, ?j)"
+        apply(rule_tac r = "r \<and>* zero p \<and>* ps p" in codeD)
+        by(simp add: sep_conj_ac)
+      from h have k2: "pc = p"
+        by(sep_drule psD, simp)
+      from h and this have k3: "mem pc = Some Bk"
+        apply(unfold zero_def)
+        by(sep_drule memD, simp)
+      from h k1 k2 k3 have stp:
+        "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, e, pc, mem)" (is "?x = ?y")
+        apply(sep_drule stD)
+        by(unfold tm.run_def, auto)
+      from h k2 
+      have "(r \<and>* zero p \<and>* ps p \<and>* st e \<and>* sg {TC i ((W0, e), W1, ?j)})  (trset_of ?x)"
+        apply(unfold stp)
+        by(sep_drule st_upd, simp add: sep_conj_ac)
+      thus "\<exists>k. (r \<and>* ps p \<and>* st e \<and>* zero p \<and>* sg {TC i ((W0, e), W1, ?j)})
+             (trset_of (tm.run (Suc k) (ft, prog, cs, pc, mem)))"
+        by(auto simp: sep_conj_ac)
+    qed
+  qed
+qed
+
+lemma hoare_if_zero_true_gen[step]: 
+  assumes "p = q"
+  shows
+  "\<lbrace>st i ** ps p ** zero q\<rbrace> 
+     i:[if_zero e]:j
+   \<lbrace>st e ** ps p ** zero q\<rbrace>"
+  by (unfold assms, rule hoare_if_zero_true)
+
+lemma hoare_if_zero_true1: 
+  "\<lbrace>st i ** ps p ** zero p\<rbrace> 
+     i:[(if_zero e; c)]:j
+   \<lbrace>st e ** ps p ** zero p\<rbrace>"
+ proof(unfold tassemble_to.simps, rule tm.code_exI, simp add: sep_conj_ac 
+        tm.Hoare_gen_def, clarify)  
+  fix j' ft prog cs pos mem r
+  assume h: "(r \<and>* ps p \<and>* st i \<and>* zero p \<and>* j' :[ c ]: j \<and>* i :[ if_zero e ]: j') 
+    (trset_of (ft, prog, cs, pos, mem))"
+  from tm.frame_rule[OF hoare_if_zero_true]
+  have "\<And> r. \<lbrace>st i \<and>* ps p \<and>* zero p \<and>* r\<rbrace>  i :[ if_zero e ]: j' \<lbrace>st e \<and>* ps p \<and>* zero p \<and>* r\<rbrace>"
+    by(simp add: sep_conj_ac)
+  from this[unfolded tm.Hoare_gen_def tassemble_to.simps, rule_format, of "j' :[ c ]: j"] h
+  have "\<exists> k. (r \<and>* zero p \<and>* ps p \<and>* st e \<and>* i :[ if_zero e ]: j' \<and>* j' :[ c ]: j)
+    (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))"
+    by(auto simp: sep_conj_ac)
+  thus "\<exists>k. (r \<and>* ps p \<and>* st e \<and>* zero p \<and>* j' :[ c ]: j \<and>* i :[ if_zero e ]: j')  
+    (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))"
+    by(simp add: sep_conj_ac)
+qed
+
+lemma hoare_if_zero_true1_gen[step]: 
+  assumes "p = q"
+  shows
+  "\<lbrace>st i ** ps p ** zero q\<rbrace> 
+     i:[(if_zero e; c)]:j
+   \<lbrace>st e ** ps p ** zero q\<rbrace>"
+  by (unfold assms, rule hoare_if_zero_true1)
+
+lemma hoare_if_zero_false: 
+  "\<lbrace>st i ** ps p ** tm p Oc\<rbrace> 
+     i:[if_zero e]:j
+   \<lbrace>st j ** ps p ** tm p Oc\<rbrace>"
+proof(unfold if_zero_def, intro t_hoare_local, rule t_hoare_label_last, simp, simp)
+  fix l
+  show "\<lbrace>st i \<and>* ps p \<and>* tm p Oc\<rbrace>  i :[ \<guillemotright> ((W0, e), W1, l) ]: l
+        \<lbrace>st l \<and>* ps p \<and>* tm p Oc\<rbrace>"
+  proof(unfold tassemble_to.simps, simp only:sep_conj_cond,
+      intro tm.code_condI, simp add: sep_conj_ac)
+    let ?j = "Suc i"
+    show "\<lbrace>ps p \<and>* st i \<and>* tm p Oc\<rbrace>  sg {TC i ((W0, e), W1, ?j)} 
+          \<lbrace>ps p \<and>* st ?j \<and>* tm p Oc\<rbrace>"
+    proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify)
+      fix ft prog cs pc mem r
+      assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p Oc \<and>* sg {TC i ((W0, e), W1, ?j)})
+        (trset_of (ft, prog, cs, pc, mem))"
+      from h have k1: "prog i = Some ((W0, e), W1, ?j)"
+        apply(rule_tac r = "r \<and>* tm p Oc \<and>* ps p" in codeD)
+        by(simp add: sep_conj_ac)
+      from h have k2: "pc = p"
+        by(sep_drule psD, simp)
+      from h and this have k3: "mem pc = Some Oc"
+        by(sep_drule memD, simp)
+      from h k1 k2 k3 have stp:
+        "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, ?j, pc, mem)" (is "?x = ?y")
+        apply(sep_drule stD)
+        by(unfold tm.run_def, auto)
+      from h k2 
+      have "(r \<and>* tm p Oc \<and>* ps p \<and>* st ?j \<and>* sg {TC i ((W0, e), W1, ?j)})  (trset_of ?x)"
+        apply(unfold stp)
+        by(sep_drule st_upd, simp add: sep_conj_ac)
+      thus "\<exists>k. (r \<and>* ps p \<and>* st ?j \<and>* tm p Oc \<and>* sg {TC i ((W0, e), W1, ?j)})
+             (trset_of (tm.run (Suc k) (ft, prog, cs, pc, mem)))"
+        by(auto simp: sep_conj_ac)
+    qed
+  qed
+qed
+
+lemma hoare_if_zero_false_gen[step]: 
+  assumes "p = q"
+  shows
+  "\<lbrace>st i ** ps p ** tm q Oc\<rbrace> 
+     i:[if_zero e]:j
+   \<lbrace>st j ** ps p ** tm q Oc\<rbrace>"
+  by (unfold assms, rule hoare_if_zero_false)
+
+
+definition "jmp e = \<guillemotright>((W0, e), (W1, e))"
+
+lemma hoare_jmp: 
+  "\<lbrace>st i \<and>* ps p \<and>* tm p v\<rbrace>  i:[jmp e]:j \<lbrace>st e \<and>* ps p \<and>* tm p v\<rbrace>"
+proof(unfold jmp_def tm.Hoare_gen_def tassemble_to.simps sep_conj_ac, clarify)
+  fix ft prog cs pos mem r
+  assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v \<and>* <(j = i + 1)> \<and>* sg {TC i ((W0, e), W1, e)})
+    (trset_of (ft, prog, cs, pos, mem))"
+  from h have k1: "prog i = Some ((W0, e), W1, e)"
+    apply(rule_tac r = "r \<and>* <(j = i + 1)> \<and>* tm p v \<and>* ps p" in codeD)
+    by(simp add: sep_conj_ac)
+  from h have k2: "p = pos"
+    by(sep_drule psD, simp)
+  from h k2 have k3: "mem pos = Some v"
+    by(sep_drule memD, simp)
+  from h k1 k2 k3 have 
+    stp: "tm.run 1 (ft, prog, cs, pos, mem) = (ft, prog, e, pos, mem)" (is "?x = ?y")
+    apply(sep_drule stD)
+    by(unfold tm.run_def, cases "mem pos", simp, cases v, auto)
+  from h k2 
+  have "(r \<and>* ps p \<and>* st e \<and>* tm p v \<and>* <(j = i + 1)> \<and>* 
+           sg {TC i ((W0, e), W1, e)}) (trset_of ?x)"
+    apply(unfold stp)
+    by(sep_drule st_upd, simp add: sep_conj_ac)
+  thus "\<exists> k. (r \<and>* ps p \<and>* st e \<and>* tm p v \<and>* <(j = i + 1)> \<and>* sg {TC i ((W0, e), W1, e)})
+             (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))"
+    apply (rule_tac x = 0 in exI)
+    by auto
+qed
+
+lemma hoare_jmp_gen[step]: 
+  assumes "p = q"
+  shows "\<lbrace>st i \<and>* ps p \<and>* tm q v\<rbrace>  i:[jmp e]:j \<lbrace>st e \<and>* ps p \<and>* tm q v\<rbrace>"
+  by (unfold assms, rule hoare_jmp)
+
+lemma hoare_jmp1: 
+  "\<lbrace>st i \<and>* ps p \<and>* tm p v\<rbrace> 
+     i:[(jmp e; c)]:j
+   \<lbrace>st e \<and>* ps p \<and>* tm p v\<rbrace>"
+proof(unfold  tassemble_to.simps, rule tm.code_exI, simp 
+              add: sep_conj_ac tm.Hoare_gen_def, clarify)
+  fix j' ft prog cs pos mem r
+  assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v \<and>* j' :[ c ]: j \<and>* i :[ jmp e ]: j') 
+    (trset_of (ft, prog, cs, pos, mem))"
+  from tm.frame_rule[OF hoare_jmp]
+  have "\<And> r. \<lbrace>st i \<and>* ps p \<and>* tm p v \<and>* r\<rbrace>  i :[ jmp e ]: j' \<lbrace>st e \<and>* ps p \<and>* tm p v \<and>* r\<rbrace>"
+    by(simp add: sep_conj_ac)
+  from this[unfolded tm.Hoare_gen_def tassemble_to.simps, rule_format, of "j' :[ c ]: j"] h
+  have "\<exists> k. (r \<and>* tm p v \<and>* ps p \<and>* st e \<and>* i :[ jmp e ]: j' \<and>* j' :[ c ]: j)
+    (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))"
+    by(auto simp: sep_conj_ac)
+  thus "\<exists>k. (r \<and>* ps p \<and>* st e \<and>* tm p v \<and>* j' :[ c ]: j \<and>* i :[ jmp e ]: j')  
+    (trset_of (tm.run (Suc k) (ft, prog, cs, pos, mem)))"
+    by(simp add: sep_conj_ac)
+qed
+
+
+lemma hoare_jmp1_gen[step]: 
+  assumes "p = q"
+  shows "\<lbrace>st i \<and>* ps p \<and>* tm q v\<rbrace> 
+            i:[(jmp e; c)]:j
+         \<lbrace>st e \<and>* ps p \<and>* tm q v\<rbrace>"
+  by (unfold assms, rule hoare_jmp1)
+
+
+lemma condI: 
+  assumes h1: b
+  and h2: "b \<Longrightarrow> p s"
+  shows "(<b> \<and>* p) s"
+  by (metis (full_types) cond_true_eq1 h1 h2)
+
+lemma condE:
+  assumes "(<b> \<and>* p) s"
+  obtains "b" and "p s"
+proof(atomize_elim)
+  from condD[OF assms]
+  show "b \<and> p s" .
+qed
+
+
+section {* Tactics *}
+
+ML {*
+  val trace_step = Attrib.setup_config_bool @{binding trace_step} (K false)
+  val trace_fwd = Attrib.setup_config_bool @{binding trace_fwd} (K false)
+*}
+
+
+ML {*
+  val tracing  = (fn ctxt => fn str =>
+                   if (Config.get ctxt trace_step) then tracing str else ())
+  fun not_pred p = fn s => not (p s)
+  fun break_sep_conj (Const (@{const_name sep_conj},_) $ t1 $ t2 $ _) =
+         (break_sep_conj t1) @ (break_sep_conj t2)
+    | break_sep_conj (Const (@{const_name sep_conj},_) $ t1 $ t2) =
+            (break_sep_conj t1) @ (break_sep_conj t2)
+                   (* dig through eta exanded terms: *)
+    | break_sep_conj (Abs (_, _, t $ Bound 0)) = break_sep_conj t
+    | break_sep_conj t = [t];
+
+  val empty_env = (Vartab.empty, Vartab.empty)
+
+  fun match_env ctxt pat trm env = 
+            Pattern.match (ctxt |> Proof_Context.theory_of) (pat, trm) env
+
+  fun match ctxt pat trm = match_env ctxt pat trm empty_env;
+
+  val inst = Envir.subst_term;
+
+  fun term_of_thm thm = thm |>  prop_of |> HOLogic.dest_Trueprop
+
+  fun get_cmd ctxt code = 
+      let val pat = term_of @{cpat "_:[(?cmd)]:_"}
+          val pat1 = term_of @{cpat "?cmd::tpg"}
+          val env = match ctxt pat code
+      in inst env pat1 end
+
+  fun is_seq_term (Const (@{const_name TSeq}, _) $ _ $ _) = true
+    | is_seq_term _ = false
+
+  fun get_hcmd  (Const (@{const_name TSeq}, _) $ hcmd $ _) = hcmd
+    | get_hcmd hcmd = hcmd
+
+  fun last [a]  = a |
+      last (a::b) = last b
+
+  fun but_last [a] = [] |
+      but_last (a::b) = a::(but_last b)
+
+  fun foldr f [] = (fn x => x) |
+      foldr f (x :: xs) = (f x) o  (foldr f xs)
+
+  fun concat [] = [] |
+      concat (x :: xs) = x @ concat xs
+
+  fun match_any ctxt pats tm = 
+              fold 
+                 (fn pat => fn b => (b orelse Pattern.matches 
+                          (ctxt |> Proof_Context.theory_of) (pat, tm))) 
+                 pats false
+
+  fun is_ps_term (Const (@{const_name ps}, _) $ _) = true
+    | is_ps_term _ = false
+
+  fun string_of_term ctxt t = t |> Syntax.pretty_term ctxt |> Pretty.str_of
+  fun string_of_cterm ctxt ct = ct |> term_of |> string_of_term ctxt
+  fun pterm ctxt t =
+          t |> string_of_term ctxt |> tracing ctxt
+  fun pcterm ctxt ct = ct |> string_of_cterm ctxt |> tracing ctxt
+  fun string_for_term ctxt t =
+       Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN)
+                   (print_mode_value ())) (Syntax.string_of_term ctxt) t
+         |> String.translate (fn c => if Char.isPrint c then str c else "")
+         |> Sledgehammer_Util.simplify_spaces  
+  fun string_for_cterm ctxt ct = ct |> term_of |> string_for_term ctxt
+  fun attemp tac = fn i => fn st => (tac i st) handle exn => Seq.empty
+  fun try_tac tac = fn i => fn st => (tac i st) handle exn => (Seq.single st)       
+ (* aux end *) 
+*}
+
+ML {* (* Functions specific to Hoare triples *)
+  fun get_pre ctxt t = 
+    let val pat = term_of @{cpat "\<lbrace>?P\<rbrace> ?c \<lbrace>?Q\<rbrace>"} 
+        val env = match ctxt pat t 
+    in inst env (term_of @{cpat "?P::tresource set \<Rightarrow> bool"}) end
+
+  fun can_process ctxt t = ((get_pre ctxt t; true) handle _ => false)
+
+  fun get_post ctxt t = 
+    let val pat = term_of @{cpat "\<lbrace>?P\<rbrace> ?c \<lbrace>?Q\<rbrace>"} 
+        val env = match ctxt pat t 
+    in inst env (term_of @{cpat "?Q::tresource set \<Rightarrow> bool"}) end;
+
+  fun get_mid ctxt t = 
+    let val pat = term_of @{cpat "\<lbrace>?P\<rbrace> ?c \<lbrace>?Q\<rbrace>"} 
+        val env = match ctxt pat t 
+    in inst env (term_of @{cpat "?c::tresource set \<Rightarrow> bool"}) end;
+
+  fun is_pc_term (Const (@{const_name st}, _) $ _) = true
+    | is_pc_term _ = false
+
+  fun mk_pc_term x =
+     Const (@{const_name st}, @{typ "nat \<Rightarrow> tresource set \<Rightarrow> bool"}) $ Free (x, @{typ "nat"})
+
+  val sconj_term = term_of @{cterm "sep_conj::tassert \<Rightarrow> tassert \<Rightarrow> tassert"}
+
+  fun mk_ps_term x =
+     Const (@{const_name ps}, @{typ "int \<Rightarrow> tresource set \<Rightarrow> bool"}) $ Free (x, @{typ "int"})
+
+  fun atomic tac  = ((SOLVED' tac) ORELSE' (K all_tac))
+
+  fun pure_sep_conj_ac_tac ctxt = 
+         (auto_tac (ctxt |> Simplifier.map_simpset (fn ss => ss addsimps @{thms sep_conj_ac}))
+          |> SELECT_GOAL)
+
+
+  fun potential_facts ctxt prop = Facts.could_unify (Proof_Context.facts_of ctxt) 
+                                       ((Term.strip_all_body prop) |> Logic.strip_imp_concl);
+
+  fun some_fact_tac ctxt = SUBGOAL (fn (goal, i) => 
+                                      (Method.insert_tac (potential_facts ctxt goal) i) THEN
+                                      (pure_sep_conj_ac_tac ctxt i));
+
+  fun sep_conj_ac_tac ctxt = 
+     (SOLVED' (auto_tac (ctxt |> Simplifier.map_simpset (fn ss => ss addsimps @{thms sep_conj_ac}))
+       |> SELECT_GOAL)) ORELSE' (atomic (some_fact_tac ctxt))
+*}
+
+ML {*
+type HoareTriple = {
+  binding: binding,
+  can_process: Proof.context -> term -> bool,
+  get_pre: Proof.context -> term -> term,
+  get_mid: Proof.context -> term -> term,
+  get_post: Proof.context -> term -> term,
+  is_pc_term: term -> bool,
+  mk_pc_term: string -> term,
+  sconj_term: term,
+  sep_conj_ac_tac: Proof.context -> int -> tactic,
+  hoare_seq1: thm,
+  hoare_seq2: thm,
+  pre_stren: thm,
+  post_weaken: thm,
+  frame_rule: thm
+}
+
+  val tm_triple = {binding = @{binding "tm_triple"}, 
+                   can_process = can_process,
+                   get_pre = get_pre,
+                   get_mid = get_mid,
+                   get_post = get_post,
+                   is_pc_term = is_pc_term,
+                   mk_pc_term = mk_pc_term,
+                   sconj_term = sconj_term,
+                   sep_conj_ac_tac = sep_conj_ac_tac,
+                   hoare_seq1 = @{thm t_hoare_seq1},
+                   hoare_seq2 = @{thm t_hoare_seq2},
+                   pre_stren = @{thm tm.pre_stren},
+                   post_weaken = @{thm tm.post_weaken},
+                   frame_rule = @{thm tm.frame_rule}
+                  }:HoareTriple
+*}
+
+ML {*
+  val _ = data_slot "HoareTriples" "HoareTriple list" "[]"
+*}
+
+ML {*
+  val _ = HoareTriples_store [tm_triple]
+*}
+
+ML {* (* aux1 functions *)
+
+fun focus_params t ctxt =
+  let
+    val (xs, Ts) =
+      split_list (Term.variant_frees t (Term.strip_all_vars t));  (*as they are printed :-*)
+    (* val (xs', ctxt') = variant_fixes xs ctxt; *)
+    (* val ps = xs' ~~ Ts; *)
+    val ps = xs ~~ Ts
+    val (_, ctxt'') = ctxt |> Variable.add_fixes xs
+  in ((xs, ps), ctxt'') end
+
+fun focus_concl ctxt t =
+  let
+    val ((xs, ps), ctxt') = focus_params t ctxt
+    val t' = Term.subst_bounds (rev (map Free ps), Term.strip_all_body t);
+  in (t' |> Logic.strip_imp_concl, ctxt') end
+
+  fun get_concl ctxt (i, state) = 
+              nth (Thm.prems_of state) (i - 1) 
+                            |> focus_concl ctxt |> (fn (x, _) => x |> HOLogic.dest_Trueprop)
+ (* aux1 end *)
+*}
+
+ML {*
+  fun indexing xs = upto (0, length xs - 1) ~~ xs
+  fun select_idxs idxs ps = 
+      map_index (fn (i, e) => if (member (op =) idxs i) then [e] else []) ps |> flat
+  fun select_out_idxs idxs ps = 
+      map_index (fn (i, e) => if (member (op =) idxs i) then [] else [e]) ps |> flat
+  fun match_pres ctxt mf env ps qs = 
+      let  fun sel_match mf env [] qs = [(env, [])]
+             | sel_match mf env (p::ps) qs = 
+                  let val pm = map (fn (i, q) => [(i, 
+                                      let val _ = tracing ctxt "Matching:"
+                                          val _ = [p, q] |>
+                                            (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt
+                                          val r = mf p q env 
+                                      in r end)]
+                                      handle _ => (
+                                      let val _ = tracing ctxt "Failed matching:"
+                                          val _ = [p, q] |>
+                                            (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt
+                                      in [] end)) qs |> flat
+                      val r = pm |> map (fn (i, env') => 
+                                let val qs' = filter_out (fn (j, q) => j = i) qs
+                                in  sel_match mf env' ps qs' |> 
+                                      map (fn (env'', idxs) => (env'', i::idxs)) end) 
+                        |> flat
+            in r end
+   in sel_match mf env ps (indexing qs) end
+
+  fun provable tac ctxt goal = 
+          let 
+              val _ = tracing ctxt "Provable trying to prove:"
+              val _ = [goal] |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt
+          in
+             (Goal.prove ctxt [] [] goal (fn {context, ...} => tac context 1); true)
+                        handle exn => false
+          end
+  fun make_sense tac ctxt thm_assms env  = 
+                thm_assms |>  map (inst env) |> forall (provable tac ctxt)
+*}
+
+ML {*
+  fun triple_for ctxt goal = 
+    filter (fn trpl => (#can_process trpl) ctxt goal) (HoareTriples.get (Proof_Context.theory_of ctxt)) |> hd
+
+  fun step_terms_for thm goal ctxt = 
+    let
+       val _ = tracing ctxt "This is the new version of step_terms_for!"
+       val _ = tracing ctxt "Tring to find triple processor: TP"
+       val TP = triple_for ctxt goal
+       val _ = #binding TP |> Binding.name_of |> tracing ctxt
+       fun mk_sep_conj tms = foldr (fn tm => fn rtm => 
+              ((#sconj_term TP)$tm$rtm)) (but_last tms) (last tms)
+       val thm_concl = thm |> prop_of 
+                 |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop 
+       val thm_assms = thm |> prop_of 
+           |> Logic.strip_imp_prems 
+       val cmd_pat = thm_concl |> #get_mid TP ctxt |> get_cmd ctxt 
+       val cmd = goal |> #get_mid TP ctxt |> get_cmd ctxt
+       val _ = tracing ctxt "matching command ... "
+       val _ = tracing ctxt "cmd_pat = "
+       val _ = pterm ctxt cmd_pat
+       val (hcmd, env1, is_last) =  (cmd, match ctxt cmd_pat cmd, true)
+             handle exn => (cmd |> get_hcmd, match ctxt cmd_pat (cmd |> get_hcmd), false)
+       val _ = tracing ctxt "hcmd ="
+       val _ = pterm ctxt hcmd
+       val _ = tracing ctxt "match command succeed! "
+       val _ = tracing ctxt "pres ="
+       val pres = goal |> #get_pre TP ctxt |> break_sep_conj 
+       val _ = pres |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt
+       val _ = tracing ctxt "pre_pats ="
+       val pre_pats = thm_concl |> #get_pre TP ctxt |> inst env1 |> break_sep_conj
+       val _ = pre_pats |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt
+       val _ = tracing ctxt "post_pats ="
+       val post_pats = thm_concl |> #get_post TP ctxt |> inst env1 |> break_sep_conj
+       val _ = post_pats |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt
+       val _ = tracing ctxt "Calculating sols"
+       val sols = match_pres ctxt (match_env ctxt) env1 pre_pats pres 
+       val _ = tracing ctxt "End calculating sols, sols ="
+       val _ = tracing ctxt (@{make_string} sols)
+       val _ = tracing ctxt "Calulating env2 and idxs"
+       val (env2, idxs) = filter (fn (env, idxs) => make_sense (#sep_conj_ac_tac TP) 
+                             ctxt thm_assms env) sols |> hd
+       val _ = tracing ctxt "End calculating env2 and idxs"
+       val _ = tracing ctxt "mterms ="
+       val mterms = select_idxs idxs pres |> map (inst env2) 
+       val _ = mterms |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt
+       val _ = tracing ctxt "nmterms = "
+       val nmterms = select_out_idxs idxs pres |> map (inst env2) 
+       val _ = nmterms |> (pretty_terms ctxt) |> Pretty.str_of |> tracing ctxt
+       val pre_cond = pre_pats |> map (inst env2) |> mk_sep_conj
+       val post_cond = post_pats |> map (inst env2) |> mk_sep_conj 
+       val post_cond_npc  = 
+               post_cond |> break_sep_conj |> filter (not_pred (#is_pc_term TP)) 
+               |> (fn x => x @ nmterms) |> mk_sep_conj |> cterm_of (Proof_Context.theory_of ctxt)
+       fun mk_frame cond rest  = 
+             if rest = [] then cond else ((#sconj_term TP)$ cond) $ (mk_sep_conj rest)
+       val pre_cond_frame = mk_frame pre_cond nmterms |> cterm_of (Proof_Context.theory_of ctxt)
+       fun post_cond_frame j' = post_cond |> break_sep_conj |> filter (not_pred (#is_pc_term TP)) 
+               |> (fn x => [#mk_pc_term TP j']@x) |> mk_sep_conj
+               |> (fn x => mk_frame x nmterms)
+               |> cterm_of (Proof_Context.theory_of ctxt)
+       val need_frame = (nmterms <> [])
+    in 
+         (post_cond_npc,
+          pre_cond_frame, 
+          post_cond_frame, need_frame, is_last)       
+    end
+*}
+
+ML {*
+  fun step_tac ctxt thm i state = 
+     let  
+       val _ = tracing ctxt "This is the new version of step_tac"
+       val (goal, ctxt) = nth (Thm.prems_of state) (i - 1) 
+                  |> focus_concl ctxt 
+                  |> (apfst HOLogic.dest_Trueprop)
+       val _ = tracing ctxt "step_tac: goal = "
+       val _ = goal |> pterm ctxt
+       val _ = tracing ctxt "Start to calculate intermediate terms ... "
+       val (post_cond_npc, pre_cond_frame, post_cond_frame, need_frame, is_last) 
+                        = step_terms_for thm goal ctxt
+       val _ = tracing ctxt "Tring to find triple processor: TP"
+       val TP = triple_for ctxt goal
+       val _ = #binding TP |> Binding.name_of |> tracing ctxt
+       fun mk_sep_conj tms = foldr (fn tm => fn rtm => 
+              ((#sconj_term TP)$tm$rtm)) (but_last tms) (last tms)
+       val _ = tracing ctxt "Calculate intermediate terms finished! "
+       val post_cond_npc_str = post_cond_npc |> string_for_cterm ctxt
+       val pre_cond_frame_str = pre_cond_frame |> string_for_cterm ctxt
+       val _ = tracing ctxt "step_tac: post_cond_npc = "
+       val _ = post_cond_npc |> pcterm ctxt
+       val _ = tracing ctxt "step_tac: pre_cond_frame = "
+       val _ = pre_cond_frame |> pcterm ctxt
+       fun tac1 i state = 
+             if is_last then (K all_tac) i state else
+              res_inst_tac ctxt [(("q", 0), post_cond_npc_str)] 
+                                          (#hoare_seq1 TP) i state
+       fun tac2 i state = res_inst_tac ctxt [(("p", 0), pre_cond_frame_str)] 
+                                          (#pre_stren TP) i state
+       fun foc_tac post_cond_frame ctxt i state  =
+           let
+               val goal = get_concl ctxt (i, state)
+               val pc_term = goal |> #get_post TP ctxt |> break_sep_conj 
+                                |> filter (#is_pc_term TP) |> hd
+               val (_$Free(j', _)) = pc_term
+               val psd = post_cond_frame j'
+               val str_psd = psd |> string_for_cterm ctxt
+               val _ = tracing ctxt "foc_tac: psd = "
+               val _ = psd |> pcterm ctxt
+           in 
+               res_inst_tac ctxt [(("q", 0), str_psd)] 
+                                          (#post_weaken TP) i state
+           end
+     val frame_tac = if need_frame then (rtac (#frame_rule TP)) else (K all_tac)
+     val print_tac = if (Config.get ctxt trace_step) then Tactical.print_tac else (K all_tac)
+     val tac = (tac1 THEN' (K (print_tac "tac1 success"))) THEN' 
+               (tac2 THEN' (K (print_tac "tac2 success"))) THEN' 
+               ((foc_tac post_cond_frame ctxt) THEN' (K (print_tac "foc_tac success"))) THEN' 
+               (frame_tac  THEN' (K (print_tac "frame_tac success"))) THEN' 
+               (((rtac thm) THEN_ALL_NEW (#sep_conj_ac_tac TP ctxt)) THEN' (K (print_tac "rtac thm success"))) THEN' 
+               (K (ALLGOALS (atomic (#sep_conj_ac_tac TP ctxt)))) THEN'
+               (* (#sep_conj_ac_tac TP ctxt) THEN' (#sep_conj_ac_tac TP ctxt) THEN'  *)
+               (K prune_params_tac)
+   in 
+        tac i state
+   end
+
+  fun unfold_cell_tac ctxt = (Local_Defs.unfold_tac ctxt @{thms one_def zero_def})
+  fun fold_cell_tac ctxt = (Local_Defs.fold_tac ctxt @{thms one_def zero_def})
+*}
+
+ML {*
+  fun sg_step_tac thms ctxt =
+     let val sg_step_tac' =  (map (fn thm  => attemp (step_tac ctxt thm)) thms)
+                               (* @ [attemp (goto_tac ctxt)]  *)
+                              |> FIRST'
+         val sg_step_tac'' = (K (unfold_cell_tac ctxt)) THEN' sg_step_tac' THEN' (K (fold_cell_tac ctxt))
+     in
+         sg_step_tac' ORELSE' sg_step_tac''
+     end
+  fun steps_tac thms ctxt i = REPEAT (sg_step_tac thms ctxt i) THEN (prune_params_tac)
+*}
+
+method_setup hstep = {* 
+  Attrib.thms >> (fn thms => fn ctxt =>
+                    (SIMPLE_METHOD' (fn i => 
+                       sg_step_tac (thms@(StepRules.get ctxt)) ctxt i)))
+  *} 
+  "One step symbolic execution using step theorems."
+
+method_setup hsteps = {* 
+  Attrib.thms >> (fn thms => fn ctxt =>
+                    (SIMPLE_METHOD' (fn i => 
+                       steps_tac (thms@(StepRules.get ctxt)) ctxt i)))
+  *} 
+  "Sequential symbolic execution using step theorems."
+
+
+ML {*
+  fun goto_tac ctxt thm i state = 
+     let  
+       val (goal, ctxt) = nth (Thm.prems_of state) (i - 1) 
+                             |> focus_concl ctxt |> (apfst HOLogic.dest_Trueprop)
+       val _ = tracing ctxt "goto_tac: goal = "
+       val _ = goal |> string_of_term ctxt |> tracing ctxt
+       val (post_cond_npc, pre_cond_frame, post_cond_frame, need_frame, is_last) 
+                        = step_terms_for thm goal ctxt
+       val _ = tracing ctxt "Tring to find triple processor: TP"
+       val TP = triple_for ctxt goal
+       val _ = #binding TP |> Binding.name_of |> tracing ctxt
+       val _ = tracing ctxt "aaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaaa"
+       val post_cond_npc_str = post_cond_npc |> string_for_cterm ctxt
+       val pre_cond_frame_str = pre_cond_frame |> string_for_cterm ctxt
+       val _ = tracing ctxt "goto_tac: post_cond_npc = "
+       val _ = post_cond_npc_str |> tracing ctxt
+       val _ = tracing ctxt "goto_tac: pre_cond_frame = "
+       val _ = pre_cond_frame_str |> tracing ctxt
+       fun tac1 i state = 
+             if is_last then (K all_tac) i state else
+              res_inst_tac ctxt [] 
+                                          (#hoare_seq2 TP) i state
+       fun tac2 i state = res_inst_tac ctxt [(("p", 0), pre_cond_frame_str)] 
+                                          (#pre_stren TP) i state
+       fun foc_tac post_cond_frame ctxt i state  =
+           let
+               val goal = get_concl ctxt (i, state)
+               val pc_term = goal |> #get_post TP ctxt |> break_sep_conj 
+                                |> filter (#is_pc_term TP) |> hd
+               val (_$Free(j', _)) = pc_term
+               val psd = post_cond_frame j'
+               val str_psd = psd |> string_for_cterm ctxt
+               val _ = tracing ctxt "goto_tac: psd = "
+               val _ = str_psd |> tracing ctxt
+           in 
+               res_inst_tac ctxt [(("q", 0), str_psd)] 
+                                          (#post_weaken TP) i state
+           end
+     val frame_tac = if need_frame then (rtac (#frame_rule TP)) else (K all_tac)
+     val _ = tracing ctxt "goto_tac: starting to apply tacs"
+     val print_tac = if (Config.get ctxt trace_step) then Tactical.print_tac else (K all_tac)
+     val tac = (tac1 THEN' (K (print_tac "tac1 success"))) THEN' 
+               (tac2 THEN' (K (print_tac "tac2 success"))) THEN' 
+               ((foc_tac post_cond_frame ctxt) THEN' (K (print_tac "foc_tac success"))) THEN' 
+               (frame_tac THEN' (K (print_tac "frame_tac success"))) THEN' 
+               ((((rtac thm) THEN_ALL_NEW (#sep_conj_ac_tac TP ctxt))) THEN'
+                 (K (print_tac "rtac success"))
+               ) THEN' 
+               (K (ALLGOALS (atomic (#sep_conj_ac_tac TP ctxt)))) THEN'
+               (K prune_params_tac)
+   in 
+        tac i state
+   end
+*}
+
+ML {*
+  fun sg_goto_tac thms ctxt =
+     let val sg_goto_tac' =  (map (fn thm  => attemp (goto_tac ctxt thm)) thms)
+                              |> FIRST'
+         val sg_goto_tac'' = (K (unfold_cell_tac ctxt)) THEN' sg_goto_tac' THEN' (K (fold_cell_tac ctxt))
+     in
+         sg_goto_tac' ORELSE' sg_goto_tac''
+     end
+  fun gotos_tac thms ctxt i = REPEAT (sg_goto_tac thms ctxt i) THEN (prune_params_tac)
+*}
+
+method_setup hgoto = {* 
+  Attrib.thms >> (fn thms => fn ctxt =>
+                    (SIMPLE_METHOD' (fn i => 
+                       sg_goto_tac (thms@(StepRules.get ctxt)) ctxt i)))
+  *} 
+  "One step symbolic execution using goto theorems."
+
+subsection {* Tactic for forward reasoning *}
+
+ML {*
+fun mk_msel_rule ctxt conclusion idx term =
+let 
+  val cjt_count = term |> break_sep_conj |> length
+  fun variants nctxt names = fold_map Name.variant names nctxt;
+
+  val (state, nctxt0) = Name.variant "s" (Variable.names_of ctxt);
+
+  fun sep_conj_prop cjts =
+        FunApp.fun_app_free
+          (FunApp.fun_app_foldr SepConj.sep_conj_term cjts) state
+        |> HOLogic.mk_Trueprop;
+
+  (* concatenate string and string of an int *)
+  fun conc_str_int str int = str ^ Int.toString int;
+
+  (* make the conjunct names *)
+  val (cjts, _) = ListExtra.range 1 cjt_count
+                  |> map (conc_str_int "a") |> variants nctxt0;
+
+ fun skel_sep_conj names (Const (@{const_name sep_conj}, _) $ t1 $ t2 $ y) =
+     (let val nm1 = take (length (break_sep_conj t1)) names 
+          val nm2 = drop (length (break_sep_conj t1)) names
+          val t1' = skel_sep_conj nm1 t1 
+          val t2' = skel_sep_conj nm2 t2 
+      in (SepConj.sep_conj_term $ t1' $ t2' $ y) end)
+  | skel_sep_conj names (Const (@{const_name sep_conj}, _) $ t1 $ t2) =
+     (let val nm1 = take (length (break_sep_conj t1)) names 
+          val nm2 = drop (length (break_sep_conj t1)) names
+          val t1' = skel_sep_conj nm1 t1 
+          val t2' = skel_sep_conj nm2 t2 
+     in (SepConj.sep_conj_term $ t1' $ t2') end)
+   | skel_sep_conj names (Abs (x, y, t $ Bound 0)) = 
+                  let val t' = (skel_sep_conj names t) 
+                      val ty' = t' |> type_of |> domain_type
+                  in (Abs (x, ty', (t' $ Bound 0))) end
+  | skel_sep_conj names t = Free (hd names, SepConj.sep_conj_term |> type_of |> domain_type);
+  val _ = tracing ctxt "xxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxx"
+  val oskel = skel_sep_conj cjts term;
+  val _ = tracing ctxt "yyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyy"
+  val ttt = oskel |> type_of
+  val _ = tracing ctxt "zzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzzz"
+  val orig = FunApp.fun_app_free oskel state |> HOLogic.mk_Trueprop
+  val _ = tracing ctxt "uuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuuu"
+  val is_selected = member (fn (x, y) => x = y) idx
+  val all_idx = ListExtra.range 0 cjt_count
+  val selected_idx = idx
+  val unselected_idx = filter_out is_selected all_idx
+  val selected = map (nth cjts) selected_idx
+  val unselected = map (nth cjts) unselected_idx
+
+  fun fun_app_foldr f [a,b] = FunApp.fun_app_free (FunApp.fun_app_free f a) b
+  | fun_app_foldr f [a] = Free (a, SepConj.sep_conj_term |> type_of |> domain_type)
+  | fun_app_foldr f (x::xs) = (FunApp.fun_app_free f x) $ (fun_app_foldr f xs)
+  | fun_app_foldr _ _ = raise Fail "fun_app_foldr";
+
+  val reordered_skel = 
+      if unselected = [] then (fun_app_foldr SepConj.sep_conj_term selected)
+          else (SepConj.sep_conj_term $ (fun_app_foldr SepConj.sep_conj_term selected)
+                        $ (fun_app_foldr SepConj.sep_conj_term unselected))
+
+  val reordered =  FunApp.fun_app_free reordered_skel state  |> HOLogic.mk_Trueprop
+  val goal = Logic.mk_implies
+               (if conclusion then (orig, reordered) else (reordered, orig));
+  val rule =
+   Goal.prove ctxt [] [] goal (fn _ => 
+        auto_tac (ctxt |> Simplifier.map_simpset (fn ss => ss addsimps @{thms sep_conj_ac})))
+         |> Drule.export_without_context
+in
+   rule
+end
+*}
+
+lemma fwd_rule: 
+  assumes "\<And> s . U s \<longrightarrow> V s"
+  shows "(U ** RR) s \<Longrightarrow> (V ** RR) s"
+  by (metis assms sep_globalise)
+
+ML {*
+  fun sg_sg_fwd_tac ctxt thm pos i state = 
+  let  
+
+  val tracing  = (fn str =>
+                   if (Config.get ctxt trace_fwd) then Output.tracing str else ())
+  fun pterm t =
+          t |> string_of_term ctxt |> tracing
+  fun pcterm ct = ct |> string_of_cterm ctxt |> tracing
+
+  fun atm thm = 
+  let
+  (* val thm = thm |> Drule.forall_intr_vars *)
+  val res =  thm |> cprop_of |> Object_Logic.atomize
+  val res' = Raw_Simplifier.rewrite_rule [res] thm
+  in res' end
+
+  fun find_idx ctxt pats terms = 
+     let val result = 
+              map (fn pat => (find_index (fn trm => ((match ctxt pat trm; true)
+                                              handle _ => false)) terms)) pats
+     in (assert_all (fn x => x >= 0) result (K "match of precondition failed"));
+         result
+     end
+
+  val goal = nth (Drule.cprems_of state) (i - 1) |> term_of
+  val _ = tracing "goal = "
+  val _ = goal |> pterm
+  
+  val ctxt_orig = ctxt
+
+  val ((ps, goal), ctxt) = Variable.focus goal ctxt_orig
+  
+  val prems = goal |> Logic.strip_imp_prems 
+
+  val cprem = nth prems (pos - 1)
+  val (_ $ (the_prem $ _)) = cprem
+  val cjts = the_prem |> break_sep_conj
+  val thm_prems = thm |> cprems_of |> hd |> Thm.dest_arg |> Thm.dest_fun
+  val thm_assms = thm |> cprems_of |> tl |> map term_of
+  val thm_cjts = thm_prems |> term_of |> break_sep_conj
+  val thm_trm = thm |> prop_of
+
+  val _ = tracing "cjts = "
+  val _ = cjts |> map pterm
+  val _ = tracing "thm_cjts = "
+  val _ = thm_cjts |> map pterm
+
+  val _ = tracing "Calculating sols"
+  val sols = match_pres ctxt (match_env ctxt) empty_env thm_cjts cjts
+  val _ = tracing "End calculating sols, sols ="
+  val _ = tracing (@{make_string} sols)
+  val _ = tracing "Calulating env2 and idxs"
+  val (env2, idx) = filter (fn (env, idxs) => make_sense sep_conj_ac_tac ctxt thm_assms env) sols |> hd
+  val ([thm'_trm], ctxt') = thm_trm |> inst env2 |> single 
+                            |> (fn trms => Variable.import_terms true trms ctxt)
+  val thm'_prem  = Logic.strip_imp_prems thm'_trm |> hd 
+  val thm'_concl = Logic.strip_imp_concl thm'_trm 
+  val thm'_prem = (Goal.prove ctxt' [] [thm'_prem] thm'_concl 
+                  (fn {context, prems = [prem]} =>  
+                      (rtac (prem RS thm)  THEN_ALL_NEW (sep_conj_ac_tac ctxt)) 1))
+  val [thm'] = Variable.export ctxt' ctxt_orig [thm'_prem]
+  val trans_rule = 
+       mk_msel_rule ctxt true idx the_prem
+  val _ = tracing "trans_rule = "
+  val _ = trans_rule |> cprop_of |> pcterm
+  val app_rule = 
+      if (length cjts = length thm_cjts) then thm' else
+       ((thm' |> atm) RS @{thm fwd_rule})
+  val _ = tracing "app_rule = "
+  val _ = app_rule |> cprop_of |> pcterm
+  val print_tac = if (Config.get ctxt trace_fwd) then Tactical.print_tac else (K all_tac)
+  val the_tac = (dtac trans_rule THEN' (K (print_tac "dtac1 success"))) THEN'
+                ((dtac app_rule THEN' (K (print_tac "dtac2 success"))))
+in
+  (the_tac i state) handle _ => no_tac state
+end
+*}
+
+ML {*
+  fun sg_fwd_tac ctxt thm i state = 
+  let  
+    val goal = nth (Drule.cprems_of state) (i - 1)          
+    val prems = goal |> term_of |> Term.strip_all_body |> Logic.strip_imp_prems 
+    val posx = ListExtra.range 1 (length prems)
+  in
+      ((map (fn pos => attemp (sg_sg_fwd_tac ctxt thm pos)) posx) |> FIRST') i state
+  end
+
+  fun fwd_tac ctxt thms i state =
+       ((map (fn thm => sg_fwd_tac ctxt thm) thms) |> FIRST') i state
+*}
+
+method_setup fwd = {* 
+  Attrib.thms >> (fn thms => fn ctxt =>
+                    (SIMPLE_METHOD' (fn i => 
+                       fwd_tac ctxt (thms@(FwdRules.get ctxt))  i)))
+  *} 
+  "Forward derivation of separation implication"
+
+text {* Testing the fwd tactic *}
+
+lemma ones_abs:
+  assumes "(ones u v \<and>* ones w x) s" "w = v + 1"
+  shows "ones u x s"
+  using assms(1) unfolding assms(2)
+proof(induct u v arbitrary: x s rule:ones_induct)
+  case (Base i j x s)
+  thus ?case by (auto elim!:condE)
+next
+  case (Step i j x s)
+  hence h: "\<And> x s. (ones (i + 1) j \<and>* ones (j + 1) x) s \<longrightarrow> ones (i + 1) x s"
+    by metis
+  hence "(ones (i + 1) x \<and>* one i) s"
+    by (rule fwd_rule, insert Step(3), auto simp:sep_conj_ac)
+  thus ?case
+    by (smt condD ones.simps sep_conj_commute)
+qed
+
+lemma one_abs: "(one m) s \<Longrightarrow> (ones m m) s"
+ by (smt cond_true_eq2 ones.simps)
+
+lemma ones_reps_abs: 
+  assumes "ones m n s"
+          "m \<le> n"
+  shows "(reps m n [nat (n - m)]) s"
+  using assms
+  by simp
+
+lemma reps_reps'_abs: 
+  assumes "(reps m n xs \<and>* zero u) s" "u = n + 1" "xs \<noteq> []"
+  shows "(reps' m u xs) s"
+  unfolding assms using assms
+  by (unfold reps'_def, simp)
+
+lemma reps'_abs:
+  assumes "(reps' m n xs \<and>* reps' u v ys) s" "u = n + 1"
+  shows "(reps' m v (xs @ ys)) s"
+  apply (unfold reps'_append, rule_tac x = u in EXS_intro)
+  by (insert assms, simp)
+
+lemmas abs_ones = one_abs ones_abs
+
+lemmas abs_reps' = ones_reps_abs reps_reps'_abs reps'_abs
+
+
+section {* Modular TM programming and verification *}
+
+definition "right_until_zero = 
+                 (TL start exit. 
+                  TLabel start;
+                     if_zero exit;
+                     move_right;
+                     jmp start;
+                  TLabel exit
+                 )"
+
+lemma ones_false [simp]: "j < i - 1 \<Longrightarrow> (ones i j) = sep_false"
+  by (simp add:pasrt_def)
+  
+lemma hoare_right_until_zero: 
+  "\<lbrace>st i ** ps u ** ones u (v - 1) ** zero v \<rbrace> 
+     i:[right_until_zero]:j
+   \<lbrace>st j ** ps v ** ones u (v - 1) ** zero v \<rbrace>"
+proof(unfold right_until_zero_def, 
+      intro t_hoare_local t_hoare_label, clarify, 
+      rule t_hoare_label_last, simp, simp)
+  fix la
+  let ?body = "i :[ (if_zero la ; move_right ; jmp i) ]: la"
+  let ?j = la
+  show "\<lbrace>st i \<and>* ps u \<and>* ones u (v - 1) \<and>* zero v\<rbrace>  ?body
+        \<lbrace>st ?j \<and>* ps v \<and>* ones u (v - 1) \<and>* zero v\<rbrace>" (is "?P u (v - 1) (ones u (v - 1))")
+  proof(induct "u" "v - 1" rule:ones_induct)
+    case (Base k)
+    moreover have "\<lbrace>st i \<and>* ps v \<and>* zero v\<rbrace> ?body
+                   \<lbrace>st ?j \<and>* ps v \<and>* zero v\<rbrace>" by hsteps
+    ultimately show ?case by (auto intro!:tm.pre_condI simp:sep_conj_cond)
+  next
+    case (Step k)
+    moreover have "\<lbrace>st i \<and>* ps k \<and>* (one k \<and>* ones (k + 1) (v - 1)) \<and>* zero v\<rbrace> 
+                     i :[ (if_zero ?j ; move_right ; jmp i) ]: ?j
+                   \<lbrace>st ?j \<and>* ps v \<and>* (one k \<and>* ones (k + 1) (v - 1)) \<and>* zero v\<rbrace>"
+    proof -
+      have s1: "\<lbrace>st i \<and>* ps k \<and>* (one k \<and>* ones (k + 1) (v - 1)) \<and>* zero v\<rbrace>
+                          ?body 
+                \<lbrace>st i \<and>* ps (k + 1) \<and>* one k \<and>* ones (k + 1) (v - 1) \<and>* zero v\<rbrace>"
+      proof(cases "k + 1 \<ge> v")
+        case True
+        with Step(1) have "v = k + 1" by arith
+        thus ?thesis
+          apply(simp add: one_def)
+          by hsteps
+      next
+        case False
+        hence eq_ones: "ones (k + 1) (v - 1) = 
+                         (one (k + 1) \<and>* ones ((k + 1) + 1) (v - 1))"
+          by simp
+        show ?thesis
+          apply(simp only: eq_ones)
+          by hsteps
+      qed
+      note Step(2)[step]
+      have s2: "\<lbrace>st i \<and>* ps (k + 1) \<and>* one k \<and>* ones (k + 1) (v - 1) \<and>* zero v\<rbrace>
+                        ?body
+                \<lbrace>st ?j \<and>* ps v \<and>* one k \<and>* ones (k + 1) (v - 1) \<and>* zero v\<rbrace>"
+        by hsteps
+      from tm.sequencing [OF s1 s2, step] 
+      show ?thesis 
+        by (auto simp:sep_conj_ac)
+    qed
+    ultimately show ?case by simp
+  qed
+qed
+
+lemma hoare_right_until_zero_gen[step]: 
+  assumes "u = v" "w = x - 1"
+  shows  "\<lbrace>st i ** ps u ** ones v w ** zero x \<rbrace> 
+              i:[right_until_zero]:j
+          \<lbrace>st j ** ps x ** ones v w ** zero x \<rbrace>"
+  by (unfold assms, rule hoare_right_until_zero)
+
+definition "left_until_zero = 
+                 (TL start exit. 
+                  TLabel start;
+                    if_zero exit;
+                    move_left;
+                    jmp start;
+                  TLabel exit
+                 )"
+
+lemma hoare_left_until_zero: 
+  "\<lbrace>st i ** ps v ** zero u ** ones (u + 1) v \<rbrace> 
+     i:[left_until_zero]:j
+   \<lbrace>st j ** ps u ** zero u ** ones (u + 1) v \<rbrace>"
+proof(unfold left_until_zero_def, 
+      intro t_hoare_local t_hoare_label, clarify, 
+      rule t_hoare_label_last, simp+)
+  fix la
+  let ?body = "i :[ (if_zero la ; move_left ; jmp i) ]: la"
+  let ?j = la
+  show "\<lbrace>st i \<and>* ps v \<and>* zero u \<and>* ones (u + 1) v\<rbrace> ?body
+        \<lbrace>st ?j \<and>* ps u \<and>* zero u \<and>* ones (u + 1) v\<rbrace>"
+  proof(induct "u+1" v  rule:ones_rev_induct)
+    case (Base k)
+    thus ?case
+      by (simp add:sep_conj_cond, intro tm.pre_condI, simp, hstep)
+  next
+    case (Step k)
+    have "\<lbrace>st i \<and>* ps k \<and>* zero u \<and>* ones (u + 1) (k - 1) \<and>* one k\<rbrace> 
+               ?body
+          \<lbrace>st ?j \<and>* ps u \<and>* zero u \<and>* ones (u + 1) (k - 1) \<and>* one k\<rbrace>"
+    proof(rule tm.sequencing[where q = 
+           "st i \<and>* ps (k - 1) \<and>* zero u \<and>* ones (u + 1) (k - 1) \<and>* one k"])
+      show "\<lbrace>st i \<and>* ps k \<and>* zero u \<and>* ones (u + 1) (k - 1) \<and>* one k\<rbrace> 
+                ?body
+            \<lbrace>st i \<and>* ps (k - 1) \<and>* zero u \<and>* ones (u + 1) (k - 1) \<and>* one k\<rbrace>"
+      proof(induct "u + 1" "k - 1" rule:ones_rev_induct)
+        case Base with Step(1) have "k = u + 1" by arith
+        thus ?thesis
+          by (simp, hsteps)
+      next
+        case Step
+        show ?thesis
+          apply (unfold ones_rev[OF Step(1)], simp)
+          apply (unfold one_def)
+          by hsteps
+      qed
+    next
+      note Step(2) [step]
+      show "\<lbrace>st i \<and>* ps (k - 1) \<and>* zero u \<and>* ones (u + 1) (k - 1) \<and>* one k\<rbrace> 
+                ?body
+            \<lbrace>st ?j \<and>* ps u \<and>* zero u \<and>* ones (u + 1) (k - 1) \<and>* one k\<rbrace>" by hsteps
+    qed
+    thus ?case by (unfold ones_rev[OF Step(1)], simp)
+  qed
+qed
+
+lemma hoare_left_until_zero_gen[step]: 
+  assumes "u = x" "w = v + 1"
+  shows  "\<lbrace>st i ** ps u ** zero v ** ones w x \<rbrace> 
+               i:[left_until_zero]:j
+          \<lbrace>st j ** ps v ** zero v ** ones w x \<rbrace>"
+  by (unfold assms, rule hoare_left_until_zero)
+
+definition "right_until_one = 
+                 (TL start exit. 
+                  TLabel start;
+                     if_one exit;
+                     move_right;
+                     jmp start;
+                  TLabel exit
+                 )"
+
+lemma hoare_right_until_one: 
+  "\<lbrace>st i ** ps u ** zeros u (v - 1) ** one v \<rbrace> 
+     i:[right_until_one]:j
+   \<lbrace>st j ** ps v ** zeros u (v - 1) ** one v \<rbrace>"
+proof(unfold right_until_one_def, 
+      intro t_hoare_local t_hoare_label, clarify, 
+      rule t_hoare_label_last, simp+)
+  fix la
+  let ?body = "i :[ (if_one la ; move_right ; jmp i) ]: la"
+  let ?j = la
+  show "\<lbrace>st i \<and>* ps u \<and>* zeros u (v - 1) \<and>* one v\<rbrace> ?body
+       \<lbrace>st ?j \<and>* ps v \<and>* zeros u (v - 1) \<and>* one v\<rbrace>"
+  proof(induct u "v - 1" rule:zeros_induct)
+    case (Base k)
+    thus ?case
+      by (simp add:sep_conj_cond, intro tm.pre_condI, simp, hsteps)
+  next
+    case (Step k)
+    have "\<lbrace>st i \<and>* ps k \<and>* zero k \<and>* zeros (k + 1) (v - 1) \<and>* one v\<rbrace> 
+            ?body
+          \<lbrace>st ?j \<and>* ps v \<and>* zero k \<and>* zeros (k + 1) (v - 1) \<and>* one v\<rbrace>"
+    proof(rule tm.sequencing[where q = 
+           "st i \<and>* ps (k + 1) \<and>* zero k \<and>* zeros (k + 1) (v - 1) \<and>* one v"])
+      show "\<lbrace>st i \<and>* ps k \<and>* zero k \<and>* zeros (k + 1) (v - 1) \<and>* one v\<rbrace> 
+               ?body
+           \<lbrace>st i \<and>* ps (k + 1) \<and>* zero k \<and>* zeros (k + 1) (v - 1) \<and>* one v\<rbrace>"
+      proof(induct "k + 1" "v - 1" rule:zeros_induct)
+        case Base
+        with Step(1) have eq_v: "k + 1 = v" by arith
+        from Base show ?thesis
+          apply (simp add:sep_conj_cond, intro tm.pre_condI, simp)
+          apply (hstep, clarsimp)
+          by hsteps
+      next
+        case Step
+        thus ?thesis
+          by (simp, hsteps)
+      qed
+    next
+      note Step(2)[step]
+        show "\<lbrace>st i \<and>* ps (k + 1) \<and>* zero k \<and>* zeros (k + 1) (v - 1) \<and>* one v\<rbrace> 
+                ?body
+              \<lbrace>st ?j \<and>* ps v \<and>* zero k \<and>* zeros (k + 1) (v - 1) \<and>* one v\<rbrace>"
+          by hsteps
+    qed
+    thus ?case by (auto simp: sep_conj_ac Step(1))
+  qed
+qed
+
+lemma hoare_right_until_one_gen[step]: 
+  assumes "u = v" "w = x - 1"
+  shows
+  "\<lbrace>st i ** ps u ** zeros v w ** one x \<rbrace> 
+     i:[right_until_one]:j
+   \<lbrace>st j **  ps x ** zeros v w ** one x \<rbrace>"
+  by (unfold assms, rule hoare_right_until_one)
+
+definition "left_until_one = 
+                 (TL start exit. 
+                  TLabel start;
+                    if_one exit;
+                    move_left;
+                    jmp start;
+                  TLabel exit
+                 )"
+
+lemma hoare_left_until_one: 
+  "\<lbrace>st i ** ps v ** one u ** zeros (u + 1) v \<rbrace> 
+     i:[left_until_one]:j
+   \<lbrace>st j ** ps u ** one u ** zeros (u + 1) v \<rbrace>"
+proof(unfold left_until_one_def, 
+      intro t_hoare_local t_hoare_label, clarify, 
+      rule t_hoare_label_last, simp+)
+  fix la
+  let ?body = "i :[ (if_one la ; move_left ; jmp i) ]: la"
+  let ?j = la
+  show "\<lbrace>st i \<and>* ps v \<and>* one u \<and>* zeros (u + 1) v\<rbrace> ?body
+        \<lbrace>st ?j \<and>* ps u \<and>* one u \<and>* zeros (u + 1) v\<rbrace>"
+  proof(induct u v rule: ones'.induct)
+    fix ia ja
+    assume h: "\<not> ja < ia \<Longrightarrow>
+             \<lbrace>st i \<and>* ps (ja - 1) \<and>* one ia \<and>* zeros (ia + 1) (ja - 1)\<rbrace> ?body
+             \<lbrace>st ?j \<and>* ps ia \<and>* one ia \<and>* zeros (ia + 1) (ja - 1)\<rbrace>"
+    show "\<lbrace>st i \<and>* ps ja \<and>* one ia \<and>* zeros (ia + 1) ja\<rbrace>  ?body
+      \<lbrace>st ?j \<and>* ps ia \<and>* one ia \<and>* zeros (ia + 1) ja\<rbrace>"
+    proof(cases "ja < ia")
+      case False
+      note lt = False
+      from h[OF this] have [step]: 
+        "\<lbrace>st i \<and>* ps (ja - 1) \<and>* one ia \<and>* zeros (ia + 1) (ja - 1)\<rbrace> ?body
+         \<lbrace>st ?j \<and>* ps ia \<and>* one ia \<and>* zeros (ia + 1) (ja - 1)\<rbrace>" .
+      show ?thesis
+      proof(cases "ja = ia")
+        case True 
+        moreover
+        have "\<lbrace>st i \<and>* ps ja \<and>* one ja\<rbrace> ?body \<lbrace>st ?j \<and>* ps ja \<and>* one ja\<rbrace>" 
+          by hsteps
+        ultimately show ?thesis by auto
+      next
+        case False
+        with lt have k1: "ia < ja" by auto       
+        from zeros_rev[of "ja" "ia + 1"] this
+        have eq_zeros: "zeros (ia + 1) ja = (zeros (ia + 1) (ja - 1) \<and>* zero ja)" 
+          by simp        
+        have s1: "\<lbrace>st i \<and>* ps ja \<and>* one ia \<and>* zeros (ia + 1) (ja - 1) \<and>* zero ja\<rbrace>
+                      ?body
+                  \<lbrace>st i \<and>* ps (ja - 1) \<and>* one ia \<and>* zeros (ia + 1) (ja - 1) \<and>* zero ja\<rbrace>"
+        proof(cases "ia + 1 \<ge> ja")
+          case True
+          from k1 True have "ja = ia + 1" by arith
+          moreover have "\<lbrace>st i \<and>* ps (ia + 1) \<and>* one (ia + 1 - 1) \<and>* zero (ia + 1)\<rbrace>  
+            i :[ (if_one ?j ; move_left ; jmp i) ]: ?j 
+                \<lbrace>st i \<and>* ps (ia + 1 - 1) \<and>* one (ia + 1 - 1) \<and>* zero (ia + 1)\<rbrace>"
+            by (hsteps)
+          ultimately show ?thesis
+            by (simp)
+        next
+          case False
+          from zeros_rev[of "ja - 1" "ia + 1"] False
+          have k: "zeros (ia + 1) (ja - 1) = 
+                      (zeros (ia + 1) (ja - 1 - 1) \<and>* zero (ja - 1))"
+            by auto
+          show ?thesis
+            apply (unfold k, simp)
+            by hsteps
+        qed      
+        have s2: "\<lbrace>st i \<and>* ps (ja - 1) \<and>* one ia \<and>* zeros (ia + 1) (ja - 1) \<and>* zero ja\<rbrace>
+                      ?body
+                  \<lbrace>st ?j \<and>* ps ia \<and>* one ia \<and>* zeros (ia + 1) (ja - 1) \<and>* zero ja\<rbrace>"
+          by hsteps
+        from tm.sequencing[OF s1 s2, step]
+        show ?thesis 
+          apply (unfold eq_zeros)
+          by hstep
+      qed (* ccc *)
+    next
+      case True
+      thus ?thesis by (auto intro:tm.hoare_sep_false)
+    qed
+  qed
+qed
+
+lemma hoare_left_until_one_gen[step]: 
+  assumes "u = x" "w = v + 1"
+  shows  "\<lbrace>st i ** ps u ** one v ** zeros w x \<rbrace> 
+              i:[left_until_one]:j
+          \<lbrace>st j ** ps v ** one v ** zeros w x \<rbrace>"
+  by (unfold assms, rule hoare_left_until_one)
+
+definition "left_until_double_zero = 
+            (TL start exit.
+              TLabel start;
+              if_zero exit;
+              left_until_zero;
+              move_left;
+              if_one start;
+              TLabel exit)"
+
+declare ones.simps[simp del]
+
+lemma reps_simps3: "ks \<noteq> [] \<Longrightarrow> 
+  reps i j (k # ks) = (ones i (i + int k) ** zero (i + int k + 1) ** reps (i + int k + 2) j ks)"
+by(case_tac ks, simp, simp add: reps.simps)
+
+lemma cond_eqI:
+  assumes h: "b \<Longrightarrow> r = s"
+  shows "(<b> ** r) = (<b> ** s)"
+proof(cases b)
+  case True
+  from h[OF this] show ?thesis by simp
+next
+  case False
+  thus ?thesis
+    by (unfold sep_conj_def set_ins_def pasrt_def, auto)
+qed
+
+lemma reps_rev: "ks \<noteq> [] 
+       \<Longrightarrow> reps i j (ks @ [k]) =  (reps i (j - int (k + 1) - 1 ) ks \<and>* 
+                                          zero (j - int (k + 1)) \<and>* ones (j - int k) j)"
+proof(induct ks arbitrary: i j)
+  case Nil
+  thus ?case by simp
+next
+  case (Cons a ks)
+  show ?case
+  proof(cases "ks = []")
+    case True
+    thus ?thesis
+    proof -
+      have eq_cond: "(j = i + int a + 2 + int k) = (-2 + (j - int k) = i + int a)" by auto
+      have "(<(-2 + (j - int k) = i + int a)> \<and>*
+            one i \<and>* ones (i + 1) (i + int a) \<and>*
+            zero (i + int a + 1) \<and>* one (i + int a + 2) \<and>* ones (3 + (i + int a)) (i + int a + 2 + int k))
+            =
+            (<(-2 + (j - int k) = i + int a)> \<and>* one i \<and>* ones (i + 1) (i + int a) \<and>*
+            zero (j - (1 + int k)) \<and>* one (j - int k) \<and>* ones (j - int k + 1) j)"
+        (is "(<?X> \<and>* ?L) = (<?X> \<and>* ?R)")
+      proof(rule cond_eqI)
+        assume h: "-2 + (j - int k) = i + int a"
+        hence eqs:  "i + int a + 1 = j - (1 + int k)" 
+                    "i + int a + 2 = j - int k"
+                    "3 + (i + int a) = j - int k + 1"
+                    "(i + int a + 2 + int k) = j"
+        by auto
+        show "?L = ?R"
+          by (unfold eqs, auto simp:sep_conj_ac)
+      qed
+      with True
+      show ?thesis
+        apply (simp del:ones_simps reps.simps)
+        apply (simp add:sep_conj_cond eq_cond)
+        by (auto simp:sep_conj_ac)
+    qed
+  next
+    case False
+    from Cons(1)[OF False, of "i + int a + 2" j] this
+    show ?thesis
+      by(simp add: reps_simps3 sep_conj_ac)
+  qed
+qed
+
+lemma hoare_if_one_reps:
+  assumes nn: "ks \<noteq> []"
+  shows "\<lbrace>st i ** ps v ** reps u v ks\<rbrace> 
+           i:[if_one e]:j
+        \<lbrace>st e ** ps v ** reps u v ks\<rbrace>"
+proof(rule rev_exhaust[of ks])
+  assume "ks = []" with nn show ?thesis by simp
+next
+  fix y ys
+  assume eq_ks: "ks = ys @ [y]"
+  show " \<lbrace>st i \<and>* ps v \<and>* reps u v ks\<rbrace>  i :[ if_one e ]: j \<lbrace>st e \<and>* ps v \<and>* reps u v ks\<rbrace>"
+  proof(cases "ys = []")
+    case False
+    have "\<lbrace>st i \<and>* ps v \<and>* reps u v (ys @ [y])\<rbrace>  i :[ if_one e ]: j \<lbrace>st e \<and>* ps v \<and>* reps u v (ys @ [y])\<rbrace>"
+      apply(unfold reps_rev[OF False], simp del:ones_simps add:ones_rev)
+      by hstep
+    thus ?thesis
+      by (simp add:eq_ks)
+  next
+    case True
+    with eq_ks
+    show ?thesis
+      apply (simp del:ones_simps add:ones_rev sep_conj_cond, intro tm.pre_condI, simp)
+      by hstep
+  qed
+qed
+
+lemma hoare_if_one_reps_gen[step]:
+  assumes nn: "ks \<noteq> []" "u = w"
+  shows "\<lbrace>st i ** ps u ** reps v w ks\<rbrace> 
+           i:[if_one e]:j
+        \<lbrace>st e ** ps u ** reps v w ks\<rbrace>"
+  by (unfold `u = w`, rule hoare_if_one_reps[OF `ks \<noteq> []`])
+
+lemma hoare_if_zero_ones_false[step]:
+  assumes "\<not> w < u" "v = w"
+  shows  "\<lbrace>st i \<and>* ps v \<and>* ones u w\<rbrace> 
+             i :[if_zero e]: j
+          \<lbrace>st j \<and>* ps v \<and>* ones u w\<rbrace>"
+  by (unfold `v = w` ones_rev[OF `\<not> w < u`], hstep)
+
+lemma hoare_left_until_double_zero_nil[step]:
+  assumes "u = v"
+  shows   "\<lbrace>st i ** ps u ** zero v\<rbrace> 
+                  i:[left_until_double_zero]:j
+           \<lbrace>st j ** ps u ** zero v\<rbrace>"
+  apply (unfold `u = v` left_until_double_zero_def, 
+      intro t_hoare_local t_hoare_label, clarsimp, 
+      rule t_hoare_label_last, simp+)
+  by (hsteps)
+
+lemma hoare_if_zero_reps_false:
+  assumes nn: "ks \<noteq> []"
+  shows "\<lbrace>st i ** ps v ** reps u v ks\<rbrace> 
+           i:[if_zero e]:j
+        \<lbrace>st j ** ps v ** reps u v ks\<rbrace>"
+proof(rule rev_exhaust[of ks])
+  assume "ks = []" with nn show ?thesis by simp
+next
+  fix y ys
+  assume eq_ks: "ks = ys @ [y]"
+  show " \<lbrace>st i \<and>* ps v \<and>* reps u v ks\<rbrace>  i :[ if_zero e ]: j \<lbrace>st j \<and>* ps v \<and>* reps u v ks\<rbrace>"
+  proof(cases "ys = []")
+    case False
+    have "\<lbrace>st i \<and>* ps v \<and>* reps u v (ys @ [y])\<rbrace>  i :[ if_zero e ]: j \<lbrace>st j \<and>* ps v \<and>* reps u v (ys @ [y])\<rbrace>"
+      apply(unfold reps_rev[OF False], simp del:ones_simps add:ones_rev)
+      by hstep
+    thus ?thesis
+      by (simp add:eq_ks)
+  next
+    case True
+    with eq_ks
+    show ?thesis
+      apply (simp del:ones_simps add:ones_rev sep_conj_cond, intro tm.pre_condI, simp)
+      by hstep
+  qed
+qed
+
+lemma hoare_if_zero_reps_false_gen[step]:
+  assumes "ks \<noteq> []" "u = w"
+  shows "\<lbrace>st i ** ps u ** reps v w ks\<rbrace> 
+           i:[if_zero e]:j
+        \<lbrace>st j ** ps u ** reps v w ks\<rbrace>"
+  by (unfold `u = w`, rule hoare_if_zero_reps_false[OF `ks \<noteq> []`])
+
+
+lemma hoare_if_zero_reps_false1:
+  assumes nn: "ks \<noteq> []"
+  shows "\<lbrace>st i ** ps u ** reps u v ks\<rbrace> 
+           i:[if_zero e]:j
+        \<lbrace>st j ** ps u ** reps u v ks\<rbrace>"
+proof -
+  from nn obtain y ys where eq_ys: "ks = y#ys"
+    by (metis neq_Nil_conv)
+  show ?thesis
+    apply (unfold eq_ys)
+    by (case_tac ys, (simp, hsteps)+)
+qed
+
+lemma hoare_if_zero_reps_false1_gen[step]:
+  assumes nn: "ks \<noteq> []"
+  and h: "u = w"
+  shows "\<lbrace>st i ** ps u ** reps w v ks\<rbrace> 
+           i:[if_zero e]:j
+        \<lbrace>st j ** ps u ** reps w v ks\<rbrace>"
+  by (unfold h, rule hoare_if_zero_reps_false1[OF `ks \<noteq> []`])
+
+lemma hoare_left_until_double_zero: 
+  assumes h: "ks \<noteq> []"
+  shows   "\<lbrace>st i ** ps v ** zero u ** zero (u + 1) ** reps (u+2) v ks\<rbrace> 
+                  i:[left_until_double_zero]:j
+           \<lbrace>st j ** ps u ** zero u ** zero (u + 1) ** reps (u+2) v ks\<rbrace>"
+proof(unfold left_until_double_zero_def, 
+      intro t_hoare_local t_hoare_label, clarsimp, 
+      rule t_hoare_label_last, simp+)
+  fix la
+  let ?body = "i :[ (if_zero la ; left_until_zero ; move_left ; if_one i) ]: j"
+  let ?j = j
+  show "\<lbrace>st i \<and>* ps v \<and>* zero u \<and>* zero (u + 1) \<and>* reps (u + 2) v ks\<rbrace> 
+           ?body
+        \<lbrace>st ?j \<and>* ps u \<and>* zero u \<and>* zero (u + 1) \<and>* reps (u + 2) v ks\<rbrace>"
+    using h
+  proof(induct ks arbitrary: v rule:rev_induct)
+    case Nil
+    with h show ?case by auto
+  next
+    case (snoc k ks)
+    show ?case
+    proof(cases "ks = []")
+      case True
+      have eq_ones: 
+        "ones (u + 2) (u + 2 + int k) = (ones (u + 2) (u + 1 + int k) \<and>* one (u + 2 + int k))"
+        by (smt ones_rev)
+      have eq_ones': "(one (u + 2) \<and>* ones (3 + u) (u + 2 + int k)) = 
+            (one (u + 2 + int k) \<and>* ones (u + 2) (u + 1 + int k))"
+        by (smt eq_ones ones.simps sep.mult_commute)
+      thus ?thesis
+        apply (insert True, simp del:ones_simps add:sep_conj_cond)
+        apply (rule tm.pre_condI, simp del:ones_simps, unfold eq_ones)
+        apply hsteps
+        apply (rule_tac p = "st j' \<and>* ps (u + 2 + int k) \<and>* zero u \<and>* 
+                             zero (u + 1) \<and>* ones (u + 2) (u + 2 + int k)" 
+                  in tm.pre_stren)
+        by (hsteps)
+    next
+      case False
+      from False have spt: "splited (ks @ [k]) ks [k]" by (unfold splited_def, auto)
+      show ?thesis
+        apply (unfold reps_splited[OF spt], simp del:ones_simps add:sep_conj_cond)
+        apply (rule tm.pre_condI, simp del:ones_simps)
+        apply (rule_tac q = "st i \<and>*
+               ps (1 + (u + int (reps_len ks))) \<and>*
+               zero u \<and>*
+               zero (u + 1) \<and>*
+               reps (u + 2) (1 + (u + int (reps_len ks))) ks \<and>*
+               zero (u + 2 + int (reps_len ks)) \<and>*
+               ones (3 + (u + int (reps_len ks))) (3 + (u + int (reps_len ks)) + int k)" in
+               tm.sequencing)
+        apply hsteps[1]
+        by (hstep snoc(1))
+    qed 
+  qed
+qed
+
+lemma hoare_left_until_double_zero_gen[step]: 
+  assumes h1: "ks \<noteq> []"
+      and h: "u = y" "w = v + 1" "x = v + 2"
+  shows   "\<lbrace>st i ** ps u ** zero v ** zero w ** reps x y ks\<rbrace> 
+                  i:[left_until_double_zero]:j
+           \<lbrace>st j ** ps v ** zero v ** zero w ** reps x y ks\<rbrace>"
+  by (unfold h, rule hoare_left_until_double_zero[OF h1])
+
+lemma hoare_jmp_reps1:
+  assumes "ks \<noteq> []"
+  shows  "\<lbrace> st i \<and>* ps u \<and>* reps u v ks\<rbrace>
+                 i:[jmp e]:j
+          \<lbrace> st e \<and>* ps u \<and>* reps u v ks\<rbrace>"
+proof -
+  from assms obtain k ks' where Cons:"ks = k#ks'"
+    by (metis neq_Nil_conv)
+  thus ?thesis
+  proof(cases "ks' = []")
+    case True with Cons
+    show ?thesis
+      apply(simp add:sep_conj_cond reps.simps, intro tm.pre_condI, simp add:ones_simps)
+      by (hgoto hoare_jmp_gen)
+  next
+    case False
+    show ?thesis
+      apply (unfold `ks = k#ks'` reps_simp3[OF False], simp add:ones_simps)
+      by (hgoto hoare_jmp[where p = u])
+  qed
+qed
+
+lemma hoare_jmp_reps1_gen[step]:
+  assumes "ks \<noteq> []" "u = v"
+  shows  "\<lbrace> st i \<and>* ps u \<and>* reps v w ks\<rbrace>
+                 i:[jmp e]:j
+          \<lbrace> st e \<and>* ps u \<and>* reps v w ks\<rbrace>"
+  by (unfold assms, rule hoare_jmp_reps1[OF `ks \<noteq> []`])
+
+lemma hoare_jmp_reps:
+      "\<lbrace> st i \<and>* ps u \<and>* reps u v ks \<and>* tm (v + 1) x \<rbrace>
+                 i:[(jmp e; c)]:j
+       \<lbrace> st e \<and>* ps u \<and>* reps u v ks \<and>* tm (v + 1) x \<rbrace>"
+proof(cases "ks")
+  case Nil
+  thus ?thesis
+    by (simp add:sep_conj_cond, intro tm.pre_condI, simp, hsteps)
+next
+  case (Cons k ks')
+  thus ?thesis
+  proof(cases "ks' = []")
+    case True with Cons
+    show ?thesis
+      apply(simp add:sep_conj_cond, intro tm.pre_condI, simp)
+      by (hgoto hoare_jmp[where p = u])
+  next
+    case False
+    show ?thesis
+      apply (unfold `ks = k#ks'` reps_simp3[OF False], simp)
+      by (hgoto hoare_jmp[where p = u])
+  qed
+qed
+
+definition "shift_right =
+            (TL start exit.
+              TLabel start;
+                 if_zero exit;
+                 write_zero;
+                 move_right;
+                 right_until_zero;
+                 write_one;
+                 move_right;
+                 jmp start;
+              TLabel exit
+            )"
+
+lemma hoare_shift_right_cons:
+  assumes h: "ks \<noteq> []"
+  shows "\<lbrace>st i \<and>* ps u ** reps u v ks \<and>* zero (v + 1) \<and>* zero (v + 2) \<rbrace> 
+            i:[shift_right]:j
+         \<lbrace>st j ** ps (v + 2) ** zero u ** reps (u + 1) (v + 1) ks ** zero (v + 2) \<rbrace>"
+proof(unfold shift_right_def, intro t_hoare_local t_hoare_label, clarify, 
+      rule t_hoare_label_last, auto)
+  fix la
+  have eq_ones: "\<And> u k. (one (u + int k + 1) \<and>* ones (u + 1) (u + int k)) = 
+                                   (one (u + 1) \<and>* ones (2 + u) (u + 1 + int k))"
+    by (smt cond_true_eq2 ones.simps ones_rev sep.mult_assoc sep.mult_commute 
+               sep.mult_left_commute sep_conj_assoc sep_conj_commute 
+               sep_conj_cond1 sep_conj_cond2 sep_conj_cond3 sep_conj_left_commute
+               sep_conj_trivial_strip2)
+  show "\<lbrace>st i \<and>* ps u \<and>* reps u v ks \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace> 
+         i :[ (if_zero la ;
+               write_zero ; move_right ; right_until_zero ; write_one ; move_right ; jmp i) ]: la
+         \<lbrace>st la \<and>* ps (v + 2) \<and>* zero u \<and>* reps (u + 1) (v + 1) ks \<and>* zero (v + 2)\<rbrace>"
+    using h
+  proof(induct ks arbitrary:i u v)
+    case (Cons k ks)
+    thus ?case 
+    proof(cases "ks = []")
+      let ?j = la
+      case True
+      let ?body = "i :[ (if_zero ?j ;
+                      write_zero ;
+                      move_right ; 
+                      right_until_zero ; 
+                      write_one ; move_right ; jmp i) ]: ?j"
+      have first_interation: 
+           "\<lbrace>st i \<and>* ps u \<and>* one u \<and>* ones (u + 1) (u + int k) \<and>* zero (u + int k + 1) \<and>* 
+                                                                             zero (u + int k + 2)\<rbrace> 
+                ?body
+            \<lbrace>st i \<and>*
+             ps (u + int k + 2) \<and>*
+             one (u + int k + 1) \<and>* ones (u + 1) (u + int k) \<and>* zero u \<and>* zero (u + int k + 2)\<rbrace>"
+        apply (hsteps)
+        by (simp add:sep_conj_ac, sep_cancel+, smt)
+      hence "\<lbrace>st i \<and>* ps u \<and>* one u \<and>* ones (u + 1) (u + int k) \<and>* zero (u + int k + 1) \<and>* 
+                                                                             zero (u + int k + 2)\<rbrace> 
+                   ?body
+             \<lbrace>st ?j \<and>* ps (u + int k + 2) \<and>* zero u \<and>* one (u + 1) \<and>* 
+                         ones (2 + u) (u + 1 + int k) \<and>* zero (u + int k + 2)\<rbrace>"
+      proof(rule tm.sequencing)
+        show "\<lbrace>st i \<and>*
+               ps (u + int k + 2) \<and>*
+               one (u + int k + 1) \<and>* ones (u + 1) (u + int k) \<and>* zero u \<and>* zero (u + int k + 2)\<rbrace> 
+                      ?body
+              \<lbrace>st ?j \<and>*
+               ps (u + int k + 2) \<and>*
+               zero u \<and>* one (u + 1) \<and>* ones (2 + u) (u + 1 + int k) \<and>* zero (u + int k + 2)\<rbrace>"
+          apply (hgoto hoare_if_zero_true_gen)
+          by (simp add:sep_conj_ac eq_ones)
+      qed
+      with True 
+      show ?thesis
+        by (simp, simp only:sep_conj_cond, intro tm.pre_condI, auto simp:sep_conj_ac)
+    next
+      case False
+      let ?j = la
+      let ?body = "i :[ (if_zero ?j ;
+                        write_zero ;
+                        move_right ; right_until_zero ; 
+                        write_one ; move_right ; jmp i) ]: ?j"
+      have eq_ones': 
+         "(one (u + int k + 1) \<and>*
+           ones (u + 1) (u + int k) \<and>*
+           zero u \<and>* reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* zero (v + 2))
+                   =
+           (zero u \<and>*
+             ones (u + 1) (u + int k) \<and>*
+             one (u + int k + 1) \<and>* reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* zero (v + 2))"
+        by (simp add:eq_ones sep_conj_ac)
+      have "\<lbrace>st i \<and>* ps u \<and>* one u \<and>* ones (u + 1) (u + int k) \<and>* zero (u + int k + 1) \<and>* 
+                 reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace> 
+                    ?body
+            \<lbrace>st i \<and>* ps (u + int k + 2) \<and>* zero u \<and>* ones (u + 1) (u + int k) \<and>* 
+                 one (u + int k + 1) \<and>* reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace>"
+        apply (hsteps)
+        by (auto simp:sep_conj_ac, sep_cancel+, smt)
+      hence "\<lbrace>st i \<and>* ps u \<and>* one u \<and>* ones (u + 1) (u + int k) \<and>* zero (u + int k + 1) \<and>* 
+                 reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace> 
+                     ?body
+            \<lbrace>st ?j \<and>* ps (v + 2) \<and>* zero u \<and>* one (u + 1) \<and>* ones (2 + u) (u + 1 + int k) \<and>*
+                 zero (2 + (u + int k)) \<and>* reps (3 + (u + int k)) (v + 1) ks \<and>* zero (v + 2)\<rbrace>"
+      proof(rule tm.sequencing)
+        have eq_ones': 
+          "\<And> u k. (one (u + int k + 1) \<and>* ones (u + 1) (u + int k) \<and>* zero (u + int k + 2)) =
+             (one (u + 1) \<and>* zero (2 + (u + int k)) \<and>* ones (2 + u) (u + 1 + int k))"
+          by (smt eq_ones sep.mult_assoc sep_conj_commute)
+        show "\<lbrace>st i \<and>* ps (u + int k + 2) \<and>* zero u \<and>*
+                    ones (u + 1) (u + int k) \<and>* one (u + int k + 1) \<and>* reps (u + int k + 2) v ks \<and>* 
+                    zero (v + 1) \<and>* zero (v + 2)\<rbrace> 
+                      ?body
+              \<lbrace>st ?j \<and>* ps (v + 2) \<and>* zero u \<and>* one (u + 1) \<and>* ones (2 + u) (u + 1 + int k) \<and>*
+                      zero (2 + (u + int k)) \<and>* reps (3 + (u + int k)) (v + 1) ks \<and>* zero (v + 2)\<rbrace>"
+          apply (hsteps Cons.hyps)
+          by (simp add:sep_conj_ac eq_ones, sep_cancel+, smt)
+      qed
+      thus ?thesis
+        by (unfold reps_simp3[OF False], auto simp:sep_conj_ac)
+    qed 
+  qed auto
+qed
+
+lemma hoare_shift_right_cons_gen[step]:
+  assumes h: "ks \<noteq> []"
+  and h1: "u = v" "x = w + 1" "y = w + 2"
+  shows "\<lbrace>st i \<and>* ps u ** reps v w ks \<and>* zero x \<and>* zero y \<rbrace> 
+            i:[shift_right]:j
+         \<lbrace>st j ** ps y ** zero v ** reps (v + 1) x ks ** zero y\<rbrace>"
+  by (unfold h1, rule hoare_shift_right_cons[OF h])
+
+lemma shift_right_nil [step]: 
+  assumes "u = v"
+  shows
+       "\<lbrace> st i \<and>* ps u \<and>* zero v \<rbrace>
+               i:[shift_right]:j
+        \<lbrace> st j \<and>* ps u \<and>* zero v \<rbrace>"
+  by (unfold assms shift_right_def, intro t_hoare_local t_hoare_label, clarify, 
+          rule t_hoare_label_last, simp+, hstep)
+
+
+text {*
+  @{text "clear_until_zero"} is useful to implement @{text "drag"}.
+*}
+
+definition "clear_until_zero = 
+             (TL start exit.
+              TLabel start;
+                 if_zero exit;
+                 write_zero;
+                 move_right;
+                 jmp start;
+              TLabel exit)"
+
+lemma  hoare_clear_until_zero[step]: 
+         "\<lbrace>st i ** ps u ** ones u v ** zero (v + 1)\<rbrace>
+              i :[clear_until_zero]: j
+          \<lbrace>st j ** ps (v + 1) ** zeros u v ** zero (v + 1)\<rbrace> "
+proof(unfold clear_until_zero_def, intro t_hoare_local, rule t_hoare_label,
+    rule t_hoare_label_last, simp+)
+  let ?body = "i :[ (if_zero j ; write_zero ; move_right ; jmp i) ]: j"
+  show "\<lbrace>st i \<and>* ps u \<and>* ones u v \<and>* zero (v + 1)\<rbrace> ?body 
+        \<lbrace>st j \<and>* ps (v + 1) \<and>* zeros u v \<and>* zero (v + 1)\<rbrace>"
+  proof(induct u v rule: zeros.induct)
+    fix ia ja
+    assume h: "\<not> ja < ia \<Longrightarrow>
+             \<lbrace>st i \<and>* ps (ia + 1) \<and>* ones (ia + 1) ja \<and>* zero (ja + 1)\<rbrace>  ?body
+             \<lbrace>st j \<and>* ps (ja + 1) \<and>* zeros (ia + 1) ja \<and>* zero (ja + 1)\<rbrace>"
+    show "\<lbrace>st i \<and>* ps ia \<and>* ones ia ja \<and>* zero (ja + 1)\<rbrace> ?body
+           \<lbrace>st j \<and>* ps (ja + 1) \<and>* zeros ia ja \<and>* zero (ja + 1)\<rbrace>"
+    proof(cases "ja < ia")
+      case True
+      thus ?thesis
+        by (simp add: ones.simps zeros.simps sep_conj_ac, simp only:sep_conj_cond,
+               intro tm.pre_condI, simp, hsteps)
+    next
+      case False
+      note h[OF False, step]
+      from False have ones_eq: "ones ia ja = (one ia \<and>* ones (ia + 1) ja)"
+        by(simp add: ones.simps)
+      from False have zeros_eq: "zeros ia ja = (zero ia \<and>* zeros (ia + 1) ja)"
+        by(simp add: zeros.simps)
+      have s1: "\<lbrace>st i \<and>* ps ia \<and>* one ia \<and>* ones (ia + 1) ja \<and>* zero (ja + 1)\<rbrace>  ?body 
+                 \<lbrace>st i \<and>* ps (ia + 1) \<and>* zero ia \<and>* ones (ia + 1) ja \<and>* zero (ja + 1)\<rbrace>"
+      proof(cases "ja < ia + 1")
+        case True
+        from True False have "ja = ia" by auto
+        thus ?thesis
+          apply(simp add: ones.simps)
+          by (hsteps)
+      next
+        case False
+        from False have "ones (ia + 1) ja = (one (ia + 1) \<and>* ones (ia + 1 + 1) ja)"
+          by(simp add: ones.simps)
+        thus ?thesis
+          by (simp, hsteps)
+      qed
+      have s2: "\<lbrace>st i \<and>* ps (ia + 1) \<and>* zero ia \<and>* ones (ia + 1) ja \<and>* zero (ja + 1)\<rbrace>
+                ?body
+                \<lbrace>st j \<and>* ps (ja + 1) \<and>* zero ia \<and>* zeros (ia + 1) ja \<and>* zero (ja + 1)\<rbrace>"
+        by hsteps
+      from tm.sequencing[OF s1 s2] have 
+        "\<lbrace>st i \<and>* ps ia \<and>* one ia \<and>* ones (ia + 1) ja \<and>* zero (ja + 1)\<rbrace>  ?body
+        \<lbrace>st j \<and>* ps (ja + 1) \<and>* zero ia \<and>* zeros (ia + 1) ja \<and>* zero (ja + 1)\<rbrace>" .
+      thus ?thesis
+        unfolding ones_eq zeros_eq by(simp add: sep_conj_ac)
+    qed
+  qed
+qed
+
+lemma  hoare_clear_until_zero_gen[step]: 
+  assumes "u = v" "x = w + 1"
+  shows "\<lbrace>st i ** ps u ** ones v w ** zero x\<rbrace>
+              i :[clear_until_zero]: j
+        \<lbrace>st j ** ps x ** zeros v w ** zero x\<rbrace>"
+  by (unfold assms, rule hoare_clear_until_zero)
+
+definition "shift_left = 
+            (TL start exit.
+              TLabel start;
+                 if_zero exit;
+                 move_left;
+                 write_one;
+                 right_until_zero;
+                 move_left;
+                 write_zero;
+                 move_right;
+                 move_right;
+                 jmp start;
+              TLabel exit)
+           "
+
+declare ones_simps[simp del]
+
+lemma hoare_move_left_reps[step]:
+  assumes "ks \<noteq> []" "u = v"
+  shows 
+    "\<lbrace>st i ** ps u ** reps v w ks\<rbrace> 
+         i:[move_left]:j
+     \<lbrace>st j ** ps (u - 1) **  reps v w ks\<rbrace>"
+proof -
+  from `ks \<noteq> []` obtain y ys where eq_ks: "ks = y#ys"
+    by (metis neq_Nil_conv)
+  show ?thesis
+    apply (unfold assms eq_ks)
+    apply (case_tac ys, simp)
+    my_block
+      have "(ones v (v + int y)) = (one v \<and>* ones (v + 1) (v + int y))"
+        by (smt ones_step_simp)
+    my_block_end
+    apply (unfold this, hsteps)
+    by (simp add:this, hsteps)
+qed
+
+lemma hoare_shift_left_cons:
+  assumes h: "ks \<noteq> []"
+  shows "\<lbrace>st i \<and>* ps u \<and>* tm (u - 1) x \<and>* reps u v ks \<and>* zero (v + 1) \<and>* zero (v + 2) \<rbrace> 
+                                   i:[shift_left]:j
+         \<lbrace>st j \<and>* ps (v + 2) \<and>* reps (u - 1) (v - 1) ks \<and>* zero v \<and>* zero (v + 1) \<and>* zero (v + 2) \<rbrace>"
+proof(unfold shift_left_def, intro t_hoare_local t_hoare_label, clarify, 
+      rule t_hoare_label_last, simp+, clarify, prune)
+  show " \<lbrace>st i \<and>* ps u \<and>* tm (u - 1) x \<and>* reps u v ks \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace> 
+             i :[ (if_zero j ;
+                   move_left ;
+                   write_one ;
+                   right_until_zero ;
+                   move_left ; write_zero ; 
+                   move_right ; move_right ; jmp i) ]: j
+         \<lbrace>st j \<and>* ps (v + 2) \<and>* reps (u - 1) (v - 1) ks \<and>* zero v \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace>"
+    using h
+  proof(induct ks arbitrary:i u v x)
+    case (Cons k ks)
+    thus ?case 
+    proof(cases "ks = []")
+      let ?body = "i :[ (if_zero j ;  move_left ; write_one ; right_until_zero ;
+                   move_left ; write_zero ; move_right ; move_right ; jmp i) ]: j"
+      case True 
+      have "\<lbrace>st i \<and>* ps u \<and>* tm (u - 1) x \<and>* (one u \<and>* ones (u + 1) (u + int k)) \<and>* 
+                                          zero (u + int k + 1) \<and>* zero (u + int k + 2)\<rbrace> 
+                         ?body
+            \<lbrace>st j \<and>* ps (u + int k + 2) \<and>* (one (u - 1) \<and>* ones u (u - 1 + int k)) \<and>*
+                       zero (u + int k) \<and>* zero (u + int k + 1) \<and>* zero (u + int k + 2)\<rbrace>"
+      apply(rule tm.sequencing [where q = "st i \<and>* ps (u + int k + 2) \<and>*
+                (one (u - 1) \<and>* ones u (u - 1 + int k)) \<and>*
+                zero (u + int k) \<and>* zero (u + int k + 1) \<and>* zero (u + int k + 2)"])
+          apply (hsteps)
+          apply (rule_tac p = "st j' \<and>* ps (u - 1) \<and>* ones (u - 1) (u + int k) \<and>*
+                                zero (u + int k + 1) \<and>* zero (u + int k + 2)" 
+            in tm.pre_stren)
+          apply (hsteps)
+          my_block
+            have "(ones (u - 1) (u + int k)) = (ones (u - 1) (u + int k - 1) \<and>* one (u + int k))"
+              by (smt ones_rev)
+          my_block_end
+          apply (unfold this)
+          apply hsteps
+          apply (simp add:sep_conj_ac, sep_cancel+)
+          apply (smt ones.simps sep.mult_assoc sep_conj_commuteI)
+          apply (simp add:sep_conj_ac)+
+          apply (sep_cancel+)
+          apply (smt ones.simps sep.mult_left_commute sep_conj_commuteI this)
+          by hstep
+        with True show ?thesis
+        by (simp add:ones_simps, simp only:sep_conj_cond, intro tm.pre_condI, simp)
+    next 
+      case False
+      let ?body = "i :[ (if_zero j ; move_left ; write_one ;right_until_zero ; move_left ; 
+                                write_zero ; move_right ; move_right ; jmp i) ]: j"
+      have "\<lbrace>st i \<and>* ps u \<and>* tm (u - 1) x \<and>* one u \<and>* ones (u + 1) (u + int k) \<and>* 
+                zero (u + int k + 1) \<and>* reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace> 
+                ?body
+            \<lbrace>st j \<and>* ps (v + 2) \<and>* one (u - 1) \<and>* ones u (u - 1 + int k) \<and>*
+                        zero (u + int k) \<and>* reps (1 + (u + int k)) (v - 1) ks \<and>*
+                                              zero v \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace>"
+        apply (rule tm.sequencing[where q = "st i \<and>* ps (u + int k + 2) \<and>* 
+                  zero (u + int k + 1) \<and>* reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* 
+                  zero (v + 2) \<and>* one (u - 1) \<and>* ones u (u - 1 + int k) \<and>* zero (u + int k)"])
+        apply (hsteps)
+        apply (rule_tac p = "st j' \<and>* ps (u - 1) \<and>* 
+                               ones (u - 1) (u + int k) \<and>*
+                               zero (u + int k + 1) \<and>* 
+                               reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* zero (v + 2)
+            " in tm.pre_stren)
+        apply hsteps
+        my_block
+          have "(ones (u - 1) (u + int k)) = 
+            (ones (u - 1) (u + int k - 1) \<and>* one (u + int k))"
+            by (smt ones_rev)
+        my_block_end
+        apply (unfold this)
+        apply (hsteps)
+        apply (sep_cancel+)
+        apply (smt ones.simps sep.mult_assoc sep_conj_commuteI)
+        apply (sep_cancel+)
+        apply (smt ones.simps this)
+        my_block
+          have eq_u: "1 + (u + int k) = u + int k + 1" by simp
+          from Cons.hyps[OF `ks \<noteq> []`, of i "u + int k + 2" Bk v, folded zero_def] 
+          have "\<lbrace>st i \<and>* ps (u + int k + 2) \<and>* zero (u + int k + 1) \<and>*
+                            reps (u + int k + 2) v ks \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace> 
+                               ?body
+                      \<lbrace>st j \<and>* ps (v + 2) \<and>*  reps (1 + (u + int k)) (v - 1) ks \<and>* 
+                                                zero v \<and>* zero (v + 1) \<and>* zero (v + 2)\<rbrace>"
+          by (simp add:eq_u)
+        my_block_end my_note hh[step] = this 
+        by hsteps
+      thus ?thesis
+        by (unfold reps_simp3[OF False], auto simp:sep_conj_ac ones_simps)
+    qed
+  qed auto
+qed
+
+lemma hoare_shift_left_cons_gen[step]:
+  assumes h: "ks \<noteq> []"
+          "v = u - 1" "w = u" "y = x + 1" "z = x + 2"
+  shows "\<lbrace>st i \<and>* ps u \<and>* tm v vv \<and>* reps w x ks \<and>* tm y Bk \<and>* tm z Bk\<rbrace> 
+                                   i:[shift_left]:j
+         \<lbrace>st j \<and>* ps z \<and>* reps v (x - 1) ks \<and>* zero x \<and>* zero y \<and>* zero z \<rbrace>"
+  by (unfold assms, fold zero_def, rule hoare_shift_left_cons[OF `ks \<noteq> []`])
+
+definition "bone c1 c2 = (TL exit l_one.
+                                if_one l_one;
+                                  (c1;
+                                   jmp exit);
+                                TLabel l_one;
+                                      c2;
+                                TLabel exit
+                              )"
+
+lemma hoare_bone_1_out:
+  assumes h: 
+        "\<And> i j . \<lbrace>st i \<and>* ps u \<and>* zero u \<and>* p \<rbrace>
+                         i:[c1]:j
+                  \<lbrace>st e \<and>* q \<rbrace>
+        "
+  shows "\<lbrace>st i \<and>* ps u \<and>* zero u \<and>* p \<rbrace>
+              i:[(bone c1 c2)]:j
+         \<lbrace>st e \<and>* q \<rbrace>
+        "
+apply (unfold bone_def, intro t_hoare_local)
+apply hsteps
+apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension)
+apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension)
+by (rule h)
+
+lemma hoare_bone_1:
+  assumes h: 
+        "\<And> i j . \<lbrace>st i \<and>* ps u \<and>* zero u \<and>* p \<rbrace>
+                         i:[c1]:j
+                  \<lbrace>st j \<and>* ps v \<and>* tm v x \<and>* q \<rbrace>
+        "
+  shows "\<lbrace>st i \<and>* ps u \<and>* zero u \<and>* p \<rbrace>
+              i:[(bone c1 c2)]:j
+         \<lbrace>st j \<and>* ps v \<and>* tm v x \<and>* q \<rbrace>
+        "
+proof -
+  note h[step]
+  show ?thesis
+    apply (unfold bone_def, intro t_hoare_local)
+    apply (rule t_hoare_label_last, auto)
+    apply hsteps
+    apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension)
+    by hsteps
+qed
+
+lemma hoare_bone_2:
+  assumes h: 
+        "\<And> i j . \<lbrace>st i \<and>* ps u \<and>* one u \<and>* p \<rbrace>
+                         i:[c2]:j
+                  \<lbrace>st j \<and>* q \<rbrace>
+        "
+  shows "\<lbrace>st i \<and>* ps u \<and>* one u \<and>* p \<rbrace>
+              i:[(bone c1 c2)]:j
+         \<lbrace>st j \<and>* q \<rbrace>
+        "
+apply (unfold bone_def, intro t_hoare_local)
+apply (rule_tac q = "st la \<and>* ps u \<and>* one u \<and>* p" in tm.sequencing)
+apply hsteps
+apply (subst tassemble_to.simps(2), intro tm.code_exI, intro tm.code_extension1)
+apply (subst tassemble_to.simps(2), intro tm.code_exI, intro tm.code_extension1)
+apply (subst tassemble_to.simps(2), intro tm.code_exI)
+apply (subst tassemble_to.simps(4), intro tm.code_condI, simp)
+apply (subst tassemble_to.simps(2), intro tm.code_exI)
+apply (subst tassemble_to.simps(4), simp add:sep_conj_cond, rule tm.code_condI, simp)
+by (rule h)
+
+lemma hoare_bone_2_out:
+  assumes h: 
+        "\<And> i j . \<lbrace>st i \<and>* ps u \<and>* one u \<and>* p \<rbrace>
+                         i:[c2]:j
+                  \<lbrace>st e \<and>* q \<rbrace>
+        "
+  shows "\<lbrace>st i \<and>* ps u \<and>* one u \<and>* p \<rbrace>
+              i:[(bone c1 c2)]:j
+         \<lbrace>st e \<and>* q \<rbrace>
+        "
+apply (unfold bone_def, intro t_hoare_local)
+apply (rule_tac q = "st la \<and>* ps u \<and>* one u \<and>* p" in tm.sequencing)
+apply hsteps
+apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension1)
+apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension1)
+apply (subst tassemble_to.simps(2), intro tm.code_exI)
+apply (subst tassemble_to.simps(4), rule tm.code_condI, simp)
+apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension)
+by (rule h)
+
+definition "bzero c1 c2 = (TL exit l_zero.
+                                if_zero l_zero;
+                                  (c1;
+                                   jmp exit);
+                                TLabel l_zero;
+                                      c2;
+                                TLabel exit
+                              )"
+
+lemma hoare_bzero_1:
+  assumes h[step]: 
+        "\<And> i j . \<lbrace>st i \<and>* ps u \<and>* one u \<and>* p \<rbrace>
+                         i:[c1]:j
+                 \<lbrace>st j \<and>* ps v \<and>* tm v x \<and>* q \<rbrace>
+        "
+  shows "\<lbrace>st i \<and>* ps u \<and>* one u \<and>* p \<rbrace>
+              i:[(bzero c1 c2)]:j
+         \<lbrace>st j \<and>* ps v \<and>* tm v x \<and>* q \<rbrace>
+        "
+apply (unfold bzero_def, intro t_hoare_local)
+apply hsteps
+apply (rule_tac c = " ((c1 ; jmp l) ; TLabel la ; c2 ; TLabel l)" in t_hoare_label_last, auto)
+apply (subst tassemble_to.simps(2), intro tm.code_exI, intro tm.code_extension)
+by hsteps
+
+lemma hoare_bzero_1_out:
+  assumes h[step]: 
+        "\<And> i j . \<lbrace>st i \<and>* ps u \<and>* one u \<and>* p \<rbrace>
+                         i:[c1]:j
+                 \<lbrace>st e \<and>* q \<rbrace>
+        "
+  shows "\<lbrace>st i \<and>* ps u \<and>* one u \<and>* p \<rbrace>
+              i:[(bzero c1 c2)]:j
+         \<lbrace>st e \<and>* q \<rbrace>
+        "
+apply (unfold bzero_def, intro t_hoare_local)
+apply hsteps
+apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension)
+apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension)
+by (rule h)
+
+lemma hoare_bzero_2:
+  assumes h: 
+        "\<And> i j. \<lbrace>st i \<and>* ps u \<and>* zero u \<and>* p \<rbrace>
+                         i:[c2]:j
+                 \<lbrace>st j \<and>* q \<rbrace>
+        "
+  shows "\<lbrace>st i \<and>* ps u \<and>* zero u \<and>* p \<rbrace>
+              i:[(bzero c1 c2)]:j
+         \<lbrace>st j \<and>* q \<rbrace>
+        "
+  apply (unfold bzero_def, intro t_hoare_local)
+  apply (rule_tac q = "st la \<and>* ps u \<and>* zero u \<and>* p" in tm.sequencing)
+  apply hsteps
+  apply (subst tassemble_to.simps(2), intro tm.code_exI, intro tm.code_extension1)
+  apply (subst tassemble_to.simps(2), intro tm.code_exI, intro tm.code_extension1)
+  apply (subst tassemble_to.simps(2), intro tm.code_exI)
+  apply (subst tassemble_to.simps(4))
+  apply (rule tm.code_condI, simp)
+  apply (subst tassemble_to.simps(2))
+  apply (rule tm.code_exI)
+  apply (subst tassemble_to.simps(4), simp add:sep_conj_cond)
+  apply (rule tm.code_condI, simp)
+  by (rule h)
+
+lemma hoare_bzero_2_out:
+  assumes h: 
+        "\<And> i j . \<lbrace>st i \<and>* ps u \<and>* zero u \<and>* p \<rbrace>
+                         i:[c2]:j
+                  \<lbrace>st e \<and>* q \<rbrace>
+        "
+  shows "\<lbrace>st i \<and>* ps u \<and>* zero u \<and>* p\<rbrace>
+              i:[(bzero c1 c2)]:j
+         \<lbrace>st e \<and>* q \<rbrace>
+        "
+apply (unfold bzero_def, intro t_hoare_local)
+apply (rule_tac q = "st la \<and>* ps u \<and>* zero u \<and>* p" in tm.sequencing)
+apply hsteps
+apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension1)
+apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension1)
+apply (subst tassemble_to.simps(2), intro tm.code_exI)
+apply (subst tassemble_to.simps(4), rule tm.code_condI, simp)
+apply (subst tassemble_to.simps(2), intro tm.code_exI, rule tm.code_extension)
+by (rule h)
+
+definition "skip_or_set = bone (write_one; move_right; move_right)
+                               (right_until_zero; move_right)"
+
+lemma reps_len_split: 
+  assumes "xs \<noteq> []" "ys \<noteq> []"
+  shows "reps_len (xs @ ys) = reps_len xs + reps_len ys + 1"
+  using assms
+proof(induct xs arbitrary:ys)
+  case (Cons x1 xs1)
+  show ?case
+  proof(cases "xs1 = []")
+    case True
+    thus ?thesis
+      by (simp add:reps_len_cons[OF `ys \<noteq> []`] reps_len_sg)
+  next
+    case False
+    hence " xs1 @ ys \<noteq> []" by simp
+    thus ?thesis
+      apply (simp add:reps_len_cons[OF `xs1@ys \<noteq> []`] reps_len_cons[OF `xs1 \<noteq> []`])
+      by (simp add: Cons.hyps[OF `xs1 \<noteq> []` `ys \<noteq> []`])
+  qed
+qed auto
+
+lemma hoare_skip_or_set_set:
+  "\<lbrace> st i \<and>* ps u \<and>* zero u \<and>* zero (u + 1) \<and>* tm (u + 2) x\<rbrace>
+         i:[skip_or_set]:j
+   \<lbrace> st j \<and>* ps (u + 2) \<and>* one u \<and>* zero (u + 1) \<and>* tm (u + 2) x\<rbrace>"
+  apply(unfold skip_or_set_def)
+  apply(rule_tac q = "st j \<and>* ps (u + 2) \<and>* tm (u + 2) x \<and>* one u \<and>* zero (u + 1)" 
+    in tm.post_weaken)
+  apply(rule hoare_bone_1)
+  apply hsteps
+  by (auto simp:sep_conj_ac, sep_cancel+, smt)
+
+lemma hoare_skip_or_set_set_gen[step]:
+  assumes "u = v" "w = v + 1" "x = v + 2"
+  shows "\<lbrace>st i \<and>* ps u \<and>* zero v \<and>* zero w \<and>* tm x xv\<rbrace>
+                   i:[skip_or_set]:j
+         \<lbrace>st j \<and>* ps x \<and>* one v \<and>* zero w \<and>* tm x xv\<rbrace>"
+  by (unfold assms, rule hoare_skip_or_set_set)
+
+lemma hoare_skip_or_set_skip:
+  "\<lbrace> st i \<and>* ps u \<and>* reps u v [k] \<and>* zero (v + 1)\<rbrace>
+         i:[skip_or_set]:j
+   \<lbrace> st j \<and>*  ps (v + 2) \<and>* reps u v [k] \<and>* zero (v + 1)\<rbrace>"
+proof -
+   show ?thesis
+     apply(unfold skip_or_set_def, unfold reps.simps, simp add:sep_conj_cond)
+     apply(rule tm.pre_condI, simp)
+     apply(rule_tac p = "st i \<and>* ps u \<and>* one u \<and>* ones (u + 1) (u + int k) \<and>* 
+                             zero (u + int k + 1)" 
+                   in tm.pre_stren)
+     apply (rule_tac q = "st j \<and>* ps (u + int k + 2) \<and>* 
+                          one u \<and>* ones (u + 1) (u + int k) \<and>* zero (u + int k + 1)
+              " in tm.post_weaken)
+     apply (rule hoare_bone_2)
+     apply (rule_tac p = " st i \<and>* ps u \<and>* ones u (u + int k) \<and>* zero (u + int k + 1) 
+       " in tm.pre_stren)
+     apply hsteps
+     apply (simp add:sep_conj_ac, sep_cancel+, auto simp:sep_conj_ac ones_simps)
+     by (sep_cancel+, smt)
+ qed
+
+lemma hoare_skip_or_set_skip_gen[step]:
+  assumes "u = v" "x = w + 1"
+  shows  "\<lbrace> st i \<and>* ps u \<and>* reps v w [k] \<and>* zero x\<rbrace>
+                  i:[skip_or_set]:j
+          \<lbrace> st j \<and>*  ps (w + 2) \<and>* reps v w [k] \<and>* zero x\<rbrace>"
+  by (unfold assms, rule hoare_skip_or_set_skip)
+
+
+definition "if_reps_z e = (move_right;
+                              bone (move_left; jmp e) (move_left)
+                             )"
+
+lemma hoare_if_reps_z_true:
+  assumes h: "k = 0"
+  shows 
+   "\<lbrace>st i \<and>* ps u \<and>* reps u v [k] \<and>* zero (v + 1)\<rbrace> 
+      i:[if_reps_z e]:j 
+    \<lbrace>st e \<and>* ps u \<and>* reps u v [k] \<and>* zero (v + 1)\<rbrace>"
+  apply (unfold reps.simps, simp add:sep_conj_cond)
+  apply (rule tm.pre_condI, simp add:h)
+  apply (unfold if_reps_z_def)
+  apply (simp add:ones_simps)
+  apply (hsteps)
+  apply (rule_tac p = "st j' \<and>* ps (u + 1) \<and>* zero (u + 1) \<and>* one u" in tm.pre_stren)
+  apply (rule hoare_bone_1_out)
+  by (hsteps)
+
+lemma hoare_if_reps_z_true_gen[step]:
+  assumes "k = 0" "u = v" "x = w + 1"
+  shows "\<lbrace>st i \<and>* ps u \<and>* reps v w [k] \<and>* zero x\<rbrace> 
+                  i:[if_reps_z e]:j 
+         \<lbrace>st e \<and>* ps u \<and>* reps v w [k] \<and>* zero x\<rbrace>"
+  by (unfold assms, rule hoare_if_reps_z_true, simp)
+
+lemma hoare_if_reps_z_false:
+  assumes h: "k \<noteq> 0"
+  shows 
+   "\<lbrace>st i \<and>* ps u \<and>* reps u v [k]\<rbrace> 
+      i:[if_reps_z e]:j 
+    \<lbrace>st j \<and>* ps u \<and>* reps u v [k]\<rbrace>"
+proof -
+  from h obtain k' where "k = Suc k'" by (metis not0_implies_Suc)
+  show ?thesis
+    apply (unfold `k = Suc k'`)
+    apply (simp add:sep_conj_cond, rule tm.pre_condI, simp)
+    apply (unfold if_reps_z_def)
+    apply (simp add:ones_simps)
+    apply hsteps
+    apply (rule_tac p = "st j' \<and>* ps (u + 1) \<and>* one (u + 1) \<and>* one u \<and>* 
+                          ones (2 + u) (u + (1 + int k'))" in tm.pre_stren)
+    apply (rule_tac hoare_bone_2)
+    by (hsteps)
+qed
+
+lemma hoare_if_reps_z_false_gen[step]:
+  assumes h: "k \<noteq> 0" "u = v"
+  shows 
+   "\<lbrace>st i \<and>* ps u \<and>* reps v w [k]\<rbrace> 
+      i:[if_reps_z e]:j 
+    \<lbrace>st j \<and>* ps u \<and>* reps v w [k]\<rbrace>"
+  by (unfold assms, rule hoare_if_reps_z_false[OF `k \<noteq> 0`])
+
+definition "if_reps_nz e = (move_right;
+                              bzero (move_left; jmp e) (move_left)
+                           )"
+
+lemma EXS_postI: 
+  assumes "\<lbrace>P\<rbrace> 
+            c
+           \<lbrace>Q x\<rbrace>"
+  shows "\<lbrace>P\<rbrace> 
+          c
+        \<lbrace>EXS x. Q x\<rbrace>"
+by (metis EXS_intro assms tm.hoare_adjust)
+
+lemma hoare_if_reps_nz_true:
+  assumes h: "k \<noteq> 0"
+  shows 
+   "\<lbrace>st i \<and>* ps u \<and>* reps u v [k]\<rbrace> 
+      i:[if_reps_nz e]:j 
+    \<lbrace>st e \<and>* ps u \<and>* reps u v [k]\<rbrace>"
+proof -
+  from h obtain k' where "k = Suc k'" by (metis not0_implies_Suc)
+  show ?thesis
+    apply (unfold `k = Suc k'`)
+    apply (unfold reps.simps, simp add:sep_conj_cond, rule tm.pre_condI, simp)
+    apply (unfold if_reps_nz_def)
+    apply (simp add:ones_simps)
+    apply hsteps
+    apply (rule_tac p = "st j' \<and>* ps (u + 1) \<and>* one (u + 1) \<and>* one u \<and>*
+                            ones (2 + u) (u + (1 + int k'))" in tm.pre_stren)
+    apply (rule hoare_bzero_1_out)
+    by hsteps
+qed
+
+
+lemma hoare_if_reps_nz_true_gen[step]:
+  assumes h: "k \<noteq> 0" "u = v"
+  shows 
+   "\<lbrace>st i \<and>* ps u \<and>* reps v w [k]\<rbrace> 
+      i:[if_reps_nz e]:j 
+    \<lbrace>st e \<and>* ps u \<and>* reps v w [k]\<rbrace>"
+  by (unfold assms, rule hoare_if_reps_nz_true[OF `k\<noteq> 0`])
+
+lemma hoare_if_reps_nz_false:
+  assumes h: "k = 0"
+  shows 
+   "\<lbrace>st i \<and>* ps u \<and>* reps u v [k] \<and>* zero (v + 1)\<rbrace> 
+      i:[if_reps_nz e]:j 
+    \<lbrace>st j \<and>* ps u \<and>* reps u v [k] \<and>* zero (v + 1)\<rbrace>"
+  apply (simp add:h sep_conj_cond)
+  apply (rule tm.pre_condI, simp)
+  apply (unfold if_reps_nz_def)
+  apply (simp add:ones_simps)
+  apply (hsteps)
+  apply (rule_tac p = "st j' \<and>* ps (u + 1) \<and>*  zero (u + 1) \<and>* one u" in tm.pre_stren)
+  apply (rule hoare_bzero_2)
+  by (hsteps)
+
+lemma hoare_if_reps_nz_false_gen[step]:
+  assumes h: "k = 0" "u = v" "x = w + 1"
+  shows 
+   "\<lbrace>st i \<and>* ps u \<and>* reps v w [k] \<and>* zero x\<rbrace> 
+      i:[if_reps_nz e]:j 
+    \<lbrace>st j \<and>* ps u \<and>* reps v w [k] \<and>* zero x\<rbrace>"
+  by (unfold assms, rule hoare_if_reps_nz_false, simp)
+
+definition "skip_or_sets n = tpg_fold (replicate n skip_or_set)"
+
+
+
+lemma hoare_skip_or_sets_set:
+  shows "\<lbrace>st i \<and>* ps u \<and>* zeros u (u + int (reps_len (replicate (Suc n) 0))) \<and>* 
+                                  tm (u + int (reps_len (replicate (Suc n) 0)) + 1) x\<rbrace> 
+            i:[skip_or_sets (Suc n)]:j 
+         \<lbrace>st j \<and>* ps (u + int (reps_len (replicate (Suc n) 0)) + 1) \<and>* 
+                     reps' u  (u + int (reps_len (replicate (Suc n) 0))) (replicate (Suc n) 0) \<and>*
+                                 tm (u + int (reps_len (replicate (Suc n) 0)) + 1) x \<rbrace>"
+proof(induct n arbitrary:i j u x)
+  case 0
+  from 0 show ?case
+    apply (simp add:reps'_def reps_len_def reps_ctnt_len_def reps_sep_len_def reps.simps)
+    apply (unfold skip_or_sets_def, simp add:tpg_fold_sg)
+    apply hsteps
+    by (auto simp:sep_conj_ac, smt cond_true_eq2 ones.simps sep_conj_left_commute)
+next
+    case (Suc n)
+    { fix n
+      have "listsum (replicate n (Suc 0)) = n"
+        by (induct n, auto)
+    } note eq_sum = this
+    have eq_len: "\<And>n. n \<noteq> 0 \<Longrightarrow> reps_len (replicate (Suc n) 0) = reps_len (replicate n 0) + 2"
+      by (simp add:reps_len_def reps_sep_len_def reps_ctnt_len_def)
+    have eq_zero: "\<And> u v. (zeros u (u + int (v + 2))) = 
+           (zeros u (u + (int v)) \<and>* zero (u + (int v) + 1) \<and>* zero (u + (int v) + 2))"
+      by (smt sep.mult_assoc zeros_rev)
+    hence eq_z: 
+      "zeros u (u + int (reps_len (replicate (Suc (Suc n)) 0)))  = 
+       (zeros u (u + int (reps_len (replicate (Suc n) 0))) \<and>*
+       zero ((u + int (reps_len (replicate (Suc n) 0))) + 1) \<and>* 
+       zero ((u + int (reps_len (replicate (Suc n) 0))) + 2))
+      " by (simp only:eq_len)
+    have hh: "\<And>x. (replicate (Suc (Suc n)) x) = (replicate (Suc n) x) @ [x]"
+      by (metis replicate_Suc replicate_append_same)
+    have hhh: "replicate (Suc n) skip_or_set \<noteq> []" "[skip_or_set] \<noteq> []" by auto
+    have eq_code: 
+          "(i :[ skip_or_sets (Suc (Suc n)) ]: j) = 
+           (i :[ (skip_or_sets (Suc n); skip_or_set) ]: j)"
+    proof(unfold skip_or_sets_def)
+      show "i :[ tpg_fold (replicate (Suc (Suc n)) skip_or_set) ]: j =
+               i :[ (tpg_fold (replicate (Suc n) skip_or_set) ; skip_or_set) ]: j"
+        apply (insert tpg_fold_app[OF hhh, of i j], unfold hh)
+        by (simp only:tpg_fold_sg)
+    qed
+    have "Suc n \<noteq> 0" by simp
+    show ?case 
+      apply (unfold eq_z eq_code)
+      apply (hstep Suc(1))
+      apply (unfold eq_len[OF `Suc n \<noteq> 0`])
+      apply hstep
+      apply (auto simp:sep_conj_ac)[1]
+      apply (sep_cancel+, prune) 
+      apply (fwd abs_ones)
+      apply ((fwd abs_reps')+, simp add:int_add_ac)
+      by (metis replicate_append_same)
+  qed
+
+lemma hoare_skip_or_sets_set_gen[step]:
+  assumes h: "p2 = p1" 
+             "p3 = p1 + int (reps_len (replicate (Suc n) 0))"
+             "p4 = p3 + 1"
+  shows "\<lbrace>st i \<and>* ps p1 \<and>* zeros p2 p3 \<and>* tm p4 x\<rbrace> 
+            i:[skip_or_sets (Suc n)]:j 
+         \<lbrace>st j \<and>* ps p4 \<and>* reps' p2  p3 (replicate (Suc n) 0) \<and>* tm p4 x\<rbrace>"
+  apply (unfold h)
+  by (rule hoare_skip_or_sets_set)
+
+declare reps.simps[simp del]
+
+lemma hoare_skip_or_sets_skip:
+  assumes h: "n < length ks"
+  shows "\<lbrace>st i \<and>* ps u \<and>* reps' u v (take n ks) \<and>* reps' (v + 1) w [ks!n] \<rbrace> 
+            i:[skip_or_sets (Suc n)]:j 
+         \<lbrace>st j \<and>* ps (w+1) \<and>* reps' u v (take n ks) \<and>* reps' (v + 1) w [ks!n]\<rbrace>"
+  using h
+proof(induct n arbitrary: i j u v w ks)
+  case 0
+  show ?case 
+    apply (subst (1 5) reps'_def, simp add:sep_conj_cond, intro tm.pre_condI, simp)
+    apply (unfold skip_or_sets_def, simp add:tpg_fold_sg)
+    apply (unfold reps'_def, simp del:reps.simps)
+    apply hsteps
+    by (sep_cancel+, smt+)
+next
+  case (Suc n)
+  from `Suc n < length ks` have "n < length ks" by auto
+  note h =  Suc(1) [OF this]
+  show ?case 
+    my_block
+      from `Suc n < length ks` 
+      have eq_take: "take (Suc n) ks = take n ks @ [ks!n]"
+        by (metis not_less_eq not_less_iff_gr_or_eq take_Suc_conv_app_nth)
+    my_block_end
+    apply (unfold this)
+    apply (subst reps'_append, simp add:sep_conj_exists, rule tm.precond_exI)
+    my_block
+      have "(i :[ skip_or_sets (Suc (Suc n)) ]: j) = 
+                 (i :[ (skip_or_sets (Suc n); skip_or_set )]: j)"
+      proof -
+        have eq_rep: 
+          "(replicate (Suc (Suc n)) skip_or_set) = ((replicate (Suc n) skip_or_set) @ [skip_or_set])"
+          by (metis replicate_Suc replicate_append_same)
+        have "replicate (Suc n) skip_or_set \<noteq> []" "[skip_or_set] \<noteq> []" by auto
+        from tpg_fold_app[OF this]
+        show ?thesis
+          by (unfold skip_or_sets_def eq_rep, simp del:replicate.simps add:tpg_fold_sg)
+      qed
+    my_block_end
+    apply (unfold this)
+    my_block
+       fix i j m 
+       have "\<lbrace>st i \<and>* ps u \<and>* (reps' u (m - 1) (take n ks) \<and>* reps' m v [ks ! n])\<rbrace> 
+                            i :[ (skip_or_sets (Suc n)) ]: j
+             \<lbrace>st j \<and>* ps (v + 1) \<and>* (reps' u (m - 1) (take n ks) \<and>* reps' m v [ks ! n])\<rbrace>"
+                  apply (rule h[THEN tm.hoare_adjust])
+                  by (sep_cancel+, auto)
+    my_block_end my_note h_c1 = this
+    my_block
+      fix j' j m 
+      have "\<lbrace>st j' \<and>* ps (v + 1) \<and>* reps' (v + 1) w [ks ! Suc n]\<rbrace> 
+                          j' :[ skip_or_set ]: j
+            \<lbrace>st j \<and>* ps (w + 1) \<and>* reps' (v + 1) w [ks ! Suc n]\<rbrace>"
+        apply (unfold reps'_def, simp)
+        apply (rule hoare_skip_or_set_skip[THEN tm.hoare_adjust])
+        by (sep_cancel+, smt)+
+    my_block_end
+    apply (hstep h_c1 this)+ 
+    by ((fwd abs_reps'), simp, sep_cancel+)
+qed
+
+lemma hoare_skip_or_sets_skip_gen[step]:
+  assumes h: "n < length ks" "u = v" "x = w + 1"
+  shows "\<lbrace>st i \<and>* ps u \<and>* reps' v w (take n ks) \<and>* reps' x y [ks!n] \<rbrace> 
+            i:[skip_or_sets (Suc n)]:j 
+         \<lbrace>st j \<and>* ps (y+1) \<and>* reps' v w (take n ks) \<and>* reps' x y [ks!n]\<rbrace>"
+  by (unfold assms, rule hoare_skip_or_sets_skip[OF `n < length ks`])
+
+lemma fam_conj_interv_simp:
+    "(fam_conj {(ia::int)<..} p) = ((p (ia + 1)) \<and>* fam_conj {ia + 1 <..} p)"
+by (smt Collect_cong fam_conj_insert_simp greaterThan_def 
+        greaterThan_eq_iff greaterThan_iff insertI1 
+        insert_compr lessThan_iff mem_Collect_eq)
+
+lemma zeros_fam_conj:
+  assumes "u \<le> v"
+  shows "(zeros u v \<and>* fam_conj {v<..} zero) = fam_conj {u - 1<..} zero"
+proof -
+  have "{u - 1<..v} ## {v <..}" by (auto simp:set_ins_def)
+  from fam_conj_disj_simp[OF this, symmetric]
+  have "(fam_conj {u - 1<..v} zero \<and>* fam_conj {v<..} zero) = fam_conj ({u - 1<..v} + {v<..}) zero" .
+  moreover 
+  from `u \<le> v` have eq_set: "{u - 1 <..} = {u - 1 <..v} + {v <..}" by (auto simp:set_ins_def)
+  moreover have "fam_conj {u - 1<..v} zero = zeros u v"
+  proof -
+    have "({u - 1<..v}) = ({u .. v})" by auto
+    moreover {
+      fix u v 
+      assume "u  \<le> (v::int)"
+      hence "fam_conj {u .. v} zero = zeros u v"
+      proof(induct rule:ones_induct)
+        case (Base i j)
+        thus ?case by auto
+      next
+        case (Step i j)
+        thus ?case
+        proof(cases "i = j") 
+          case True
+          show ?thesis
+            by (unfold True, simp add:fam_conj_simps)
+        next
+          case False 
+          with `i \<le> j` have hh: "i + 1 \<le> j" by auto
+          hence eq_set: "{i..j} = (insert i {i + 1 .. j})"
+            by (smt simp_from_to)
+          have "i \<notin> {i + 1 .. j}" by simp
+          from fam_conj_insert_simp[OF this, folded eq_set]
+          have "fam_conj {i..j} zero = (zero i \<and>* fam_conj {i + 1..j} zero)" .
+          with Step(2)[OF hh] Step
+          show ?thesis by simp
+        qed
+      qed
+    } 
+    moreover note this[OF `u  \<le> v`]
+    ultimately show ?thesis by simp
+  qed
+  ultimately show ?thesis by smt
+qed
+
+declare replicate.simps [simp del]
+
+lemma hoare_skip_or_sets_comb:
+  assumes "length ks \<le> n"
+  shows "\<lbrace>st i \<and>* ps u \<and>* reps u v ks \<and>* fam_conj {v<..} zero\<rbrace> 
+                i:[skip_or_sets (Suc n)]:j 
+         \<lbrace>st j \<and>* ps ((v + int (reps_len (list_ext n ks)) - int (reps_len ks))+ 2) \<and>* 
+          reps' u (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1) (list_ext n ks) \<and>*
+          fam_conj {(v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1)<..} zero \<rbrace>"
+proof(cases "ks = []")
+  case True
+  show ?thesis
+    apply (subst True, simp only:reps.simps sep_conj_cond)
+    apply (rule tm.pre_condI, simp)
+    apply (rule_tac p = "st i \<and>* ps (v + 1) \<and>*
+            zeros (v + 1) (v + 1 + int (reps_len (replicate (Suc n) 0))) \<and>*
+            tm (v + 2 + int (reps_len (replicate (Suc n) 0))) Bk \<and>* 
+            fam_conj {(v + 2 + int (reps_len (replicate (Suc n) 0)))<..} zero
+      " in tm.pre_stren)
+    apply hsteps
+    apply (auto simp:sep_conj_ac)[1]
+    apply (auto simp:sep_conj_ac)[2]
+    my_block
+      from True have "(list_ext n ks) = (replicate (Suc n) 0)"
+        by (metis append_Nil diff_zero list.size(3) list_ext_def)
+    my_block_end my_note le_red = this
+    my_block
+      from True have "(reps_len ks) = 0"
+        by (metis reps_len_nil)
+    my_block_end
+    apply (unfold this le_red, simp)
+    my_block
+      have "v + 2 + int (reps_len (replicate (Suc n) 0)) = 
+            v + int (reps_len (replicate (Suc n) 0)) + 2" by smt
+    my_block_end my_note eq_len = this
+    apply (unfold this)
+    apply (sep_cancel+)
+    apply (fold zero_def)
+    apply (subst fam_conj_interv_simp, simp)
+    apply (simp only:int_add_ac)
+    apply (simp only:sep_conj_ac, sep_cancel+)
+    my_block
+      have "v + 1 \<le> (2 + (v + int (reps_len (replicate (Suc n) 0))))" by simp
+      from zeros_fam_conj[OF this]
+      have "(fam_conj {v<..} zero) = (zeros (v + 1) (2 + (v + int (reps_len (replicate (Suc n) 0)))) \<and>*
+                                        fam_conj {2 + (v + int (reps_len (replicate (Suc n) 0)))<..} zero)"
+        by simp
+    my_block_end
+    apply (subst (asm) this, simp only:int_add_ac, sep_cancel+)
+    by (smt cond_true_eq2 sep.mult_assoc sep.mult_commute 
+            sep.mult_left_commute sep_conj_assoc sep_conj_commute 
+         sep_conj_left_commute zeros.simps zeros_rev)
+next 
+  case False
+  show ?thesis
+    my_block
+      have "(i:[skip_or_sets (Suc n)]:j) = 
+              (i:[(skip_or_sets (length ks);  skip_or_sets (Suc n - length ks))]:j)"
+        apply (unfold skip_or_sets_def)
+        my_block
+          have "(replicate (Suc n) skip_or_set) = 
+                   (replicate (length ks) skip_or_set @ (replicate (Suc n - length ks) skip_or_set))"
+            by (smt assms replicate_add)
+        my_block_end
+        apply (unfold this, rule tpg_fold_app, simp add:False)
+        by (insert `length ks \<le> n`, simp)
+    my_block_end
+    apply (unfold this)
+    my_block
+      from False have "length ks = (Suc (length ks - 1))" by simp
+    my_block_end
+    apply (subst (1) this)
+    my_block
+      from False
+      have "(reps u v ks \<and>* fam_conj {v<..} zero) =
+            (reps' u (v + 1) ks \<and>* fam_conj {v+1<..} zero)"
+        apply (unfold reps'_def, simp)
+        by (subst fam_conj_interv_simp, simp add:sep_conj_ac)
+    my_block_end
+    apply (unfold this) 
+    my_block
+      fix i j
+      have "\<lbrace>st i \<and>* ps u \<and>* reps' u (v + 1) ks \<rbrace> 
+                i :[ skip_or_sets (Suc (length ks - 1))]: j
+            \<lbrace>st j \<and>* ps (v + 2) \<and>* reps' u (v + 1) ks \<rbrace>"
+        my_block
+          have "ks = take (length ks - 1) ks @ [ks!(length ks - 1)]"
+            by (smt False drop_0 drop_eq_Nil id_take_nth_drop)  
+        my_block_end my_note eq_ks = this
+        apply (subst (1) this)
+        apply (unfold reps'_append, simp add:sep_conj_exists, rule tm.precond_exI)
+        my_block
+          from False have "(length ks - Suc 0) < length ks"
+            by (smt `length ks = Suc (length ks - 1)`)
+        my_block_end
+        apply hsteps
+        apply (subst eq_ks, unfold reps'_append, simp only:sep_conj_exists)
+        by (rule_tac x = m in EXS_intro, simp add:sep_conj_ac, sep_cancel+, smt)
+    my_block_end
+    apply (hstep this)
+    my_block
+      fix u n
+      have "(fam_conj {u <..} zero) = 
+         (zeros (u + 1) (u + int n + 1) \<and>* tm (u + int n + 2) Bk \<and>* fam_conj {(u + int n + 2)<..} zero)"
+        my_block
+          have "u + 1 \<le> (u + int n + 2)" by auto
+          from zeros_fam_conj[OF this, symmetric]
+          have "fam_conj {u<..} zero = (zeros (u + 1) (u + int n + 2) \<and>* fam_conj {u + int n + 2<..} zero)"
+            by simp
+        my_block_end
+        apply (subst this)
+        my_block
+          have "(zeros (u + 1) (u + int n + 2)) = 
+                   ((zeros (u + 1) (u + int n + 1) \<and>* tm (u + int n + 2) Bk))"
+            by (smt zero_def zeros_rev)
+        my_block_end
+        by (unfold this, auto simp:sep_conj_ac)
+    my_block_end
+    apply (subst (1) this[of _ "(reps_len (replicate (Suc (n - length ks)) 0))"])
+    my_block
+      from `length ks \<le> n`
+      have "Suc n - length ks = Suc (n - length ks)" by auto 
+    my_block_end my_note eq_suc = this
+    apply (subst this)
+    apply hsteps
+    apply (simp add: sep_conj_ac this, sep_cancel+)
+    apply (fwd abs_reps')+
+    my_block
+      have "(int (reps_len (replicate (Suc (n - length ks)) 0))) =
+            (int (reps_len (list_ext n ks)) - int (reps_len ks) - 1)"
+        apply (unfold list_ext_def eq_suc)
+        my_block
+          have "replicate (Suc (n - length ks)) 0 \<noteq> []" by simp
+        my_block_end
+        by (unfold reps_len_split[OF False this], simp)
+    my_block_end
+    apply (unfold this)
+    my_block
+      from `length ks \<le> n`
+      have "(ks @ replicate (Suc (n - length ks)) 0) =  (list_ext n ks)"
+        by (unfold list_ext_def, simp)
+    my_block_end
+    apply (unfold this, simp)
+    apply (subst fam_conj_interv_simp, unfold zero_def, simp, simp add:int_add_ac sep_conj_ac)
+    by (sep_cancel+, smt)
+qed
+
+lemma hoare_skip_or_sets_comb_gen:
+  assumes "length ks \<le> n" "u = v" "w = x"
+  shows "\<lbrace>st i \<and>* ps u \<and>* reps v w ks \<and>* fam_conj {x<..} zero\<rbrace> 
+                i:[skip_or_sets (Suc n)]:j 
+         \<lbrace>st j \<and>* ps ((x + int (reps_len (list_ext n ks)) - int (reps_len ks))+ 2) \<and>* 
+          reps' u (x + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1) (list_ext n ks) \<and>*
+          fam_conj {(x + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1)<..} zero \<rbrace>"
+  by (unfold assms, rule hoare_skip_or_sets_comb[OF `length ks \<le> n`])
+
+definition "locate n = (skip_or_sets (Suc n);
+                        move_left;
+                        move_left;
+                        left_until_zero;
+                        move_right
+                       )"
+
+lemma list_ext_tail_expand:
+  assumes h: "length ks \<le> a"
+  shows "list_ext a ks = take a (list_ext a ks) @ [(list_ext a ks)!a]"
+proof -
+  let ?l = "list_ext a ks"
+  from h have eq_len: "length ?l = Suc a"
+    by (smt list_ext_len_eq)
+  hence "?l \<noteq> []" by auto
+  hence "?l = take (length ?l - 1) ?l @ [?l!(length ?l - 1)]"
+    by (metis `length (list_ext a ks) = Suc a` diff_Suc_1 le_refl 
+                    lessI take_Suc_conv_app_nth take_all)
+  from this[unfolded eq_len]
+  show ?thesis by simp
+qed
+
+lemma reps'_nn_expand:
+  assumes "xs \<noteq> []"
+  shows "(reps' u v xs) = (reps u (v - 1) xs \<and>* zero v)"
+  by (metis assms reps'_def)
+
+lemma sep_conj_st1: "(p \<and>* st t \<and>* q) = (st t \<and>* p \<and>* q)"
+  by (simp only:sep_conj_ac)
+
+lemma sep_conj_st2: "(p \<and>* st t) = (st t \<and>* p)"
+  by (simp only:sep_conj_ac)
+
+lemma sep_conj_st3: "((st t \<and>* p) \<and>* r) = (st t \<and>* p \<and>* r)"
+  by (simp only:sep_conj_ac)
+
+lemma sep_conj_st4: "(EXS x. (st t \<and>* r x)) = ((st t) \<and>* (EXS x. r x))"
+  apply (unfold pred_ex_def, default+)
+  apply (safe)
+  apply (sep_cancel, auto)
+  by (auto elim!: sep_conjE intro!:sep_conjI)
+
+lemmas sep_conj_st = sep_conj_st1 sep_conj_st2 sep_conj_st3 sep_conj_st4
+
+lemma sep_conj_cond3 : "(<cond1> \<and>* <cond2>) = <(cond1 \<and> cond2)>"
+  by (smt cond_eqI cond_true_eq sep_conj_commute sep_conj_empty)
+
+lemma sep_conj_cond4 : "(<cond1> \<and>* <cond2> \<and>* r) = (<(cond1 \<and> cond2)> \<and>* r)"
+  by (metis Hoare_gen.sep_conj_cond3 Hoare_tm.sep_conj_cond3)
+
+lemmas sep_conj_cond = sep_conj_cond3 sep_conj_cond4 sep_conj_cond 
+
+lemma hoare_left_until_zero_reps: 
+  "\<lbrace>st i ** ps v ** zero u ** reps (u + 1) v [k]\<rbrace> 
+        i:[left_until_zero]:j
+   \<lbrace>st j ** ps u ** zero u ** reps (u + 1) v [k]\<rbrace>"
+  apply (unfold reps.simps, simp only:sep_conj_cond)
+  apply (rule tm.pre_condI, simp)
+  by hstep
+
+lemma hoare_left_until_zero_reps_gen[step]: 
+  assumes "u = x" "w = v + 1"
+  shows "\<lbrace>st i ** ps u ** zero v ** reps w x [k]\<rbrace> 
+                i:[left_until_zero]:j
+         \<lbrace>st j ** ps v ** zero v ** reps w x [k]\<rbrace>"
+  by (unfold assms, rule hoare_left_until_zero_reps)
+
+lemma reps_lenE:
+  assumes "reps u v ks s"
+  shows "( <(v = u + int (reps_len ks) - 1)> \<and>* reps u v ks ) s"
+proof(rule condI)
+  from reps_len_correct[OF assms] show "v = u + int (reps_len ks) - 1" .
+next
+  from assms show "reps u v ks s" .
+qed
+
+lemma condI1: 
+  assumes "p s" "b"
+  shows "(<b> \<and>* p) s"
+proof(rule condI[OF assms(2)])
+  from  assms(1) show "p s" .
+qed
+
+lemma hoare_locate_set:
+  assumes "length ks \<le> n"
+  shows "\<lbrace>st i \<and>* zero (u - 1) \<and>* ps u \<and>* reps u v ks \<and>* fam_conj {v<..} zero\<rbrace> 
+                i:[locate n]:j 
+         \<lbrace>st j \<and>* zero (u - 1) \<and>* 
+             (EXS m w. ps m \<and>* reps' u (m - 1) (take n (list_ext n ks)) \<and>* 
+                         reps m w [(list_ext n ks)!n] \<and>* fam_conj {w<..} zero)\<rbrace>"
+proof(cases "take n (list_ext n ks) = []")
+  case False
+  show ?thesis
+    apply (unfold locate_def)
+    apply (hstep hoare_skip_or_sets_comb_gen)
+    apply (subst (3) list_ext_tail_expand[OF `length ks \<le> n`])
+    apply (subst (1) reps'_append, simp add:sep_conj_exists)
+    apply (rule tm.precond_exI)
+    apply (subst (1) reps'_nn_expand[OF False]) 
+    apply (rule_tac p = "st j' \<and>* <(m = u + int (reps_len (take n (list_ext n ks))) + 1)> \<and>*
+            ps (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 2) \<and>*
+            ((reps u (m - 1 - 1) (take n (list_ext n ks)) \<and>* zero (m - 1)) \<and>*
+             reps' m (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1)
+              [list_ext n ks ! n]) \<and>*
+            fam_conj
+             {v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1<..}
+             zero \<and>*
+            zero (u - 1)" 
+      in tm.pre_stren)
+    my_block
+      have "[list_ext n ks ! n] \<noteq> []" by simp
+      from reps'_nn_expand[OF this]
+      have "(reps' m (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1) [list_ext n ks ! n]) =
+                (reps m (v + (int (reps_len (list_ext n ks)) - int (reps_len ks))) [list_ext n ks ! n] \<and>*
+                   zero (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1))" 
+        by simp
+    my_block_end 
+    apply (subst this)
+    my_block
+      have "(fam_conj {v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 1<..} zero) =
+             (zero (v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 2) \<and>* 
+              fam_conj {v + int (reps_len (list_ext n ks)) - int (reps_len ks) + 2<..} zero)"
+        by (subst fam_conj_interv_simp, smt)
+    my_block_end
+    apply (unfold this) 
+    apply (simp only:sep_conj_st)
+    apply hsteps
+    apply (auto simp:sep_conj_ac)[1]
+    apply (sep_cancel+)
+    apply (rule_tac x = m in EXS_intro)
+    apply (rule_tac x = "m + int (list_ext n ks ! n)" in EXS_intro)
+    apply (simp add:sep_conj_ac del:ones_simps, sep_cancel+)
+    apply (subst (asm) sep_conj_cond)+
+    apply (erule_tac condE, clarsimp, simp add:sep_conj_ac int_add_ac)
+    apply (fwd abs_reps')
+    apply (fwd abs_reps')
+    apply (simp add:sep_conj_ac int_add_ac)
+    apply (sep_cancel+)
+    apply (subst (asm) reps'_def, subst fam_conj_interv_simp, subst fam_conj_interv_simp, 
+           simp add:sep_conj_ac int_add_ac)
+    my_block
+      fix s
+      assume h: "(reps (1 + (u + int (reps_len (take n (list_ext n ks)))))
+             (v + (- int (reps_len ks) + int (reps_len (list_ext n ks)))) [list_ext n ks ! n]) s"
+      (is "?P s")
+      from reps_len_correct[OF this] list_ext_tail_expand[OF `length ks \<le> n`]
+      have hh: "v + (- int (reps_len ks) + 
+                    int (reps_len (take n (list_ext n ks) @ [list_ext n ks ! n]))) =
+                  1 + (u + int (reps_len (take n (list_ext n ks)))) + 
+                       int (reps_len [list_ext n ks ! n]) - 1"
+        by metis
+      have "[list_ext n ks ! n] \<noteq> []" by simp
+      from hh[unfolded reps_len_split[OF False this]]
+      have "v  =  u + (int (reps_len ks)) - 1"
+        by simp
+      from condI1[where p = ?P, OF h this]
+      have "(<(v = u + int (reps_len ks) - 1)> \<and>*
+             reps (1 + (u + int (reps_len (take n (list_ext n ks)))))
+             (v + (- int (reps_len ks) + int (reps_len (list_ext n ks)))) [list_ext n ks ! n]) s" .
+    my_block_end
+    apply (fwd this, (subst (asm) sep_conj_cond)+, erule condE, simp add:sep_conj_ac int_add_ac
+              reps_len_sg)
+    apply (fwd reps_lenE, (subst (asm) sep_conj_cond)+, erule condE, simp add:sep_conj_ac int_add_ac
+            reps_len_sg)
+    by (fwd reps_lenE, (subst (asm) sep_conj_cond)+, erule condE, simp add:sep_conj_ac)
+next
+  case True
+  show ?thesis
+    apply (unfold locate_def)
+    apply (hstep hoare_skip_or_sets_comb)
+    apply (subst (3) list_ext_tail_expand[OF `length ks \<le> n`])
+    apply (subst (1) reps'_append, simp add:sep_conj_exists)
+    apply (rule tm.precond_exI)
+    my_block
+      from True `length ks \<le> n`
+      have "ks = []" "n = 0"
+        apply (metis le0 le_antisym length_0_conv less_nat_zero_code list_ext_len take_eq_Nil)
+        by (smt True length_take list.size(3) list_ext_len)
+    my_block_end
+    apply (unfold True this)
+    apply (simp add:reps'_def list_ext_def reps.simps reps_len_def reps_sep_len_def 
+                 reps_ctnt_len_def
+      del:ones_simps)
+    apply (subst sep_conj_cond)+
+    apply (rule tm.pre_condI, simp del:ones_simps)
+    apply (subst fam_conj_interv_simp, simp add:sep_conj_st del:ones_simps)
+    apply (hsteps)
+    apply (auto simp:sep_conj_ac)[1]
+    apply (sep_cancel+)
+    apply (rule_tac x = "(v + int (listsum (replicate (Suc 0) (Suc 0))))" in EXS_intro)+
+    apply (simp only:sep_conj_ac, sep_cancel+)
+    apply (auto)
+    apply (subst fam_conj_interv_simp)
+    apply (subst fam_conj_interv_simp)
+    by smt
+qed
+
+lemma hoare_locate_set_gen[step]:
+  assumes "length ks \<le> n" 
+           "u = v - 1" "v = w" "x = y"
+  shows "\<lbrace>st i \<and>* zero u \<and>* ps v \<and>* reps w x ks \<and>* fam_conj {y<..} zero\<rbrace> 
+                i:[locate n]:j 
+         \<lbrace>st j \<and>* zero u \<and>* 
+             (EXS m w. ps m \<and>* reps' v (m - 1) (take n (list_ext n ks)) \<and>* 
+                         reps m w [(list_ext n ks)!n] \<and>* fam_conj {w<..} zero)\<rbrace>"
+  by (unfold assms, rule hoare_locate_set[OF `length ks \<le> n`])
+
+lemma hoare_locate_skip: 
+  assumes h: "n < length ks"
+  shows 
+   "\<lbrace>st i \<and>* ps u \<and>* zero (u - 1) \<and>* reps' u (v - 1) (take n ks) \<and>* reps' v w [ks!n] \<and>* tm (w + 1) x\<rbrace> 
+      i:[locate n]:j 
+    \<lbrace>st j \<and>* ps v \<and>* zero (u - 1) \<and>* reps' u (v - 1) (take n ks) \<and>* reps' v w [ks!n] \<and>* tm (w + 1) x\<rbrace>"
+proof -
+  show ?thesis
+    apply (unfold locate_def)
+    apply hsteps
+    apply (subst (2 4) reps'_def, simp add:reps.simps sep_conj_cond del:ones_simps)
+    apply (intro tm.pre_condI, simp del:ones_simps)
+    apply hsteps
+    apply (case_tac "(take n ks) = []", simp add:reps'_def sep_conj_cond del:ones_simps)
+    apply (rule tm.pre_condI, simp del:ones_simps)
+    apply hsteps
+    apply (simp del:ones_simps add:reps'_def)
+    by hsteps
+qed
+
+
+lemma hoare_locate_skip_gen[step]: 
+  assumes "n < length ks"
+          "v = u - 1" "w = u" "x = y - 1" "z' = z + 1"
+  shows 
+   "\<lbrace>st i \<and>* ps u \<and>* tm v Bk \<and>* reps' w x (take n ks) \<and>* reps' y z [ks!n] \<and>* tm z' vx\<rbrace> 
+      i:[locate n]:j 
+    \<lbrace>st j \<and>* ps y \<and>* tm v Bk \<and>* reps' w x (take n ks) \<and>* reps' y z [ks!n] \<and>* tm z' vx\<rbrace>"
+  by (unfold assms, fold zero_def, rule hoare_locate_skip[OF `n < length ks`])
+
+definition "Inc a = locate a; 
+                    right_until_zero; 
+                    move_right;
+                    shift_right;
+                    move_left;
+                    left_until_double_zero;
+                    write_one;
+                    left_until_double_zero;
+                    move_right;
+                    move_right
+                    "
+
+lemma ones_int_expand: "(ones m (m + int k)) = (one m \<and>* ones (m + 1) (m + int k))"
+  by (simp add:ones_simps)
+
+lemma reps_splitedI:
+  assumes h1: "(reps u v ks1 \<and>* zero (v + 1) \<and>* reps (v + 2) w ks2) s"
+  and h2: "ks1 \<noteq> []"
+  and h3: "ks2 \<noteq> []"
+  shows "(reps u w (ks1 @ ks2)) s"
+proof - 
+  from h2 h3
+  have "splited (ks1 @ ks2) ks1 ks2" by (auto simp:splited_def)
+  from h1 obtain s1 where 
+    "(reps u v ks1) s1" by (auto elim:sep_conjE)
+  from h1 obtain s2 where 
+    "(reps (v + 2) w ks2) s2" by (auto elim:sep_conjE)
+  from reps_len_correct[OF `(reps u v ks1) s1`] 
+  have eq_v: "v = u + int (reps_len ks1) - 1" .
+  from reps_len_correct[OF `(reps (v + 2) w ks2) s2`]
+  have eq_w: "w = v + 2 + int (reps_len ks2) - 1" .
+  from h1
+  have "(reps u (u + int (reps_len ks1) - 1) ks1 \<and>*
+         zero (u + int (reps_len ks1)) \<and>* reps (u + int (reps_len ks1) + 1) w ks2) s"
+    apply (unfold eq_v eq_w[unfolded eq_v])
+    by (sep_cancel+, smt)
+  thus ?thesis
+    by(unfold reps_splited[OF `splited (ks1 @ ks2) ks1 ks2`], simp)
+qed
+
+lemma reps_sucI:
+  assumes h: "(reps u v (xs@[x]) \<and>* one (1 + v)) s"
+  shows "(reps u (1 + v) (xs@[Suc x])) s"
+proof(cases "xs = []")
+  case True
+  from h obtain s' where "(reps u v (xs@[x])) s'" by (auto elim:sep_conjE)
+  from reps_len_correct[OF this] have " v = u + int (reps_len (xs @ [x])) - 1" .
+  with True have eq_v: "v = u + int x" by (simp add:reps_len_sg)
+  have eq_one1: "(one (1 + (u + int x)) \<and>* ones (u + 1) (u + int x)) = ones (u + 1) (1 + (u + int x))"
+    by (smt ones_rev sep.mult_commute)
+  from h show ?thesis
+    apply (unfold True, simp add:eq_v reps.simps sep_conj_cond sep_conj_ac ones_rev)
+    by (sep_cancel+, simp add: eq_one1, smt)
+next
+  case False
+  from h obtain s1 s2 where hh: "(reps u v (xs@[x])) s1" "s = s1 + s2" "s1 ## s2" "one (1 + v) s2"
+    by (auto elim:sep_conjE)
+  from hh(1)[unfolded reps_rev[OF False]]
+  obtain s11 s12 s13 where hhh:
+    "(reps u (v - int (x + 1) - 1) xs) s11"
+    "(zero (v - int (x + 1))) s12" "(ones (v - int x) v) s13"
+    "s11 ## (s12 + s13)" "s12 ## s13" "s1 = s11 + s12 + s13"
+    apply (atomize_elim)
+    apply(elim sep_conjE)+
+    apply (rule_tac x = xa in exI)
+    apply (rule_tac x = xaa in exI)
+    apply (rule_tac x = ya in exI)
+    apply (intro conjI, assumption+)
+    by (auto simp:set_ins_def)
+  show ?thesis
+  proof(rule reps_splitedI[where v = "(v - int (x + 1) - 1)"])
+    show "(reps u (v - int (x + 1) - 1) xs \<and>* zero (v - int (x + 1) - 1 + 1) \<and>* 
+                                    reps (v - int (x + 1) - 1 + 2) (1 + v) [Suc x]) s"
+    proof(rule sep_conjI)
+      from hhh(1) show "reps u (v - int (x + 1) - 1) xs s11" .
+    next
+      show "(zero (v - int (x + 1) - 1 + 1) \<and>* reps (v - int (x + 1) - 1 + 2) (1 + v) [Suc x]) (s12 + (s13 + s2))"
+      proof(rule sep_conjI)
+        from hhh(2) show "zero (v - int (x + 1) - 1 + 1) s12" by smt
+      next
+        from hh(4) hhh(3)
+        show "reps (v - int (x + 1) - 1 + 2) (1 + v) [Suc x] (s13 + s2)"
+          apply (simp add:reps.simps del:ones_simps add:ones_rev)
+          by (smt hh(3) hh(4) hhh(4) hhh(5) hhh(6) sep_add_disjD sep_conjI sep_disj_addI1)
+      next
+        show "s12 ## s13 + s2" 
+          by (metis hh(3) hhh(4) hhh(5) hhh(6) sep_add_commute sep_add_disjD 
+              sep_add_disjI2 sep_disj_addD sep_disj_addD1 sep_disj_addI1 sep_disj_commuteI)
+      next
+        show "s12 + (s13 + s2) = s12 + (s13 + s2)" by metis 
+      qed
+    next
+      show "s11 ## s12 + (s13 + s2)"
+        by (metis hh(3) hhh(4) hhh(5) hhh(6) sep_disj_addD1 sep_disj_addI1 sep_disj_addI3)
+    next
+      show "s = s11 + (s12 + (s13 + s2))"
+        by (smt hh(2) hh(3) hhh(4) hhh(5) hhh(6) sep_add_assoc sep_add_commute 
+             sep_add_disjD sep_add_disjI2 sep_disj_addD1 sep_disj_addD2 
+              sep_disj_addI1 sep_disj_addI3 sep_disj_commuteI)
+    qed
+  next
+    from False show "xs \<noteq> []" .
+  next
+    show "[Suc x] \<noteq> []" by simp
+  qed
+qed
+
+lemma cond_expand: "(<cond> \<and>* p) s = (cond \<and> (p s))"
+  by (metis (full_types) condD pasrt_def sep_conj_commuteI 
+             sep_conj_sep_emptyI sep_empty_def sep_globalise)
+
+lemma ones_rev1:
+  assumes "\<not> (1 + n) < m"
+  shows "(ones m n \<and>* one (1 + n)) = (ones m (1 + n))"
+  by (insert ones_rev[OF assms, simplified], simp)
+
+lemma reps_one_abs:
+  assumes "(reps u v [k] \<and>* one w) s"
+          "w = v + 1"
+  shows "(reps u w [Suc k]) s"
+  using assms unfolding assms
+  apply (simp add:reps.simps sep_conj_ac)
+  apply (subst (asm) sep_conj_cond)+
+  apply (erule condE, simp)
+  by (simp add:ones_rev sep_conj_ac, sep_cancel+, smt)
+
+lemma reps'_reps_abs:
+  assumes "(reps' u v xs \<and>* reps w x ys) s"
+          "w = v + 1"  "ys \<noteq> []"
+  shows "(reps u x (xs@ys)) s"
+proof(cases "xs = []")
+  case False
+  with assms
+  have h: "splited (xs@ys) xs ys" by (simp add:splited_def)
+  from assms(1)[unfolded assms(2)]
+  show ?thesis
+    apply (unfold reps_splited[OF h])
+    apply (insert False, unfold reps'_def, simp)
+    apply (fwd reps_lenE, (subst (asm) sep_conj_cond)+)
+    by (erule condE, simp)
+next
+  case True
+  with assms
+  show ?thesis
+    apply (simp add:reps'_def)
+    by (erule condE, simp)
+qed
+
+lemma reps_one_abs1:
+  assumes "(reps u v (xs@[k]) \<and>* one w) s"
+          "w = v + 1"
+  shows "(reps u w (xs@[Suc k])) s"
+proof(cases "xs = []")
+  case True
+  with assms reps_one_abs
+  show ?thesis by simp
+next
+  case False
+  hence "splited (xs@[k]) xs [k]" by (simp add:splited_def)
+  from assms(1)[unfolded reps_splited[OF this] assms(2)]
+  show ?thesis
+    apply (fwd reps_one_abs)
+    apply (fwd reps_reps'_abs) 
+    apply (fwd reps'_reps_abs)
+    by (simp add:assms)
+qed
+  
+lemma tm_hoare_inc00: 
+  assumes h: "a < length ks \<and> ks ! a = v"
+  shows "\<lbrace>st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u ia ks \<and>* fam_conj {ia<..} zero\<rbrace> 
+    i :[ Inc a ]: j
+    \<lbrace>st j \<and>*
+     ps u \<and>*
+     zero (u - 2) \<and>*
+     zero (u - 1) \<and>*
+     reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1) (ks[a := Suc v]) \<and>*
+     fam_conj {(ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1)<..} zero\<rbrace>"
+  (is "\<lbrace> ?P \<rbrace> ?code \<lbrace> ?Q \<rbrace>")
+proof -
+  from h have "a < length ks" "ks ! a = v" by auto
+  from list_nth_expand[OF `a < length ks`]
+  have eq_ks: "ks = take a ks @ [ks!a] @ drop (Suc a) ks" .
+  from `a < length ks` have "ks \<noteq> []" by auto
+  hence "(reps u ia ks \<and>* zero (ia + 1)) = reps' u (ia + 1) ks"
+    by (simp add:reps'_def)
+  also from eq_ks have "\<dots> = reps' u (ia + 1) (take a ks @ [ks!a] @ drop (Suc a) ks)" by simp
+  also have "\<dots>  = (EXS m. reps' u (m - 1) (take a ks) \<and>* 
+                     reps' m (ia + 1) (ks ! a # drop (Suc a) ks))"
+    by (simp add:reps'_append)
+  also have "\<dots> = (EXS m. reps' u (m - 1) (take a ks) \<and>* 
+                     reps' m (ia + 1) ([ks ! a] @ drop (Suc a) ks))"
+    by simp
+  also have "\<dots> = (EXS m ma. reps' u (m - 1) (take a ks) \<and>*
+                       reps' m (ma - 1) [ks ! a] \<and>* reps' ma (ia + 1) (drop (Suc a) ks))"
+    by (simp only:reps'_append sep_conj_exists)
+  finally have eq_s: "(reps u ia ks \<and>* zero (ia + 1)) = \<dots>" .
+  { fix m ma
+    have eq_u: "-1 + u = u - 1" by smt
+    have " \<lbrace>st i \<and>*
+            ps u \<and>*
+            zero (u - 2) \<and>*
+            zero (u - 1) \<and>*
+            (reps' u (m - 1) (take a ks) \<and>*
+             reps' m (ma - 1) [ks ! a] \<and>* reps' ma (ia + 1) (drop (Suc a) ks)) \<and>*
+            fam_conj {ia + 1<..} zero\<rbrace> 
+           i :[ Inc a ]: j
+           \<lbrace>st j \<and>*
+            ps u \<and>*
+            zero (u - 2) \<and>*
+            zero (u - 1) \<and>*
+            reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1) (ks[a := Suc v]) \<and>*
+            fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1<..} zero\<rbrace>"
+    proof(cases "(drop (Suc a) ks) = []")
+      case True
+      { fix hc
+        have eq_1: "(1 + (m + int (ks ! a))) = (m + int (ks ! a) + 1)" by simp
+        have eq_2: "take a ks @ [Suc (ks ! a)] = ks[a := Suc v]"
+          apply (subst (3) eq_ks, unfold True, simp)
+          by (metis True append_Nil2 eq_ks h upd_conv_take_nth_drop)
+        assume h: "(fam_conj {1 + (m + int (ks ! a))<..} zero \<and>* 
+                      reps u (1 + (m + int (ks ! a))) (take a ks @ [Suc (ks ! a)])) hc"
+        hence "(fam_conj {m + int (ks ! a) + 1<..} zero \<and>* reps u (m + int (ks ! a) + 1) (ks[a := Suc v])) hc"
+          by (unfold eq_1 eq_2 , sep_cancel+)
+      } note eq_fam = this
+      show ?thesis
+        apply (unfold Inc_def, subst (3) reps'_def, simp add:True sep_conj_cond)
+        apply (intro tm.pre_condI, simp)
+        apply (subst fam_conj_interv_simp, simp, subst (3) zero_def)
+        apply hsteps
+        apply (subst reps'_sg, simp add:sep_conj_cond del:ones_simps)
+        apply (rule tm.pre_condI, simp del:ones_simps)
+        apply hsteps
+        apply (rule_tac p = "
+          st j' \<and>* ps (1 + (m + int (ks ! a))) \<and>* zero (u - 1) \<and>* zero (u - 2) \<and>*
+                   reps u (1 + (m + int (ks ! a))) (take a ks @ [Suc (ks!a)]) 
+                            \<and>* fam_conj {1 + (m + int (ks ! a))<..} zero
+          " in tm.pre_stren)
+        apply (hsteps)
+        apply (simp add:sep_conj_ac list_ext_lt[OF `a < length ks`], sep_cancel+)
+        apply (fwd eq_fam, sep_cancel+)
+        apply (simp del:ones_simps add:sep_conj_ac)
+        apply (sep_cancel+, prune)
+        apply ((fwd abs_reps')+, simp)
+        apply (fwd reps_one_abs abs_reps')+
+        apply (subst (asm) reps'_def, simp)
+        by (subst fam_conj_interv_simp, simp add:sep_conj_ac)
+    next 
+      case False
+      then obtain k' ks' where eq_drop: "drop (Suc a) ks = [k']@ks'"
+        by (metis append_Cons append_Nil list.exhaust)
+      from `a < length ks`
+      have eq_ks: "ks[a := Suc v] = (take a ks @ [Suc (ks ! a)]) @ (drop (Suc a) ks)"
+        apply (fold `ks!a = v`)
+        by (metis append_Cons append_Nil append_assoc upd_conv_take_nth_drop)
+      show ?thesis
+        apply (unfold Inc_def)
+        apply (unfold Inc_def eq_drop reps'_append, simp add:sep_conj_exists del:ones_simps)
+        apply (rule tm.precond_exI, subst (2) reps'_sg)
+        apply (subst sep_conj_cond)+
+        apply (subst (1) ones_int_expand)
+        apply (rule tm.pre_condI, simp del:ones_simps)
+        apply hsteps
+        (* apply (hsteps hoare_locate_skip_gen[OF `a < length ks`]) *)
+        apply (subst reps'_sg, simp add:sep_conj_cond del:ones_simps)
+        apply (rule tm.pre_condI, simp del:ones_simps)
+        apply hsteps
+        apply (rule_tac p = "st j' \<and>*
+                ps (2 + (m + int (ks ! a))) \<and>*
+                reps (2 + (m + int (ks ! a))) ia (drop (Suc a) ks) \<and>* zero (ia + 1) \<and>* zero (ia + 2) \<and>*
+                reps u (m + int (ks ! a)) (take a ks @ [ks!a]) \<and>* zero (1 + (m + int (ks ! a))) \<and>*
+                zero (u - 2) \<and>* zero (u - 1) \<and>* fam_conj {ia + 2<..} zero
+          " in tm.pre_stren)
+        apply (hsteps hoare_shift_right_cons_gen[OF False]
+                hoare_left_until_double_zero_gen[OF False])
+        apply (rule_tac p = 
+          "st j' \<and>* ps (1 + (m + int (ks ! a))) \<and>*
+          zero (u - 2) \<and>* zero (u - 1) \<and>* 
+          reps u (1 + (m + int (ks ! a))) (take a ks @ [Suc (ks ! a)]) \<and>*
+          zero (2 + (m + int (ks ! a))) \<and>*
+          reps (3 + (m + int (ks ! a))) (ia + 1) (drop (Suc a) ks) \<and>* fam_conj {ia + 1<..} zero
+          " in tm.pre_stren)
+        apply (hsteps)
+        apply (simp add:sep_conj_ac, sep_cancel+)
+        apply (unfold list_ext_lt[OF `a < length ks`], simp)
+        apply (fwd abs_reps')+ 
+        apply(fwd reps'_reps_abs)
+        apply (subst eq_ks, simp)
+        apply (sep_cancel+) 
+        apply (thin_tac "mb = 4 + (m + (int (ks ! a) + int k'))")
+        apply (thin_tac "ma = 2 + (m + int (ks ! a))", prune)
+        apply (simp add: int_add_ac sep_conj_ac, sep_cancel+)
+        apply (fwd reps_one_abs1, subst fam_conj_interv_simp, simp, sep_cancel+, smt)
+        apply (fwd abs_ones)+
+        apply (fwd abs_reps')
+        apply (fwd abs_reps')
+        apply (fwd abs_reps')
+        apply (fwd abs_reps')
+        apply (unfold eq_drop, simp add:int_add_ac sep_conj_ac)
+        apply (sep_cancel+)
+        apply (fwd  reps'_abs[where xs = "take a ks"])
+        apply (fwd reps'_abs[where xs = "[k']"])
+        apply (unfold reps'_def, simp add:int_add_ac sep_conj_ac)
+        apply (sep_cancel+)
+        by (subst (asm) fam_conj_interv_simp, simp, smt)
+      qed
+  } note h = this
+  show ?thesis
+    apply (subst fam_conj_interv_simp)
+    apply (rule_tac p = "st i \<and>*  ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* 
+                              (reps u ia ks \<and>* zero (ia + 1)) \<and>* fam_conj {ia + 1<..} zero" 
+      in tm.pre_stren)
+    apply (unfold eq_s, simp only:sep_conj_exists)
+    apply (intro tm.precond_exI h)
+    by (sep_cancel+, unfold eq_s, simp)
+qed
+
+declare ones_simps [simp del]
+
+lemma tm_hoare_inc01:
+  assumes "length ks \<le> a \<and> v = 0"
+  shows 
+   "\<lbrace>st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u ia ks \<and>* fam_conj {ia<..} zero\<rbrace> 
+       i :[ Inc a ]: j
+    \<lbrace>st j \<and>*
+     ps u \<and>*
+     zero (u - 2) \<and>*
+     zero (u - 1) \<and>*
+     reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1) ((list_ext a ks)[a := Suc v]) \<and>*
+     fam_conj {(ia + int (reps_len (list_ext a ks)) - int (reps_len ks) + 1)<..} zero\<rbrace>"
+proof -
+  from assms have "length ks \<le> a" "v = 0" by auto
+  show ?thesis
+    apply (rule_tac p = "
+      st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* (reps u ia ks \<and>* <(ia = u + int (reps_len ks) - 1)>) \<and>* 
+             fam_conj {ia<..} zero
+      " in tm.pre_stren)
+    apply (subst sep_conj_cond)+
+    apply (rule tm.pre_condI, simp)
+    apply (unfold Inc_def)
+    apply hstep
+    (* apply (hstep hoare_locate_set_gen[OF `length ks \<le> a`]) *)
+    apply (simp only:sep_conj_exists)
+    apply (intro tm.precond_exI)
+    my_block
+      fix m w
+      have "reps m w [list_ext a ks ! a] =
+            (ones m (m + int (list_ext a ks ! a)) \<and>* <(w = m + int (list_ext a ks ! a))>)"
+        by (simp add:reps.simps)
+    my_block_end
+    apply (unfold this)
+    apply (subst sep_conj_cond)+
+    apply (rule tm.pre_condI, simp)
+    apply (subst fam_conj_interv_simp)
+    apply (hsteps)
+    apply (subst fam_conj_interv_simp, simp)
+    apply (hsteps)
+    apply (rule_tac p = "st j' \<and>* ps (m + int (list_ext a ks ! a) + 1) \<and>*
+                           zero (u - 2) \<and>* zero (u - 1) \<and>* 
+                           reps u (m + int (list_ext a ks ! a) + 1) 
+                                ((take a (list_ext a ks))@[Suc (list_ext a ks ! a)]) \<and>*
+                           fam_conj {(m + int (list_ext a ks ! a) + 1)<..} zero
+                         " in tm.pre_stren)
+    apply (hsteps)
+    apply (simp add:sep_conj_ac, sep_cancel+)
+    apply (unfold `v = 0`, prune)
+    my_block
+      from `length ks \<le> a` have "list_ext a ks ! a = 0"
+        by (metis le_refl list_ext_tail)
+      from `length ks \<le> a` have "a < length (list_ext a ks)"
+        by (metis list_ext_len)
+      from reps_len_suc[OF this] 
+      have eq_len: "int (reps_len (list_ext a ks)) = 
+                        int (reps_len (list_ext a ks[a := Suc (list_ext a ks ! a)])) - 1" 
+        by smt
+      fix m w hc
+      assume h: "(fam_conj {m + int (list_ext a ks ! a) + 1<..} zero \<and>*
+                 reps u (m + int (list_ext a ks ! a) + 1) (take a (list_ext a ks) @ [Suc (list_ext a ks ! a)]))
+                 hc"
+      then obtain s where 
+        "(reps u (m + int (list_ext a ks ! a) + 1) (take a (list_ext a ks) @ [Suc (list_ext a ks ! a)])) s"
+        by (auto dest!:sep_conjD)
+      from reps_len_correct[OF this]
+      have "m  = u + int (reps_len (take a (list_ext a ks) @ [Suc (list_ext a ks ! a)])) 
+                        - int (list_ext a ks ! a) - 2" by smt
+      from h [unfolded this]
+      have "(fam_conj {u + int (reps_len (list_ext a ks))<..} zero \<and>*
+           reps u (u + int (reps_len (list_ext a ks))) (list_ext a ks[a := Suc 0]))
+           hc"
+        apply (unfold eq_len, simp)
+        my_block
+          from `a < length (list_ext a ks)`
+          have "take a (list_ext a ks) @ [Suc (list_ext a ks ! a)] = 
+                list_ext a ks[a := Suc (list_ext a ks ! a)]"
+            by (smt `list_ext a ks ! a = 0` assms length_take list_ext_tail_expand list_update_length)
+        my_block_end
+        apply (unfold this)
+        my_block
+          have "-1 + (u + int (reps_len (list_ext a ks[a := Suc (list_ext a ks ! a)]))) = 
+                u + (int (reps_len (list_ext a ks[a := Suc (list_ext a ks ! a)])) - 1)" by simp
+        my_block_end
+        apply (unfold this)
+        apply (sep_cancel+)
+        by (unfold `(list_ext a ks ! a) = 0`, simp)
+    my_block_end
+    apply (rule this, assumption)
+    apply (simp only:sep_conj_ac, sep_cancel+)+
+    apply (fwd abs_reps')+
+    apply (fwd reps_one_abs) 
+    apply (fwd reps'_reps_abs)
+    apply (simp add:int_add_ac sep_conj_ac)
+    apply (sep_cancel+)
+    apply (subst fam_conj_interv_simp, simp add:sep_conj_ac, smt)
+    apply (fwd reps_lenE, (subst (asm) sep_conj_cond)+, erule condE, simp)
+    by (sep_cancel+)
+qed
+
+definition "Dec a e  = (TL continue. 
+                          (locate a; 
+                           if_reps_nz continue;
+                           left_until_double_zero;
+                           move_right;
+                           move_right;
+                           jmp e);
+                          (TLabel continue;
+                           right_until_zero; 
+                           move_left;
+                           write_zero;
+                           move_right;
+                           move_right;
+                           shift_left;
+                           move_left;
+                           move_left;
+                           move_left;
+                           left_until_double_zero;
+                           move_right;
+                           move_right))"
+
+lemma cond_eq: "((<b> \<and>* p) s) = (b \<and> (p s))"
+proof
+  assume "(<b> \<and>* p) s"
+  from condD[OF this] show " b \<and> p s" .
+next
+  assume "b \<and> p s"
+  hence b and "p s" by auto
+  from `b` have "(<b>) 0" by (auto simp:pasrt_def)
+  moreover have "s = 0 + s" by auto
+  moreover have "0 ## s" by auto
+  moreover note `p s`
+  ultimately show "(<b> \<and>* p) s" by (auto intro!:sep_conjI)
+qed
+
+lemma tm_hoare_dec_fail00:
+  assumes "a < length ks \<and> ks ! a = 0"
+  shows "\<lbrace>st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u ia ks \<and>* fam_conj {ia<..} zero\<rbrace>  
+             i :[ Dec a e ]: j
+         \<lbrace>st e \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>*
+          reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks)) (list_ext a ks[a := 0]) \<and>*
+          fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) <..} zero\<rbrace>"
+proof -
+  from assms have "a < length ks" "ks!a = 0" by auto
+  from list_nth_expand[OF `a < length ks`]
+  have eq_ks: "ks = take a ks @ [ks ! a] @ drop (Suc a) ks" .
+  show ?thesis
+  proof(cases " drop (Suc a) ks = []")
+    case False
+    then obtain k' ks' where eq_drop: "drop (Suc a) ks = [k']@ks'"
+      by (metis append_Cons append_Nil list.exhaust)
+    show ?thesis
+      apply (unfold Dec_def, intro t_hoare_local)
+      apply (subst tassemble_to.simps(2), rule tm.code_exI, rule tm.code_extension)
+      apply (subst (1) eq_ks)
+      my_block
+        have "(reps u ia (take a ks @ [ks ! a] @ drop (Suc a) ks) \<and>* fam_conj {ia<..} zero) = 
+              (reps' u (ia + 1) ((take a ks @ [ks ! a]) @ drop (Suc a) ks) \<and>* fam_conj {ia + 1<..} zero)"
+          apply (subst fam_conj_interv_simp)
+          by (unfold reps'_def, simp add:sep_conj_ac)
+      my_block_end
+      apply (unfold this)
+      apply (subst reps'_append)
+      apply (unfold eq_drop)
+      apply (subst (2) reps'_append)
+      apply (simp only:sep_conj_exists, intro tm.precond_exI)
+      apply (subst (2) reps'_def, simp add:reps.simps ones_simps)
+      apply (subst reps'_append, simp only:sep_conj_exists, intro tm.precond_exI)
+      apply hstep
+      (* apply (hstep hoare_locate_skip_gen[OF `a < length ks`]) *)
+      my_block
+        fix m mb
+        have "(reps' mb (m - 1) [ks ! a]) = (reps mb (m - 2) [ks!a] \<and>* zero (m - 1))"
+          by (simp add:reps'_def, smt)
+      my_block_end
+      apply (unfold this)
+      apply hstep
+      (* apply (hstep hoare_if_reps_nz_false_gen[OF `ks!a = 0`]) *)
+      apply (simp only:reps.simps(2), simp add:`ks!a = 0`)
+      apply (rule_tac p = "st j'b \<and>*
+        ps mb \<and>*
+        reps u mb ((take a ks)@[ks ! a]) \<and>* <(m - 2 = mb)> \<and>*
+        zero (m - 1) \<and>*
+        zero (u - 1) \<and>*
+        one m \<and>*
+        zero (u - 2) \<and>*
+        ones (m + 1) (m + int k') \<and>*
+        <(-2 + ma = m + int k')> \<and>* zero (ma - 1) \<and>* reps' ma (ia + 1) ks' \<and>* fam_conj {ia + 1<..} zero"
+        in tm.pre_stren)
+      apply hsteps
+      apply (simp add:sep_conj_ac, sep_cancel+) 
+      apply (subgoal_tac "m + int k' = ma - 2", simp)
+      apply (fwd abs_ones)+
+      apply (subst (asm) sep_conj_cond)+
+      apply (erule condE, auto)
+      apply (fwd abs_reps')+
+      apply (subgoal_tac "ma = m + int k' + 2", simp)
+      apply (fwd abs_reps')+
+      my_block
+        from `a < length ks`
+        have "list_ext a ks = ks" by (auto simp:list_ext_def)
+      my_block_end
+      apply (simp add:this)
+      apply (subst eq_ks, simp add:eq_drop `ks!a = 0`)
+      apply (subst (asm) reps'_def, simp)
+      apply (subst fam_conj_interv_simp, simp add:sep_conj_ac, sep_cancel+)
+      apply (metis append_Cons assms eq_Nil_appendI eq_drop eq_ks list_update_id)
+      apply (clarsimp)
+      apply (subst (asm) sep_conj_cond)+
+      apply (erule condE, clarsimp)
+      apply (subst (asm) sep_conj_cond)+
+      apply (erule condE, clarsimp)
+      apply (simp add:sep_conj_ac, sep_cancel+)
+      apply (fwd abs_reps')+
+      by (fwd reps'_reps_abs, simp add:`ks!a = 0`)
+  next 
+    case True
+    show ?thesis
+      apply (unfold Dec_def, intro t_hoare_local)
+      apply (subst tassemble_to.simps(2), rule tm.code_exI, rule tm.code_extension)
+      apply (subst (1) eq_ks, unfold True, simp)
+      my_block
+        have "(reps u ia (take a ks @ [ks ! a]) \<and>* fam_conj {ia<..} zero) = 
+              (reps' u (ia + 1) (take a ks @ [ks ! a]) \<and>* fam_conj {ia + 1<..} zero)"
+          apply (unfold reps'_def, subst fam_conj_interv_simp)
+          by (simp add:sep_conj_ac)
+      my_block_end
+      apply (subst (1) this)
+      apply (subst reps'_append)
+      apply (simp only:sep_conj_exists, intro tm.precond_exI)
+      apply (subst fam_conj_interv_simp, simp) 
+      my_block
+        have "(zero (2 + ia)) = (tm (2 + ia) Bk)"
+          by (simp add:zero_def)
+      my_block_end my_note eq_z = this
+      apply hstep
+      (* apply (hstep hoare_locate_skip_gen[OF `a < length ks`]) *)
+      my_block
+        fix m 
+        have "(reps' m (ia + 1) [ks ! a]) = (reps m ia [ks!a] \<and>* zero (ia + 1))"
+          by (simp add:reps'_def)
+      my_block_end
+      apply (unfold this, prune)
+      apply hstep
+      (* apply (hstep hoare_if_reps_nz_false_gen[OF `ks!a = 0`]) *)
+      apply (simp only:reps.simps(2), simp add:`ks!a = 0`)
+      apply (rule_tac p = "st j'b \<and>* ps m \<and>* (reps u m ((take a ks)@[ks!a]) \<and>* <(ia = m)>) 
+                              \<and>* zero (ia + 1) \<and>* zero (u - 1) \<and>*  
+                              zero (2 + ia) \<and>* zero (u - 2) \<and>* fam_conj {2 + ia<..} zero"
+        in tm.pre_stren)
+      apply hsteps
+      apply (simp add:sep_conj_ac)
+      apply ((subst (asm) sep_conj_cond)+, erule condE, simp)
+      my_block
+        from `a < length ks`  have "list_ext a ks = ks" by (metis list_ext_lt) 
+      my_block_end
+      apply (unfold this, simp)
+      apply (subst fam_conj_interv_simp)
+      apply (subst fam_conj_interv_simp, simp)
+      apply (simp only:sep_conj_ac, sep_cancel+)
+      apply (subst eq_ks, unfold True `ks!a = 0`, simp)
+      apply (metis True append_Nil2 assms eq_ks list_update_same_conv) 
+      apply (simp add:sep_conj_ac)
+      apply (subst (asm) sep_conj_cond)+
+      apply (erule condE, simp, thin_tac "ia = m")
+      apply (fwd abs_reps')+
+      apply (simp add:sep_conj_ac int_add_ac, sep_cancel+)
+      apply (unfold reps'_def, simp)
+      by (metis sep.mult_commute)
+  qed
+qed
+
+lemma tm_hoare_dec_fail01:
+  assumes "length ks \<le> a"
+  shows "\<lbrace>st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u ia ks \<and>* fam_conj {ia<..} zero\<rbrace>  
+                       i :[ Dec a e ]: j
+         \<lbrace>st e \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>*
+          reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks)) (list_ext a ks[a := 0]) \<and>*
+          fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) <..} zero\<rbrace>"
+  apply (unfold Dec_def, intro t_hoare_local)
+  apply (subst tassemble_to.simps(2), rule tm.code_exI, rule tm.code_extension)
+  apply (rule_tac p = "st i \<and>* ps u \<and>* zero (u - 2) \<and>*
+                       zero (u - 1) \<and>* reps u ia ks \<and>* fam_conj {ia<..} zero \<and>* 
+                       <(ia = u + int (reps_len ks) - 1)>" in tm.pre_stren)
+  apply hstep
+  (* apply (hstep hoare_locate_set_gen[OF `length ks \<le> a`]) *)
+  apply (simp only:sep_conj_exists, intro tm.precond_exI)
+  my_block
+    from assms
+    have "list_ext a ks ! a = 0" by (metis le_refl list_ext_tail) 
+  my_block_end my_note is_z = this
+  apply (subst fam_conj_interv_simp)
+  apply hstep
+  (* apply (hstep hoare_if_reps_nz_false_gen[OF is_z]) *)
+  apply (unfold is_z)
+  apply (subst (1) reps.simps)
+  apply (rule_tac p = "st j'b \<and>* ps m \<and>*  reps u m (take a (list_ext a ks) @ [0]) \<and>* zero (w + 1) \<and>*
+                         <(w = m + int 0)> \<and>* zero (u - 1) \<and>* 
+                         fam_conj {w + 1<..} zero \<and>* zero (u - 2) \<and>* 
+                         <(ia = u + int (reps_len ks) - 1)>" in tm.pre_stren)
+  my_block
+    have "(take a (list_ext a ks)) @ [0] \<noteq> []" by simp
+  my_block_end
+  apply hsteps
+  (* apply (hsteps hoare_left_until_double_zero_gen[OF this]) *)
+  apply (simp add:sep_conj_ac)
+  apply prune
+  apply (subst (asm) sep_conj_cond)+
+  apply (elim condE, simp add:sep_conj_ac, prune)
+  my_block
+    fix m w ha
+    assume h1: "w = m \<and> ia = u + int (reps_len ks) - 1"
+      and  h: "(ps u \<and>*
+              st e \<and>*
+              zero (u - 1) \<and>*
+              zero (m + 1) \<and>*
+              fam_conj {m + 1<..} zero \<and>* zero (u - 2) \<and>* reps u m (take a (list_ext a ks) @ [0])) ha"
+    from h1 have eq_w: "w = m" and eq_ia: "ia = u + int (reps_len ks) - 1" by auto
+    from h obtain s' where "reps u m (take a (list_ext a ks) @ [0]) s'"
+      by (auto dest!:sep_conjD)
+    from reps_len_correct[OF this] 
+    have eq_m: "m = u + int (reps_len (take a (list_ext a ks) @ [0])) - 1" .
+    from h[unfolded eq_m, simplified]
+    have "(ps u \<and>*
+                st e \<and>*
+                zero (u - 1) \<and>*
+                zero (u - 2) \<and>*
+                fam_conj {u + (-1 + int (reps_len (list_ext a ks)))<..} zero \<and>*
+                reps u (u + (-1 + int (reps_len (list_ext a ks)))) (list_ext a ks[a := 0])) ha"
+      apply (sep_cancel+)
+      apply (subst fam_conj_interv_simp, simp)
+      my_block
+        from `length ks \<le> a` have "list_ext a ks[a := 0] = list_ext a ks"
+          by (metis is_z list_update_id)
+      my_block_end
+      apply (unfold this)
+      my_block
+        from `length ks \<le> a` is_z 
+        have "take a (list_ext a ks) @ [0] = list_ext a ks"
+          by (metis list_ext_tail_expand)
+      my_block_end
+      apply (unfold this)
+      by (simp add:sep_conj_ac, sep_cancel+, smt)
+  my_block_end
+  apply (rule this, assumption)
+  apply (sep_cancel+)[1]
+  apply (subst (asm) sep_conj_cond)+
+  apply (erule condE, prune, simp)
+  my_block
+    fix s m
+    assume "(reps' u (m - 1) (take a (list_ext a ks)) \<and>* ones m m \<and>* zero (m + 1)) s"
+    hence "reps' u (m + 1) (take a (list_ext a ks) @ [0]) s"
+      apply (unfold reps'_append)
+      apply (rule_tac x = m in EXS_intro)
+      by (subst (2) reps'_def, simp add:reps.simps)
+  my_block_end
+  apply (rotate_tac 1, fwd this)
+  apply (subst (asm) reps'_def, simp add:sep_conj_ac)
+  my_block
+    fix s
+    assume h: "(st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* 
+                   reps u ia ks \<and>* fam_conj {ia<..} zero) s"
+    then obtain s' where "reps u ia ks s'" by (auto dest!:sep_conjD)
+    from reps_len_correct[OF this] have eq_ia: "ia = u + int (reps_len ks) - 1" .
+    from h
+    have "(st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u ia ks \<and>*
+           fam_conj {ia<..} zero \<and>* <(ia = u + int (reps_len ks) - 1)>) s"
+      by (unfold eq_ia, simp)
+  my_block_end
+  by (rule this, assumption)
+
+lemma t_hoare_label1: 
+      "(\<And>l. l = i \<Longrightarrow> \<lbrace>st l \<and>* p\<rbrace>  l :[ c l ]: j \<lbrace>st k \<and>* q\<rbrace>) \<Longrightarrow>
+             \<lbrace>st l \<and>* p \<rbrace> 
+               i:[(TLabel l; c l)]:j
+             \<lbrace>st k \<and>* q\<rbrace>"
+by (unfold tassemble_to.simps, intro tm.code_exI tm.code_condI, clarify, auto)
+
+lemma tm_hoare_dec_fail1:
+  assumes "a < length ks \<and> ks ! a = 0 \<or> length ks \<le> a"
+  shows "\<lbrace>st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u ia ks \<and>* fam_conj {ia<..} zero\<rbrace>  
+                      i :[ Dec a e ]: j
+         \<lbrace>st e \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>*
+          reps u (ia + int (reps_len (list_ext a ks)) - int (reps_len ks)) (list_ext a ks[a := 0]) \<and>*
+         fam_conj {ia + int (reps_len (list_ext a ks)) - int (reps_len ks) <..} zero\<rbrace>"
+  using assms
+proof
+  assume "a < length ks \<and> ks ! a = 0"
+  thus ?thesis
+    by (rule tm_hoare_dec_fail00)
+next
+  assume "length ks \<le> a"
+  thus ?thesis
+    by (rule tm_hoare_dec_fail01)
+qed
+
+lemma shift_left_nil_gen[step]:
+  assumes "u = v"
+  shows "\<lbrace>st i \<and>* ps u \<and>* zero v\<rbrace> 
+              i :[shift_left]:j
+         \<lbrace>st j \<and>* ps u \<and>* zero v\<rbrace>"
+ apply(unfold assms shift_left_def, intro t_hoare_local t_hoare_label, clarify, 
+                 rule t_hoare_label_last, simp, clarify, prune, simp)
+ by hstep
+
+lemma tm_hoare_dec_suc1: 
+  assumes "a < length ks \<and> ks ! a = Suc v"
+  shows "\<lbrace>st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u ia ks \<and>* fam_conj {ia<..} zero\<rbrace> 
+                    i :[ Dec a e ]: j
+         \<lbrace>st j \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>*
+             reps u (ia - 1) (list_ext a ks[a := v]) \<and>*
+             fam_conj {ia - 1<..} zero\<rbrace>"
+proof -
+  from assms have "a < length ks" " ks ! a = Suc v" by auto
+  from list_nth_expand[OF `a < length ks`]
+  have eq_ks: "ks = take a ks @ [ks ! a] @ drop (Suc a) ks" .
+  show ?thesis
+  proof(cases " drop (Suc a) ks = []")
+    case False
+    then obtain k' ks' where eq_drop: "drop (Suc a) ks = [k']@ks'"
+      by (metis append_Cons append_Nil list.exhaust)
+    show ?thesis
+      apply (unfold Dec_def, intro t_hoare_local)
+      apply (subst tassemble_to.simps(2), rule tm.code_exI)
+      apply (subst (1) eq_ks)
+      my_block
+        have "(reps u ia (take a ks @ [ks ! a] @ drop (Suc a) ks) \<and>* fam_conj {ia<..} zero) = 
+              (reps' u (ia + 1) ((take a ks @ [ks ! a]) @ drop (Suc a) ks) \<and>* fam_conj {ia + 1<..} zero)"
+          apply (subst fam_conj_interv_simp)
+          by (unfold reps'_def, simp add:sep_conj_ac)
+      my_block_end
+      apply (unfold this)
+      apply (subst reps'_append)
+      apply (unfold eq_drop)
+      apply (subst (2) reps'_append)
+      apply (simp only:sep_conj_exists, intro tm.precond_exI)
+      apply (subst (2) reps'_def, simp add:reps.simps ones_simps)
+      apply (subst reps'_append, simp only:sep_conj_exists, intro tm.precond_exI)
+      apply (rule_tac q =
+       "st l \<and>*
+        ps mb \<and>*
+        zero (u - 1) \<and>*
+        reps' u (mb - 1) (take a ks) \<and>*
+        reps' mb (m - 1) [ks ! a] \<and>*
+        one m \<and>*
+        zero (u - 2) \<and>*
+        ones (m + 1) (m + int k') \<and>*
+        <(-2 + ma = m + int k')> \<and>* zero (ma - 1) \<and>* reps' ma (ia + 1) ks' \<and>* fam_conj {ia + 1<..} zero"
+      in tm.sequencing)
+      apply (rule tm.code_extension)
+      apply hstep
+      (* apply (hstep hoare_locate_skip_gen[OF `a < length ks`]) *)
+      apply (subst (2) reps'_def, simp)
+      my_block
+        fix i j l m mb
+        from `ks!a = (Suc v)` have "ks!a \<noteq> 0" by simp
+        from hoare_if_reps_nz_true[OF this, where u = mb and v = "m - 2"]
+        have "\<lbrace>st i \<and>* ps mb \<and>* reps mb (-2 + m) [ks ! a]\<rbrace>  
+                        i :[ if_reps_nz l ]: j
+              \<lbrace>st l \<and>* ps mb \<and>* reps mb (-2 + m) [ks ! a]\<rbrace>"
+          by smt
+      my_block_end
+      apply (hgoto this)
+      apply (simp add:sep_conj_ac, sep_cancel+)
+      apply (subst reps'_def, simp add:sep_conj_ac)
+      apply (rule tm.code_extension1)
+      apply (rule t_hoare_label1, simp, prune)
+      apply (subst (2) reps'_def, simp add:reps.simps)
+      apply (rule_tac p = "st j' \<and>* ps mb \<and>* zero (u - 1) \<and>* reps' u (mb - 1) (take a ks) \<and>*
+        ((ones mb (mb + int (ks ! a)) \<and>* <(-2 + m = mb + int (ks ! a))>) \<and>* zero (mb + int (ks ! a) + 1)) \<and>*
+          one (mb + int (ks ! a) + 2) \<and>* zero (u - 2) \<and>* 
+          ones (mb + int (ks ! a) + 3) (mb + int (ks ! a) + int k' + 2) \<and>*
+        <(-2 + ma = m + int k')> \<and>* zero (ma - 1) \<and>* reps' ma (ia + 1) ks' \<and>* fam_conj {ia + 1<..} zero
+        " in tm.pre_stren)
+      apply hsteps 
+      (* apply (simp add:sep_conj_ac) *)
+      apply (unfold `ks!a = Suc v`)
+      my_block
+        fix mb
+        have "(ones mb (mb + int (Suc v))) = (ones mb (mb + int v) \<and>* one (mb + int (Suc v)))"
+          by (simp add:ones_rev)
+      my_block_end
+      apply (unfold this, prune)
+      apply hsteps
+      apply (rule_tac p = "st j'a \<and>* 
+               ps (mb + int (Suc v) + 2) \<and>* zero (mb + int (Suc v) + 1) \<and>*
+               reps (mb + int (Suc v) + 2) ia (drop (Suc a) ks) \<and>* zero (ia + 1) \<and>* zero (ia + 2) \<and>*
+        zero (mb + int (Suc v)) \<and>*
+        ones mb (mb + int v) \<and>*
+        zero (u - 1) \<and>*
+        reps' u (mb - 1) (take a ks) \<and>*
+        zero (u - 2) \<and>* fam_conj {ia + 2<..} zero
+        " in tm.pre_stren) 
+      apply hsteps
+      (* apply (hsteps hoare_shift_left_cons_gen[OF False]) *)
+      apply (rule_tac p = "st j'a \<and>* ps (ia - 1) \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* 
+                           reps u (ia - 1) (take a ks @ [v] @ drop (Suc a) ks) \<and>*
+                           zero ia \<and>* zero (ia + 1) \<and>* zero (ia + 2) \<and>*
+                           fam_conj {ia + 2<..} zero
+        " in tm.pre_stren)
+      apply hsteps
+      apply (simp add:sep_conj_ac)
+      apply (subst fam_conj_interv_simp)
+      apply (subst fam_conj_interv_simp)
+      apply (subst fam_conj_interv_simp)
+      apply (simp add:sep_conj_ac)
+      apply (sep_cancel+)
+      my_block
+        have "take a ks @ v # drop (Suc a) ks = list_ext a ks[a := v]"
+        proof -
+          from `a < length ks` have "list_ext a ks = ks" by (metis list_ext_lt)
+          hence "list_ext a ks[a:=v] = ks[a:=v]" by simp
+          moreover from `a < length ks` have "ks[a:=v] = take a ks @ v # drop (Suc a) ks"
+            by (metis upd_conv_take_nth_drop)
+          ultimately show ?thesis by metis
+        qed
+      my_block_end
+      apply (unfold this, sep_cancel+, smt)
+      apply (simp add:sep_conj_ac)
+      apply (fwd abs_reps')+
+      apply (simp add:sep_conj_ac int_add_ac)
+      apply (sep_cancel+)
+      apply (subst (asm) reps'_def, simp add:sep_conj_ac)
+      apply (subst (asm) sep_conj_cond)+
+      apply (erule condE, clarsimp)
+      apply (simp add:sep_conj_ac, sep_cancel+)
+      apply (fwd abs_ones)+
+      apply (fwd abs_reps')+
+      apply (subst (asm) reps'_def, simp)
+      apply (subst (asm) fam_conj_interv_simp)
+      apply (simp add:sep_conj_ac int_add_ac eq_drop reps'_def)
+      apply (subst (asm) sep_conj_cond)+
+      apply (erule condE, clarsimp)
+      by (simp add:sep_conj_ac int_add_ac)
+  next
+    case True
+    show ?thesis
+      apply (unfold Dec_def, intro t_hoare_local)
+      apply (subst tassemble_to.simps(2), rule tm.code_exI)
+      apply (subst (1) eq_ks, simp add:True)
+      my_block
+        have "(reps u ia (take a ks @ [ks ! a]) \<and>* fam_conj {ia<..} zero) = 
+              (reps' u (ia + 1) (take a ks @ [ks ! a]) \<and>* fam_conj {ia + 1<..} zero)"
+          apply (subst fam_conj_interv_simp)
+          by (unfold reps'_def, simp add:sep_conj_ac)
+      my_block_end
+      apply (unfold this)
+      apply (subst reps'_append)
+      apply (simp only:sep_conj_exists, intro tm.precond_exI)
+      apply (rule_tac q = "st l \<and>* ps m \<and>* zero (u - 1) \<and>* reps' u (m - 1) (take a ks) \<and>*
+            reps' m (ia + 1) [ks ! a] \<and>* zero (2 + ia) \<and>* zero (u - 2) \<and>* fam_conj {2 + ia<..} zero"
+             in tm.sequencing)
+      apply (rule tm.code_extension)
+      apply (subst fam_conj_interv_simp, simp)
+      apply hsteps
+      (* apply (hstep hoare_locate_skip_gen[OF `a < length ks`]) *)
+      my_block
+        fix m
+        have "(reps' m (ia + 1) [ks ! a]) = 
+              (reps m ia [ks!a] \<and>* zero (ia + 1))"
+          by (unfold reps'_def, simp)
+      my_block_end
+      apply (unfold this)
+      my_block
+        fix i j l m
+        from `ks!a = (Suc v)` have "ks!a \<noteq> 0" by simp
+      my_block_end
+      apply (hgoto hoare_if_reps_nz_true_gen)
+      apply (rule tm.code_extension1)
+      apply (rule t_hoare_label1, simp)
+      apply (thin_tac "la = j'", prune)
+      apply (subst (1) reps.simps)
+      apply (subst sep_conj_cond)+
+      apply (rule tm.pre_condI, simp)
+      apply hsteps
+      apply (unfold `ks!a = Suc v`)
+      my_block
+        fix m
+        have "(ones m (m + int (Suc v))) = (ones m (m + int v) \<and>* one (m + int (Suc v)))"
+          by (simp add:ones_rev)
+      my_block_end
+      apply (unfold this)
+      apply hsteps 
+      apply (rule_tac p = "st j'a \<and>* ps (m + int v) \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* 
+                           reps u (m + int v) (take a ks @ [v]) \<and>* zero (m + (1 + int v)) \<and>*
+                           zero (2 + (m + int v)) \<and>* zero (3 + (m + int v)) \<and>*
+                           fam_conj {3 + (m + int v)<..} zero
+        " in tm.pre_stren)
+      apply hsteps
+      apply (simp add:sep_conj_ac, sep_cancel+)
+      my_block
+        have "take a ks @ [v] = list_ext a ks[a := v]"
+        proof -
+          from True `a < length ks` have "ks = take a ks @ [ks!a]"
+            by (metis append_Nil2 eq_ks)
+          hence "ks[a:=v] = take a ks @ [v]"
+            by (metis True `a < length ks` upd_conv_take_nth_drop)
+          moreover from `a < length ks` have "list_ext a ks = ks"
+            by (metis list_ext_lt)
+          ultimately show ?thesis by simp
+        qed
+      my_block_end my_note eq_l = this
+      apply (unfold this)
+      apply (subst fam_conj_interv_simp)
+      apply (subst fam_conj_interv_simp)
+      apply (subst fam_conj_interv_simp)
+      apply (simp add:sep_conj_ac, sep_cancel, smt)
+      apply (simp add:sep_conj_ac int_add_ac)+
+      apply (sep_cancel+)
+      apply (fwd abs_reps')+
+      apply (fwd reps'_reps_abs)
+      by (simp add:eq_l)
+  qed
+qed
+
+definition "cfill_until_one = (TL start exit.
+                                TLabel start;
+                                  if_one exit;
+                                  write_one;
+                                  move_left;
+                                  jmp start;
+                                TLabel exit
+                               )"
+
+lemma hoare_cfill_until_one:
+   "\<lbrace>st i \<and>* ps v \<and>* one (u - 1) \<and>* zeros u v\<rbrace> 
+              i :[ cfill_until_one ]: j
+    \<lbrace>st j \<and>* ps (u - 1) \<and>* ones (u - 1) v \<rbrace>"
+proof(induct u v rule:zeros_rev_induct)
+  case (Base x y)
+  thus ?case
+    apply (subst sep_conj_cond)+
+    apply (rule tm.pre_condI, simp add:ones_simps)
+    apply (unfold cfill_until_one_def, intro t_hoare_local t_hoare_label, rule t_hoare_label_last, simp+)
+    by hstep
+next
+  case (Step x y)
+  show ?case
+    apply (rule_tac q = "st i \<and>* ps (y - 1) \<and>* one (x - 1) \<and>* zeros x (y - 1) \<and>* one y" in tm.sequencing)
+    apply (subst cfill_until_one_def, intro t_hoare_local t_hoare_label, rule t_hoare_label_last, simp+)
+    apply hsteps
+    my_block
+      fix i j l
+      have "\<lbrace>st i \<and>* ps (y - 1) \<and>* one (x - 1) \<and>* zeros x (y - 1)\<rbrace>  
+              i :[ jmp l ]: j
+            \<lbrace>st l \<and>* ps (y - 1) \<and>* one (x - 1) \<and>* zeros x (y - 1)\<rbrace>"
+        apply (case_tac "(y - 1) < x", simp add:zeros_simps)
+        apply (subst sep_conj_cond)+
+        apply (rule tm.pre_condI, simp)
+        apply hstep
+        apply (drule_tac zeros_rev, simp)
+        by hstep
+    my_block_end
+    apply (hstep this)
+    (* The next half *)
+    apply (hstep Step(2), simp add:sep_conj_ac, sep_cancel+)
+    by (insert Step(1), simp add:ones_rev sep_conj_ac)
+qed
+
+definition "cmove = (TL start exit.
+                       TLabel start;
+                         left_until_zero;
+                         left_until_one;
+                         move_left;
+                         if_zero exit;
+                         move_right;
+                         write_zero;
+                         right_until_one;
+                         right_until_zero;
+                         write_one;
+                         jmp start;
+                     TLabel exit
+                    )"
+
+declare zeros.simps [simp del] zeros_simps[simp del]
+
+lemma hoare_cmove:
+  assumes "w \<le> k"
+  shows "\<lbrace>st i \<and>* ps (v + 2 + int w) \<and>* zero (u - 1) \<and>* 
+              reps u (v - int w) [k - w] \<and>* zeros (v - int w + 1) (v + 1) \<and>*
+              one (v + 2) \<and>* ones (v + 3) (v + 2 + int w) \<and>* zeros (v + 3 + int w)  (v + int(reps_len [k]) + 1)\<rbrace>
+                                 i :[cmove]: j
+          \<lbrace>st j \<and>* ps (u - 1) \<and>* zero (u - 1) \<and>* reps u u [0] \<and>* zeros (u + 1) (v + 1) \<and>*
+                                                                  reps (v + 2) (v + int (reps_len [k]) + 1) [k]\<rbrace>"
+  using assms
+proof(induct "k - w" arbitrary: w)
+  case (0 w)
+  hence "w = k" by auto
+  show ?case
+    apply (simp add: `w = k` del:zeros.simps zeros_simps)
+    apply (unfold cmove_def, intro t_hoare_local t_hoare_label, rule t_hoare_label_last, simp+)
+    apply (simp add:reps_len_def reps_sep_len_def reps_ctnt_len_def del:zeros_simps zeros.simps)
+    apply (rule_tac p = "st i \<and>* ps (v + 2 + int k) \<and>* zero (u - 1) \<and>*
+                         reps u u [0] \<and>* zeros (u + 1) (v + 1) \<and>*
+                         ones (v + 2) (v + 2 + int k) \<and>* zeros (v + 3 + int k) (2 + (v + int k)) \<and>*
+                         <(u = v - int k)>" 
+      in tm.pre_stren)
+    my_block
+      fix i j
+      have "\<lbrace>st i \<and>* ps (v + 2 + int k) \<and>* zeros (u + 1) (v + 1) \<and>* ones (v + 2) (v + 2 + int k) 
+                                                             \<and>* <(u = v - int k)>\<rbrace>
+                  i :[ left_until_zero ]: j
+            \<lbrace>st j \<and>* ps (v + 1) \<and>* zeros (u + 1) (v + 1) \<and>* ones (v + 2) (v + 2 + int k)
+                                                             \<and>* <(u = v - int k)>\<rbrace>"
+        apply (subst sep_conj_cond)+
+        apply (rule tm.pre_condI, simp)
+        my_block
+          have "(zeros (v - int k + 1) (v + 1)) = (zeros (v - int k + 1) v \<and>* zero (v + 1))"
+            by (simp only:zeros_rev, smt)
+        my_block_end
+        apply (unfold this)
+        by hsteps
+    my_block_end
+    apply (hstep this)
+    my_block
+      fix i j 
+      have "\<lbrace>st i \<and>* ps (v + 1) \<and>* reps u u [0] \<and>* zeros (u + 1) (v + 1)\<rbrace> 
+                i :[left_until_one]:j 
+            \<lbrace>st j \<and>* ps u \<and>* reps u u [0] \<and>* zeros (u + 1) (v + 1)\<rbrace>"
+        apply (simp add:reps.simps ones_simps)
+        by hsteps
+    my_block_end
+    apply (hsteps this)
+    apply ((subst (asm) sep_conj_cond)+, erule condE, clarsimp)
+    apply (fwd abs_reps')+
+    apply (simp only:sep_conj_ac int_add_ac, sep_cancel+)
+    apply (simp add:int_add_ac sep_conj_ac zeros_simps)
+    apply (simp add:sep_conj_ac int_add_ac, sep_cancel+)
+    apply (fwd reps_lenE)
+    apply (subst (asm) sep_conj_cond)+
+    apply (erule condE, clarsimp)
+    apply (subgoal_tac "v  = u + int k + int (reps_len [0]) - 1", clarsimp)
+    apply (simp add:reps_len_sg)
+    apply (fwd abs_ones)+
+    apply (fwd abs_reps')+
+    apply (simp add:sep_conj_ac int_add_ac)
+    apply (sep_cancel+)
+    apply (simp add:reps.simps, smt)
+    by (clarsimp)
+next
+  case (Suc k' w)
+  from `Suc k' = k - w` `w \<le> k` 
+  have h: "k' = k - (Suc w)" "Suc w \<le> k" by auto
+  show ?case
+    apply (rule tm.sequencing[OF _ Suc(1)[OF h(1, 2)]])
+    apply (unfold cmove_def, intro t_hoare_local t_hoare_label, rule t_hoare_label_last, simp+)
+    apply (simp add:reps_len_def reps_sep_len_def reps_ctnt_len_def del:zeros_simps zeros.simps)
+    my_block
+      fix i j
+      have "\<lbrace>st i \<and>* ps (v + 2 + int w) \<and>* zeros (v - int w + 1) (v + 1) \<and>*
+                               one (v + 2) \<and>* ones (v + 3) (v + 2 + int w) \<rbrace> 
+                    i :[left_until_zero]: j
+            \<lbrace>st j \<and>* ps (v + 1) \<and>* zeros (v - int w + 1) (v + 1) \<and>*
+                               one (v + 2) \<and>* ones (v + 3) (v + 2 + int w) \<rbrace>"
+        my_block
+          have "(one (v + 2) \<and>* ones (v + 3) (v + 2 + int w)) = 
+                 ones (v + 2) (v + 2 + int w)"
+            by (simp only:ones_simps, smt)
+        my_block_end
+        apply (unfold this)
+        my_block
+          have "(zeros (v - int w + 1) (v + 1)) = (zeros (v - int w + 1) v \<and>*  zero (v + 1))"
+            by (simp only:zeros_rev, simp)
+        my_block_end
+        apply (unfold this)
+        by hsteps
+    my_block_end
+    apply (hstep this)
+    my_block
+      fix i j
+      have "\<lbrace>st i \<and>* ps (v + 1) \<and>* reps u (v - int w) [k - w] \<and>* zeros (v - int w + 1) (v + 1)\<rbrace> 
+                 i :[left_until_one]: j 
+            \<lbrace>st j \<and>* ps (v - int w) \<and>* reps u (v - int w) [k - w] \<and>* zeros (v - int w + 1) (v + 1)\<rbrace>"
+        apply (simp add:reps.simps ones_rev)
+        apply (subst sep_conj_cond)+
+        apply (rule tm.pre_condI, clarsimp)
+        apply (subgoal_tac "u + int (k - w) = v - int w", simp)
+        defer
+        apply simp
+        by hsteps
+    my_block_end
+    apply (hstep this)
+    my_block
+      have "(reps u (v - int w) [k - w]) = (reps u (v - (1 + int w)) [k - Suc w] \<and>* one (v - int w))"
+        apply (subst (1 2) reps.simps)
+        apply (subst sep_conj_cond)+
+        my_block
+          have "((v - int w = u + int (k - w))) =
+                (v - (1 + int w) = u + int (k - Suc w))"
+            apply auto
+            apply (smt Suc.prems h(2))
+            by (smt Suc.prems h(2))
+        my_block_end
+        apply (simp add:this)
+        my_block
+          fix b p q
+          assume "(b \<Longrightarrow> (p::tassert) = q)"
+          have "(<b> \<and>* p) = (<b> \<and>* q)"
+            by (metis `b \<Longrightarrow> p = q` cond_eqI)
+        my_block_end
+        apply (rule this)
+        my_block
+          assume "v - (1 + int w) = u + int (k - Suc w)"
+          hence "v = 1 + int w +  u + int (k - Suc w)" by auto
+        my_block_end
+        apply (simp add:this)
+        my_block
+          have "\<not> (u + int (k - w)) < u" by auto
+        my_block_end
+        apply (unfold ones_rev[OF this])
+        my_block
+          from Suc (2, 3) have "(u + int (k - w) - 1) = (u + int (k - Suc w))"
+            by auto
+        my_block_end
+        apply (unfold this)
+        my_block
+          from Suc (2, 3) have "(u + int (k - w)) =  (1 + (u + int (k - Suc w)))"
+            by auto
+        my_block_end
+        by (unfold this, simp)
+    my_block_end
+    apply (unfold this)
+    my_block
+      fix i j
+      have "\<lbrace>st i \<and>* ps (v - int w) \<and>*
+                        (reps u (v - (1 + int w)) [k - Suc w] \<and>* one (v - int w))\<rbrace> 
+                 i :[ move_left]: j
+            \<lbrace>st j \<and>* ps (v - (1 + int w)) \<and>*
+                        (reps u (v - (1 + int w)) [k - Suc w] \<and>* one (v - int w))\<rbrace>"
+        apply (simp add:reps.simps ones_rev)
+        apply (subst sep_conj_cond)+
+        apply (rule tm.pre_condI, clarsimp)
+        apply (subgoal_tac " u + int (k - Suc w) = v - (1 + int w)", simp)
+        defer
+        apply simp
+        apply hsteps
+        by (simp add:sep_conj_ac, sep_cancel+, smt)
+    my_block_end
+    apply (hstep this)
+    my_block
+      fix i' j'
+      have "\<lbrace>st i' \<and>* ps (v - (1 + int w)) \<and>* reps u (v - (1 + int w)) [k - Suc w]\<rbrace> 
+               i' :[ if_zero j ]: j'
+            \<lbrace>st j' \<and>* ps (v - (1 + int w)) \<and>* reps u (v - (1 + int w)) [k - Suc w]\<rbrace>"
+        apply (simp add:reps.simps ones_rev)
+        apply (subst sep_conj_cond)+
+        apply (rule tm.pre_condI, clarsimp)
+        apply (subgoal_tac " u + int (k - Suc w) = v - (1 + int w)", simp)
+        defer
+        apply simp
+        by hstep
+    my_block_end
+    apply (hstep this)
+    my_block
+      fix i j
+      have "\<lbrace>st i \<and>* ps (v - (1 + int w)) \<and>*  reps u (v - (1 + int w)) [k - Suc w]\<rbrace> 
+                i :[ move_right ]: j 
+            \<lbrace>st j \<and>* ps (v - int w) \<and>*  reps u (v - (1 + int w)) [k - Suc w] \<rbrace>"
+        apply (simp add:reps.simps ones_rev)
+        apply (subst sep_conj_cond)+
+        apply (rule tm.pre_condI, clarsimp)
+        apply (subgoal_tac " u + int (k - Suc w) =  v - (1 + int w)", simp)
+        defer
+        apply simp
+        by hstep
+    my_block_end
+    apply (hsteps this)
+    my_block
+      fix i j
+      have "\<lbrace>st i \<and>* ps (v - int w) \<and>*  one (v + 2) \<and>* 
+                         zero (v - int w) \<and>* zeros (v - int w + 1) (v + 1)\<rbrace> 
+                 i :[right_until_one]: j
+            \<lbrace>st j \<and>* ps (v + 2) \<and>*  one (v + 2) \<and>*  zero (v - int w) \<and>* zeros (v - int w + 1) (v + 1)\<rbrace>"
+        my_block
+          have "(zero (v - int w) \<and>* zeros (v - int w + 1) (v + 1)) = 
+                    (zeros (v - int w) (v + 1))"
+            by (simp add:zeros_simps)
+        my_block_end
+        apply (unfold this)
+        by hsteps
+    my_block_end
+    apply (hstep this)
+    my_block
+      from Suc(2, 3) have "w < k" by auto
+      hence "(zeros (v + 3 + int w) (2 + (v + int k))) = 
+                  (zero (v + 3 + int w) \<and>* zeros (4 + (v + int w)) (2 + (v + int k)))"
+        by (simp add:zeros_simps)
+    my_block_end
+    apply (unfold this)
+    my_block
+      fix i j
+      have "\<lbrace>st i \<and>* ps (v + 2) \<and>* zero (v + 3 + int w) \<and>* zeros (4 + (v + int w)) (2 + (v + int k)) \<and>* 
+                                                        one (v + 2) \<and>* ones (v + 3) (v + 2 + int w)\<rbrace>
+                i :[right_until_zero]: j
+            \<lbrace>st j \<and>* ps (v + 3 + int w) \<and>* zero (v + 3 + int w) \<and>* zeros (4 + (v + int w)) (2 + (v + int k)) \<and>* 
+                                                        one (v + 2) \<and>* ones (v + 3) (v + 2 + int w)\<rbrace>"
+        my_block
+          have "(one (v + 2) \<and>* ones (v + 3) (v + 2 + int w)) =
+                (ones (v + 2) (v + 2 + int w))"
+            by (simp add:ones_simps, smt)
+        my_block_end
+        apply (unfold this)
+        by hsteps
+    my_block_end
+    apply (hsteps this, simp only:sep_conj_ac)
+    apply (sep_cancel+, simp add:sep_conj_ac)
+    my_block
+      fix s
+      assume "(zero (v - int w) \<and>* zeros (v - int w + 1) (v + 1)) s"
+      hence "zeros (v - int w) (v + 1) s"
+        by (simp add:zeros_simps)
+    my_block_end
+    apply (fwd this)
+    my_block
+      fix s
+      assume "(one (v + 3 + int w) \<and>* ones (v + 3) (v + 2 + int w)) s"
+      hence "ones (v + 3) (3 + (v + int w)) s"
+        by (simp add:ones_rev sep_conj_ac, smt)
+    my_block_end
+    apply (fwd this)
+    by (simp add:sep_conj_ac, smt)
+qed
+
+definition "cinit = (right_until_zero; move_right; write_one)"
+
+definition "copy = (cinit; cmove; move_right; move_right; right_until_one; move_left; move_left; cfill_until_one)"
+
+lemma hoare_copy:
+  shows
+   "\<lbrace>st i \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u v [k] \<and>* zero (v + 1) \<and>*
+                                                     zeros (v + 2) (v + int(reps_len [k]) + 1)\<rbrace>
+                                  i :[copy]: j
+    \<lbrace>st j \<and>* ps u \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>* reps u v [k] \<and>* zero (v + 1) \<and>* 
+                                                      reps (v + 2) (v + int (reps_len [k]) + 1) [k]\<rbrace>"
+  apply (unfold copy_def)
+  my_block
+    fix i j
+    have 
+       "\<lbrace>st i \<and>* ps u \<and>* reps u v [k] \<and>* zero (v + 1) \<and>* zeros (v + 2) (v + int(reps_len [k]) + 1)\<rbrace>
+                      i :[cinit]: j
+        \<lbrace>st j \<and>* ps (v + 2) \<and>* reps u v [k] \<and>* zero (v + 1) \<and>*
+                                           one (v + 2) \<and>* zeros (v + 3) (v + int(reps_len [k]) + 1)\<rbrace>"
+      apply (unfold cinit_def)
+      apply (simp add:reps.simps)
+      apply (subst sep_conj_cond)+
+      apply (rule tm.pre_condI, simp)
+      apply hsteps
+      apply (simp add:sep_conj_ac)
+      my_block
+        have "(zeros (u + int k + 2) (u + int k + int (reps_len [k]) + 1)) = 
+              (zero (u + int k + 2) \<and>*  zeros (u + int k + 3) (u + int k + int (reps_len [k]) + 1))"
+          by (smt reps_len_sg zeros_step_simp)
+      my_block_end
+      apply (unfold this)
+      apply hstep
+      by (simp add:sep_conj_ac, sep_cancel+, smt)
+  my_block_end
+  apply (hstep this)
+  apply (rule_tac p = "st j' \<and>* ps (v + 2) \<and>* reps u v [k] \<and>* zero (v + 1) \<and>*
+          one (v + 2) \<and>* zeros (v + 3) (v + int (reps_len [k]) + 1) \<and>* zero (u - 2) \<and>* zero (u - 1) \<and>*
+            <(v = u + int (reps_len [k]) - 1)>
+    " in tm.pre_stren)
+  my_block
+    fix i j
+    from hoare_cmove[where w = 0 and k = k and i = i and j = j and v = v and u = u]
+    have "\<lbrace>st i \<and>* ps (v + 2) \<and>* zero (u - 1) \<and>* reps u v [k] \<and>* zero (v + 1) \<and>*
+                                            one (v + 2) \<and>* zeros (v + 3) (v + int(reps_len [k]) + 1)\<rbrace>
+                      i :[cmove]: j
+          \<lbrace>st j \<and>* ps (u - 1) \<and>* zero (u - 1) \<and>* reps u u [0] \<and>* zeros (u + 1) (v + 1) \<and>*
+                                                       reps (v + 2) (v + int (reps_len [k]) + 1) [k]\<rbrace>"
+      by (auto simp:ones_simps zeros_simps)
+  my_block_end
+  apply (hstep this)
+  apply (hstep, simp)
+  my_block
+    have "reps u u [0] = one u" by (simp add:reps.simps ones_simps)
+  my_block_end my_note eq_repsz = this
+  apply (unfold this)
+  apply (hstep)
+  apply (subst reps.simps, simp add: ones_simps)
+  apply hsteps
+  apply (subst sep_conj_cond)+
+  apply (rule tm.pre_condI, simp del:zeros.simps zeros_simps)
+  apply (thin_tac "int (reps_len [k]) = 1 + int k \<and> v = u + int (reps_len [k]) - 1")
+  my_block
+    have "(zeros (u + 1) (u + int k + 1)) = (zeros (u + 1) (u + int k) \<and>* zero (u + int k + 1))"
+      by (simp only:zeros_rev, smt)
+  my_block_end
+  apply (unfold this)
+  apply (hstep, simp)
+  my_block
+    fix i j
+    from hoare_cfill_until_one[where v = "u + int k" and u = "u + 1"]
+    have "\<lbrace>st i \<and>* ps (u + int k) \<and>* one u \<and>* zeros (u + 1) (u + int k)\<rbrace> 
+              i :[ cfill_until_one ]: j
+          \<lbrace>st j \<and>* ps u \<and>* ones u (u + int k) \<rbrace>"
+      by simp
+  my_block_end
+  apply (hstep this, simp add:sep_conj_ac reps.simps ones_simps)
+  apply (simp add:sep_conj_ac reps.simps ones_simps)
+  apply (subst sep_conj_cond)+
+  apply (subst (asm) sep_conj_cond)+
+  apply (rule condI)
+  apply (erule condE, simp)
+  apply (simp add: reps_len_def reps_sep_len_def reps_ctnt_len_def)
+  apply (sep_cancel+)
+  by (erule condE, simp)
+
+end 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/My_block.thy	Thu Mar 06 13:28:38 2014 +0000
@@ -0,0 +1,36 @@
+theory My_block
+imports Main
+keywords 
+  "my_block" :: prf_script and 
+  "my_block_end" :: prf_script and 
+  "my_note" :: prf_decl
+begin
+
+ML {*
+  val my_block = Proof.assert_backward #> Proof.enter_forward #> Proof.begin_block
+*}
+
+ML {*
+  val my_block_end = Proof.end_block #> Proof.enter_backward
+*}
+
+ML {*
+  val _ = Outer_Syntax.command @{command_spec "my_block"} "begin my block"
+        (Scan.succeed (Toplevel.print o (Toplevel.proof my_block)))
+  val _ = Outer_Syntax.command @{command_spec "my_block_end"} "begin my block"
+        (Scan.succeed (Toplevel.print o (Toplevel.proof my_block_end)))
+*}
+
+ML {*
+  fun my_note x = 
+     Proof.assert_backward #> Proof.enter_forward #> Proof.note_thmss_cmd x #> Proof.enter_backward
+*}
+
+ML {*
+val _ =
+  Outer_Syntax.command @{command_spec "my_note"} "define facts"
+    (Parse_Spec.name_facts >> (Toplevel.print oo (Toplevel.proof o my_note)));
+*}
+
+end
+