updated turing_basic by Jian
authorChristian Urban <christian dot urban at kcl dot ac dot uk>
Fri, 18 Jan 2013 13:56:35 +0000
changeset 50 816e84ca16d6
parent 49 b388dceee892
child 51 6725c9c026f6
updated turing_basic by Jian
Paper/Paper.thy
Paper/document/root.bib
Paper/document/root.tex
paper.pdf
thys/turing_basic.thy
turing_basic.thy
--- a/Paper/Paper.thy	Fri Jan 18 13:03:09 2013 +0000
+++ b/Paper/Paper.thy	Fri Jan 18 13:56:35 2013 +0000
@@ -7,6 +7,9 @@
 hide_const (open) s 
 *)
 
+
+hide_const (open) Divides.adjust
+
 abbreviation
   "update2 p a \<equiv> update a p"
 
@@ -31,6 +34,8 @@
 
 section {* Introduction *}
 
+
+
 text {*
 
 %\noindent
@@ -46,7 +51,7 @@
 
 
 \noindent
-Suppose you want to mechanise a proof whether a predicate @{term P}, say, is
+Suppose you want to mechanise a proof about whether a predicate @{term P}, say, is
 decidable or not. Decidability of @{text P} usually amounts to showing
 whether \mbox{@{term "P \<or> \<not>P"}} holds. But this does \emph{not} work
 in Isabelle/HOL and other HOL theorem provers, since they are based on classical logic
@@ -61,7 +66,7 @@
 
 The only satisfying way out of this problem in a theorem prover based
 on classical logic is to formalise a theory of computability. Norrish
-provided such a formalisation for the HOL4 theorem prover. He choose
+provided such a formalisation for the HOL4. He choose
 the $\lambda$-calculus as the starting point for his formalisation of
 computability theory, because of its ``simplicity'' \cite[Page
 297]{Norrish11}.  Part of his formalisation is a clever infrastructure
@@ -71,7 +76,7 @@
  to have formalisations for more operational models of
 computations, such as Turing machines or register machines.  One
 reason is that many proofs in the literature use them.  He noted
-however that in the context of theorem provers \cite[Page 310]{Norrish11}:
+however that \cite[Page 310]{Norrish11}:
 
 \begin{quote}
 \it``If register machines are unappealing because of their 
@@ -85,7 +90,7 @@
 of register machines) and recursive functions. To see the difficulties
 involved with this work, one has to understand that Turing machine
 programs can be completely \emph{unstructured}, behaving 
-similar to Basic's infamous goto. This precludes in the
+similar to Basic's infamous goto \cite{Dijkstra68}. This precludes in the
 general case a compositional Hoare-style reasoning about Turing
 programs.  We provide such Hoare-rules for when it is possible to
 reason in a compositional manner (which is fortunately quite often), but also tackle 
@@ -152,9 +157,8 @@
 \noindent
 In this paper we follow the approach by Boolos et al \cite{Boolos87},
 which goes back to Post \cite{Post36}, where all Turing machines
-operate on tapes that contain only \emph{blank} or \emph{occupied} cells
-(represented by @{term Bk} and @{term Oc}, respectively, in our
-formalisation). Traditionally the content of a cell can be any
+operate on tapes that contain only \emph{blank} or \emph{occupied} cells. 
+Traditionally the content of a cell can be any
 character from a finite alphabet. Although computationally equivalent,
 the more restrictive notion of Turing machines in \cite{Boolos87} makes
 the reasoning more uniform. In addition some proofs \emph{about} Turing
@@ -186,8 +190,7 @@
 section {* Turing Machines *}
 
 text {* \noindent
-  Turing machines can be thought of as having a read-write-unit, also
-  referred to as \emph{head},
+  Turing machines can be thought of as having a \emph{head},
   ``gliding'' over a potentially infinite tape. Boolos et
   al~\cite{Boolos87} only consider tapes with cells being either blank
   or occupied, which we represent by a datatype having two
@@ -198,7 +201,7 @@
   convention that the head, abbreviated @{term hd}, of the right-list is
   the cell on which the head of the Turing machine currently operates. This can
   be pictured as follows:
-
+  %
   \begin{center}
   \begin{tikzpicture}
   \draw[very thick] (-3.0,0)   -- ( 3.0,0);
@@ -233,12 +236,16 @@
   the Turing machine can perform:
 
   \begin{center}
-  \begin{tabular}{rcl@ {\hspace{5mm}}l}
-  @{text "a"} & $::=$  & @{term "W0"} & write blank (@{term Bk})\\
-  & $\mid$ & @{term "W1"} & write occupied (@{term Oc})\\
-  & $\mid$ & @{term L} & move left\\
-  & $\mid$ & @{term R} & move right\\
-  & $\mid$ & @{term Nop} & do-nothing operation\\
+  \begin{tabular}[t]{@ {}rcl@ {\hspace{2mm}}l}
+  @{text "a"} & $::=$  & @{term "W0"} & (write blank, @{term Bk})\\
+  & $\mid$ & @{term "W1"} & (write occupied, @{term Oc})\\
+  \end{tabular}
+  \begin{tabular}[t]{rcl@ {\hspace{2mm}}l}
+  & $\mid$ & @{term L} & (move left)\\
+  & $\mid$ & @{term R} & (move right)\\
+  \end{tabular}
+  \begin{tabular}[t]{rcl@ {\hspace{2mm}}l@ {}}
+  & $\mid$ & @{term Nop} & (do-nothing operation)\\
   \end{tabular}
   \end{center}
 
@@ -263,15 +270,14 @@
   The first two clauses replace the head of the right-list
   with a new @{term Bk} or @{term Oc}, respectively. To see that
   these two clauses make sense in case where @{text r} is the empty
-  list, one has to know that the tail function, @{term tl}, is defined in
-  Isabelle/HOL
+  list, one has to know that the tail function, @{term tl}, is defined
   such that @{term "tl [] == []"} holds. The third clause 
   implements the move of the head one step to the left: we need
   to test if the left-list @{term l} is empty; if yes, then we just prepend a 
   blank cell to the right-list; otherwise we have to remove the
   head from the left-list and prepend it to the right-list. Similarly
   in the fourth clause for a right move action. The @{term Nop} operation
-  leaves the the tape unchanged (last clause).
+  leaves the the tape unchanged.
 
   %Note that our treatment of the tape is rather ``unsymmetric''---we
   %have the convention that the head of the right-list is where the
@@ -291,23 +297,24 @@
   %definition above, and by using the @{term update} function
   %cover uniformly all cases including corner cases.
 
-  Next we need to define the \emph{states} of a Turing machine.  Given
-  how little is usually said about how to represent them in informal
-  presentations, it might be surprising that in a theorem prover we
-  have to select carefully a representation. If we use the naive
-  representation where a Turing machine consists of a finite set of
-  states, then we will have difficulties composing two Turing
-  machines: we would need to combine two finite sets of states,
-  possibly renaming states apart whenever both machines share
-  states.\footnote{The usual disjoint union operation in Isabelle/HOL
-  cannot be used as it does not preserve types.} This renaming can be
-  quite cumbersome to reason about. Therefore we made the choice of
+  Next we need to define the \emph{states} of a Turing machine.  
+  %Given
+  %how little is usually said about how to represent them in informal
+  %presentations, it might be surprising that in a theorem prover we
+  %have to select carefully a representation. If we use the naive
+  %representation where a Turing machine consists of a finite set of
+  %states, then we will have difficulties composing two Turing
+  %machines: we would need to combine two finite sets of states,
+  %possibly renaming states apart whenever both machines share
+  %states.\footnote{The usual disjoint union operation in Isabelle/HOL
+  %cannot be used as it does not preserve types.} This renaming can be
+  %quite cumbersome to reason about. 
+  We followed the choice made by \cite{AspertiRicciotti12} 
   representing a state by a natural number and the states of a Turing
-  machine will always consist of the initial segment of natural
-  numbers starting from @{text 0} up to the number of states of the
-  machine. In doing so we can compose two Turing machine by
+  machine by the initial segment of natural numbers starting from @{text 0}.
+  In doing so we can compose two Turing machine by
   shifting the states of one by an appropriate amount to a higher
-  segment and adjusting some ``next states'' in the other.
+  segment and adjusting some ``next states'' in the other. {\it composition here?}
 
   An \emph{instruction} @{term i} of a Turing machine is a pair consisting of 
   an action and a natural number (the next state). A \emph{program} @{term p} of a Turing
@@ -333,7 +340,7 @@
   similarly the second component determines what should be done in
   case of reading @{term Oc}. We have the convention that the first
   state is always the \emph{starting state} of the Turing machine.
-  The zeroth state is special in that it will be used as the
+  The @{text 0}-state is special in that it will be used as the
   ``halting state''.  There are no instructions for the @{text
   0}-state, but it will always perform a @{term Nop}-operation and
   remain in the @{text 0}-state.  Unlike Asperti and Riccioti
@@ -352,9 +359,9 @@
   \begin{tabular}{l@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
   \multicolumn{3}{l}{@{thm fetch.simps(1)[where b=DUMMY]}}\\
   @{thm (lhs) fetch.simps(2)} & @{text "\<equiv>"} & @{text "case nth_of p (2 * s) of"}\\
-  \multicolumn{3}{@ {\hspace{1.4cm}}l}{@{text "None \<Rightarrow> (Nop, 0) | Some i \<Rightarrow> i"}}\\
+  \multicolumn{3}{@ {\hspace{4cm}}l}{@{text "None \<Rightarrow> (Nop, 0) | Some i \<Rightarrow> i"}}\\
   @{thm (lhs) fetch.simps(3)} & @{text "\<equiv>"} & @{text "case nth_of p (2 * s + 1) of"}\\
-  \multicolumn{3}{@ {\hspace{1.4cm}}l}{@{text "None \<Rightarrow> (Nop, 0) | Some i \<Rightarrow> i"}}
+  \multicolumn{3}{@ {\hspace{4cm}}l}{@{text "None \<Rightarrow> (Nop, 0) | Some i \<Rightarrow> i"}}
   \end{tabular}
   \end{center}
 
@@ -365,7 +372,7 @@
   (@{term None}-case). In doing so we slightly deviate from the description
   in \cite{Boolos87}: if their Turing machines transition to a non-existing
   state, then the computation is halted. We will transition in such cases
-  to the @{text 0}-state. However, with introducing the
+  to the @{text 0}-state.\footnote{\it However, with introducing the
   notion of \emph{well-formed} Turing machine programs we will later exclude such
   cases and make the  @{text 0}-state the only ``halting state''. A program 
   @{term p} is said to be well-formed if it satisfies
@@ -384,6 +391,7 @@
   state; the second that @{text p} has a @{term Bk} and @{term Oc} instruction for every 
   state, and the third that every next-state is one of the states mentioned in
   the program or being the @{text 0}-state.
+  }
 
   A \emph{configuration} @{term c} of a Turing machine is a state together with 
   a tape. This is written as @{text "(s, (l, r))"}. If we have a 
@@ -485,11 +493,8 @@
 
   \begin{center}
   \begin{tabular}{@ {}l@ {\hspace{1mm}}c@ {\hspace{1mm}}l@ {}}
-  @{thm (lhs) shift.simps} @{text "\<equiv>"}\\
-  \hspace{4mm}@{thm (rhs) shift.simps}\\
-  @{thm (lhs) adjust.simps} @{text "\<equiv>"}\\
-  \hspace{4mm}@{text "map (\<lambda> (a, s)."}\\
-  \hspace{14mm}@{text "(a, if s = 0 then length p div 2 + 1 else s)) p"}\\
+  @{thm (lhs) shift.simps} @{text "\<equiv>"} @{thm (rhs) shift.simps}\\
+  @{thm (lhs) adjust.simps} @{text "\<equiv>"} @{thm (rhs) adjust.simps}\\
   \end{tabular}
   \end{center}
 
@@ -503,7 +508,7 @@
   of two Turing machine programs @{text "p\<^isub>1"} and @{text "p\<^isub>2"}
 
   \begin{center}
-  @{thm tm_comp.simps[THEN eq_reflection]}
+  @{thm tm_comp.simps[where ?p1.0="p\<^isub>1" and ?p2.0="p\<^isub>2", THEN eq_reflection]}
   \end{center}
 
   \noindent
--- a/Paper/document/root.bib	Fri Jan 18 13:03:09 2013 +0000
+++ b/Paper/document/root.bib	Fri Jan 18 13:56:35 2013 +0000
@@ -66,3 +66,12 @@
   pages =        {103--105}
 }
 
+@article{Dijkstra68,
+  author    = {E.~W.~Dijkstra},
+  title     = {{G}o to {S}tatement {C}onsidered {H}armful},
+  journal   = {Communications of the ACM},
+  volume    = {11},
+  number    = {3},
+  year      = {1968},
+  pages     = {147-148}
+}
\ No newline at end of file
--- a/Paper/document/root.tex	Fri Jan 18 13:03:09 2013 +0000
+++ b/Paper/document/root.tex	Fri Jan 18 13:56:35 2013 +0000
@@ -10,6 +10,10 @@
 \usepackage{tikz}
 \usepackage{pgf}
 
+%% for testing
+\usepackage{endnotes}
+\let\footnote=\endnote
+
 % urls in roman style, theory text in math-similar italics
 \urlstyle{rm}
 \isabellestyle{it}
Binary file paper.pdf has changed
--- a/thys/turing_basic.thy	Fri Jan 18 13:03:09 2013 +0000
+++ b/thys/turing_basic.thy	Fri Jan 18 13:56:35 2013 +0000
@@ -71,7 +71,7 @@
 fun step :: "config \<Rightarrow> tprog \<Rightarrow> config"
   where 
   "step (s, l, r) (p, off) = 
-  (let (a, s') = fetch p (s - off) (read r) in (s', update a (l, r)))"
+     (let (a, s') = fetch p (s - off) (read r) in (s', update a (l, r)))"
 
 fun steps :: "config \<Rightarrow> tprog \<Rightarrow> nat \<Rightarrow> config"
   where
@@ -93,7 +93,6 @@
                     (\<forall>(a, s) \<in> set p. s \<le> length p div 2
                                              + off \<and> s \<ge> off))"
 
-(* FIXME: needed? *)
 lemma halt_lemma: 
   "\<lbrakk>wf LE; \<forall>n. (\<not> P (f n) \<longrightarrow> (f (Suc n), (f n)) \<in> LE)\<rbrakk> \<Longrightarrow> \<exists>n. P (f n)"
 by (metis wf_iff_no_infinite_down_chain)
--- a/turing_basic.thy	Fri Jan 18 13:03:09 2013 +0000
+++ b/turing_basic.thy	Fri Jan 18 13:56:35 2013 +0000
@@ -1,738 +1,614 @@
-theory turing_basic
-imports Main
-begin
-
-section {* Basic definitions of Turing machine *}
-
-(* Title: Turing machine's definition and its charater
-   Author: Xu Jian <xujian817@hotmail.com>
-   Maintainer: Xu Jian
-*)
-
-text {*
-  Actions of Turing machine (Abbreviated TM in the following* ).
-*}
-
-datatype taction = 
-  -- {* Write zero *}
-  W0 | 
-  -- {* Write one *}
-  W1 | 
-  -- {* Move left *}
-  L | 
-  -- {* Move right *}
-  R | 
-  -- {* Do nothing *}
-  Nop
-
-text {*
-  Tape contents in every block.
-*}
-
-datatype block = 
-  -- {* Blank *}
-  Bk | 
-  -- {* Occupied *}
-  Oc
-
-text {*
-  Tape is represented as a pair of lists $(L_{left}, L_{right})$,
-  where $L_left$, named {\em left list}, is used to represent
-  the tape to the left of RW-head and
-  $L_{right}$, named {\em right list}, is used to represent the tape
-  under and to the right of RW-head.
-*}
-
-type_synonym tape = "block list \<times> block list"
-
-text {* The state of turing machine.*}
-type_synonym tstate = nat
-
-text {*
-  Turing machine instruction is represented as a 
-  pair @{text "(action, next_state)"},
-  where @{text "action"} is the action to take at the current state 
-  and @{text "next_state"} is the next state the machine is getting into
-  after the action.
-*}
-type_synonym tinst = "taction \<times> tstate"
-
-text {*
-  Program of Turing machine is represented as a list of Turing instructions
-  and the execution of the program starts from the head of the list.
-  *}
-type_synonym tprog = "tinst list"
-
-
-text {*
-  Turing machine configuration, which consists of the current state 
-  and the tape.
-*}
-type_synonym t_conf = "tstate \<times> tape"
-
-fun nth_of ::  "'a list \<Rightarrow> nat \<Rightarrow> 'a option"
-  where
-  "nth_of xs n = (if n < length xs then Some (xs!n)
-                  else None)"
-
-text {*
-  The function used to fetech instruction out of Turing program.
-  *}
-
-fun fetch :: "tprog \<Rightarrow> tstate \<Rightarrow> block \<Rightarrow> tinst"
-  where
-  "fetch p s b = (if s = 0 then (Nop, 0) else
-                  case b of 
-                     Bk \<Rightarrow> case nth_of p (2 * (s - 1)) of
-                          Some i \<Rightarrow> i
-                        | None \<Rightarrow> (Nop, 0) 
-                   | Oc \<Rightarrow> case nth_of p (2 * (s - 1) +1) of
-                          Some i \<Rightarrow> i
-                        | None \<Rightarrow> (Nop, 0))"
-
-
-fun new_tape :: "taction \<Rightarrow> tape \<Rightarrow> tape"
-where 
-   "new_tape action (leftn, rightn) = (case action of
-                                         W0 \<Rightarrow> (leftn, Bk#(tl rightn)) |
-                                         W1 \<Rightarrow> (leftn, Oc#(tl rightn)) |
-                                         L  \<Rightarrow>  (if leftn = [] then (tl leftn, Bk#rightn)
-                                               else (tl leftn, (hd leftn) # rightn)) |
-                                         R  \<Rightarrow> if rightn = [] then (Bk#leftn,tl rightn) 
-                                               else ((hd rightn)#leftn, tl rightn) |
-                                         Nop \<Rightarrow> (leftn, rightn)
-                                       )"
-
-text {*
-  The one step function used to transfer Turing machine configuration.
-*}
-fun tstep :: "t_conf \<Rightarrow> tprog \<Rightarrow> t_conf"
-  where
-  "tstep c p = (let (s, l, r) = c in 
-                     let (ac, ns) = (fetch p s (case r of [] \<Rightarrow> Bk |     
-                                                               x # xs \<Rightarrow> x)) in
-                       (ns, new_tape ac (l, r)))"
-
-text {*
-  The many-step function.
-*}
-fun steps :: "t_conf \<Rightarrow> tprog \<Rightarrow> nat \<Rightarrow> t_conf"
-  where
-  "steps c p 0 = c" |
-  "steps c p (Suc n) = steps (tstep c p) p n"
-
-lemma tstep_red: "steps c p (Suc n) = tstep (steps c p n) p"
-proof(induct n arbitrary: c)
-  fix c
-  show "steps c p (Suc 0) = tstep (steps c p 0) p" by(simp add: steps.simps)
-next
-  fix n c
-  assume ind: "\<And> c. steps c p (Suc n) = tstep (steps c p n) p"
-  have "steps (tstep c p) p (Suc n) = tstep (steps (tstep c p) p n) p"
-    by(rule ind)
-  thus "steps c p (Suc (Suc n)) = tstep (steps c p (Suc n)) p" by(simp add: steps.simps)
-qed
-
-declare Let_def[simp] option.split[split]
-
-definition 
-  "iseven n \<equiv> \<exists> x. n = 2 * x"
-
-
-text {*
-  The following @{text "t_correct"} function is used to specify the wellformedness of Turing
-  machine.
-*}
-fun t_correct :: "tprog \<Rightarrow> bool"
-  where
-  "t_correct p = (length p \<ge> 2 \<and> iseven (length p) \<and> 
-                   list_all (\<lambda> (acn, s). s \<le> length p div 2) p)"
-
-declare t_correct.simps[simp del]
-
-lemma allimp: "\<lbrakk>\<forall>x. P x \<longrightarrow> Q x; \<forall>x. P x\<rbrakk> \<Longrightarrow> \<forall>x. Q x"
-by(auto elim: allE)
-
-lemma halt_lemma: "\<lbrakk>wf LE; \<forall> n. (\<not> P (f n) \<longrightarrow> (f (Suc n), (f n)) \<in> LE)\<rbrakk> \<Longrightarrow> \<exists> n. P (f n)"
-apply(rule exCI, drule allimp, auto)
-apply(drule_tac f = f  in wf_inv_image, simp add: inv_image_def)
-apply(erule wf_induct, auto)
-done
-
-lemma steps_add: "steps c t (x + y) = steps (steps c t x) t y"
-by(induct x arbitrary: c, auto simp: steps.simps tstep_red)
-
-lemma listall_set: "list_all p t \<Longrightarrow> \<forall> a \<in> set t. p a"
-by(induct t, auto)
-
-lemma fetch_ex: "\<exists>b a. fetch T aa ab = (b, a)"
-by(simp add: fetch.simps)
-definition exponent :: "'a \<Rightarrow> nat \<Rightarrow> 'a list" ("_\<^bsup>_\<^esup>" [0, 0]100)
-  where "exponent x n = replicate n x"
-
-text {* 
-  @{text "tinres l1 l2"} means left list @{text "l1"} is congruent with
-  @{text "l2"} with respect to the execution of Turing machine. 
-  Appending Blank to the right of eigther one does not affect the 
-  outcome of excution. 
-*}
-
-definition tinres :: "block list \<Rightarrow> block list \<Rightarrow> bool"
-  where
-  "tinres bx by = (\<exists> n. bx = by@Bk\<^bsup>n\<^esup> \<or> by = bx @ Bk\<^bsup>n\<^esup>)"
-
-lemma exp_zero: "a\<^bsup>0\<^esup> = []"
-by(simp add: exponent_def)
-lemma exp_ind_def: "a\<^bsup>Suc x \<^esup> = a # a\<^bsup>x\<^esup>"
-by(simp add: exponent_def)
-
-text {*
-  The following lemma shows the meaning of @{text "tinres"} with respect to 
-  one step execution.
-  *}
-lemma tinres_step: 
-  "\<lbrakk>tinres l l'; tstep (ss, l, r) t = (sa, la, ra); tstep (ss, l', r) t = (sb, lb, rb)\<rbrakk>
-    \<Longrightarrow> tinres la lb \<and> ra = rb \<and> sa = sb"
-apply(auto simp: tstep.simps fetch.simps new_tape.simps 
-        split: if_splits taction.splits list.splits
-                 block.splits)
-apply(case_tac [!] "t ! (2 * (ss - Suc 0))", 
-     auto simp: exponent_def tinres_def split: if_splits taction.splits list.splits
-                 block.splits)
-apply(case_tac [!] "t ! (2 * (ss - Suc 0) + Suc 0)", 
-     auto simp: exponent_def tinres_def split: if_splits taction.splits list.splits
-                 block.splits)
-done
-
-declare tstep.simps[simp del] steps.simps[simp del]
-
-text {*
-  The following lemma shows the meaning of @{text "tinres"} with respect to 
-  many step execution.
-  *}
-lemma tinres_steps: 
-  "\<lbrakk>tinres l l'; steps (ss, l, r) t stp = (sa, la, ra); steps (ss, l', r) t stp = (sb, lb, rb)\<rbrakk>
-    \<Longrightarrow> tinres la lb \<and> ra = rb \<and> sa = sb"
-apply(induct stp arbitrary: sa la ra sb lb rb, simp add: steps.simps)
-apply(simp add: tstep_red)
-apply(case_tac "(steps (ss, l, r) t stp)")
-apply(case_tac "(steps (ss, l', r) t stp)")
-proof -
-  fix stp sa la ra sb lb rb a b c aa ba ca
-  assume ind: "\<And>sa la ra sb lb rb. \<lbrakk>steps (ss, l, r) t stp = (sa, la, ra); 
-          steps (ss, l', r) t stp = (sb, lb, rb)\<rbrakk> \<Longrightarrow> tinres la lb \<and> ra = rb \<and> sa = sb"
-  and h: " tinres l l'" "tstep (steps (ss, l, r) t stp) t = (sa, la, ra)"
-         "tstep (steps (ss, l', r) t stp) t = (sb, lb, rb)" "steps (ss, l, r) t stp = (a, b, c)" 
-         "steps (ss, l', r) t stp = (aa, ba, ca)"
-  have "tinres b ba \<and> c = ca \<and> a = aa"
-    apply(rule_tac ind, simp_all add: h)
-    done
-  thus "tinres la lb \<and> ra = rb \<and> sa = sb"
-    apply(rule_tac l = b and l' = ba and r = c  and ss = a   
-            and t = t in tinres_step)
-    using h
-    apply(simp, simp, simp)
-    done
-qed
-
-text {*
-  The following function @{text "tshift tp n"} is used to shift Turing programs 
-  @{text "tp"} by @{text "n"} when it is going to be combined with others.
-  *}
-
-fun tshift :: "tprog \<Rightarrow> nat \<Rightarrow> tprog"
-  where
-  "tshift tp off = (map (\<lambda> (action, state). (action, (if state = 0 then 0
-                                                      else state + off))) tp)"
-
-text {*
-  When two Turing programs are combined, the end state (state @{text "0"}) of the one 
-  at the prefix position needs to be connected to the start state 
-  of the one at postfix position. If @{text "tp"} is the Turing program
-  to be at the prefix, @{text "change_termi_state tp"} is the transformed Turing program.
-  *}
-fun change_termi_state :: "tprog \<Rightarrow> tprog"
-  where
-  "change_termi_state t = 
-       (map (\<lambda> (acn, ns). if ns = 0 then (acn, Suc ((length t) div 2)) else (acn, ns)) t)"
-
-text {*
-  @{text "t_add tp1 tp2"} is the combined Truing program.
-*}
-
-fun t_add :: "tprog \<Rightarrow> tprog \<Rightarrow> tprog" ("_ |+| _" [0, 0] 100)
-  where
-  "t_add t1 t2 = ((change_termi_state t1) @ (tshift t2 ((length t1) div 2)))"
-
-text {*
-  Tests whether the current configuration is at state @{text "0"}.
-*}
-definition isS0 :: "t_conf \<Rightarrow> bool"
-  where
-  "isS0 c = (let (s, l, r) = c in s = 0)"
-
-declare tstep.simps[simp del] steps.simps[simp del] 
-        t_add.simps[simp del] fetch.simps[simp del]
-        new_tape.simps[simp del]
-
-
-text {*
-  Single step execution starting from state @{text "0"} will not make any progress.
-*}
-lemma tstep_0: "tstep (0, tp) p = (0, tp)"
-apply(simp add: tstep.simps fetch.simps new_tape.simps)
-done
-
-
-text {*
-  Many step executions starting from state @{text "0"} will not make any progress.
-*}
-
-lemma steps_0: "steps (0, tp) p stp = (0, tp)"
-apply(induct stp)
-apply(simp add: steps.simps)
-apply(simp add: tstep_red tstep_0)
-done
-
-lemma s_keep_step: "\<lbrakk>a \<le> length A div 2; tstep (a, b, c) A = (s, l, r); t_correct A\<rbrakk>
-  \<Longrightarrow> s \<le> length A div 2"
-apply(simp add: tstep.simps fetch.simps t_correct.simps iseven_def 
-  split: if_splits block.splits list.splits)
-apply(case_tac [!] a, auto simp: list_all_length)
-apply(erule_tac x = "2 * nat" in allE, auto)
-apply(erule_tac x = "2 * nat" in allE, auto)
-apply(erule_tac x = "Suc (2 * nat)" in allE, auto)
-done
-
-lemma s_keep: "\<lbrakk>steps (Suc 0, tp) A stp = (s, l, r);  t_correct A\<rbrakk> \<Longrightarrow> s \<le> length A div 2"
-proof(induct stp arbitrary: s l r)
-  case 0 thus "?case" by(auto simp: t_correct.simps steps.simps)
-next
-  fix stp s l r
-  assume ind: "\<And>s l r. \<lbrakk>steps (Suc 0, tp) A stp = (s, l, r); t_correct A\<rbrakk> \<Longrightarrow> s \<le> length A div 2"
-  and h1: "steps (Suc 0, tp) A (Suc stp) = (s, l, r)"
-  and h2: "t_correct A"
-  from h1 h2 show "s \<le> length A div 2"
-  proof(simp add: tstep_red, cases "(steps (Suc 0, tp) A stp)", simp)
-    fix a b c 
-    assume h3: "tstep (a, b, c) A = (s, l, r)"
-    and h4: "steps (Suc 0, tp) A stp = (a, b, c)"
-    have "a \<le> length A div 2"
-      using h2 h4
-      by(rule_tac l = b and r = c in ind, auto)
-    thus "?thesis"
-      using h3 h2
-      by(simp add: s_keep_step)
-  qed
-qed
-
-lemma t_merge_fetch_pre:
-  "\<lbrakk>fetch A s b = (ac, ns); s \<le> length A div 2; t_correct A; s \<noteq> 0\<rbrakk> \<Longrightarrow> 
-  fetch (A |+| B) s b = (ac, if ns = 0 then Suc (length A div 2) else ns)"
-apply(subgoal_tac "2 * (s - Suc 0) < length A \<and> Suc (2 * (s - Suc 0)) < length A")
-apply(auto simp: fetch.simps t_add.simps split: if_splits block.splits)
-apply(simp_all add: nth_append change_termi_state.simps)
-done
-
-lemma [simp]:  "\<lbrakk>\<not> a \<le> length A div 2; t_correct A\<rbrakk> \<Longrightarrow> fetch A a b = (Nop, 0)"
-apply(auto simp: fetch.simps del: nth_of.simps split: block.splits)
-apply(case_tac [!] a, auto simp: t_correct.simps iseven_def)
-done
-
-lemma  [elim]: "\<lbrakk>t_correct A; \<not> isS0 (tstep (a, b, c) A)\<rbrakk> \<Longrightarrow> a \<le> length A div 2"
-apply(rule_tac classical, auto simp: tstep.simps new_tape.simps isS0_def)
-done
-
-lemma [elim]: "\<lbrakk>t_correct A; \<not> isS0 (tstep (a, b, c) A)\<rbrakk> \<Longrightarrow> 0 < a"
-apply(rule_tac classical, simp add: tstep_0 isS0_def)
-done
-
-
-lemma t_merge_pre_eq_step: "\<lbrakk>tstep (a, b, c) A = cf; t_correct A; \<not> isS0 cf\<rbrakk> 
-        \<Longrightarrow> tstep (a, b, c) (A |+| B) = cf"
-apply(subgoal_tac "a \<le> length A div 2 \<and> a \<noteq> 0")
-apply(simp add: tstep.simps)
-apply(case_tac "fetch A a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp)
-apply(drule_tac B = B in t_merge_fetch_pre, simp, simp, simp, simp add: isS0_def, auto)
-done
-
-lemma t_merge_pre_eq:  "\<lbrakk>steps (Suc 0, tp) A stp = cf; \<not> isS0 cf; t_correct A\<rbrakk>
-    \<Longrightarrow> steps (Suc 0, tp) (A |+| B) stp = cf"
-proof(induct stp arbitrary: cf)
-  case 0 thus "?case" by(simp add: steps.simps)
-next
-  fix stp cf
-  assume ind: "\<And>cf. \<lbrakk>steps (Suc 0, tp) A stp = cf; \<not> isS0 cf; t_correct A\<rbrakk> 
-                 \<Longrightarrow> steps (Suc 0, tp) (A |+| B) stp = cf"
-  and h1: "steps (Suc 0, tp) A (Suc stp) = cf"
-  and h2: "\<not> isS0 cf"
-  and h3: "t_correct A"
-  from h1 h2 h3 show "steps (Suc 0, tp) (A |+| B) (Suc stp) = cf"
-  proof(simp add: tstep_red, cases "steps (Suc 0, tp) (A) stp", simp)
-    fix a b c
-    assume h4: "tstep (a, b, c) A = cf"
-    and h5: "steps (Suc 0, tp) A stp = (a, b, c)"
-    have "steps (Suc 0, tp) (A |+| B) stp = (a, b, c)"
-    proof(cases a)
-      case 0 thus "?thesis"
-        using h4 h2
-        apply(simp add: tstep_0, cases cf, simp add: isS0_def)
-        done
-    next
-      case (Suc n) thus "?thesis"
-        using h5 h3
-        apply(rule_tac ind, auto simp: isS0_def)
-        done
-    qed
-    thus "tstep (steps (Suc 0, tp) (A |+| B) stp) (A |+| B) = cf"
-      using h4 h5 h3 h2
-      apply(simp)
-      apply(rule t_merge_pre_eq_step, auto)
-      done
-  qed
-qed
-
-declare nth.simps[simp del] tshift.simps[simp del] change_termi_state.simps[simp del]
-
-lemma [simp]: "length (change_termi_state A) = length A"
-by(simp add: change_termi_state.simps)
-
-lemma first_halt_point: "steps (Suc 0, tp) A stp = (0, tp')
- \<Longrightarrow> \<exists>stp. \<not> isS0 (steps (Suc 0, tp) A stp) \<and> steps (Suc 0, tp) A (Suc stp) = (0, tp')"
-proof(induct stp)
-  case 0  thus "?case" by(simp add: steps.simps)
-next
-  case (Suc n) 
-  fix stp
-  assume ind: "steps (Suc 0, tp) A stp = (0, tp') \<Longrightarrow> 
-       \<exists>stp. \<not> isS0 (steps (Suc 0, tp) A stp) \<and> steps (Suc 0, tp) A (Suc stp) = (0, tp')"
-    and h: "steps (Suc 0, tp) A (Suc stp) = (0, tp')"
-  from h show "\<exists>stp. \<not> isS0 (steps (Suc 0, tp) A stp) \<and> steps (Suc 0, tp) A (Suc stp) = (0, tp')"
-  proof(simp add: tstep_red, cases "steps (Suc 0, tp) A stp", simp, case_tac a)
-    fix a b c
-    assume g1: "a = (0::nat)"
-    and g2: "tstep (a, b, c) A = (0, tp')"
-    and g3: "steps (Suc 0, tp) A stp = (a, b, c)"
-    have "steps (Suc 0, tp) A stp = (0, tp')"
-      using g2 g1 g3
-      by(simp add: tstep_0)
-    hence "\<exists> stp. \<not> isS0 (steps (Suc 0, tp) A stp) \<and> steps (Suc 0, tp) A (Suc stp) = (0, tp')"
-      by(rule ind)
-    thus "\<exists>stp. \<not> isS0 (steps (Suc 0, tp) A stp) \<and> tstep (steps (Suc 0, tp) A stp) A = (0, tp')" 
-      apply(simp add: tstep_red)
-      done
-  next
-    fix a b c nat
-    assume g1: "steps (Suc 0, tp) A stp = (a, b, c)"
-    and g2: "steps (Suc 0, tp) A (Suc stp) = (0, tp')" "a= Suc nat"
-    thus "\<exists>stp. \<not> isS0 (steps (Suc 0, tp) A stp) \<and> tstep (steps (Suc 0, tp) A stp) A = (0, tp')"
-      apply(rule_tac x = stp in exI)
-      apply(simp add: isS0_def tstep_red)
-      done
-  qed
-qed 
-   
-lemma t_merge_pre_halt_same': 
-  "\<lbrakk>\<not> isS0 (steps (Suc 0, tp) A stp) ; steps (Suc 0, tp) A (Suc stp) = (0, tp'); t_correct A\<rbrakk>
-  \<Longrightarrow> steps (Suc 0, tp) (A |+| B) (Suc stp) = (Suc (length A div 2), tp')"    
-proof(simp add: tstep_red, cases "steps (Suc 0, tp) A stp", simp)
-  fix a b c 
-  assume h1: "\<not> isS0 (a, b, c)"
-  and h2: "tstep (a, b, c) A = (0, tp')"
-  and h3: "t_correct A"
-  and h4: "steps (Suc 0, tp) A stp = (a, b, c)"
-  have "steps (Suc 0, tp) (A |+| B) stp = (a, b, c)"
-    using h1 h4 h3
-    apply(rule_tac  t_merge_pre_eq, auto)
-    done
-  moreover have "tstep (a, b, c) (A |+| B) = (Suc (length A div 2), tp')"
-    using h2 h3 h1 h4 
-    apply(simp add: tstep.simps)
-    apply(case_tac " fetch A a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp)
-    apply(drule_tac B = B in t_merge_fetch_pre, auto simp: isS0_def intro: s_keep)
-    done
-  ultimately show "tstep (steps (Suc 0, tp) (A |+| B) stp) (A |+| B) = (Suc (length A div 2), tp')"
-    by(simp)
-qed
-
-text {*
-  When Turing machine @{text "A"} and @{text "B"} are combined and the execution
-  of @{text "A"} can termination within @{text "stp"} steps, 
-  the combined machine @{text "A |+| B"} will eventually get into the starting 
-  state of machine @{text "B"}.
-*}
-lemma t_merge_pre_halt_same: "
-  \<lbrakk>steps (Suc 0, tp) A stp = (0, tp'); t_correct A; t_correct B\<rbrakk>
-     \<Longrightarrow> \<exists> stp. steps (Suc 0, tp) (A |+| B) stp = (Suc (length A div 2), tp')"
-proof -
-  assume a_wf: "t_correct A"
-  and b_wf: "t_correct B"
-  and a_ht: "steps (Suc 0, tp) A stp = (0, tp')"
-  have halt_point: "\<exists> stp. \<not> isS0 (steps (Suc 0, tp) A stp) \<and> steps (Suc 0, tp) A (Suc stp) = (0, tp')"
-    using a_ht
-    by(erule_tac first_halt_point)
-  then obtain stp' where "\<not> isS0 (steps (Suc 0, tp) A stp') \<and> steps (Suc 0, tp) A (Suc stp') = (0, tp')"..
-  hence "steps (Suc 0, tp) (A |+| B) (Suc stp') = (Suc (length A div 2), tp')"
-    using a_wf
-    apply(rule_tac t_merge_pre_halt_same', auto)
-    done
-  thus "?thesis" ..
-qed
-
-lemma fetch_0: "fetch p 0 b = (Nop, 0)"
-by(simp add: fetch.simps)
-
-lemma [simp]: "length (tshift B x) = length B"
-by(simp add: tshift.simps)
-
-lemma [simp]: "t_correct A \<Longrightarrow> 2 * (length A div 2) = length A"
-apply(simp add: t_correct.simps iseven_def, auto)
-done
-
-lemma t_merge_fetch_snd: 
-  "\<lbrakk>fetch B a b = (ac, ns); t_correct A; t_correct B; a > 0 \<rbrakk>
-  \<Longrightarrow> fetch (A |+| B) (a + length A div 2) b
-  = (ac, if ns = 0 then 0 else ns + length A div 2)"
-apply(auto simp: fetch.simps t_add.simps split: if_splits block.splits)
-apply(case_tac [!] a, simp_all)
-apply(simp_all add: nth_append change_termi_state.simps tshift.simps)
-done
-
-lemma t_merge_snd_eq_step: 
-  "\<lbrakk>tstep (s, l, r) B = (s', l', r'); t_correct A; t_correct B; s > 0\<rbrakk>
-    \<Longrightarrow> tstep (s + length A div 2, l, r) (A |+| B) = 
-       (if s' = 0 then 0 else s' + length A div 2, l' ,r') "
-apply(simp add: tstep.simps)
-apply(cases "fetch B s (case r of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)")
-apply(auto simp: t_merge_fetch_snd)
-apply(frule_tac [!] t_merge_fetch_snd, auto)
-done 
-
-text {*
-  Relates the executions of TM @{text "B"}, one is when @{text "B"} is executed alone,
-  the other is the execution when @{text "B"} is in the combined TM.
-*}
-lemma t_merge_snd_eq_steps: 
-  "\<lbrakk>t_correct A; t_correct B; steps (s, l, r) B stp = (s', l', r'); s > 0\<rbrakk>
-  \<Longrightarrow> steps (s + length A div 2, l, r) (A |+| B) stp = 
-      (if s' = 0 then 0 else s' + length A div 2, l', r')"
-proof(induct stp arbitrary: s' l' r')
-  case 0 thus "?case" 
-    by(simp add: steps.simps)
-next
-  fix stp s' l' r'
-  assume ind: "\<And>s' l' r'. \<lbrakk>t_correct A; t_correct B; steps (s, l, r) B stp = (s', l', r'); 0 < s\<rbrakk>
-                   \<Longrightarrow> steps (s + length A div 2, l, r) (A |+| B) stp = 
-                          (if s' = 0 then 0 else s' + length A div 2, l', r')"
-  and h1: "steps (s, l, r) B (Suc stp) = (s', l', r')"
-  and h2: "t_correct A"
-  and h3: "t_correct B"
-  and h4: "0 < s"
-  from h1 show "steps (s + length A div 2, l, r) (A |+| B) (Suc stp) 
-            = (if s' = 0 then 0 else s' + length A div 2, l', r')"
-  proof(simp only: tstep_red, cases "steps (s, l, r) B stp")
-    fix a b c 
-    assume h5: "steps (s, l, r) B stp = (a, b, c)" "tstep (steps (s, l, r) B stp) B = (s', l', r')"
-    hence h6: "(steps (s + length A div 2, l, r) (A |+| B) stp) = 
-                ((if a = 0 then 0 else a + length A div 2, b, c))"
-      using h2 h3 h4
-      by(rule_tac ind, auto)
-    thus "tstep (steps (s + length A div 2, l, r) (A |+| B) stp) (A |+| B) = 
-       (if s' = 0 then 0 else s'+ length A div 2, l', r')"
-      using h5
-    proof(auto)
-      assume "tstep (0, b, c) B = (0, l', r')" thus "tstep (0, b, c) (A |+| B) = (0, l', r')"
-        by(simp add: tstep_0)
-    next
-      assume "tstep (0, b, c) B = (s', l', r')" "0 < s'"
-      thus "tstep (0, b, c) (A |+| B) = (s' + length A div 2, l', r')"
-        by(simp add: tstep_0)
-    next
-      assume "tstep (a, b, c) B = (0, l', r')" "0 < a"
-      thus "tstep (a + length A div 2, b, c) (A |+| B) = (0, l', r')"
-        using h2 h3
-        by(drule_tac t_merge_snd_eq_step, auto)
-    next
-      assume "tstep (a, b, c) B = (s', l', r')" "0 < a" "0 < s'"
-      thus "tstep (a + length A div 2, b, c) (A |+| B) = (s' + length A div 2, l', r')"
-        using h2 h3
-        by(drule_tac t_merge_snd_eq_step, auto)
-    qed
-  qed
-qed
-
-lemma t_merge_snd_halt_eq: 
-  "\<lbrakk>steps (Suc 0, tp) B stp = (0, tp'); t_correct A; t_correct B\<rbrakk>
-  \<Longrightarrow> \<exists>stp. steps (Suc (length A div 2), tp) (A |+| B) stp = (0, tp')"
-apply(case_tac tp, cases tp', simp)
-apply(drule_tac  s = "Suc 0" in t_merge_snd_eq_steps, auto)
-done
-
-lemma t_inj: "\<lbrakk>steps (Suc 0, tp) A stpa = (0, tp1); steps (Suc 0, tp) A stpb = (0, tp2)\<rbrakk> 
-      \<Longrightarrow> tp1 = tp2"
-proof -
-  assume h1: "steps (Suc 0, tp) A stpa = (0, tp1)" 
-  and h2: "steps (Suc 0, tp) A stpb = (0, tp2)"
-  thus "?thesis"
-  proof(cases "stpa < stpb")
-    case True thus "?thesis"
-      using h1 h2
-      apply(drule_tac less_imp_Suc_add, auto)
-      apply(simp del: add_Suc_right add_Suc add: add_Suc_right[THEN sym] steps_add steps_0)
-      done
-  next
-    case False thus "?thesis"
-      using h1 h2
-      apply(drule_tac leI)
-      apply(case_tac "stpb = stpa", auto)
-      apply(subgoal_tac "stpb < stpa")
-      apply(drule_tac less_imp_Suc_add, auto)
-      apply(simp del: add_Suc_right add_Suc add: add_Suc_right[THEN sym] steps_add steps_0)
-      done
-  qed
-qed
-
-type_synonym t_assert = "tape \<Rightarrow> bool"
-
-definition t_imply :: "t_assert \<Rightarrow> t_assert \<Rightarrow> bool" ("_ \<turnstile>-> _" [0, 0] 100)
-  where
-  "t_imply a1 a2 = (\<forall> tp. a1 tp \<longrightarrow> a2 tp)"
-
-
-locale turing_merge =
-  fixes A :: "tprog" and B :: "tprog" and P1 :: "t_assert"
-  and P2 :: "t_assert"
-  and P3 :: "t_assert"
-  and P4 :: "t_assert"
-  and Q1:: "t_assert"
-  and Q2 :: "t_assert"
-  assumes 
-  A_wf : "t_correct A"
-  and B_wf : "t_correct B"
-  and A_halt : "P1 tp \<Longrightarrow> \<exists> stp. let (s, tp') = steps (Suc 0, tp) A stp in s = 0 \<and> Q1 tp'"
-  and B_halt : "P2 tp \<Longrightarrow> \<exists> stp. let (s, tp') = steps (Suc 0, tp) B stp in s = 0 \<and> Q2 tp'"
-  and A_uhalt : "P3 tp \<Longrightarrow> \<not> (\<exists> stp. isS0 (steps (Suc 0, tp) A stp))"
-  and B_uhalt: "P4 tp \<Longrightarrow> \<not> (\<exists> stp. isS0 (steps (Suc 0, tp) B stp))"
-begin
-
-
-text {*
-  The following lemma tries to derive the Hoare logic rule for sequentially combined TMs.
-  It deals with the situtation when both @{text "A"} and @{text "B"} are terminated.
-*}
-
-thm t_merge_pre_halt_same
-
-lemma t_merge_halt: 
-  assumes aimpb: "Q1 \<turnstile>-> P2"
-  shows "P1 \<turnstile>->  \<lambda> tp. (\<exists> stp tp'. steps (Suc 0, tp) (A |+| B)  stp = (0, tp') \<and> Q2 tp')"
-proof(simp add: t_imply_def, auto)
-  fix a b
-  assume h: "P1 (a, b)"
-  hence "\<exists> stp. let (s, tp') = steps (Suc 0, a, b) A stp in s = 0 \<and> Q1 tp'"
-    using A_halt by simp
-  from this obtain stp1 where "let (s, tp') = steps (Suc 0, a, b) A stp1 in s = 0 \<and> Q1 tp'" ..
-  thus "\<exists>stp aa ba. steps (Suc 0, a, b) (A |+| B) stp = (0, aa, ba) \<and> Q2 (aa, ba)"
-  proof(case_tac "steps (Suc 0, a, b) A stp1", simp, erule_tac conjE)
-    fix aa ba c
-    assume g1: "Q1 (ba, c)" 
-      and g2: "steps (Suc 0, a, b) A stp1 = (0, ba, c)"
-    hence "P2 (ba, c)"
-      using aimpb apply(simp add: t_imply_def)
-      done
-    hence "\<exists> stp. let (s, tp') = steps (Suc 0, ba, c) B stp in s = 0 \<and> Q2 tp'"
-      using B_halt by simp
-    from this obtain stp2 where "let (s, tp') = steps (Suc 0, ba, c) B stp2 in s = 0 \<and> Q2 tp'" ..
-    thus "?thesis"
-    proof(case_tac "steps (Suc 0, ba, c) B stp2", simp, erule_tac conjE)
-      fix aa bb ca
-      assume g3: " Q2 (bb, ca)" "steps (Suc 0, ba, c) B stp2 = (0, bb, ca)"
-      have "\<exists> stp. steps (Suc 0, a, b) (A |+| B) stp = (Suc (length A div 2), ba , c)"
-        using g2 A_wf B_wf
-        by(rule_tac t_merge_pre_halt_same, auto)
-      moreover have "\<exists> stp. steps (Suc (length A div 2), ba, c) (A |+| B) stp = (0, bb, ca)"
-        using g3 A_wf B_wf
-        apply(rule_tac t_merge_snd_halt_eq, auto)
-        done
-      ultimately show "\<exists>stp aa ba. steps (Suc 0, a, b) (A |+| B) stp = (0, aa, ba) \<and> Q2 (aa, ba)"
-        apply(erule_tac exE, erule_tac exE)
-        apply(rule_tac x = "stp + stpa" in exI, simp add: steps_add)
-        using g3 by simp
-    qed
-  qed
-qed
-
-lemma  t_merge_uhalt_tmp:
-  assumes B_uh: "\<forall>stp. \<not> isS0 (steps (Suc 0, b, c) B stp)"
-  and merge_ah: "steps (Suc 0, tp) (A |+| B) stpa = (Suc (length A div 2), b, c)" 
-  shows "\<forall> stp. \<not> isS0 (steps (Suc 0, tp) (A |+| B) stp)"
-  using B_uh merge_ah
-apply(rule_tac allI)
-apply(case_tac "stp > stpa")
-apply(erule_tac x = "stp - stpa" in allE)
-apply(case_tac "(steps (Suc 0, b, c) B (stp - stpa))", simp)
-proof -
-  fix stp a ba ca 
-  assume h1: "\<not> isS0 (a, ba, ca)" "stpa < stp"
-  and h2: "steps (Suc 0, b, c) B (stp - stpa) = (a, ba, ca)"
-  have "steps (Suc 0 + length A div 2, b, c) (A |+| B) (stp - stpa) = 
-      (if a = 0 then 0 else a + length A div 2, ba, ca)"
-    using A_wf B_wf h2
-    by(rule_tac t_merge_snd_eq_steps, auto)
-  moreover have "a > 0" using h1 by(simp add: isS0_def)
-  moreover have "\<exists> stpb. stp = stpa + stpb" 
-    using h1 by(rule_tac x = "stp - stpa" in exI, simp)
-  ultimately show "\<not> isS0 (steps (Suc 0, tp) (A |+| B) stp)"
-    using merge_ah
-    by(auto simp: steps_add isS0_def)
-next
-  fix stp
-  assume h: "steps (Suc 0, tp) (A |+| B) stpa = (Suc (length A div 2), b, c)" "\<not> stpa < stp"
-  hence "\<exists> stpb. stpa = stp + stpb" apply(rule_tac x = "stpa - stp" in exI, auto) done
-  thus "\<not> isS0 (steps (Suc 0, tp) (A |+| B) stp)"
-    using h
-    apply(auto)
-    apply(cases "steps (Suc 0, tp) (A |+| B) stp", simp add: steps_add isS0_def steps_0)
-    done
-qed
-
-text {*
-  The following lemma deals with the situation when TM @{text "B"} can not terminate.
-  *}
-
-lemma t_merge_uhalt: 
-  assumes aimpb: "Q1 \<turnstile>-> P4"
-  shows "P1 \<turnstile>-> \<lambda> tp. \<not> (\<exists> stp. isS0 (steps (Suc 0, tp) (A |+| B) stp))"
-proof(simp only: t_imply_def, rule_tac allI, rule_tac impI)
-  fix tp 
-  assume init_asst: "P1 tp"
-  show "\<not> (\<exists>stp. isS0 (steps (Suc 0, tp) (A |+| B) stp))"
-  proof -
-    have "\<exists> stp. let (s, tp') = steps (Suc 0, tp) A stp in s = 0 \<and> Q1 tp'"
-      using A_halt[of tp] init_asst
-      by(simp)
-    from this obtain stpx where "let (s, tp') = steps (Suc 0, tp) A stpx in s = 0 \<and> Q1 tp'" ..
-    thus "?thesis"
-    proof(cases "steps (Suc 0, tp) A stpx", simp, erule_tac conjE)
-      fix a b c
-      assume "Q1 (b, c)"
-        and h3: "steps (Suc 0, tp) A stpx = (0, b, c)"
-      hence h2: "P4 (b, c)"  using aimpb
-        by(simp add: t_imply_def)
-      have "\<exists> stp. steps (Suc 0, tp) (A |+| B) stp = (Suc (length A div 2), b, c)"
-        using h3 A_wf B_wf
-        apply(rule_tac stp = stpx in t_merge_pre_halt_same, auto)
-        done
-      from this obtain stpa where h4:"steps (Suc 0, tp) (A |+| B) stpa = (Suc (length A div 2), b, c)" ..
-      have " \<not> (\<exists> stp. isS0 (steps (Suc 0, b, c) B stp))"
-        using B_uhalt [of "(b, c)"] h2 apply simp
-        done
-      from this and h4 show "\<forall>stp. \<not> isS0 (steps (Suc 0, tp) (A |+| B) stp)"
-        by(rule_tac t_merge_uhalt_tmp, auto)
-    qed
-  qed
-qed
-end
-
-end
-
+(* Title: Turing machines
+   Author: Xu Jian <xujian817@hotmail.com>
+   Maintainer: Xu Jian
+*)
+
+theory turing_basic
+imports Main
+begin
+
+section {* Basic definitions of Turing machine *}
+
+datatype action = W0 | W1 | L | R | Nop
+
+datatype cell = Bk | Oc
+
+type_synonym tape = "cell list \<times> cell list"
+
+type_synonym state = nat
+
+type_synonym instr = "action \<times> state"
+
+type_synonym tprog = "instr list \<times> nat"
+
+type_synonym config = "state \<times> tape"
+
+fun nth_of where
+  "nth_of xs i = (if i \<ge> length xs then None
+                  else Some (xs ! i))"
+
+lemma nth_of_map [simp]:
+  shows "nth_of (map f p) n = (case (nth_of p n) of None \<Rightarrow> None | Some x \<Rightarrow> Some (f x))"
+apply(induct p arbitrary: n)
+apply(auto)
+apply(case_tac n)
+apply(auto)
+done
+
+fun 
+  fetch :: "instr list \<Rightarrow> state \<Rightarrow> cell \<Rightarrow> instr"
+where
+  "fetch p 0 b = (Nop, 0)"
+| "fetch p (Suc s) Bk = 
+     (case nth_of p (2 * s) of
+        Some i \<Rightarrow> i
+      | None \<Rightarrow> (Nop, 0))"
+|"fetch p (Suc s) Oc = 
+     (case nth_of p ((2 * s) + 1) of
+         Some i \<Rightarrow> i
+       | None \<Rightarrow> (Nop, 0))"
+
+lemma fetch_Nil [simp]:
+  shows "fetch [] s b = (Nop, 0)"
+apply(case_tac s)
+apply(auto)
+apply(case_tac b)
+apply(auto)
+done
+
+fun 
+  update :: "action \<Rightarrow> tape \<Rightarrow> tape"
+where 
+  "update W0 (l, r) = (l, Bk # (tl r))" 
+| "update W1 (l, r) = (l, Oc # (tl r))"
+| "update L (l, r) = (if l = [] then ([], Bk # r) else (tl l, (hd l) # r))" 
+| "update R (l, r) = (if r = [] then (Bk # l, []) else ((hd r) # l, tl r))" 
+| "update Nop (l, r) = (l, r)"
+
+abbreviation 
+  "read r == if (r = []) then Bk else hd r"
+
+fun step :: "config \<Rightarrow> tprog \<Rightarrow> config"
+  where 
+  "step (s, l, r) (p, off) = 
+  (let (a, s') = fetch p (s - off) (read r) in (s', update a (l, r)))"
+
+fun steps :: "config \<Rightarrow> tprog \<Rightarrow> nat \<Rightarrow> config"
+  where
+  "steps c p 0 = c" |
+  "steps c p (Suc n) = steps (step c p) p n"
+
+lemma step_red [simp]: 
+  shows "steps c p (Suc n) = step (steps c p n) p"
+by (induct n arbitrary: c) (auto)
+
+lemma steps_add [simp]: 
+  shows "steps c p (m + n) = steps (steps c p m) p n"
+by (induct m arbitrary: c) (auto)
+
+fun 
+  tm_wf :: "tprog \<Rightarrow> bool"
+where
+  "tm_wf (p, off) = (length p \<ge> 2 \<and> length p mod 2 = 0 \<and> 
+                    (\<forall>(a, s) \<in> set p. s \<le> length p div 2
+                                             + off \<and> s \<ge> off))"
+
+(* FIXME: needed? *)
+lemma halt_lemma: 
+  "\<lbrakk>wf LE; \<forall>n. (\<not> P (f n) \<longrightarrow> (f (Suc n), (f n)) \<in> LE)\<rbrakk> \<Longrightarrow> \<exists>n. P (f n)"
+by (metis wf_iff_no_infinite_down_chain)
+
+abbreviation exponent :: "'a \<Rightarrow> nat \<Rightarrow> 'a list" ("_ \<up> _" [100, 99] 100)
+  where "x \<up> n == replicate n x"
+
+consts tape_of :: "'a \<Rightarrow> cell list" ("<_>" 100)
+
+fun tape_of_nat_list :: "nat list \<Rightarrow> cell list" 
+  where 
+  "tape_of_nat_list [] = []" |
+  "tape_of_nat_list [n] = Oc\<up>(Suc n)" |
+  "tape_of_nat_list (n#ns) = Oc\<up>(Suc n) @ Bk # (tape_of_nat_list ns)"
+
+defs (overloaded)
+  tape_of_nl_abv: "<am> \<equiv> tape_of_nat_list am"
+  tape_of_nat_abv : "<(n::nat)> \<equiv> Oc\<up>(Suc n)"
+
+definition tinres :: "cell list \<Rightarrow> cell list \<Rightarrow> bool"
+  where
+  "tinres xs ys = (\<exists>n. xs = ys @ Bk \<up> n \<or> ys = xs @ Bk \<up> n)"
+
+fun 
+  shift :: "instr list \<Rightarrow> nat \<Rightarrow> instr list"
+where
+  "shift p n = (map (\<lambda> (a, s). (a, (if s = 0 then 0 else s + n))) p)"
+
+
+lemma length_shift [simp]: 
+  "length (shift p n) = length p"
+by (simp)
+
+fun 
+  adjust :: "instr list \<Rightarrow> instr list"
+where
+  "adjust p = map (\<lambda> (a, s). (a, if s = 0 then (Suc (length p div 2)) else s)) p"
+
+lemma length_adjust[simp]: 
+  shows "length (adjust p) = length p"
+by (induct p) (auto)
+
+fun
+  tm_comp :: "instr list \<Rightarrow> instr list \<Rightarrow> instr list" ("_ |+| _" [0, 0] 100)
+where
+  "tm_comp p1 p2 = ((adjust p1) @ (shift p2 ((length p1) div 2)))"
+
+fun
+  is_final :: "config \<Rightarrow> bool"
+where
+  "is_final (s, l, r) = (s = 0)"
+
+lemma is_final_steps:
+  assumes "is_final (s, l, r)"
+  shows "is_final (steps (s, l, r) (p, off) n)"
+using assms by (induct n) (auto)
+
+fun 
+  holds_for :: "(tape \<Rightarrow> bool) \<Rightarrow> config \<Rightarrow> bool" ("_ holds'_for _" [100, 99] 100)
+where
+  "P holds_for (s, l, r) = P (l, r)"  
+
+(*
+lemma step_0 [simp]: 
+  shows "step (0, (l, r)) p = (0, (l, r))"
+by simp
+
+lemma steps_0 [simp]: 
+  shows "steps (0, (l, r)) p n = (0, (l, r))"
+by (induct n) (simp_all)
+*)
+
+lemma is_final_holds[simp]:
+  assumes "is_final c"
+  shows "Q holds_for (steps c (p, off) n) = Q holds_for c"
+using assms 
+apply(induct n)
+apply(case_tac [!] c)
+apply(auto)
+done
+
+type_synonym assert = "tape \<Rightarrow> bool"
+
+definition assert_imp :: "assert \<Rightarrow> assert \<Rightarrow> bool" ("_ \<mapsto> _" [0, 0] 100)
+  where
+  "assert_imp P Q = (\<forall>l r. P (l, r) \<longrightarrow> Q (l, r))"
+
+lemma holds_for_imp:
+  assumes "P holds_for c"
+  and "P \<mapsto> Q"
+  shows "Q holds_for c"
+using assms unfolding assert_imp_def by (case_tac c, auto)
+
+lemma test:
+  assumes "is_final (steps (1, (l, r)) p n1)"
+  and     "is_final (steps (1, (l, r)) p n2)"
+  shows "Q holds_for (steps (1, (l, r)) p n1) \<longleftrightarrow> Q holds_for (steps (1, (l, r)) p n2)"
+proof -
+  obtain n3 where "n1 = n2 + n3 \<or> n2 = n1 + n3"
+    by (metis le_iff_add nat_le_linear)
+  with assms show "Q holds_for (steps (1, (l, r)) p n1) \<longleftrightarrow> Q holds_for (steps (1, (l, r)) p n2)"  
+    by(case_tac p) (auto)
+qed
+
+definition
+  Hoare :: "assert \<Rightarrow> tprog \<Rightarrow> assert \<Rightarrow> bool" ("({(1_)}/ (_)/ {(1_)})" 50)
+where
+  "{P} p {Q} \<equiv> 
+     (\<forall>l r. P (l, r) \<longrightarrow> (\<exists>n. is_final (steps (1, (l, r)) p n) \<and> Q holds_for (steps (1, (l, r)) p n)))"
+
+lemma HoareI:
+  assumes "\<And>l r. P (l, r) \<Longrightarrow> \<exists>n. is_final (steps (1, (l, r)) p n) \<and> Q holds_for (steps (1, (l, r)) p n)"
+  shows "{P} p {Q}"
+unfolding Hoare_def using assms by auto
+
+text {*
+{P1} A {Q1}   {P2} B {Q2}  Q1 \<mapsto> P2
+-----------------------------------
+    {P1} A |+| B {Q2}
+*}
+
+lemma step_0 [simp]: 
+  shows "step (0, (l, r)) p = (0, (l, r))"
+by (case_tac p, simp)
+
+lemma steps_0 [simp]: 
+  shows "steps (0, (l, r)) p n = (0, (l, r))"
+by (induct n) (simp_all)
+
+declare steps.simps[simp del]
+
+lemma before_final: 
+  assumes "steps (1, tp) A n = (0, tp')"
+  obtains n' where "\<not> is_final (steps (1, tp) A n')" 
+        and "steps (1, tp) A (Suc n') = (0, tp')"
+proof -
+  from assms have "\<exists> n'. \<not> is_final (steps (1, tp) A n') \<and> 
+               steps (1, tp) A (Suc n') = (0, tp')"
+  proof(induct n arbitrary: tp', simp add: steps.simps)
+    fix n tp'
+    assume ind: 
+      "\<And>tp'. steps (1, tp) A n = (0, tp') \<Longrightarrow>
+      \<exists>n'. \<not> is_final (steps (1, tp) A n') \<and> steps (1, tp) A (Suc n') = (0, tp')"
+    and h: " steps (1, tp) A (Suc n) = (0, tp')"
+    from h show  "\<exists>n'. \<not> is_final (steps (1, tp) A n') \<and> steps (1, tp) A (Suc n') = (0, tp')"
+    proof(simp add: step_red del: steps.simps, 
+                     case_tac "(steps (Suc 0, tp) A n)", case_tac "a = 0", simp)
+      fix a b c
+      assume " steps (Suc 0, tp) A n = (0, tp')"
+      hence "\<exists>n'. \<not> is_final (steps (1, tp) A n') \<and> steps (1, tp) A (Suc n') = (0, tp')"
+        apply(rule_tac ind, simp)
+        done
+      thus "\<exists>n'. \<not> is_final (steps (Suc 0, tp) A n') \<and> step (steps (Suc 0, tp) A n') A = (0, tp')"
+        apply(simp)
+        done
+    next
+      fix a b c
+      assume "steps (Suc 0, tp) A n = (a, b, c)"
+             "step (steps (Suc 0, tp) A n) A = (0, tp')"
+        "a \<noteq> 0"
+      thus "\<exists>n'. \<not> is_final (steps (Suc 0, tp) A n') \<and> 
+        step (steps (Suc 0, tp) A n') A = (0, tp')"
+        apply(rule_tac x = n in exI, simp)
+        done
+    qed
+  qed
+  thus "(\<And>n'. \<lbrakk>\<not> is_final (steps (1, tp) A n'); 
+    steps (1, tp) A (Suc n') = (0, tp')\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow> thesis"
+    by auto
+qed
+
+declare tm_comp.simps [simp del] adjust.simps[simp del] shift.simps[simp del]
+
+lemma length_comp:
+"length (A |+| B) = length A + length B"
+apply(auto simp: tm_comp.simps)
+done
+
+lemma tmcomp_fetch_in_first:
+  assumes "case (fetch A a x) of (ac, ns) \<Rightarrow> ns \<noteq> 0"
+  shows "fetch (A |+| B) a x = fetch A a x"
+using assms
+apply(case_tac a, case_tac [!] x, 
+auto simp: length_comp tm_comp.simps length_adjust nth_append)
+apply(simp_all add: adjust.simps)
+done
+
+
+lemma is_final_eq: "is_final (ba, tp) = (ba = 0)"
+apply(case_tac tp, simp add: is_final.simps)
+done
+
+lemma t_merge_pre_eq_step: 
+  assumes step: "step (a, b, c) (A, 0) = cf"
+  and     tm_wf: "tm_wf (A, 0)" 
+  and     unfinal: "\<not> is_final cf"
+  shows "step (a, b, c) (A |+| B, 0) = cf"
+proof -
+  have "fetch (A |+| B) a (read c) = fetch A a (read c)"
+  proof(rule_tac tmcomp_fetch_in_first)
+    from step and unfinal show "case fetch A a (read c) of (ac, ns) \<Rightarrow> ns \<noteq> 0"
+      apply(auto simp: is_final.simps)
+      apply(case_tac "fetch A a (read c)", simp_all add: is_final_eq)
+      done
+  qed      
+  thus "?thesis"
+    using step
+    apply(auto simp: step.simps is_final.simps)
+    done
+qed
+
+declare tm_wf.simps[simp del] step.simps[simp del]
+
+lemma t_merge_pre_eq:  
+  "\<lbrakk>steps (Suc 0, tp) (A, 0) stp = cf; \<not> is_final cf; tm_wf (A, 0)\<rbrakk>
+  \<Longrightarrow> steps (Suc 0, tp) (A |+| B, 0) stp = cf"
+proof(induct stp arbitrary: cf, simp add: steps.simps)
+  fix stp cf
+  assume ind: "\<And>cf. \<lbrakk>steps (Suc 0, tp) (A, 0) stp = cf; \<not> is_final cf; tm_wf (A, 0)\<rbrakk> 
+    \<Longrightarrow> steps (Suc 0, tp) (A |+| B, 0) stp = cf"
+  and h: "steps (Suc 0, tp) (A, 0) (Suc stp) = cf"
+      "\<not> is_final cf" "tm_wf (A, 0)"
+  from h show "steps (Suc 0, tp) (A |+| B, 0) (Suc stp) = cf"
+  proof(simp add: step_red, case_tac "(steps (Suc 0, tp) (A, 0) stp)", simp)
+    fix a b c
+    assume g: "steps (Suc 0, tp) (A, 0) stp = (a, b, c)"
+      "step (a, b, c) (A, 0) = cf"
+    have "(steps (Suc 0, tp) (A |+| B, 0) stp) = (a, b, c)"
+    proof(rule ind, simp_all add: h g)
+      show "0 < a"
+        using g h
+        apply(simp add: step_red)
+        apply(case_tac a, auto simp: step_0)
+        apply(case_tac "steps (Suc 0, tp) (A, 0) stp", simp)
+        done
+    qed
+    thus "step (steps (Suc 0, tp) (A |+| B, 0) stp) (A |+| B, 0) = cf"
+      apply(simp)
+      apply(rule_tac t_merge_pre_eq_step, simp_all add: g h)
+      done
+  qed
+qed
+
+lemma tmcomp_fetch_in_first2:
+  assumes "fetch A a x = (ac, 0)"
+          "tm_wf (A, 0)"
+          "a \<le> length A div 2" "a > 0"
+  shows "fetch (A |+| B) a x = (ac, Suc (length A div 2))"
+using assms
+apply(case_tac a, case_tac [!] x, 
+auto simp: length_comp tm_comp.simps length_adjust nth_append)
+apply(simp_all add: adjust.simps)
+done
+
+lemma tmcomp_exec_after_first:
+  "\<lbrakk>0 < a; step (a, b, c) (A, 0) = (0, tp'); tm_wf (A, 0); 
+       a \<le> length A div 2\<rbrakk>
+       \<Longrightarrow> step (a, b, c) (A |+| B, 0) = (Suc (length A div 2), tp')"
+apply(simp add: step.simps, auto)
+apply(case_tac "fetch A a Bk", simp add: tmcomp_fetch_in_first2)
+apply(case_tac "fetch A a (hd c)", simp add: tmcomp_fetch_in_first2)
+done
+
+lemma step_nothalt_pre: "\<lbrakk>step (aa, ba, ca) (A, 0) = (a, b, c);  0 < a\<rbrakk> \<Longrightarrow> 0 < aa"
+apply(case_tac "aa = 0", simp add: step_0, simp)
+done
+
+lemma nth_in_set: 
+  "\<lbrakk> A ! i = x; i <  length A\<rbrakk> \<Longrightarrow> x \<in> set A"
+by auto
+
+lemma step_nothalt: 
+  "\<lbrakk>step (aa, ba, ca) (A, 0) = (a, b, c); 0 < a; tm_wf (A, 0)\<rbrakk> \<Longrightarrow> 
+  a \<le> length A div 2"
+apply(simp add: step.simps)
+apply(case_tac aa, case_tac [!] aa, auto split: if_splits simp: tm_wf.simps)
+apply(case_tac "A ! (2 * nat)", simp)
+apply(erule_tac x = "(aa, a)" in ballE, simp_all add: nth_in_set)
+apply(case_tac "hd ca", auto split: if_splits simp: tm_wf.simps)
+apply(case_tac "A ! (2 * nat)", simp)
+apply(erule_tac x = "(aa, a)" in ballE, simp_all add: nth_in_set)
+apply(case_tac "A ! (Suc (2 * nat))")
+apply(erule_tac x = "(aa,bb)" in ballE, simp_all add: nth_in_set)
+done
+
+lemma steps_in_range: 
+  " \<lbrakk>0 < a; steps (Suc 0, tp) (A, 0) stp = (a, b, c); tm_wf (A, 0)\<rbrakk>
+  \<Longrightarrow> a \<le> length A div 2"
+proof(induct stp arbitrary: a b c)
+  fix a b c
+  assume h: "0 < a" "steps (Suc 0, tp) (A, 0) 0 = (a, b, c)" 
+            "tm_wf (A, 0)"
+  thus "a \<le> length A div 2"
+    apply(simp add: steps.simps tm_wf.simps, auto)
+    done
+next
+  fix stp a b c
+  assume ind: "\<And>a b c. \<lbrakk>0 < a; steps (Suc 0, tp) (A, 0) stp = (a, b, c); 
+    tm_wf (A, 0)\<rbrakk> \<Longrightarrow> a \<le> length A div 2"
+  and h: "0 < a" "steps (Suc 0, tp) (A, 0) (Suc stp) = (a, b, c)" "tm_wf (A, 0)"
+  from h show "a \<le> length A div 2"
+  proof(simp add: step_red, case_tac "(steps (Suc 0, tp) (A, 0) stp)", simp)
+    fix aa ba ca
+    assume g: "step (aa, ba, ca) (A, 0) = (a, b, c)" 
+           "steps (Suc 0, tp) (A, 0) stp = (aa, ba, ca)"
+    hence "aa \<le> length A div 2"
+      apply(rule_tac ind, auto simp: h step_nothalt_pre)
+      done
+    thus "?thesis"
+      using g h
+      apply(rule_tac step_nothalt, auto)
+      done
+  qed
+qed
+
+lemma t_merge_pre_halt_same: 
+  assumes a_ht: "steps (1, tp) (A, 0) n = (0, tp')"
+  and a_wf: "tm_wf (A, 0)"
+  obtains n' where "steps (1, tp) (A |+| B, 0) n' = (Suc (length A div 2), tp')"
+proof -
+  assume a: "\<And>n. steps (1, tp) (A |+| B, 0) n = (Suc (length A div 2), tp') \<Longrightarrow> thesis"
+  obtain stp' where "\<not> is_final (steps (1, tp) (A, 0) stp')" and 
+                          "steps (1, tp) (A, 0) (Suc stp') = (0, tp')"
+  using a_ht before_final by blast
+  then have "steps (1, tp) (A |+| B, 0) (Suc stp') = (Suc (length A div 2), tp')"
+  proof(simp add: step_red)
+    assume "\<not> is_final (steps (Suc 0, tp) (A, 0) stp')"
+           " step (steps (Suc 0, tp) (A, 0) stp') (A, 0) = (0, tp')"
+    moreover hence "(steps (Suc 0, tp) (A |+| B, 0) stp') = (steps (Suc 0, tp) (A, 0) stp')"
+      apply(rule_tac t_merge_pre_eq)
+      apply(simp_all add: a_wf a_ht)
+      done
+    ultimately show "step (steps (Suc 0, tp) (A |+| B, 0) stp') (A |+| B, 0) = (Suc (length A div 2), tp')"
+      apply(case_tac " steps (Suc 0, tp) (A, 0) stp'", simp)
+      apply(rule tmcomp_exec_after_first, simp_all add: a_wf)
+      apply(erule_tac steps_in_range, auto simp: a_wf)
+      done
+  qed
+  with a show thesis by blast
+qed
+
+lemma tm_comp_fetch_second_zero:
+  "\<lbrakk>fetch B sa' x = (a, 0); tm_wf (A, 0); tm_wf (B, 0); sa' > 0\<rbrakk>
+     \<Longrightarrow> fetch (A |+| B) (sa' + (length A div 2)) x = (a, 0)"
+apply(case_tac x)
+apply(case_tac [!] sa',
+  auto simp: fetch.simps length_comp length_adjust nth_append tm_comp.simps
+             tm_wf.simps shift.simps split: if_splits)
+done 
+
+lemma tm_comp_fetch_second_inst:
+  "\<lbrakk>sa > 0; s > 0;  tm_wf (A, 0); tm_wf (B, 0); fetch B sa x = (a, s)\<rbrakk>
+     \<Longrightarrow> fetch (A |+| B) (sa + length A div 2) x = (a, s + length A div 2)"
+apply(case_tac x)
+apply(case_tac [!] sa,
+  auto simp: fetch.simps length_comp length_adjust nth_append tm_comp.simps
+             tm_wf.simps shift.simps split: if_splits)
+done 
+
+lemma t_merge_second_same:
+  assumes a_wf: "tm_wf (A, 0)"
+  and b_wf: "tm_wf (B, 0)"
+  and steps: "steps (Suc 0, l, r) (B, 0) stp = (s, l', r')"
+  shows "steps (Suc (length A div 2), l, r)  (A |+| B, 0) stp
+       = (if s = 0 then 0
+          else s + length A div 2, l', r')"
+using a_wf b_wf steps
+proof(induct stp arbitrary: s l' r', simp add: steps.simps, simp)
+  fix stpa sa l'a r'a
+  assume ind: "\<And>s l' r'. steps (Suc 0, l, r) (B, 0) stpa = (s, l', r') \<Longrightarrow>
+    steps (Suc (length A div 2), l, r) (A |+| B, 0) stpa = 
+                (if s = 0 then 0 else s + length A div 2, l', r')"
+  and h: "step (steps (Suc 0, l, r) (B, 0) stpa) (B, 0) = (sa, l'a, r'a)"
+  obtain sa' l'' r'' where a: "(steps (Suc 0, l, r) (B, 0) stpa) = (sa', l'', r'')"
+    apply(case_tac "steps (Suc 0, l, r) (B, 0) stpa", auto)
+    done
+  from this have b: "steps (Suc (length A div 2), l, r) (A |+| B, 0) stpa = 
+                (if sa' = 0 then 0 else sa' + length A div 2, l'', r'')"
+    apply(erule_tac ind)
+    done
+  from a b h show 
+    "(sa = 0 \<longrightarrow> step (steps (Suc (length A div 2), l, r) (A |+| B, 0) stpa) (A |+| B, 0) = (0, l'a, r'a)) \<and>
+    (0 < sa \<longrightarrow> step (steps (Suc (length A div 2), l, r) (A |+| B, 0) stpa) (A |+| B, 0) = (sa + length A div 2, l'a, r'a))"
+  proof(case_tac "sa' = 0", auto)
+    assume "step (sa', l'', r'') (B, 0) = (0, l'a, r'a)" "0 < sa'"
+    thus "step (sa' + length A div 2, l'', r'') (A |+| B, 0) = (0, l'a, r'a)"
+      using a_wf b_wf
+      apply(simp add:  step.simps)
+      apply(case_tac "fetch B sa' (read r'')", auto)
+      apply(simp_all add: step.simps tm_comp_fetch_second_zero)
+      done
+  next
+    assume "step (sa', l'', r'') (B, 0) = (sa, l'a, r'a)" "0 < sa'" "0 < sa"
+    thus "step (sa' + length A div 2, l'', r'') (A |+| B, 0) = (sa + length A div 2, l'a, r'a)"
+      using a_wf b_wf
+      apply(simp add: step.simps)
+      apply(case_tac "fetch B sa' (read r'')", auto)
+      apply(simp_all add: step.simps tm_comp_fetch_second_inst)
+      done
+  qed
+qed
+
+lemma t_merge_second_halt_same:
+  "\<lbrakk>tm_wf (A, 0); tm_wf (B, 0); 
+   steps (1, l, r) (B, 0) stp = (0, l', r')\<rbrakk>
+     \<Longrightarrow> steps (Suc (length A div 2), l, r)  (A |+| B, 0) stp
+       = (0, l', r')"
+using t_merge_second_same[where s = "0"]
+apply(auto)
+done
+
+lemma Hoare_plus_halt: 
+  assumes aimpb: "Q1 \<mapsto> P2"
+  and A_wf : "tm_wf (A, 0)"
+  and B_wf : "tm_wf (B, 0)"
+  and A_halt : "{P1} (A, 0) {Q1}"
+  and B_halt : "{P2} (B, 0) {Q2}"
+  shows "{P1} (A |+| B, 0) {Q2}"
+proof(rule HoareI)
+  fix l r
+  assume h: "P1 (l, r)"
+  then obtain n1 where a: "is_final (steps (1, l, r) (A, 0) n1)" and b: "Q1 holds_for (steps (1, l, r) (A, 0) n1)"
+    using A_halt unfolding Hoare_def by auto
+  then obtain l' r' where "steps (1, l, r) (A, 0) n1 = (0, l', r')" and c: "Q1 holds_for (0, l', r')"
+    by(case_tac "steps (1, l, r) (A, 0) n1", auto)
+  then obtain stpa where d: "steps (1, l, r) (A |+| B, 0) stpa = (Suc (length A div 2), l', r')"
+    using A_wf
+    by(rule_tac t_merge_pre_halt_same, auto)
+  from c aimpb have "P2 holds_for (0, l', r')"
+    by(rule holds_for_imp)
+  from this have "P2 (l', r')" by auto
+  from this obtain n2 where e: "is_final (steps (1, l', r') (B, 0) n2)" and f: "Q2 holds_for (steps (1, l', r') (B, 0) n2)"
+    using B_halt unfolding Hoare_def
+    by auto
+  then obtain l'' r'' where "steps (1, l', r') (B, 0) n2 = (0, l'', r'')" and g: "Q2 holds_for (0, l'', r'')"
+    by(case_tac "steps (1, l', r') (B, 0) n2", auto)
+  from this have "steps (Suc (length A div 2), l', r')  (A |+| B, 0) n2
+    = (0, l'', r'')"
+    apply(rule_tac t_merge_second_halt_same, auto simp: A_wf B_wf)
+    done
+  thus "\<exists>n. is_final (steps (1, l, r) (A |+| B, 0) n) \<and> Q2 holds_for (steps (1, l, r) (A |+| B, 0) n)"
+    using d g
+    apply(rule_tac x = "stpa + n2" in exI)
+    apply(simp add: steps_add)
+    done
+qed
+
+definition
+  Hoare_unhalt :: "assert \<Rightarrow> tprog \<Rightarrow> bool" ("({(1_)}/ (_))" 50)
+where
+  "{P} p \<equiv> 
+     (\<forall>l r. P (l, r) \<longrightarrow> (\<forall> n . \<not> (is_final (steps (1, (l, r)) p n))))"
+
+lemma Hoare_unhalt_I:
+  assumes "\<And>l r. P (l, r) \<Longrightarrow> \<forall> n. \<not> is_final (steps (1, (l, r)) p n)"
+  shows "{P} p"
+unfolding Hoare_unhalt_def using assms by auto
+
+lemma Hoare_plus_unhalt: 
+  assumes aimpb: "Q1 \<mapsto> P2"
+  and A_wf : "tm_wf (A, 0)"
+  and B_wf : "tm_wf (B, 0)"
+  and A_halt : "{P1} (A, 0) {Q1}"
+  and B_uhalt : "{P2} (B, 0)"
+  shows "{P1} (A |+| B, 0)"
+proof(rule_tac Hoare_unhalt_I)
+  fix l r
+  assume h: "P1 (l, r)"
+  then obtain n1 where a: "is_final (steps (1, l, r) (A, 0) n1)" and b: "Q1 holds_for (steps (1, l, r) (A, 0) n1)"
+    using A_halt unfolding Hoare_def by auto
+  then obtain l' r' where "steps (1, l, r) (A, 0) n1 = (0, l', r')" and c: "Q1 holds_for (0, l', r')"
+    by(case_tac "steps (1, l, r) (A, 0) n1", auto)
+  then obtain stpa where d: "steps (1, l, r) (A |+| B, 0) stpa = (Suc (length A div 2), l', r')"
+    using A_wf
+    by(rule_tac t_merge_pre_halt_same, auto)
+  from c aimpb have "P2 holds_for (0, l', r')"
+    by(rule holds_for_imp)
+  from this have "P2 (l', r')" by auto
+  from this have e: "\<forall> n. \<not> is_final (steps (Suc 0, l', r') (B, 0) n)  "
+    using B_uhalt unfolding Hoare_unhalt_def
+    by auto
+  from e show "\<forall>n. \<not> is_final (steps (1, l, r) (A |+| B, 0) n)"
+  proof(rule_tac allI, case_tac "n > stpa")
+    fix n
+    assume h2: "stpa < n"
+    hence "\<not> is_final (steps (Suc 0, l', r') (B, 0) (n - stpa))"
+      using e
+      apply(erule_tac x = "n - stpa" in allE) by simp
+    then obtain s'' l'' r'' where f: "steps (Suc 0, l', r') (B, 0) (n - stpa) = (s'', l'', r'')" and g: "s'' \<noteq> 0"
+      apply(case_tac "steps (Suc 0, l', r') (B, 0) (n - stpa)", auto)
+      done
+    have k: "steps (Suc (length A div 2), l', r') (A |+| B, 0) (n - stpa) = (s''+ length A div 2, l'', r'') "
+      using A_wf B_wf f g
+      apply(drule_tac t_merge_second_same, auto)
+      done
+    show "\<not> is_final (steps (1, l, r) (A |+| B, 0) n)"
+    proof -
+      have "\<not> is_final (steps (1, l, r) (A |+| B, 0) (stpa + (n  - stpa)))"
+        using d k A_wf
+        apply(simp only: steps_add d, simp add: tm_wf.simps)
+        done
+      thus "\<not> is_final (steps (1, l, r) (A |+| B, 0) n)"
+        using h2 by simp
+    qed
+  next
+    fix n
+    assume h2: "\<not> stpa < n"
+    with d show "\<not> is_final (steps (1, l, r) (A |+| B, 0) n)"
+      apply(auto)
+      apply(subgoal_tac "\<exists> d. stpa = n + d", auto)
+      apply(case_tac "(steps (Suc 0, l, r) (A |+| B, 0) n)", simp)
+      apply(rule_tac x = "stpa - n" in exI, simp)
+      done
+  qed
+qed
+        
+end
+