utm/UTM.thy
author zhang
Mon, 15 Oct 2012 13:23:52 +0000
changeset 371 48b231495281
parent 370 1ce04eb1c8ad
permissions -rw-r--r--
Some illustration added together with more explanations.

theory UTM
imports Main uncomputable recursive abacus UF GCD 
begin

section {* Wang coding of input arguments *}

text {*
  The direct compilation of the universal function @{text "rec_F"} can not give us UTM, because @{text "rec_F"} is of arity 2,
  where the first argument represents the Godel coding of the TM being simulated and the second argument represents the right number (in Wang's coding) of the TM tape. 
  (Notice, left number is always @{text "0"} at the very beginning). However, UTM needs to simulate the execution of any TM which may
  very well take many input arguments. Therefore, a initialization TM needs to run before the TM compiled from @{text "rec_F"}, and the sequential 
  composition of these two TMs will give rise to the UTM we are seeking. The purpose of this initialization TM is to transform the multiple 
  input arguments of the TM being simulated into Wang's coding, so that it can be consumed by the TM compiled from @{text "rec_F"} as the second
  argument. 

  However, this initialization TM (named @{text "t_wcode"}) can not be constructed by compiling from any resurve function, because every recursive 
  function takes a fixed number of input arguments, while @{text "t_wcode"} needs to take varying number of arguments and tranform them into 
  Wang's coding. Therefore, this section give a direct construction of @{text "t_wcode"} with just some parts being obtained from recursive functions.

\newlength{\basewidth}
\settowidth{\basewidth}{xxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxx}
\newlength{\baseheight}
\settoheight{\baseheight}{$B:R$}
\newcommand{\vsep}{5\baseheight}

The TM used to generate the Wang's code of input arguments is divided into three TMs
 executed sequentially, namely $prepare$, $mainwork$ and $adjust$¡£According to the
 convention, start state of ever TM is fixed to state $1$ while the final state is
 fixed to $0$.

The input and output of $prepare$ are illustrated respectively by Figure
\ref{prepare_input} and \ref{prepare_output}.


\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
  [tbox/.style = {draw, thick, inner sep = 5pt}]
  \node (0) {};
  \node (1) [tbox, text height = 3.5pt, right = -0.9pt of 0] {\wuhao $m$};
  \node (2) [tbox, right = -0.9pt of 1] {\wuhao $0$};
  \node (3) [tbox, right = -0.9pt of 2] {\wuhao $a_1$};
  \node (4) [tbox, right = -0.9pt of 3] {\wuhao $0$};
  \node (5) [tbox, right = -0.9pt of 4] {\wuhao $a_2$};
  \node (6) [right = -0.9pt of 5] {\ldots \ldots};
  \node (7) [tbox, right = -0.9pt of 6] {\wuhao $a_n$};
  \draw [->, >=latex, thick] (1)+(0, -4\baseheight) -- (1);
\end{tikzpicture}}
\caption{The input of TM $prepare$} \label{prepare_input}
\end{figure}

\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
  \node (0) {};
  \node (1) [draw, text height = 3.5pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
  \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $0$};
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $a_1$};
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
  \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {\wuhao $a_2$};
  \node (7) [right = -0.9pt of 6] {\ldots \ldots};
  \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {\wuhao $a_n$};
  \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {\wuhao $0$};
  \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {\wuhao $0$};
  \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {\wuhao $1$};
  \draw [->, >=latex, thick] (10)+(0, -4\baseheight) -- (10);
\end{tikzpicture}}
\caption{The output of TM $prepare$} \label{prepare_output}
\end{figure}

As shown in Figure \ref{prepare_input}, the input of $prepare$ is the same as the the input
of UTM, where $m$ is the Godel coding of the TM being interpreted and $a_1$ through $a_n$ are the $n$ input arguments of the TM under interpretation. The purpose of $purpose$ is to transform this initial tape layout to the one shown in Figure \ref{prepare_output},
which is convenient for the generation of Wang's codding of $a_1, \ldots, a_n$. The coding procedure starts from $a_n$ and ends after $a_1$ is encoded. The coding result is stored in an accumulator at the end of the tape (initially represented by the $1$ two blanks right to $a_n$ in Figure \ref{prepare_output}). In Figure \ref{prepare_output}, arguments $a_1, \ldots, a_n$ are separated by two blanks on both ends with the rest so that movement conditions can be implemented conveniently in subsequent TMs, because, by convention,
two consecutive blanks are usually used to signal the end or start of a large chunk of data. The diagram of $prepare$ is given in Figure \ref{prepare_diag}.


\begin{figure}[h!]
\centering
\scalebox{0.9}{
\begin{tikzpicture}
     \node[circle,draw] (1) {$1$};
     \node[circle,draw] (2) at ($(1)+(0.3\basewidth, 0)$) {$2$};
     \node[circle,draw] (3) at ($(2)+(0.3\basewidth, 0)$) {$3$};
     \node[circle,draw] (4) at ($(3)+(0.3\basewidth, 0)$) {$4$};
     \node[circle,draw] (5) at ($(4)+(0.3\basewidth, 0)$) {$5$};
     \node[circle,draw] (6) at ($(5)+(0.3\basewidth, 0)$) {$6$};
     \node[circle,draw] (7) at ($(6)+(0.3\basewidth, 0)$) {$7$};
     \node[circle,draw] (8) at ($(7)+(0.3\basewidth, 0)$) {$0$};


     \draw [->, >=latex] (1) edge [loop above] node[above] {$S_1:L$} (1)
     ;
     \draw [->, >=latex] (1) -- node[above] {$S_0:S_1$} (2)
     ;
     \draw [->, >=latex] (2) edge [loop above] node[above] {$S_1:R$} (2)
     ;
     \draw [->, >=latex] (2) -- node[above] {$S_0:L$} (3)
     ;
     \draw [->, >=latex] (3) edge[loop above] node[above] {$S_1:S_0$} (3)
     ;
     \draw [->, >=latex] (3) -- node[above] {$S_0:R$} (4)
     ;
     \draw [->, >=latex] (4) edge[loop above] node[above] {$S_0:R$} (4)
     ;
     \draw [->, >=latex] (4) -- node[above] {$S_0:R$} (5)
     ;
     \draw [->, >=latex] (5) edge[loop above] node[above] {$S_1:R$} (5)
     ;
     \draw [->, >=latex] (5) -- node[above] {$S_0:R$} (6)
     ;
     \draw [->, >=latex] (6) edge[bend left = 50] node[below] {$S_1:R$} (5)
     ;
     \draw [->, >=latex] (6) -- node[above] {$S_0:R$} (7)
     ;
     \draw [->, >=latex] (7) edge[loop above] node[above] {$S_0:S_1$} (7)
     ;
     \draw [->, >=latex] (7) -- node[above] {$S_1:L$} (8)
     ;
 \end{tikzpicture}}
\caption{The diagram of TM $prepare$} \label{prepare_diag}
\end{figure}

The purpose of TM $mainwork$ is to compute the Wang's encoding of $a_1, \ldots, a_n$. Every bit of $a_1, \ldots, a_n$, including the separating bits, is processed from left to right.
In order to detect the termination condition when the left most bit of $a_1$ is reached,
TM $mainwork$ needs to look ahead and consider three different situations at the start of
every iteration:
\begin{enumerate}
    \item The TM configuration for the first situation is shown in Figure \ref{mainwork_case_one_input},
        where the accumulator is stored in $r$, both of the next two bits
        to be encoded are $1$. The configuration at the end of the iteration
        is shown in Figure \ref{mainwork_case_one_output}, where the first 1-bit has been
        encoded and cleared. Notice that the accumulator has been changed to
        $(r+1) \times 2$ to reflect the encoded bit.
    \item The TM configuration for the second situation is shown in Figure
        \ref{mainwork_case_two_input},
        where the accumulator is stored in $r$, the next two bits
        to be encoded are $1$ and $0$. After the first
        $1$-bit was encoded and cleared, the second $0$-bit is difficult to detect
        and process. To solve this problem, these two consecutive bits are
        encoded in one iteration.  In this situation, only the first $1$-bit needs
        to be cleared since the second one is cleared by definition.
        The configuration at the end of the iteration
        is shown in Figure \ref{mainwork_case_two_output}.
        Notice that the accumulator has been changed to
        $(r+1) \times 4$ to reflect the two encoded bits.
    \item The third situation corresponds to the case when the last bit of $a_1$ is reached.
        The TM configurations at the start and end of the iteration are shown in
        Figure \ref{mainwork_case_three_input} and \ref{mainwork_case_three_output}
        respectively. For this situation, only the read write head needs to be moved to
        the left to prepare a initial configuration for TM $adjust$ to start with.
\end{enumerate}
The diagram of $mainwork$ is given in Figure \ref{mainwork_diag}. The two rectangular nodes
labeled with $2 \times x$ and $4 \times x$ are two TMs compiling from recursive functions
so that we do not have to design and verify two quite complicated TMs.


\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
  \node (0) {};
  \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
  \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $0$};
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $a_1$};
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
  \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {\wuhao $a_2$};
  \node (7) [right = -0.9pt of 6] {\ldots \ldots};
  \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {\wuhao $a_i$};
  \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {\wuhao $1$};
  \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {\wuhao $1$};
  \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {\wuhao $0$};
  \node (12) [right = -0.9pt of 11] {\ldots \ldots};
  \node (13) [draw, right = -0.9pt of 12, thick, inner sep = 5pt] {\wuhao $0$};
  \node (14) [draw, text height = 3.9pt, right = -0.9pt of 13, thick, inner sep = 5pt] {\wuhao $r$};
  \draw [->, >=latex, thick] (13)+(0, -4\baseheight) -- (13);
\end{tikzpicture}}
\caption{The first situation for TM $mainwork$ to consider} \label{mainwork_case_one_input}
\end{figure}


\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
  \node (0) {};
  \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
  \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $0$};
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $a_1$};
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
  \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {\wuhao $a_2$};
  \node (7) [right = -0.9pt of 6] {\ldots \ldots};
  \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {\wuhao $a_i$};
  \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {\wuhao $1$};
  \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {\wuhao $0$};
  \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {\wuhao $0$};
  \node (12) [right = -0.9pt of 11] {\ldots \ldots};
  \node (13) [draw, right = -0.9pt of 12, thick, inner sep = 5pt] {\wuhao $0$};
  \node (14) [draw, text height = 2.7pt, right = -0.9pt of 13, thick, inner sep = 5pt] {\wuhao $(r+1) \times 2$};
  \draw [->, >=latex, thick] (13)+(0, -4\baseheight) -- (13);
\end{tikzpicture}}
\caption{The output for the first case of TM $mainwork$'s processing}
\label{mainwork_case_one_output}
\end{figure}

\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
  \node (0) {};
  \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
  \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $0$};
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $a_1$};
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
  \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {\wuhao $a_2$};
  \node (7) [right = -0.9pt of 6] {\ldots \ldots};
  \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {\wuhao $a_i$};
  \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {\wuhao $1$};
  \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {\wuhao $0$};
  \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {\wuhao $1$};
  \node (12) [draw, right = -0.9pt of 11, thick, inner sep = 5pt] {\wuhao $0$};
  \node (13) [right = -0.9pt of 12] {\ldots \ldots};
  \node (14) [draw, right = -0.9pt of 13, thick, inner sep = 5pt] {\wuhao $0$};
  \node (15) [draw, text height = 3.9pt, right = -0.9pt of 14, thick, inner sep = 5pt] {\wuhao $r$};
  \draw [->, >=latex, thick] (14)+(0, -4\baseheight) -- (14);
\end{tikzpicture}}
\caption{The second situation for TM $mainwork$ to consider} \label{mainwork_case_two_input}
\end{figure}

\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
  \node (0) {};
  \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
  \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $0$};
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $a_1$};
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
  \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {\wuhao $a_2$};
  \node (7) [right = -0.9pt of 6] {\ldots \ldots};
  \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {\wuhao $a_i$};
  \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {\wuhao $1$};
  \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {\wuhao $0$};
  \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {\wuhao $0$};
  \node (12) [draw, right = -0.9pt of 11, thick, inner sep = 5pt] {\wuhao $0$};
  \node (13) [right = -0.9pt of 12] {\ldots \ldots};
  \node (14) [draw, right = -0.9pt of 13, thick, inner sep = 5pt] {\wuhao $0$};
  \node (15) [draw, text height = 2.7pt, right = -0.9pt of 14, thick, inner sep = 5pt] {\wuhao $(r+1) \times 4$};
  \draw [->, >=latex, thick] (14)+(0, -4\baseheight) -- (14);
\end{tikzpicture}}
\caption{The output for the second case of TM $mainwork$'s processing}
\label{mainwork_case_two_output}
\end{figure}

\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
  \node (0) {};
  \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
  \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $0$};
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $1$};
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
  \node (6) [right = -0.9pt of 5] {\ldots \ldots};
  \node (7) [draw, right = -0.9pt of 6, thick, inner sep = 5pt] {\wuhao $0$};
  \node (8) [draw, text height = 3.9pt, right = -0.9pt of 7, thick, inner sep = 5pt] {\wuhao $r$};
  \draw [->, >=latex, thick] (7)+(0, -4\baseheight) -- (7);
\end{tikzpicture}}
\caption{The third situation for TM $mainwork$ to consider} \label{mainwork_case_three_input}
\end{figure}

\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
  \node (0) {};
  \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
  \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $0$};
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $1$};
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
  \node (6) [right = -0.9pt of 5] {\ldots \ldots};
  \node (7) [draw, right = -0.9pt of 6, thick, inner sep = 5pt] {\wuhao $0$};
  \node (8) [draw, text height = 3.9pt, right = -0.9pt of 7, thick, inner sep = 5pt] {\wuhao $r$};
  \draw [->, >=latex, thick] (3)+(0, -4\baseheight) -- (3);
\end{tikzpicture}}
\caption{The output for the third case of TM $mainwork$'s processing}
\label{mainwork_case_three_output}
\end{figure}

\begin{figure}[h!]
\centering
\scalebox{0.9}{
\begin{tikzpicture}
     \node[circle,draw] (1) {$1$};
     \node[circle,draw] (2) at ($(1)+(0.3\basewidth, 0)$) {$2$};
     \node[circle,draw] (3) at ($(2)+(0.3\basewidth, 0)$) {$3$};
     \node[circle,draw] (4) at ($(3)+(0.3\basewidth, 0)$) {$4$};
     \node[circle,draw] (5) at ($(4)+(0.3\basewidth, 0)$) {$5$};
     \node[circle,draw] (6) at ($(5)+(0.3\basewidth, 0)$) {$6$};
     \node[circle,draw] (7) at ($(2)+(0, -7\baseheight)$) {$7$};
     \node[circle,draw] (8) at ($(7)+(0, -7\baseheight)$) {$8$};
     \node[circle,draw] (9) at ($(8)+(0.3\basewidth, 0)$) {$9$};
     \node[circle,draw] (10) at ($(9)+(0.3\basewidth, 0)$) {$10$};
     \node[circle,draw] (11) at ($(10)+(0.3\basewidth, 0)$) {$11$};
     \node[circle,draw] (12) at ($(11)+(0.3\basewidth, 0)$) {$12$};
     \node[draw] (13) at ($(6)+(0.3\basewidth, 0)$) {$2 \times x$};
     \node[circle,draw] (14) at ($(13)+(0.3\basewidth, 0)$) {$j_1$};
     \node[draw] (15) at ($(12)+(0.3\basewidth, 0)$) {$4 \times x$};
     \node[draw] (16) at ($(15)+(0.3\basewidth, 0)$) {$j_2$};
     \node[draw] (17) at ($(7)+(0.3\basewidth, 0)$) {$0$};

     \draw [->, >=latex] (1) edge[loop left] node[above] {$S_0:L$} (1)
     ;
     \draw [->, >=latex] (1) -- node[above] {$S_1:L$} (2)
     ;
     \draw [->, >=latex] (2) -- node[above] {$S_1:R$} (3)
     ;
     \draw [->, >=latex] (2) -- node[left] {$S_1:L$} (7)
     ;
     \draw [->, >=latex] (3) edge[loop above] node[above] {$S_1:S_0$} (3)
     ;
     \draw [->, >=latex] (3) -- node[above] {$S_0:R$} (4)
     ;
     \draw [->, >=latex] (4) edge[loop above] node[above] {$S_0:R$} (4)
     ;
     \draw [->, >=latex] (4) -- node[above] {$S_1:R$} (5)
     ;
     \draw [->, >=latex] (5) edge[loop above] node[above] {$S_1:R$} (5)
     ;
     \draw [->, >=latex] (5) -- node[above] {$S_0:S_1$} (6)
     ;
     \draw [->, >=latex] (6) edge[loop above] node[above] {$S_1:L$} (6)
     ;
     \draw [->, >=latex] (6) -- node[above] {$S_0:R$} (13)
     ;
     \draw [->, >=latex] (13) -- (14)
     ;
     \draw (14) -- ($(14)+(0, 6\baseheight)$) -- ($(1) + (0, 6\baseheight)$) node [above,midway] {$S_1:L$}
     ;
     \draw [->, >=latex] ($(1) + (0, 6\baseheight)$) -- (1)
     ;
     \draw [->, >=latex] (7) -- node[above] {$S_0:R$} (17)
     ;
     \draw [->, >=latex] (7) -- node[left] {$S_1:R$} (8)
     ;
     \draw [->, >=latex] (8) -- node[above] {$S_0:R$} (9)
     ;
     \draw [->, >=latex] (9) -- node[above] {$S_0:R$} (10)
     ;
     \draw [->, >=latex] (10) -- node[above] {$S_1:R$} (11)
     ;
     \draw [->, >=latex] (10) edge[loop above] node[above] {$S_0:R$} (10)
     ;
     \draw [->, >=latex] (11) edge[loop above] node[above] {$S_1:R$} (11)
     ;
     \draw [->, >=latex] (11) -- node[above] {$S_0:S_1$} (12)
     ;
     \draw [->, >=latex] (12) -- node[above] {$S_0:R$} (15)
     ;
     \draw [->, >=latex] (12) edge[loop above] node[above] {$S_1:L$} (12)
     ;
     \draw [->, >=latex] (15) -- (16)
     ;
     \draw (16) -- ($(16)+(0, -4\baseheight)$) -- ($(1) + (0, -18\baseheight)$) node [below,midway] {$S_1:L$}
     ;
     \draw [->, >=latex] ($(1) + (0, -18\baseheight)$) -- (1)
     ;
 \end{tikzpicture}}
\caption{The diagram of TM $mainwork$} \label{mainwork_diag}
\end{figure}

The purpose of TM $adjust$ is to encode the last bit of $a_1$. The initial and final configuration
of this TM are shown in Figure \ref{adjust_initial} and \ref{adjust_final} respectively.
The diagram of TM $adjust$ is shown in Figure \ref{adjust_diag}.


\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
  \node (0) {};
  \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
  \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $0$};
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $1$};
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
  \node (6) [right = -0.9pt of 5] {\ldots \ldots};
  \node (7) [draw, right = -0.9pt of 6, thick, inner sep = 5pt] {\wuhao $0$};
  \node (8) [draw, text height = 3.9pt, right = -0.9pt of 7, thick, inner sep = 5pt] {\wuhao $r$};
  \draw [->, >=latex, thick] (3)+(0, -4\baseheight) -- (3);
\end{tikzpicture}}
\caption{Initial configuration of TM $adjust$} \label{adjust_initial}
\end{figure}

\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
  \node (0) {};
  \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
  \node (3) [draw, text height = 2.9pt, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $r+1$};
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $0$};
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
  \node (6) [right = -0.9pt of 5] {\ldots \ldots};
  \draw [->, >=latex, thick] (1)+(0, -4\baseheight) -- (1);
\end{tikzpicture}}
\caption{Final configuration of TM $adjust$} \label{adjust_final}
\end{figure}


\begin{figure}[h!]
\centering
\scalebox{0.9}{
\begin{tikzpicture}
     \node[circle,draw] (1) {$1$};
     \node[circle,draw] (2) at ($(1)+(0.3\basewidth, 0)$) {$2$};
     \node[circle,draw] (3) at ($(2)+(0.3\basewidth, 0)$) {$3$};
     \node[circle,draw] (4) at ($(3)+(0.3\basewidth, 0)$) {$4$};
     \node[circle,draw] (5) at ($(4)+(0.3\basewidth, 0)$) {$5$};
     \node[circle,draw] (6) at ($(5)+(0.3\basewidth, 0)$) {$6$};
     \node[circle,draw] (7) at ($(6)+(0.3\basewidth, 0)$) {$7$};
     \node[circle,draw] (8) at ($(4)+(0, -7\baseheight)$) {$8$};
     \node[circle,draw] (9) at ($(8)+(0.3\basewidth, 0)$) {$9$};
     \node[circle,draw] (10) at ($(9)+(0.3\basewidth, 0)$) {$10$};
     \node[circle,draw] (11) at ($(10)+(0.3\basewidth, 0)$) {$11$};
     \node[circle,draw] (12) at ($(11)+(0.3\basewidth, 0)$) {$0$};


     \draw [->, >=latex] (1) -- node[above] {$S_1:R$} (2)
     ;
     \draw [->, >=latex] (1) edge[loop above] node[above] {$S_0:S_1$} (1)
     ;
     \draw [->, >=latex] (2) -- node[above] {$S_1:R$} (3)
     ;
     \draw [->, >=latex] (3) edge[loop above] node[above] {$S_0:R$} (3)
     ;
     \draw [->, >=latex] (3) -- node[above] {$S_1:R$} (4)
     ;
     \draw [->, >=latex] (4) -- node[above] {$S_1:L$} (5)
     ;
     \draw [->, >=latex] (4) -- node[right] {$S_0:L$} (8)
     ;
     \draw [->, >=latex] (5) -- node[above] {$S_0:L$} (6)
     ;
     \draw [->, >=latex] (5) edge[loop above] node[above] {$S_1:S_0$} (5)
     ;
     \draw [->, >=latex] (6) -- node[above] {$S_1:R$} (7)
     ;
     \draw [->, >=latex] (6) edge[loop above] node[above] {$S_0:L$} (6)
     ;
     \draw (7) -- ($(7)+(0, 6\baseheight)$) -- ($(2) + (0, 6\baseheight)$) node [above,midway] {$S_0:S_1$}
     ;
     \draw [->, >=latex] ($(2) + (0, 6\baseheight)$) -- (2)
     ;
     \draw [->, >=latex] (8) edge[loop left] node[left] {$S_1:S_0$} (8)
     ;
     \draw [->, >=latex] (8) -- node[above] {$S_0:L$} (9)
     ;
     \draw [->, >=latex] (9) edge[loop above] node[above] {$S_0:L$} (9)
     ;
     \draw [->, >=latex] (9) -- node[above] {$S_1:L$} (10)
     ;
     \draw [->, >=latex] (10) edge[loop above] node[above] {$S_0:L$} (10)
     ;
     \draw [->, >=latex] (10) -- node[above] {$S_0:L$} (11)
     ;
     \draw [->, >=latex] (11) edge[loop above] node[above] {$S_1:L$} (11)
     ;
     \draw [->, >=latex] (11) -- node[above] {$S_0:R$} (12)
     ;
 \end{tikzpicture}}
\caption{Diagram of TM $adjust$} \label{adjust_diag}
\end{figure}
*}


definition rec_twice :: "recf"
  where
  "rec_twice = Cn 1 rec_mult [id 1 0, constn 2]"

definition rec_fourtimes  :: "recf"
  where
  "rec_fourtimes = Cn 1 rec_mult [id 1 0, constn 4]"

definition abc_twice :: "abc_prog"
  where
  "abc_twice = (let (aprog, ary, fp) = rec_ci rec_twice in 
                       aprog [+] dummy_abc ((Suc 0)))"

definition abc_fourtimes :: "abc_prog"
  where
  "abc_fourtimes = (let (aprog, ary, fp) = rec_ci rec_fourtimes in 
                       aprog [+] dummy_abc ((Suc 0)))"

definition twice_ly :: "nat list"
  where
  "twice_ly = layout_of abc_twice"

definition fourtimes_ly :: "nat list"
  where
  "fourtimes_ly = layout_of abc_fourtimes"

definition t_twice :: "tprog"
  where
  "t_twice = change_termi_state (tm_of (abc_twice) @ (tMp 1 (start_of twice_ly (length abc_twice) - Suc 0)))"

definition t_fourtimes :: "tprog"
  where
  "t_fourtimes = change_termi_state (tm_of (abc_fourtimes) @ 
             (tMp 1 (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)))"


definition t_twice_len :: "nat"
  where
  "t_twice_len = length t_twice div 2"

definition t_wcode_main_first_part:: "tprog"
  where
  "t_wcode_main_first_part \<equiv> 
                   [(L, 1), (L, 2), (L, 7), (R, 3),
                    (R, 4), (W0, 3), (R, 4), (R, 5),
                    (W1, 6), (R, 5), (R, 13), (L, 6),
                    (R, 0), (R, 8), (R, 9), (Nop, 8),
                    (R, 10), (W0, 9), (R, 10), (R, 11), 
                    (W1, 12), (R, 11), (R, t_twice_len + 14), (L, 12)]"

definition t_wcode_main :: "tprog"
  where
  "t_wcode_main = (t_wcode_main_first_part @ tshift t_twice 12 @ [(L, 1), (L, 1)]
                    @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])"

fun bl_bin :: "block list \<Rightarrow> nat"
  where
  "bl_bin [] = 0" 
| "bl_bin (Bk # xs) = 2 * bl_bin xs"
| "bl_bin (Oc # xs) = Suc (2 * bl_bin xs)"

declare bl_bin.simps[simp del]

type_synonym bin_inv_t = "block list \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"

fun wcode_before_double :: "bin_inv_t"
  where
  "wcode_before_double ires rs (l, r) =
     (\<exists> ln rn. l = Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
               r = Oc\<^bsup>(Suc (Suc rs))\<^esup> @ Bk\<^bsup>rn \<^esup>)"

declare wcode_before_double.simps[simp del]

fun wcode_after_double :: "bin_inv_t"
  where
  "wcode_after_double ires rs (l, r) = 
     (\<exists> ln rn. l = Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and>
         r = Oc\<^bsup>Suc (Suc (Suc 2*rs))\<^esup> @ Bk\<^bsup>rn\<^esup>)"

declare wcode_after_double.simps[simp del]

fun wcode_on_left_moving_1_B :: "bin_inv_t"
  where
  "wcode_on_left_moving_1_B ires rs (l, r) = 
     (\<exists> ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # Oc # ires \<and> 
               r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
               ml + mr > Suc 0 \<and> mr > 0)"

declare wcode_on_left_moving_1_B.simps[simp del]

fun wcode_on_left_moving_1_O :: "bin_inv_t"
  where
  "wcode_on_left_moving_1_O ires rs (l, r) = 
     (\<exists> ln rn.
               l = Oc # ires \<and> 
               r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"

declare wcode_on_left_moving_1_O.simps[simp del]

fun wcode_on_left_moving_1 :: "bin_inv_t"
  where
  "wcode_on_left_moving_1 ires rs (l, r) = 
          (wcode_on_left_moving_1_B ires rs (l, r) \<or> wcode_on_left_moving_1_O ires rs (l, r))"

declare wcode_on_left_moving_1.simps[simp del]

fun wcode_on_checking_1 :: "bin_inv_t"
  where
   "wcode_on_checking_1 ires rs (l, r) = 
    (\<exists> ln rn. l = ires \<and>
              r = Oc # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"

fun wcode_erase1 :: "bin_inv_t"
  where
"wcode_erase1 ires rs (l, r) = 
       (\<exists> ln rn. l = Oc # ires \<and> 
                 tl r = Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"

declare wcode_erase1.simps [simp del]

fun wcode_on_right_moving_1 :: "bin_inv_t"
  where
  "wcode_on_right_moving_1 ires rs (l, r) = 
       (\<exists> ml mr rn.        
             l = Bk\<^bsup>ml\<^esup> @ Oc # ires \<and> 
             r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
             ml + mr > Suc 0)"

declare wcode_on_right_moving_1.simps [simp del] 

declare wcode_on_right_moving_1.simps[simp del]

fun wcode_goon_right_moving_1 :: "bin_inv_t"
  where
  "wcode_goon_right_moving_1 ires rs (l, r) = 
      (\<exists> ml mr ln rn. 
            l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
            r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
            ml + mr = Suc rs)"

declare wcode_goon_right_moving_1.simps[simp del]

fun wcode_backto_standard_pos_B :: "bin_inv_t"
  where
  "wcode_backto_standard_pos_B ires rs (l, r) = 
          (\<exists> ln rn. l =  Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
               r =  Bk # Oc\<^bsup>(Suc (Suc rs))\<^esup> @ Bk\<^bsup>rn \<^esup>)"

declare wcode_backto_standard_pos_B.simps[simp del]

fun wcode_backto_standard_pos_O :: "bin_inv_t"
  where
   "wcode_backto_standard_pos_O ires rs (l, r) = 
        (\<exists> ml mr ln rn. 
            l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and>
            r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
            ml + mr = Suc (Suc rs) \<and> mr > 0)"

declare wcode_backto_standard_pos_O.simps[simp del]

fun wcode_backto_standard_pos :: "bin_inv_t"
  where
  "wcode_backto_standard_pos ires rs (l, r) = (wcode_backto_standard_pos_B ires rs (l, r) \<or>
                                            wcode_backto_standard_pos_O ires rs (l, r))"

declare wcode_backto_standard_pos.simps[simp del]

lemma [simp]: "<0::nat> = [Oc]"
apply(simp add: tape_of_nat_abv exponent_def tape_of_nat_list.simps)
done

lemma tape_of_Suc_nat: "<Suc (a ::nat)> = replicate a Oc @ [Oc, Oc]"
apply(simp add: tape_of_nat_abv exp_ind tape_of_nat_list.simps)
apply(simp only: exp_ind_def[THEN sym])
apply(simp only: exp_ind, simp, simp add: exponent_def)
done

lemma [simp]: "length (<a::nat>) = Suc a"
apply(simp add: tape_of_nat_abv tape_of_nat_list.simps)
done

lemma [simp]: "<[a::nat]> = <a>"
apply(simp add: tape_of_nat_abv tape_of_nl_abv exponent_def
                tape_of_nat_list.simps)
done

lemma bin_wc_eq: "bl_bin xs = bl2wc xs"
proof(induct xs)
  show " bl_bin [] = bl2wc []" 
    apply(simp add: bl_bin.simps)
    done
next
  fix a xs
  assume "bl_bin xs = bl2wc xs"
  thus " bl_bin (a # xs) = bl2wc (a # xs)"
    apply(case_tac a, simp_all add: bl_bin.simps bl2wc.simps)
    apply(simp_all add: bl2nat.simps bl2nat_double)
    done
qed

declare exp_def[simp del]

lemma bl_bin_nat_Suc:  
  "bl_bin (<Suc a>) = bl_bin (<a>) + 2^(Suc a)"
apply(simp add: tape_of_nat_abv bin_wc_eq)
apply(simp add: bl2wc.simps)
done
lemma [simp]: " rev (a\<^bsup>aa\<^esup>) = a\<^bsup>aa\<^esup>"
apply(simp add: exponent_def)
done
 
declare tape_of_nl_abv_cons[simp del]

lemma tape_of_nl_rev: "rev (<lm::nat list>) = (<rev lm>)"
apply(induct lm rule: list_tl_induct, simp)
apply(case_tac "list = []", simp add: tape_of_nl_abv tape_of_nat_list.simps)
apply(simp add: tape_of_nat_list_butlast_last tape_of_nl_abv_cons)
done
lemma [simp]: "a\<^bsup>Suc 0\<^esup> = [a]" 
by(simp add: exp_def)
lemma tape_of_nl_cons_app1: "(<a # xs @ [b]>) = (Oc\<^bsup>Suc a\<^esup> @ Bk # (<xs@ [b]>))"
apply(case_tac xs, simp add: tape_of_nl_abv tape_of_nat_list.simps)
apply(simp add: tape_of_nl_abv  tape_of_nat_list.simps)
done

lemma bl_bin_bk_oc[simp]:
  "bl_bin (xs @ [Bk, Oc]) = 
  bl_bin xs + 2*2^(length xs)"
apply(simp add: bin_wc_eq)
using bl2nat_cons_oc[of "xs @ [Bk]"]
apply(simp add: bl2nat_cons_bk bl2wc.simps)
done

lemma tape_of_nat[simp]: "(<a::nat>) = Oc\<^bsup>Suc a\<^esup>"
apply(simp add: tape_of_nat_abv)
done
lemma tape_of_nl_cons_app2: "(<c # xs @ [b]>) = (<c # xs> @ Bk # Oc\<^bsup>Suc b\<^esup>)"
proof(induct "length xs" arbitrary: xs c,
  simp add: tape_of_nl_abv  tape_of_nat_list.simps)
  fix x xs c
  assume ind: "\<And>xs c. x = length xs \<Longrightarrow> <c # xs @ [b]> = 
    <c # xs> @ Bk # Oc\<^bsup>Suc b\<^esup>"
    and h: "Suc x = length (xs::nat list)" 
  show "<c # xs @ [b]> = <c # xs> @ Bk # Oc\<^bsup>Suc b\<^esup>"
  proof(case_tac xs, simp add: tape_of_nl_abv  tape_of_nat_list.simps)
    fix a list
    assume g: "xs = a # list"
    hence k: "<a # list @ [b]> =  <a # list> @ Bk # Oc\<^bsup>Suc b\<^esup>"
      apply(rule_tac ind)
      using h
      apply(simp)
      done
    from g and k show "<c # xs @ [b]> = <c # xs> @ Bk # Oc\<^bsup>Suc b\<^esup>"
      apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
      done
  qed
qed

lemma [simp]: "length (<aa # a # list>) = Suc (Suc aa) + length (<a # list>)"
apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
done

lemma [simp]: "bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # tape_of_nat_list (a # lista) @ [Bk, Oc]) =
              bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # tape_of_nat_list (a # lista)) + 
              2* 2^(length (Oc\<^bsup>Suc aa\<^esup> @ Bk # tape_of_nat_list (a # lista)))"
using bl_bin_bk_oc[of "Oc\<^bsup>Suc aa\<^esup> @ Bk # tape_of_nat_list (a # lista)"]
apply(simp)
done

lemma [simp]: 
  "bl_bin (<aa # list>) + (4 * rs + 4) * 2 ^ (length (<aa # list>) - Suc 0)
  = bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # <list @ [0]>) + rs * (2 * 2 ^ (aa + length (<list @ [0]>)))"
apply(case_tac "list", simp add: add_mult_distrib, simp)
apply(simp add: tape_of_nl_cons_app2 add_mult_distrib)
apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
done
  
lemma tape_of_nl_app_Suc: "((<list @ [Suc ab]>)) = (<list @ [ab]>) @ [Oc]"
apply(induct list)
apply(simp add: tape_of_nl_abv tape_of_nat_list.simps exp_ind)
apply(case_tac list)
apply(simp_all add:tape_of_nl_abv tape_of_nat_list.simps exp_ind)
done

lemma [simp]: "bl_bin (Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [ab]> @ [Oc])
              = bl_bin (Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [ab]>) +
              2^(length (Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [ab]>))"
apply(simp add: bin_wc_eq)
apply(simp add: bl2nat_cons_oc bl2wc.simps)
using bl2nat_cons_oc[of "Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [ab]>"]
apply(simp)
done
lemma [simp]: "bl_bin (Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [ab]>) + (4 * 2 ^ (aa + length (<list @ [ab]>)) +
         4 * (rs * 2 ^ (aa + length (<list @ [ab]>)))) =
       bl_bin (Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [Suc ab]>) +
         rs * (2 * 2 ^ (aa + length (<list @ [Suc ab]>)))"
apply(simp add: tape_of_nl_app_Suc)
done

declare tape_of_nat[simp del]

fun wcode_double_case_inv :: "nat \<Rightarrow> bin_inv_t"
  where
  "wcode_double_case_inv st ires rs (l, r) = 
          (if st = Suc 0 then wcode_on_left_moving_1 ires rs (l, r)
          else if st = Suc (Suc 0) then wcode_on_checking_1 ires rs (l, r)
          else if st = 3 then wcode_erase1 ires rs (l, r)
          else if st = 4 then wcode_on_right_moving_1 ires rs (l, r)
          else if st = 5 then wcode_goon_right_moving_1 ires rs (l, r)
          else if st = 6 then wcode_backto_standard_pos ires rs (l, r)
          else if st = 13 then wcode_before_double ires rs (l, r)
          else False)"

declare wcode_double_case_inv.simps[simp del]

fun wcode_double_case_state :: "t_conf \<Rightarrow> nat"
  where
  "wcode_double_case_state (st, l, r) = 
   13 - st"

fun wcode_double_case_step :: "t_conf \<Rightarrow> nat"
  where
  "wcode_double_case_step (st, l, r) = 
      (if st = Suc 0 then (length l)
      else if st = Suc (Suc 0) then (length r)
      else if st = 3 then 
                 if hd r = Oc then 1 else 0
      else if st = 4 then (length r)
      else if st = 5 then (length r)
      else if st = 6 then (length l)
      else 0)"

fun wcode_double_case_measure :: "t_conf \<Rightarrow> nat \<times> nat"
  where
  "wcode_double_case_measure (st, l, r) = 
     (wcode_double_case_state (st, l, r), 
      wcode_double_case_step (st, l, r))"

definition wcode_double_case_le :: "(t_conf \<times> t_conf) set"
  where "wcode_double_case_le \<equiv> (inv_image lex_pair wcode_double_case_measure)"

lemma [intro]: "wf lex_pair"
by(auto intro:wf_lex_prod simp:lex_pair_def)

lemma wf_wcode_double_case_le[intro]: "wf wcode_double_case_le"
by(auto intro:wf_inv_image simp: wcode_double_case_le_def )
term fetch

lemma [simp]: "fetch t_wcode_main (Suc 0) Bk = (L, Suc 0)"
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
                fetch.simps nth_of.simps)
done

lemma [simp]: "fetch t_wcode_main (Suc 0) Oc = (L, Suc (Suc 0))"
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
                fetch.simps nth_of.simps)
done

lemma [simp]: "fetch t_wcode_main (Suc (Suc 0)) Oc = (R, 3)"
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
                fetch.simps nth_of.simps)
done

lemma [simp]: "fetch t_wcode_main (Suc (Suc (Suc 0))) Bk = (R, 4)"
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
                fetch.simps nth_of.simps)
done 

lemma [simp]: "fetch t_wcode_main (Suc (Suc (Suc 0))) Oc = (W0, 3)"
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
                fetch.simps nth_of.simps)
done

lemma [simp]: "fetch t_wcode_main 4 Bk = (R, 4)"
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
                fetch.simps nth_of.simps)
done

lemma [simp]: "fetch t_wcode_main 4 Oc = (R, 5)"
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
                fetch.simps nth_of.simps)
done

lemma [simp]: "fetch t_wcode_main 5 Oc = (R, 5)"
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
                fetch.simps nth_of.simps)
done

lemma [simp]: "fetch t_wcode_main 5 Bk = (W1, 6)"
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
                fetch.simps nth_of.simps)
done

lemma [simp]: "fetch t_wcode_main 6 Bk = (R, 13)"
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
                fetch.simps nth_of.simps)
done

lemma [simp]: "fetch t_wcode_main 6 Oc = (L, 6)"
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
                fetch.simps nth_of.simps)
done
lemma [elim]: "Bk\<^bsup>mr\<^esup> = [] \<Longrightarrow> mr = 0"
apply(case_tac mr, auto simp: exponent_def)
done

lemma [simp]: "wcode_on_left_moving_1 ires rs (b, []) = False"
apply(simp add: wcode_on_left_moving_1.simps wcode_on_left_moving_1_B.simps
                wcode_on_left_moving_1_O.simps, auto)
done


declare wcode_on_checking_1.simps[simp del]

lemmas wcode_double_case_inv_simps = 
  wcode_on_left_moving_1.simps wcode_on_left_moving_1_O.simps
  wcode_on_left_moving_1_B.simps wcode_on_checking_1.simps
  wcode_erase1.simps wcode_on_right_moving_1.simps
  wcode_goon_right_moving_1.simps wcode_backto_standard_pos.simps


lemma [simp]: "wcode_on_left_moving_1 ires rs (b, r) \<Longrightarrow> b \<noteq> []"
apply(simp add: wcode_double_case_inv_simps, auto)
done


lemma [elim]: "\<lbrakk>wcode_on_left_moving_1 ires rs (b, Bk # list);
                tl b = aa \<and> hd b # Bk # list = ba\<rbrakk> \<Longrightarrow> 
               wcode_on_left_moving_1 ires rs (aa, ba)"
apply(simp only: wcode_on_left_moving_1.simps wcode_on_left_moving_1_O.simps
                wcode_on_left_moving_1_B.simps)
apply(erule_tac disjE)
apply(erule_tac exE)+
apply(case_tac ml, simp)
apply(rule_tac x = "mr - Suc (Suc 0)" in exI, rule_tac x = rn in exI)
apply(case_tac mr, simp, case_tac nat, simp, simp add: exp_ind)
apply(rule_tac disjI1)
apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI, 
      simp add: exp_ind_def)
apply(erule_tac exE)+
apply(simp)
done


lemma [elim]: 
  "\<lbrakk>wcode_on_left_moving_1 ires rs (b, Oc # list); tl b = aa \<and> hd b # Oc # list = ba\<rbrakk> 
    \<Longrightarrow> wcode_on_checking_1 ires rs (aa, ba)"
apply(simp only: wcode_double_case_inv_simps)
apply(erule_tac disjE)
apply(erule_tac [!] exE)+
apply(case_tac mr, simp, simp add: exp_ind_def)
apply(rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
done


lemma [simp]: "wcode_on_checking_1 ires rs (b, []) = False" 
apply(auto simp: wcode_double_case_inv_simps)
done         
 
lemma [simp]: "wcode_on_checking_1 ires rs (b, Bk # list) = False"
apply(auto simp: wcode_double_case_inv_simps)
done         
  
lemma [elim]: "\<lbrakk>wcode_on_checking_1 ires rs (b, Oc # ba);Oc # b = aa \<and> list = ba\<rbrakk>
  \<Longrightarrow> wcode_erase1 ires rs (aa, ba)"
apply(simp only: wcode_double_case_inv_simps)
apply(erule_tac exE)+
apply(rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
done


lemma [simp]: "wcode_on_checking_1 ires rs (b, []) = False"
apply(simp add: wcode_double_case_inv_simps)
done

lemma [simp]: "wcode_on_checking_1 ires rs ([], Bk # list) = False"
apply(simp add: wcode_double_case_inv_simps)
done

lemma [simp]: "wcode_erase1 ires rs (b, []) = False"
apply(simp add: wcode_double_case_inv_simps)
done

lemma [simp]: "wcode_on_right_moving_1 ires rs (b, []) = False"
apply(simp add: wcode_double_case_inv_simps exp_ind_def)
done

lemma [simp]: "wcode_on_right_moving_1 ires rs (b, []) = False"
apply(simp add: wcode_double_case_inv_simps exp_ind_def)
done

lemma [elim]: "\<lbrakk>wcode_on_right_moving_1 ires rs (b, Bk # ba);  Bk # b = aa \<and> list = b\<rbrakk> \<Longrightarrow> 
  wcode_on_right_moving_1 ires rs (aa, ba)"
apply(simp only: wcode_double_case_inv_simps)
apply(erule_tac exE)+
apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - Suc 0" in exI,
      rule_tac x = rn in exI)
apply(simp add: exp_ind_def)
apply(case_tac mr, simp, simp add: exp_ind_def)
done

lemma [elim]: 
  "\<lbrakk>wcode_on_right_moving_1 ires rs (b, Oc # ba); Oc # b = aa \<and> list = ba\<rbrakk> 
  \<Longrightarrow> wcode_goon_right_moving_1 ires rs (aa, ba)"
apply(simp only: wcode_double_case_inv_simps)
apply(erule_tac exE)+
apply(rule_tac x = "Suc 0" in exI, rule_tac x = "rs" in exI,
      rule_tac x = "ml - Suc (Suc 0)" in exI, rule_tac x = rn in exI)
apply(case_tac mr, simp_all add: exp_ind_def)
apply(case_tac ml, simp, case_tac nat, simp, simp)
apply(simp add: exp_ind_def)
done

lemma [simp]: 
  "wcode_on_right_moving_1 ires rs (b, []) \<Longrightarrow> False"
apply(simp add: wcode_double_case_inv_simps exponent_def)
done

lemma [elim]: "\<lbrakk>wcode_erase1 ires rs (b, Bk # ba); Bk # b = aa \<and> list = ba; c = Bk # ba\<rbrakk> 
  \<Longrightarrow> wcode_on_right_moving_1 ires rs (aa, ba)"
apply(simp only: wcode_double_case_inv_simps)
apply(erule_tac exE)+
apply(rule_tac x = "Suc 0" in exI, rule_tac x = "Suc (Suc ln)" in exI, 
      rule_tac x = rn in exI, simp add: exp_ind)
done

lemma [elim]: "\<lbrakk>wcode_erase1 ires rs (aa, Oc # list);  b = aa \<and> Bk # list = ba\<rbrakk> \<Longrightarrow> 
  wcode_erase1 ires rs (aa, ba)"
apply(simp only: wcode_double_case_inv_simps)
apply(erule_tac exE)+
apply(rule_tac x = ln in exI, rule_tac x = rn in exI, auto)
done

lemma [elim]: "\<lbrakk>wcode_goon_right_moving_1 ires rs (aa, []); b = aa \<and> [Oc] = ba\<rbrakk> 
              \<Longrightarrow> wcode_backto_standard_pos ires rs (aa, ba)"
apply(simp only: wcode_double_case_inv_simps)
apply(erule_tac exE)+
apply(rule_tac disjI2)
apply(simp only:wcode_backto_standard_pos_O.simps)
apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, rule_tac x = ln in exI,
      rule_tac x = rn in exI, simp)
apply(case_tac mr, simp_all add: exponent_def)
done

lemma [elim]: 
  "\<lbrakk>wcode_goon_right_moving_1 ires rs (aa, Bk # list);  b = aa \<and> Oc # list = ba\<rbrakk>
  \<Longrightarrow> wcode_backto_standard_pos ires rs (aa, ba)"
apply(simp only: wcode_double_case_inv_simps)
apply(erule_tac exE)+
apply(rule_tac disjI2)
apply(simp only:wcode_backto_standard_pos_O.simps)
apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, rule_tac x = ln in exI,
      rule_tac x = "rn - Suc 0" in exI, simp)
apply(case_tac mr, simp, case_tac rn, simp, simp_all add: exp_ind_def)
done

lemma [elim]: "\<lbrakk>wcode_goon_right_moving_1 ires rs (b, Oc # ba);  Oc # b = aa \<and> list = ba\<rbrakk> 
  \<Longrightarrow> wcode_goon_right_moving_1 ires rs (aa, ba)"
apply(simp only: wcode_double_case_inv_simps)
apply(erule_tac exE)+
apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - Suc 0" in exI, 
      rule_tac x = ln in exI, rule_tac x = rn in exI)
apply(simp add: exp_ind_def)
apply(case_tac mr, simp, case_tac rn, simp_all add: exp_ind_def)
done

lemma [elim]: "\<lbrakk>wcode_backto_standard_pos ires rs (b, []);  Bk # b = aa\<rbrakk> \<Longrightarrow> False"
apply(auto simp: wcode_double_case_inv_simps wcode_backto_standard_pos_O.simps
                 wcode_backto_standard_pos_B.simps)
apply(case_tac mr, simp_all add: exp_ind_def)
done

lemma [elim]: "\<lbrakk>wcode_backto_standard_pos ires rs (b, Bk # ba); Bk # b = aa \<and> list = ba\<rbrakk> 
  \<Longrightarrow> wcode_before_double ires rs (aa, ba)"
apply(simp only: wcode_double_case_inv_simps wcode_backto_standard_pos_B.simps
                 wcode_backto_standard_pos_O.simps wcode_before_double.simps)
apply(erule_tac disjE)
apply(erule_tac exE)+
apply(rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
apply(auto)
apply(case_tac [!] mr, simp_all add: exp_ind_def)
done

lemma [simp]: "wcode_backto_standard_pos ires rs ([], Oc # list) = False"
apply(auto simp: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps
                 wcode_backto_standard_pos_O.simps)
done

lemma [simp]: "wcode_backto_standard_pos ires rs (b, []) = False"
apply(auto simp: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps
                 wcode_backto_standard_pos_O.simps)
apply(case_tac mr, simp, simp add: exp_ind_def)
done

lemma [elim]: "\<lbrakk>wcode_backto_standard_pos ires rs (b, Oc # list); tl b = aa; hd b # Oc # list =  ba\<rbrakk>
       \<Longrightarrow> wcode_backto_standard_pos ires rs (aa, ba)"
apply(simp only:  wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps
                 wcode_backto_standard_pos_O.simps)
apply(erule_tac disjE)
apply(simp)
apply(erule_tac exE)+
apply(case_tac ml, simp)
apply(rule_tac disjI1, rule_tac conjI)
apply(rule_tac x = ln  in exI, simp, rule_tac x = rn in exI, simp)
apply(rule_tac disjI2)
apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = ln in exI, 
      rule_tac x = rn in exI, simp)
apply(simp add: exp_ind_def)
done

declare new_tape.simps[simp del] nth_of.simps[simp del] fetch.simps[simp del]
lemma wcode_double_case_first_correctness:
  "let P = (\<lambda> (st, l, r). st = 13) in 
       let Q = (\<lambda> (st, l, r). wcode_double_case_inv st ires rs (l, r)) in 
       let f = (\<lambda> stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp) in
       \<exists> n .P (f n) \<and> Q (f (n::nat))"
proof -
  let ?P = "(\<lambda> (st, l, r). st = 13)"
  let ?Q = "(\<lambda> (st, l, r). wcode_double_case_inv st ires rs (l, r))"
  let ?f = "(\<lambda> stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp)"
  have "\<exists> n. ?P (?f n) \<and> ?Q (?f (n::nat))"
  proof(rule_tac halt_lemma2)
    show "wf wcode_double_case_le"
      by auto
  next
    show "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow>
                   ?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_double_case_le"
    proof(rule_tac allI, case_tac "?f na", simp add: tstep_red)
      fix na a b c
      show "a \<noteq> 13 \<and> wcode_double_case_inv a ires rs (b, c) \<longrightarrow>
               (case tstep (a, b, c) t_wcode_main of (st, x) \<Rightarrow> 
                   wcode_double_case_inv st ires rs x) \<and> 
                (tstep (a, b, c) t_wcode_main, a, b, c) \<in> wcode_double_case_le"
        apply(rule_tac impI, simp add: wcode_double_case_inv.simps)
        apply(auto split: if_splits simp: tstep.simps, 
              case_tac [!] c, simp_all, case_tac [!] "(c::block list)!0")
        apply(simp_all add: new_tape.simps wcode_double_case_inv.simps wcode_double_case_le_def
                                        lex_pair_def)
        apply(auto split: if_splits)
        done
    qed
  next
    show "?Q (?f 0)"
      apply(simp add: steps.simps wcode_double_case_inv.simps 
                                  wcode_on_left_moving_1.simps
                                  wcode_on_left_moving_1_B.simps)
      apply(rule_tac disjI1)
      apply(rule_tac x = "Suc m" in exI, simp add: exp_ind_def)
      apply(rule_tac x = "Suc 0" in exI, simp add: exp_ind_def)
      apply(auto)
      done
  next
    show "\<not> ?P (?f 0)"
      apply(simp add: steps.simps)
      done
  qed
  thus "let P = \<lambda>(st, l, r). st = 13;
    Q = \<lambda>(st, l, r). wcode_double_case_inv st ires rs (l, r);
    f = steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main
    in \<exists>n. P (f n) \<and> Q (f n)"
    apply(simp add: Let_def)
    done
qed
    
lemma [elim]: "t_ncorrect tp
    \<Longrightarrow> t_ncorrect (abacus.tshift tp a)"
apply(simp add: t_ncorrect.simps shift_length)
done

lemma tshift_fetch: "\<lbrakk> fetch tp a b = (aa, st'); 0 < st'\<rbrakk>
       \<Longrightarrow> fetch (abacus.tshift tp (length tp1 div 2)) a b 
          = (aa, st' + length tp1 div 2)"
apply(subgoal_tac "a > 0")
apply(auto simp: fetch.simps nth_of.simps shift_length nth_map
                 tshift.simps split: block.splits if_splits)
done

lemma t_steps_steps_eq: "\<lbrakk>steps (st, l, r) tp stp = (st', l', r');
         0 < st';  
         0 < st \<and> st \<le> length tp div 2; 
         t_ncorrect tp1;
          t_ncorrect tp\<rbrakk>
    \<Longrightarrow> t_steps (st + length tp1 div 2, l, r) (tshift tp (length tp1 div 2), 
                                                      length tp1 div 2) stp
       = (st' + length tp1 div 2, l', r')"
apply(induct stp arbitrary: st' l' r', simp add: steps.simps t_steps.simps,
      simp add: tstep_red stepn)
apply(case_tac "(steps (st, l, r) tp stp)", simp)
proof -
  fix stp st' l' r' a b c
  assume ind: "\<And>st' l' r'.
    \<lbrakk>a = st' \<and> b = l' \<and> c = r'; 0 < st'\<rbrakk>
    \<Longrightarrow> t_steps (st + length tp1 div 2, l, r) 
    (abacus.tshift tp (length tp1 div 2), length tp1 div 2) stp = 
     (st' + length tp1 div 2, l', r')"
  and h: "tstep (a, b, c) tp = (st', l', r')" "0 < st'" "t_ncorrect tp1"  "t_ncorrect tp"
  have k: "t_steps (st + length tp1 div 2, l, r) (abacus.tshift tp (length tp1 div 2),
         length tp1 div 2) stp = (a + length tp1 div 2, b, c)"
    apply(rule_tac ind, simp)
    using h
    apply(case_tac a, simp_all add: tstep.simps fetch.simps)
    done
  from h and this show "t_step (t_steps (st + length tp1 div 2, l, r) (abacus.tshift tp (length tp1 div 2), length tp1 div 2) stp)
           (abacus.tshift tp (length tp1 div 2), length tp1 div 2) =
          (st' + length tp1 div 2, l', r')"
    apply(simp add: k)
    apply(simp add: tstep.simps t_step.simps)
    apply(case_tac "fetch tp a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp)
    apply(subgoal_tac "fetch (abacus.tshift tp (length tp1 div 2)) a
                       (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) = (aa, st' + length tp1 div 2)", simp)
    apply(simp add: tshift_fetch)
    done
qed 

lemma t_tshift_lemma: "\<lbrakk> steps (st, l, r) tp stp = (st', l', r'); 
                         st' \<noteq> 0; 
                         stp > 0;
                         0 < st \<and> st \<le> length tp div 2;
                         t_ncorrect tp1;
                         t_ncorrect tp;
                         t_ncorrect tp2
                         \<rbrakk>
         \<Longrightarrow> \<exists> stp>0. steps (st + length tp1 div 2, l, r) (tp1 @ tshift tp (length tp1 div 2) @ tp2) stp 
                  = (st' + length tp1 div 2, l', r')"
proof -
  assume h: "steps (st, l, r) tp stp = (st', l', r')"
    "st' \<noteq> 0" "stp > 0"
    "0 < st \<and> st \<le> length tp div 2"
    "t_ncorrect tp1"
    "t_ncorrect tp"
    "t_ncorrect tp2"
  from h have 
    "\<exists>stp>0. t_steps (st + length tp1 div 2, l, r) (tp1 @ abacus.tshift tp (length tp1 div 2) @ tp2, 0) stp = 
                            (st' + length tp1 div 2, l', r')"
    apply(rule_tac stp = stp in turing_shift, simp_all add: shift_length)
    apply(simp add: t_steps_steps_eq)
    apply(simp add: t_ncorrect.simps shift_length)
    done
  thus "\<exists> stp>0. steps (st + length tp1 div 2, l, r) (tp1 @ tshift tp (length tp1 div 2) @ tp2) stp 
                  = (st' + length tp1 div 2, l', r')"
    apply(erule_tac exE)
    apply(rule_tac x = stp in exI, simp)
    apply(subgoal_tac "length (tp1 @ abacus.tshift tp (length tp1 div 2) @ tp2) mod 2 = 0")
    apply(simp only: steps_eq)
    using h
    apply(auto simp: t_ncorrect.simps shift_length)
    apply arith
    done
qed  
  

lemma t_twice_len_ge: "Suc 0 \<le> length t_twice div 2"
apply(simp add: t_twice_def tMp.simps shift_length)
done

lemma [intro]: "rec_calc_rel (recf.id (Suc 0) 0) [rs] rs"
  apply(rule_tac calc_id, simp_all)
  done
  
lemma [intro]: "rec_calc_rel (constn 2) [rs] 2"
using prime_rel_exec_eq[of "constn 2" "[rs]" 2]
apply(subgoal_tac "primerec (constn 2) 1", auto)
done

lemma  [intro]: "rec_calc_rel rec_mult [rs, 2] (2 * rs)"
using prime_rel_exec_eq[of "rec_mult" "[rs, 2]"  "2*rs"]
apply(subgoal_tac "primerec rec_mult (Suc (Suc 0))", auto)
done
lemma t_twice_correct: "\<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) 
            (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0)) stp =
       (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
proof(case_tac "rec_ci rec_twice")
  fix a b c
  assume h: "rec_ci rec_twice = (a, b, c)"
  have "\<exists>stp m l. steps (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\<^bsup>n\<^esup>) (tm_of abc_twice @ tMp (Suc 0) 
    (start_of twice_ly (length abc_twice) - 1)) stp = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2*rs)\<^esup> @ Bk\<^bsup>l\<^esup>)"
  proof(rule_tac t_compiled_by_rec)
    show "rec_ci rec_twice = (a, b, c)" by (simp add: h)
  next
    show "rec_calc_rel rec_twice [rs] (2 * rs)"
      apply(simp add: rec_twice_def)
      apply(rule_tac rs =  "[rs, 2]" in calc_cn, simp_all)
      apply(rule_tac allI, case_tac k, auto)
      done
  next
    show "length [rs] = Suc 0" by simp
  next
    show "layout_of (a [+] dummy_abc (Suc 0)) = layout_of (a [+] dummy_abc (Suc 0))"
      by simp
  next
    show "start_of twice_ly (length abc_twice) = 
      start_of (layout_of (a [+] dummy_abc (Suc 0))) (length (a [+] dummy_abc (Suc 0)))"
      using h
      apply(simp add: twice_ly_def abc_twice_def)
      done
  next
    show "tm_of abc_twice = tm_of (a [+] dummy_abc (Suc 0))"
      using h
      apply(simp add: abc_twice_def)
      done
  qed
  thus "\<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) 
            (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0)) stp =
       (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
    apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
    done
qed

lemma change_termi_state_fetch: "\<lbrakk>fetch ap a b = (aa, st); st > 0\<rbrakk>
       \<Longrightarrow> fetch (change_termi_state ap) a b = (aa, st)"
apply(case_tac b, auto simp: fetch.simps nth_of.simps change_termi_state.simps nth_map
                       split: if_splits block.splits)
done

lemma change_termi_state_exec_in_range:
     "\<lbrakk>steps (st, l, r) ap stp = (st', l', r'); st' \<noteq> 0\<rbrakk>
    \<Longrightarrow> steps (st, l, r) (change_termi_state ap) stp = (st', l', r')"
proof(induct stp arbitrary: st l r st' l' r', simp add: steps.simps)
  fix stp st l r st' l' r'
  assume ind: "\<And>st l r st' l' r'. 
    \<lbrakk>steps (st, l, r) ap stp = (st', l', r'); st' \<noteq> 0\<rbrakk> \<Longrightarrow>
    steps (st, l, r) (change_termi_state ap) stp = (st', l', r')"
  and h: "steps (st, l, r) ap (Suc stp) = (st', l', r')" "st' \<noteq> 0"
  from h show "steps (st, l, r) (change_termi_state ap) (Suc stp) = (st', l', r')"
  proof(simp add: tstep_red, case_tac "steps (st, l, r) ap stp", simp)
    fix a b c
    assume g: "steps (st, l, r) ap stp = (a, b, c)"
              "tstep (a, b, c) ap = (st', l', r')" "0 < st'"
    hence "steps (st, l, r) (change_termi_state ap) stp = (a, b, c)"
      apply(rule_tac ind, simp)
      apply(case_tac a, simp_all add: tstep_0)
      done
    from g and this show "tstep (steps (st, l, r) (change_termi_state ap) stp)
      (change_termi_state ap) = (st', l', r')"
      apply(simp add: tstep.simps)
      apply(case_tac "fetch ap a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp)
      apply(subgoal_tac "fetch (change_termi_state ap) a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)
                   = (aa, st')", simp)
      apply(simp add: change_termi_state_fetch)
      done
  qed
qed

lemma change_termi_state_fetch0: 
  "\<lbrakk>0 < a; a \<le> length ap div 2; t_correct ap; fetch ap a b = (aa, 0)\<rbrakk>
  \<Longrightarrow> fetch (change_termi_state ap) a b = (aa, Suc (length ap div 2))"
apply(case_tac b, auto simp: fetch.simps nth_of.simps change_termi_state.simps nth_map
                       split: if_splits block.splits)
done

lemma turing_change_termi_state: 
  "\<lbrakk>steps (Suc 0, l, r) ap stp = (0, l', r'); t_correct ap\<rbrakk>
     \<Longrightarrow> \<exists> stp. steps (Suc 0, l, r) (change_termi_state ap) stp = 
        (Suc (length ap div 2), l', r')"
apply(drule first_halt_point)
apply(erule_tac exE)
apply(rule_tac x = "Suc stp" in exI, simp add: tstep_red)
apply(case_tac "steps (Suc 0, l, r) ap stp")
apply(simp add: isS0_def change_termi_state_exec_in_range)
apply(subgoal_tac "steps (Suc 0, l, r) (change_termi_state ap) stp = (a, b, c)", simp)
apply(simp add: tstep.simps)
apply(case_tac "fetch ap a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp)
apply(subgoal_tac "fetch (change_termi_state ap) a 
  (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) = (aa, Suc (length ap div 2))", simp)
apply(rule_tac ap = ap in change_termi_state_fetch0, simp_all)
apply(rule_tac tp = "(l, r)" and l = b and r = c  and stp = stp and A = ap in s_keep, simp_all)
apply(simp add: change_termi_state_exec_in_range)
done

lemma t_twice_change_term_state:
  "\<exists> stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp
     = (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
using t_twice_correct[of ires rs n]
apply(erule_tac exE)
apply(erule_tac exE)
apply(erule_tac exE)
proof(drule_tac turing_change_termi_state)
  fix stp ln rn
  show "t_correct (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0))"
    apply(rule_tac t_compiled_correct, simp_all)
    apply(simp add: twice_ly_def)
    done
next
  fix stp ln rn
  show "\<exists>stp. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
    (change_termi_state (tm_of abc_twice @ tMp (Suc 0) 
    (start_of twice_ly (length abc_twice) - Suc 0))) stp =
    (Suc (length (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0)) div 2),
    Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>) \<Longrightarrow>
    \<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp = 
    (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
    apply(erule_tac exE)
    apply(simp add: t_twice_len_def t_twice_def)
    apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
    done
qed

lemma t_twice_append_pre:
  "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp
  = (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)
   \<Longrightarrow> \<exists> stp>0. steps (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
     (t_wcode_main_first_part @ tshift t_twice (length t_wcode_main_first_part div 2) @
      ([(L, 1), (L, 1)] @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp 
    = (Suc (t_twice_len) + length t_wcode_main_first_part div 2, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
proof(rule_tac t_tshift_lemma, simp_all add: t_twice_len_ge)
  assume "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp = 
    (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  thus "0 < stp"
    apply(case_tac stp, simp add: steps.simps t_twice_len_ge t_twice_len_def)
    using t_twice_len_ge
    apply(simp, simp)
    done
next
  show "t_ncorrect t_wcode_main_first_part"
    apply(simp add: t_ncorrect.simps t_wcode_main_first_part_def)
    done
next
  show "t_ncorrect t_twice"
    using length_tm_even[of abc_twice]
    apply(auto simp: t_ncorrect.simps t_twice_def)
    apply(arith)
    done
next
  show "t_ncorrect ((L, Suc 0) # (L, Suc 0) #
       abacus.tshift t_fourtimes (t_twice_len + 13) @ [(L, Suc 0), (L, Suc 0)])"
    using length_tm_even[of abc_fourtimes]
    apply(simp add: t_ncorrect.simps shift_length t_fourtimes_def)
    apply arith
    done
qed
  
lemma t_twice_append:
  "\<exists> stp ln rn. steps (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
     (t_wcode_main_first_part @ tshift t_twice (length t_wcode_main_first_part div 2) @
      ([(L, 1), (L, 1)] @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp 
    = (Suc (t_twice_len) + length t_wcode_main_first_part div 2, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  using t_twice_change_term_state[of ires rs n]
  apply(erule_tac exE)
  apply(erule_tac exE)
  apply(erule_tac exE)
  apply(drule_tac t_twice_append_pre)
  apply(erule_tac exE)
  apply(rule_tac x = stpa in exI, rule_tac x = ln in exI, rule_tac x = rn in exI)
  apply(simp)
  done
  
lemma [simp]: "fetch t_wcode_main (Suc (t_twice_len + length t_wcode_main_first_part div 2)) Oc
     = (L, Suc 0)"
apply(subgoal_tac "length (t_twice) mod 2 = 0")
apply(simp add: t_wcode_main_def nth_append fetch.simps t_wcode_main_first_part_def 
  nth_of.simps shift_length t_twice_len_def, auto)
apply(simp add: t_twice_def)
apply(subgoal_tac "length (tm_of abc_twice) mod 2 = 0")
apply arith
apply(rule_tac tm_even)
done

lemma wcode_jump1: 
  "\<exists> stp ln rn. steps (Suc (t_twice_len) + length t_wcode_main_first_part div 2,
                       Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>n\<^esup>)
     t_wcode_main stp 
    = (Suc 0, Bk\<^bsup>ln\<^esup> @ Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
apply(rule_tac x = "Suc 0" in exI, rule_tac x = "m" in exI, rule_tac x = n in exI)
apply(simp add: steps.simps tstep.simps exp_ind_def new_tape.simps)
apply(case_tac m, simp, simp add: exp_ind_def)
apply(simp add: exp_ind_def[THEN sym] exp_ind[THEN sym])
done

lemma wcode_main_first_part_len:
  "length t_wcode_main_first_part = 24"
  apply(simp add: t_wcode_main_first_part_def)
  done

lemma wcode_double_case: 
  shows "\<exists>stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
          (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
proof -
  have "\<exists>stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
          (13,  Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
    using wcode_double_case_first_correctness[of ires rs m n]
    apply(simp)
    apply(erule_tac exE)
    apply(case_tac "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, 
           Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main na",
          auto simp: wcode_double_case_inv.simps
                     wcode_before_double.simps)
    apply(rule_tac x = na in exI, rule_tac x = ln in exI, rule_tac x = rn in exI)
    apply(simp)
    done    
  from this obtain stpa lna rna where stp1: 
    "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa = 
    (13, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast
  have "\<exists> stp ln rn. steps (13, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp =
    (13 + t_twice_len, Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rn\<^esup>)"
    using t_twice_append[of "Bk\<^bsup>lna\<^esup> @ Oc # ires" "Suc rs" rna]
    apply(erule_tac exE)
    apply(erule_tac exE)
    apply(erule_tac exE)
    apply(simp add: wcode_main_first_part_len)
    apply(rule_tac x = stp in exI, rule_tac x = "ln + lna" in exI, 
          rule_tac x = rn in exI)
    apply(simp add: t_wcode_main_def)
    apply(simp add: exp_ind_def[THEN sym] exp_add[THEN sym])
    done
  from this obtain stpb lnb rnb where stp2: 
    "steps (13, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb =
    (13 + t_twice_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnb\<^esup>)" by blast
  have "\<exists>stp ln rn. steps (13 + t_twice_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires,
    Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnb\<^esup>) t_wcode_main stp = 
       (Suc 0,  Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rn\<^esup>)"
    using wcode_jump1[of lnb "Oc # ires" "Suc rs" rnb]
    apply(erule_tac exE)
    apply(erule_tac exE)
    apply(erule_tac exE)
    apply(rule_tac x = stp in exI, 
          rule_tac x = ln in exI, 
          rule_tac x = rn in exI, simp add:wcode_main_first_part_len t_wcode_main_def)
    apply(subgoal_tac "Bk\<^bsup>lnb\<^esup> @ Bk # Bk # Oc # ires = Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires", simp)
    apply(simp add: exp_ind_def[THEN sym] exp_ind[THEN sym])
    apply(simp)
    apply(case_tac lnb, simp, simp add: exp_ind_def[THEN sym] exp_ind)
    done               
  from this obtain stpc lnc rnc where stp3: 
    "steps (13 + t_twice_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires,
    Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnb\<^esup>) t_wcode_main stpc = 
       (Suc 0,  Bk # Bk\<^bsup>lnc\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnc\<^esup>)"
    by blast
  from stp1 stp2 stp3 show "?thesis"
    apply(rule_tac x = "stpa + stpb + stpc" in exI, rule_tac x = lnc in exI,
         rule_tac x = rnc in exI)
    apply(simp add: steps_add)
    done
qed
    

(* Begin: fourtime_case*)
fun wcode_on_left_moving_2_B :: "bin_inv_t"
  where
  "wcode_on_left_moving_2_B ires rs (l, r) =
     (\<exists> ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # Bk # Oc # ires \<and>
                 r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
                 ml + mr > Suc 0 \<and> mr > 0)"

fun wcode_on_left_moving_2_O :: "bin_inv_t"
  where
  "wcode_on_left_moving_2_O ires rs (l, r) =
     (\<exists> ln rn. l = Bk # Oc # ires \<and>
               r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"

fun wcode_on_left_moving_2 :: "bin_inv_t"
  where
  "wcode_on_left_moving_2 ires rs (l, r) = 
      (wcode_on_left_moving_2_B ires rs (l, r) \<or> 
      wcode_on_left_moving_2_O ires rs (l, r))"

fun wcode_on_checking_2 :: "bin_inv_t"
  where
  "wcode_on_checking_2 ires rs (l, r) =
       (\<exists> ln rn. l = Oc#ires \<and> 
                 r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"

fun wcode_goon_checking :: "bin_inv_t"
  where
  "wcode_goon_checking ires rs (l, r) =
       (\<exists> ln rn. l = ires \<and>
                 r = Oc # Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"

fun wcode_right_move :: "bin_inv_t"
  where
  "wcode_right_move ires rs (l, r) = 
     (\<exists> ln rn. l = Oc # ires \<and>
                 r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"

fun wcode_erase2 :: "bin_inv_t"
  where
  "wcode_erase2 ires rs (l, r) = 
        (\<exists> ln rn. l = Bk # Oc # ires \<and>
                 tl r = Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"

fun wcode_on_right_moving_2 :: "bin_inv_t"
  where
  "wcode_on_right_moving_2 ires rs (l, r) = 
        (\<exists> ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # ires \<and> 
                     r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> ml + mr > Suc 0)"

fun wcode_goon_right_moving_2 :: "bin_inv_t"
  where
  "wcode_goon_right_moving_2 ires rs (l, r) = 
        (\<exists> ml mr ln rn. l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and>
                        r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> ml + mr = Suc rs)"

fun wcode_backto_standard_pos_2_B :: "bin_inv_t"
  where
  "wcode_backto_standard_pos_2_B ires rs (l, r) = 
           (\<exists> ln rn. l = Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
                     r = Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"

fun wcode_backto_standard_pos_2_O :: "bin_inv_t"
  where
  "wcode_backto_standard_pos_2_O ires rs (l, r) = 
          (\<exists> ml mr ln rn. l = Oc\<^bsup>ml \<^esup>@ Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
                          r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
                          ml + mr = (Suc (Suc rs)) \<and> mr > 0)"

fun wcode_backto_standard_pos_2 :: "bin_inv_t"
  where
  "wcode_backto_standard_pos_2 ires rs (l, r) = 
           (wcode_backto_standard_pos_2_O ires rs (l, r) \<or> 
           wcode_backto_standard_pos_2_B ires rs (l, r))"

fun wcode_before_fourtimes :: "bin_inv_t"
  where
  "wcode_before_fourtimes ires rs (l, r) = 
          (\<exists> ln rn. l = Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
                    r = Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"

declare wcode_on_left_moving_2_B.simps[simp del] wcode_on_left_moving_2.simps[simp del]
        wcode_on_left_moving_2_O.simps[simp del] wcode_on_checking_2.simps[simp del]
        wcode_goon_checking.simps[simp del] wcode_right_move.simps[simp del]
        wcode_erase2.simps[simp del]
        wcode_on_right_moving_2.simps[simp del] wcode_goon_right_moving_2.simps[simp del]
        wcode_backto_standard_pos_2_B.simps[simp del] wcode_backto_standard_pos_2_O.simps[simp del]
        wcode_backto_standard_pos_2.simps[simp del]

lemmas wcode_fourtimes_invs = 
       wcode_on_left_moving_2_B.simps wcode_on_left_moving_2.simps
        wcode_on_left_moving_2_O.simps wcode_on_checking_2.simps
        wcode_goon_checking.simps wcode_right_move.simps
        wcode_erase2.simps
        wcode_on_right_moving_2.simps wcode_goon_right_moving_2.simps
        wcode_backto_standard_pos_2_B.simps wcode_backto_standard_pos_2_O.simps
        wcode_backto_standard_pos_2.simps

fun wcode_fourtimes_case_inv :: "nat \<Rightarrow> bin_inv_t"
  where
  "wcode_fourtimes_case_inv st ires rs (l, r) = 
           (if st = Suc 0 then wcode_on_left_moving_2 ires rs (l, r)
            else if st = Suc (Suc 0) then wcode_on_checking_2 ires rs (l, r)
            else if st = 7 then wcode_goon_checking ires rs (l, r)
            else if st = 8 then wcode_right_move ires rs (l, r)
            else if st = 9 then wcode_erase2 ires rs (l, r)
            else if st = 10 then wcode_on_right_moving_2 ires rs (l, r)
            else if st = 11 then wcode_goon_right_moving_2 ires rs (l, r)
            else if st = 12 then wcode_backto_standard_pos_2 ires rs (l, r)
            else if st = t_twice_len + 14 then wcode_before_fourtimes ires rs (l, r)
            else False)"

declare wcode_fourtimes_case_inv.simps[simp del]

fun wcode_fourtimes_case_state :: "t_conf \<Rightarrow> nat"
  where
  "wcode_fourtimes_case_state (st, l, r) = 13 - st"

fun wcode_fourtimes_case_step :: "t_conf \<Rightarrow> nat"
  where
  "wcode_fourtimes_case_step (st, l, r) = 
         (if st = Suc 0 then length l
          else if st = 9 then 
           (if hd r = Oc then 1
            else 0)
          else if st = 10 then length r
          else if st = 11 then length r
          else if st = 12 then length l
          else 0)"

fun wcode_fourtimes_case_measure :: "t_conf \<Rightarrow> nat \<times> nat"
  where
  "wcode_fourtimes_case_measure (st, l, r) = 
     (wcode_fourtimes_case_state (st, l, r), 
      wcode_fourtimes_case_step (st, l, r))"

definition wcode_fourtimes_case_le :: "(t_conf \<times> t_conf) set"
  where "wcode_fourtimes_case_le \<equiv> (inv_image lex_pair wcode_fourtimes_case_measure)"

lemma wf_wcode_fourtimes_case_le[intro]: "wf wcode_fourtimes_case_le"
by(auto intro:wf_inv_image simp: wcode_fourtimes_case_le_def)

lemma [simp]: "fetch t_wcode_main (Suc (Suc 0)) Bk = (L, 7)"
apply(simp add: t_wcode_main_def fetch.simps 
  t_wcode_main_first_part_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_main 7 Oc = (R, 8)"
apply(simp add: t_wcode_main_def fetch.simps 
  t_wcode_main_first_part_def nth_of.simps)
done
 
lemma [simp]: "fetch t_wcode_main 8 Bk = (R, 9)"
apply(simp add: t_wcode_main_def fetch.simps 
  t_wcode_main_first_part_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_main 9 Bk = (R, 10)"
apply(simp add: t_wcode_main_def fetch.simps 
  t_wcode_main_first_part_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_main 9 Oc = (W0, 9)"
apply(simp add: t_wcode_main_def fetch.simps 
  t_wcode_main_first_part_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_main 10 Bk = (R, 10)"
apply(simp add: t_wcode_main_def fetch.simps 
  t_wcode_main_first_part_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_main 10 Oc = (R, 11)"
apply(simp add: t_wcode_main_def fetch.simps 
  t_wcode_main_first_part_def nth_of.simps)
done 

lemma [simp]: "fetch t_wcode_main 11 Bk = (W1, 12)"
apply(simp add: t_wcode_main_def fetch.simps 
  t_wcode_main_first_part_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_main 11 Oc = (R, 11)"
apply(simp add: t_wcode_main_def fetch.simps 
  t_wcode_main_first_part_def nth_of.simps)
done 

lemma [simp]: "fetch t_wcode_main 12 Oc = (L, 12)"
apply(simp add: t_wcode_main_def fetch.simps 
  t_wcode_main_first_part_def nth_of.simps)
done 

lemma [simp]: "fetch t_wcode_main 12 Bk = (R, t_twice_len + 14)"
apply(simp add: t_wcode_main_def fetch.simps 
  t_wcode_main_first_part_def nth_of.simps)
done


lemma [simp]: "wcode_on_left_moving_2 ires rs (b, []) = False"
apply(auto simp: wcode_fourtimes_invs)
done

lemma [simp]: "wcode_on_checking_2 ires rs (b, []) = False"
apply(auto simp: wcode_fourtimes_invs)
done          

lemma [simp]: "wcode_goon_checking ires rs (b, []) = False"
apply(auto simp: wcode_fourtimes_invs)
done

lemma [simp]: "wcode_right_move ires rs (b, []) = False"
apply(auto simp: wcode_fourtimes_invs)
done

lemma [simp]: "wcode_erase2 ires rs (b, []) = False"
apply(auto simp: wcode_fourtimes_invs)
done

lemma [simp]: "wcode_on_right_moving_2 ires rs (b, []) = False"
apply(auto simp: wcode_fourtimes_invs exponent_def)
done

lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, []) = False"
apply(auto simp: wcode_fourtimes_invs exponent_def)
done
    
lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
apply(simp add: wcode_fourtimes_invs, auto)
done
        
lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Bk # list) \<Longrightarrow>  wcode_on_left_moving_2 ires rs (tl b, hd b # Bk # list)"
apply(simp only: wcode_fourtimes_invs)
apply(erule_tac disjE)
apply(erule_tac exE)+
apply(case_tac ml, simp)
apply(rule_tac x = "mr - (Suc (Suc 0))" in exI, rule_tac x = rn in exI, simp)
apply(case_tac mr, simp, case_tac nat, simp, simp add: exp_ind)
apply(rule_tac disjI1)
apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI,
      simp add: exp_ind_def)
apply(simp)
done

lemma [simp]: "wcode_on_checking_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
apply(auto simp: wcode_fourtimes_invs)
done

lemma  [simp]: "wcode_on_checking_2 ires rs (b, Bk # list)
       \<Longrightarrow>   wcode_goon_checking ires rs (tl b, hd b # Bk # list)"
apply(simp only: wcode_fourtimes_invs)
apply(auto)
done

lemma [simp]: "wcode_goon_checking ires rs (b, Bk # list) = False"
apply(simp add: wcode_fourtimes_invs)
done

lemma [simp]: " wcode_right_move ires rs (b, Bk # list) \<Longrightarrow> b\<noteq> []" 
apply(simp add: wcode_fourtimes_invs)
done

lemma [simp]: "wcode_right_move ires rs (b, Bk # list) \<Longrightarrow>  wcode_erase2 ires rs (Bk # b, list)"
apply(auto simp:wcode_fourtimes_invs )
apply(rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
done

lemma [simp]: "wcode_erase2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
apply(auto simp: wcode_fourtimes_invs)
done

lemma [simp]: "wcode_erase2 ires rs (b, Bk # list) \<Longrightarrow> wcode_on_right_moving_2 ires rs (Bk # b, list)"
apply(auto simp:wcode_fourtimes_invs )
apply(rule_tac x = "Suc (Suc 0)" in exI, simp add: exp_ind)
apply(rule_tac x =  "Suc (Suc ln)" in exI, simp add: exp_ind, auto)
done

lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
apply(auto simp:wcode_fourtimes_invs )
done

lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Bk # list)
       \<Longrightarrow> wcode_on_right_moving_2 ires rs (Bk # b, list)"
apply(auto simp: wcode_fourtimes_invs)
apply(rule_tac x = "Suc ml" in exI, simp add: exp_ind_def)
apply(rule_tac x = "mr - 1" in exI, case_tac mr, auto simp: exp_ind_def)
done

lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
apply(auto simp: wcode_fourtimes_invs)
done

lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Bk # list) \<Longrightarrow> 
                 wcode_backto_standard_pos_2 ires rs (b, Oc # list)"
apply(simp add: wcode_fourtimes_invs, auto)
apply(rule_tac x = ml in exI, auto)
apply(rule_tac x = "Suc 0" in exI, simp)
apply(case_tac mr, simp_all add: exp_ind_def)
apply(rule_tac x = "rn - 1" in exI, simp)
apply(case_tac rn, simp, simp add: exp_ind_def)
done
   
lemma  [simp]: "wcode_backto_standard_pos_2 ires rs (b, Bk # list) \<Longrightarrow>  b \<noteq> []"
apply(simp add: wcode_fourtimes_invs, auto)
done

lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
apply(simp add: wcode_fourtimes_invs, auto)
done

lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Oc # list) \<Longrightarrow> 
                     wcode_on_checking_2 ires rs (tl b, hd b # Oc # list)"
apply(auto simp: wcode_fourtimes_invs)
apply(case_tac [!] mr, simp_all add: exp_ind_def)
done

lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, []) \<Longrightarrow> b \<noteq> []"
apply(auto simp: wcode_fourtimes_invs)
done

lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, []) \<Longrightarrow>
              wcode_backto_standard_pos_2 ires rs (b, [Oc])"
apply(simp only: wcode_fourtimes_invs)
apply(erule_tac exE)+
apply(rule_tac disjI1)
apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, 
      rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
apply(case_tac mr, simp, simp add: exp_ind_def)
done

lemma "wcode_backto_standard_pos_2 ires rs (b, Bk # list)
       \<Longrightarrow> (\<exists>ln. b = Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires) \<and> (\<exists>rn. list = Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
apply(auto simp: wcode_fourtimes_invs)
apply(case_tac [!] mr, auto simp: exp_ind_def)
done


lemma [simp]: "wcode_on_checking_2 ires rs (b, Oc # list) \<Longrightarrow> False"
apply(simp add: wcode_fourtimes_invs)
done

lemma [simp]: "wcode_goon_checking ires rs (b, Oc # list) \<Longrightarrow>
  (b = [] \<longrightarrow> wcode_right_move ires rs ([Oc], list)) \<and>
  (b \<noteq> [] \<longrightarrow> wcode_right_move ires rs (Oc # b, list))"
apply(simp only: wcode_fourtimes_invs)
apply(erule_tac exE)+
apply(auto)
done

lemma [simp]: "wcode_right_move ires rs (b, Oc # list) = False"
apply(auto simp: wcode_fourtimes_invs)
done

lemma [simp]: " wcode_erase2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
apply(simp add: wcode_fourtimes_invs)
done

lemma [simp]: "wcode_erase2 ires rs (b, Oc # list)
       \<Longrightarrow> wcode_erase2 ires rs (b, Bk # list)"
apply(auto simp: wcode_fourtimes_invs)
done

lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
apply(simp only: wcode_fourtimes_invs)
apply(auto)
done

lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Oc # list)
       \<Longrightarrow> wcode_goon_right_moving_2 ires rs (Oc # b, list)"
apply(auto simp: wcode_fourtimes_invs)
apply(case_tac mr, simp_all add: exp_ind_def)
apply(rule_tac x = "Suc 0" in exI, auto)
apply(rule_tac x = "ml - 2" in exI)
apply(case_tac ml, simp, case_tac nat, simp_all add: exp_ind_def)
done

lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
apply(simp only:wcode_fourtimes_invs, auto)
done

lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, Bk # list)
       \<Longrightarrow> (\<exists>ln. b = Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires) \<and> (\<exists>rn. list = Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
apply(simp add: wcode_fourtimes_invs, auto)
apply(case_tac [!] mr, simp_all add: exp_ind_def)
done

lemma [simp]: "wcode_on_checking_2 ires rs (b, Oc # list) = False"
apply(simp add: wcode_fourtimes_invs)
done

lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Oc # list) \<Longrightarrow>
       wcode_goon_right_moving_2 ires rs (Oc # b, list)"
apply(simp only:wcode_fourtimes_invs, auto)
apply(rule_tac x = "Suc ml" in exI, auto simp: exp_ind_def)
apply(rule_tac x = "mr - 1" in exI)
apply(case_tac mr, case_tac rn, auto simp: exp_ind_def)
done

lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
apply(simp only: wcode_fourtimes_invs, auto)
done
 
lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, Oc # list)    
            \<Longrightarrow> wcode_backto_standard_pos_2 ires rs (tl b, hd b # Oc # list)"
apply(simp only: wcode_fourtimes_invs)
apply(erule_tac disjE)
apply(erule_tac exE)+
apply(case_tac ml, simp)
apply(rule_tac disjI2)
apply(rule_tac conjI, rule_tac x = ln in exI, simp)
apply(rule_tac x = rn in exI, simp)
apply(rule_tac disjI1)
apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, 
      rule_tac x = ln in exI, rule_tac x = rn in exI, simp add: exp_ind_def)
apply(simp)
done

lemma wcode_fourtimes_case_first_correctness:
 shows "let P = (\<lambda> (st, l, r). st = t_twice_len + 14) in 
  let Q = (\<lambda> (st, l, r). wcode_fourtimes_case_inv st ires rs (l, r)) in 
  let f = (\<lambda> stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp) in
  \<exists> n .P (f n) \<and> Q (f (n::nat))"
proof -
  let ?P = "(\<lambda> (st, l, r). st = t_twice_len + 14)"
  let ?Q = "(\<lambda> (st, l, r). wcode_fourtimes_case_inv st ires rs (l, r))"
  let ?f = "(\<lambda> stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp)"
  have "\<exists> n . ?P (?f n) \<and> ?Q (?f (n::nat))"
  proof(rule_tac halt_lemma2)
    show "wf wcode_fourtimes_case_le"
      by auto
  next
    show "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow>
                  ?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_fourtimes_case_le"
    apply(rule_tac allI,
     case_tac "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main na", simp,
     rule_tac impI)
    apply(simp add: tstep_red tstep.simps, case_tac c, simp, case_tac [2] aa, simp_all)
    
    apply(simp_all add: wcode_fourtimes_case_inv.simps new_tape.simps 
                        wcode_fourtimes_case_le_def lex_pair_def split: if_splits)
    done
  next
    show "?Q (?f 0)"
      apply(simp add: steps.simps wcode_fourtimes_case_inv.simps)
      apply(simp add: wcode_on_left_moving_2.simps wcode_on_left_moving_2_B.simps 
                      wcode_on_left_moving_2_O.simps)
      apply(rule_tac x = "Suc m" in exI, simp add: exp_ind_def)
      apply(rule_tac x ="Suc 0" in exI, auto)
      done
  next
    show "\<not> ?P (?f 0)"
      apply(simp add: steps.simps)
      done
  qed
  thus "?thesis"
    apply(erule_tac exE, simp)
    done
qed

definition t_fourtimes_len :: "nat"
  where
  "t_fourtimes_len = (length t_fourtimes div 2)"

lemma t_fourtimes_len_gr:  "t_fourtimes_len > 0"
apply(simp add: t_fourtimes_len_def t_fourtimes_def)
done

lemma t_fourtimes_correct: 
  "\<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) 
    (tm_of abc_fourtimes @ tMp (Suc 0) (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) stp =
       (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
proof(case_tac "rec_ci rec_fourtimes")
  fix a b c
  assume h: "rec_ci rec_fourtimes = (a, b, c)"
  have "\<exists>stp m l. steps (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\<^bsup>n\<^esup>) (tm_of abc_fourtimes @ tMp (Suc 0) 
    (start_of fourtimes_ly (length abc_fourtimes) - 1)) stp = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4*rs)\<^esup> @ Bk\<^bsup>l\<^esup>)"
  proof(rule_tac t_compiled_by_rec)
    show "rec_ci rec_fourtimes = (a, b, c)" by (simp add: h)
  next
    show "rec_calc_rel rec_fourtimes [rs] (4 * rs)"
      using prime_rel_exec_eq [of rec_fourtimes "[rs]" "4 * rs"]
      apply(subgoal_tac "primerec rec_fourtimes (length [rs])")
      apply(simp_all add: rec_fourtimes_def rec_exec.simps)
      apply(auto)
      apply(simp only: Nat.One_nat_def[THEN sym], auto)
      done
  next
    show "length [rs] = Suc 0" by simp
  next
    show "layout_of (a [+] dummy_abc (Suc 0)) = layout_of (a [+] dummy_abc (Suc 0))"
      by simp
  next
    show "start_of fourtimes_ly (length abc_fourtimes) = 
      start_of (layout_of (a [+] dummy_abc (Suc 0))) (length (a [+] dummy_abc (Suc 0)))"
      using h
      apply(simp add: fourtimes_ly_def abc_fourtimes_def)
      done
  next
    show "tm_of abc_fourtimes = tm_of (a [+] dummy_abc (Suc 0))"
      using h
      apply(simp add: abc_fourtimes_def)
      done
  qed
  thus "\<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) 
            (tm_of abc_fourtimes @ tMp (Suc 0) (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) stp =
       (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
    apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
    done
qed

lemma t_fourtimes_change_term_state:
  "\<exists> stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp
     = (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
using t_fourtimes_correct[of ires rs n]
apply(erule_tac exE)
apply(erule_tac exE)
apply(erule_tac exE)
proof(drule_tac turing_change_termi_state)
  fix stp ln rn
  show "t_correct (tm_of abc_fourtimes @ tMp (Suc 0) 
    (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))"
    apply(rule_tac t_compiled_correct, auto simp: fourtimes_ly_def)
    done
next
  fix stp ln rn
  show "\<exists>stp. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
    (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) 
        (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) stp =
    (Suc (length (tm_of abc_fourtimes @ tMp (Suc 0) (start_of fourtimes_ly 
    (length abc_fourtimes) - Suc 0)) div 2), Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>) \<Longrightarrow>
    \<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp =
    (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
    apply(erule_tac exE)
    apply(simp add: t_fourtimes_len_def t_fourtimes_def)
    apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
    done
qed

lemma t_fourtimes_append_pre:
  "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp
  = (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)
   \<Longrightarrow> \<exists> stp>0. steps (Suc 0 + length (t_wcode_main_first_part @ 
              tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2,
       Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
     ((t_wcode_main_first_part @ 
  tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) @ 
  tshift t_fourtimes (length (t_wcode_main_first_part @ 
  tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2) @ ([(L, 1), (L, 1)])) stp 
  = (Suc t_fourtimes_len + length (t_wcode_main_first_part @ 
  tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2,
  Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
proof(rule_tac t_tshift_lemma, auto)
  assume "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp =
    (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  thus "0 < stp"
    using t_fourtimes_len_gr
    apply(case_tac stp, simp_all add: steps.simps)
    done
next
  show "Suc 0 \<le> length t_fourtimes div 2"
    apply(simp add: t_fourtimes_def shift_length tMp.simps)
    done
next
  show "t_ncorrect (t_wcode_main_first_part @ 
    abacus.tshift t_twice (length t_wcode_main_first_part div 2) @ 
    [(L, Suc 0), (L, Suc 0)])"
    apply(simp add: t_ncorrect.simps t_wcode_main_first_part_def shift_length
                    t_twice_def)
    using tm_even[of abc_twice]
    by arith
next
  show "t_ncorrect t_fourtimes"
    apply(simp add: t_fourtimes_def steps.simps t_ncorrect.simps)
    using tm_even[of abc_fourtimes]
    by arith
next
  show "t_ncorrect [(L, Suc 0), (L, Suc 0)]"
    apply(simp add: t_ncorrect.simps)
    done
qed

lemma [simp]: "length t_wcode_main_first_part = 24"
apply(simp add: t_wcode_main_first_part_def)
done

lemma [simp]: "(26 + length t_twice) div 2 = (length t_twice) div 2 + 13"
using tm_even[of abc_twice]
apply(simp add: t_twice_def)
done

lemma [simp]: "((26 + length (abacus.tshift t_twice 12)) div 2)
             = (length (abacus.tshift t_twice 12) div 2 + 13)"
using tm_even[of abc_twice]
apply(simp add: t_twice_def)
done 

lemma [simp]: "t_twice_len + 14 =  14 + length (abacus.tshift t_twice 12) div 2"
using tm_even[of abc_twice]
apply(simp add: t_twice_def t_twice_len_def shift_length)
done

lemma t_fourtimes_append:
  "\<exists> stp ln rn. 
  steps (Suc 0 + length (t_wcode_main_first_part @ tshift t_twice
  (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, 
  Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
  ((t_wcode_main_first_part @ tshift t_twice (length t_wcode_main_first_part div 2) @
  [(L, 1), (L, 1)]) @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)]) stp 
  = (Suc t_fourtimes_len + length (t_wcode_main_first_part @ tshift t_twice
  (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires,
                                                                 Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  using t_fourtimes_change_term_state[of ires rs n]
  apply(erule_tac exE)
  apply(erule_tac exE)
  apply(erule_tac exE)
  apply(drule_tac t_fourtimes_append_pre)
  apply(erule_tac exE)
  apply(rule_tac x = stpa in exI, rule_tac x = ln in exI, rule_tac x = rn in exI)
  apply(simp add: t_twice_len_def shift_length)
  done

lemma t_wcode_main_len: "length t_wcode_main = length t_twice + length t_fourtimes + 28"
apply(simp add: t_wcode_main_def shift_length)
done
 
lemma [simp]: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) b
             = (L, Suc 0)"
using tm_even[of "abc_twice"] tm_even[of "abc_fourtimes"]
apply(case_tac b)
apply(simp_all only: fetch.simps)
apply(auto simp: nth_of.simps t_wcode_main_len t_twice_len_def
                 t_fourtimes_def t_twice_def t_fourtimes_def t_fourtimes_len_def)
apply(auto simp: t_wcode_main_def t_wcode_main_first_part_def shift_length t_twice_def nth_append 
                    t_fourtimes_def)
done

lemma wcode_jump2: 
  "\<exists> stp ln rn. steps (t_twice_len + 14 + t_fourtimes_len
  , Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, Oc\<^bsup>Suc (4 * rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>) t_wcode_main stp =
  (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4 * rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
apply(rule_tac x = "Suc 0" in exI)
apply(simp add: steps.simps shift_length)
apply(rule_tac x = lnb in exI, rule_tac x = rnb in exI)
apply(simp add: tstep.simps new_tape.simps)
done

lemma wcode_fourtimes_case:
  shows "\<exists>stp ln rn.
  steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
proof -
  have "\<exists>stp ln rn.
  steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  (t_twice_len + 14, Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
    using wcode_fourtimes_case_first_correctness[of ires rs m n]
    apply(simp add: wcode_fourtimes_case_inv.simps, auto)
    apply(rule_tac x = na in exI, rule_tac x = ln in exI,
          rule_tac x = rn in exI)
    apply(simp)
    done
  from this obtain stpa lna rna where stp1:
    "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa =
  (t_twice_len + 14, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast
  have "\<exists>stp ln rn. steps (t_twice_len + 14, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rna\<^esup>)
                     t_wcode_main stp =
          (t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires,  Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
    using t_fourtimes_append[of " Bk\<^bsup>lna\<^esup> @ Oc # ires" "rs + 1" rna]
    apply(erule_tac exE)
    apply(erule_tac exE)
    apply(erule_tac exE)
    apply(simp add: t_wcode_main_def)
    apply(rule_tac x = stp in exI, 
          rule_tac x = "ln + lna" in exI,
          rule_tac x = rn in exI, simp)
    apply(simp add: exp_ind_def[THEN sym] exp_add[THEN sym])
    done
  from this obtain stpb lnb rnb where stp2:
    "steps (t_twice_len + 14, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rna\<^esup>)
                     t_wcode_main stpb =
       (t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires,  Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>)"
    by blast
  have "\<exists>stp ln rn. steps (t_twice_len + 14 + t_fourtimes_len,
    Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires,  Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>)
    t_wcode_main stp =
    (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
    apply(rule wcode_jump2)
    done
  from this obtain stpc lnc rnc where stp3: 
    "steps (t_twice_len + 14 + t_fourtimes_len,
    Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires,  Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>)
    t_wcode_main stpc =
    (Suc 0, Bk # Bk\<^bsup>lnc\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnc\<^esup>)"
    by blast
  from stp1 stp2 stp3 show "?thesis"
    apply(rule_tac x = "stpa + stpb + stpc" in exI,
          rule_tac x = lnc in exI, rule_tac x = rnc in exI)
    apply(simp add: steps_add)
    done
qed

(**********************************************************)

fun wcode_on_left_moving_3_B :: "bin_inv_t"
  where
  "wcode_on_left_moving_3_B ires rs (l, r) = 
       (\<exists> ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # Bk # Bk # ires \<and>
                    r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
                    ml + mr > Suc 0 \<and> mr > 0 )"

fun wcode_on_left_moving_3_O :: "bin_inv_t"
  where
  "wcode_on_left_moving_3_O ires rs (l, r) = 
         (\<exists> ln rn. l = Bk # Bk # ires \<and>
                   r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"

fun wcode_on_left_moving_3 :: "bin_inv_t"
  where
  "wcode_on_left_moving_3 ires rs (l, r) = 
       (wcode_on_left_moving_3_B ires rs (l, r) \<or>  
        wcode_on_left_moving_3_O ires rs (l, r))"

fun wcode_on_checking_3 :: "bin_inv_t"
  where
  "wcode_on_checking_3 ires rs (l, r) = 
         (\<exists> ln rn. l = Bk # ires \<and>
             r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"

fun wcode_goon_checking_3 :: "bin_inv_t"
  where
  "wcode_goon_checking_3 ires rs (l, r) = 
         (\<exists> ln rn. l = ires \<and>
             r = Bk # Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"

fun wcode_stop :: "bin_inv_t"
  where
  "wcode_stop ires rs (l, r) = 
          (\<exists> ln rn. l = Bk # ires \<and>
             r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"

fun wcode_halt_case_inv :: "nat \<Rightarrow> bin_inv_t"
  where
  "wcode_halt_case_inv st ires rs (l, r) = 
          (if st = 0 then wcode_stop ires rs (l, r)
           else if st = Suc 0 then wcode_on_left_moving_3 ires rs (l, r)
           else if st = Suc (Suc 0) then wcode_on_checking_3 ires rs (l, r)
           else if st = 7 then wcode_goon_checking_3 ires rs (l, r)
           else False)"

fun wcode_halt_case_state :: "t_conf \<Rightarrow> nat"
  where
  "wcode_halt_case_state (st, l, r) = 
           (if st = 1 then 5
            else if st = Suc (Suc 0) then 4
            else if st = 7 then 3
            else 0)"

fun wcode_halt_case_step :: "t_conf \<Rightarrow> nat"
  where
  "wcode_halt_case_step (st, l, r) = 
         (if st = 1 then length l
         else 0)"

fun wcode_halt_case_measure :: "t_conf \<Rightarrow> nat \<times> nat"
  where
  "wcode_halt_case_measure (st, l, r) = 
     (wcode_halt_case_state (st, l, r), 
      wcode_halt_case_step (st, l, r))"

definition wcode_halt_case_le :: "(t_conf \<times> t_conf) set"
  where "wcode_halt_case_le \<equiv> (inv_image lex_pair wcode_halt_case_measure)"

lemma wf_wcode_halt_case_le[intro]: "wf wcode_halt_case_le"
by(auto intro:wf_inv_image simp: wcode_halt_case_le_def)

declare wcode_on_left_moving_3_B.simps[simp del] wcode_on_left_moving_3_O.simps[simp del]  
        wcode_on_checking_3.simps[simp del] wcode_goon_checking_3.simps[simp del] 
        wcode_on_left_moving_3.simps[simp del] wcode_stop.simps[simp del]

lemmas wcode_halt_invs = 
  wcode_on_left_moving_3_B.simps wcode_on_left_moving_3_O.simps
  wcode_on_checking_3.simps wcode_goon_checking_3.simps 
  wcode_on_left_moving_3.simps wcode_stop.simps

lemma [simp]: "fetch t_wcode_main 7 Bk = (R, 0)"
apply(simp add: fetch.simps t_wcode_main_def nth_append nth_of.simps
                t_wcode_main_first_part_def)
done

lemma [simp]: "wcode_on_left_moving_3 ires rs (b, [])  = False"
apply(simp only: wcode_halt_invs)
apply(simp add: exp_ind_def)
done    

lemma [simp]: "wcode_on_checking_3 ires rs (b, []) = False"
apply(simp add: wcode_halt_invs)
done
              
lemma [simp]: "wcode_goon_checking_3 ires rs (b, []) = False"
apply(simp add: wcode_halt_invs)
done 

lemma [simp]: "wcode_on_left_moving_3 ires rs (b, Bk # list)
 \<Longrightarrow> wcode_on_left_moving_3 ires rs (tl b, hd b # Bk # list)"
apply(simp only: wcode_halt_invs)
apply(erule_tac disjE)
apply(erule_tac exE)+
apply(case_tac ml, simp)
apply(rule_tac x = "mr - 2" in exI, rule_tac x = rn in exI)
apply(case_tac mr, simp, simp add: exp_ind, simp add: exp_ind[THEN sym])
apply(rule_tac disjI1)
apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, 
      rule_tac x = rn in exI, simp add: exp_ind_def)
apply(simp)
done

lemma [simp]: "wcode_goon_checking_3 ires rs (b, Bk # list) \<Longrightarrow> 
  (b = [] \<longrightarrow> wcode_stop ires rs ([Bk], list)) \<and>
  (b \<noteq> [] \<longrightarrow> wcode_stop ires rs (Bk # b, list))"
apply(auto simp: wcode_halt_invs)
done

lemma [simp]: "wcode_on_left_moving_3 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
apply(auto simp: wcode_halt_invs)
done

lemma [simp]: "wcode_on_left_moving_3 ires rs (b, Oc # list) \<Longrightarrow> 
               wcode_on_checking_3 ires rs (tl b, hd b # Oc # list)"
apply(simp add:wcode_halt_invs, auto)
apply(case_tac [!] mr, simp_all add: exp_ind_def)
done     

lemma [simp]: "wcode_on_checking_3 ires rs (b, Oc # list) = False"
apply(auto simp: wcode_halt_invs)
done

lemma [simp]: "wcode_on_left_moving_3 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
apply(simp add: wcode_halt_invs, auto)
done


lemma [simp]: "wcode_on_checking_3 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
apply(auto simp: wcode_halt_invs)
done

lemma [simp]: "wcode_on_checking_3 ires rs (b, Bk # list) \<Longrightarrow> 
  wcode_goon_checking_3 ires rs (tl b, hd b # Bk # list)"
apply(auto simp: wcode_halt_invs)
done

lemma [simp]: "wcode_goon_checking_3 ires rs (b, Oc # list) = False"
apply(simp add: wcode_goon_checking_3.simps)
done

lemma t_halt_case_correctness: 
shows "let P = (\<lambda> (st, l, r). st = 0) in 
       let Q = (\<lambda> (st, l, r). wcode_halt_case_inv st ires rs (l, r)) in 
       let f = (\<lambda> stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp) in
       \<exists> n .P (f n) \<and> Q (f (n::nat))"
proof -
  let ?P = "(\<lambda> (st, l, r). st = 0)"
  let ?Q = "(\<lambda> (st, l, r). wcode_halt_case_inv st ires rs (l, r))"
  let ?f = "(\<lambda> stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp)"
  have "\<exists> n. ?P (?f n) \<and> ?Q (?f (n::nat))"
  proof(rule_tac halt_lemma2)
    show "wf wcode_halt_case_le" by auto
  next
    show "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow> 
                    ?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_halt_case_le"
      apply(rule_tac allI, rule_tac impI, case_tac "?f na")
      apply(simp add: tstep_red tstep.simps)
      apply(case_tac c, simp, case_tac [2] aa)
      apply(simp_all split: if_splits add: new_tape.simps wcode_halt_case_le_def lex_pair_def)
      done      
  next 
    show "?Q (?f 0)"
      apply(simp add: steps.simps wcode_halt_invs)
      apply(rule_tac x = "Suc m" in exI, simp add: exp_ind_def)
      apply(rule_tac x = "Suc 0" in exI, auto)
      done
  next
    show "\<not> ?P (?f 0)"
      apply(simp add: steps.simps)
      done
  qed
  thus "?thesis"
    apply(auto)
    done
qed

declare wcode_halt_case_inv.simps[simp del]
lemma [intro]: "\<exists> xs. (<rev list @ [aa::nat]> :: block list) = Oc # xs"
apply(case_tac "rev list", simp)
apply(simp add: tape_of_nat_abv tape_of_nat_list.simps exp_ind_def)
apply(case_tac list, simp, simp)
done

lemma wcode_halt_case:
  "\<exists>stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
  t_wcode_main stp  = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  using t_halt_case_correctness[of ires rs m n]
apply(simp)
apply(erule_tac exE)
apply(case_tac "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires,
                Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main na")
apply(auto simp: wcode_halt_case_inv.simps wcode_stop.simps)
apply(rule_tac x = na in exI, rule_tac x = ln in exI, 
      rule_tac x = rn in exI, simp)
done

lemma bl_bin_one: "bl_bin [Oc] =  Suc 0"
apply(simp add: bl_bin.simps)
done

lemma t_wcode_main_lemma_pre:
  "\<lbrakk>args \<noteq> []; lm = <args::nat list>\<rbrakk> \<Longrightarrow> 
       \<exists> stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev lm @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main
                    stp
      = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin lm + rs * 2^(length lm - 1) \<^esup> @ Bk\<^bsup>rn\<^esup>)"
proof(induct "length args" arbitrary: args lm rs m n, simp)
  fix x args lm rs m n
  assume ind:
    "\<And>args lm rs m n.
    \<lbrakk>x = length args; (args::nat list) \<noteq> []; lm = <args>\<rbrakk>
    \<Longrightarrow> \<exists>stp ln rn.
    steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev lm @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
    (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin lm + rs * 2 ^ (length lm - 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  
    and h: "Suc x = length args" "(args::nat list) \<noteq> []" "lm = <args>"
  from h have "\<exists> (a::nat) xs. args = xs @ [a]"
    apply(rule_tac x = "last args" in exI)
    apply(rule_tac x = "butlast args" in exI, auto)
    done    
  from this obtain a xs where "args = xs @ [a]" by blast
  from h and this show
    "\<exists>stp ln rn.
    steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev lm @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
    (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin lm + rs * 2 ^ (length lm - 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  proof(case_tac "xs::nat list", simp)
    show "\<exists>stp ln rn.
      steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
      (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<a>) + rs * 2 ^ a\<^esup> @ Bk\<^bsup>rn\<^esup>)"
    proof(induct "a" arbitrary: m n rs ires, simp)
      fix m n rs ires
      show "\<exists>stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
        t_wcode_main stp  = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin [Oc] + rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
        apply(simp add: bl_bin_one)
        apply(rule_tac wcode_halt_case)
        done
    next
      fix a m n rs ires
      assume ind2: 
        "\<And>m n rs ires.
        \<exists>stp ln rn.
        steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
        (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<a>) + rs * 2 ^ a\<^esup> @ Bk\<^bsup>rn\<^esup>)"
      show "\<exists>stp ln rn.
        steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev (<Suc a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
        (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<Suc a>) + rs * 2 ^ Suc a\<^esup> @ Bk\<^bsup>rn\<^esup>)"
      proof -
        have "\<exists>stp ln rn.
          steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev (<Suc a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
          (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
          apply(simp add: tape_of_nat)
          using wcode_double_case[of m "Oc\<^bsup>a\<^esup> @ Bk # Bk # ires" rs n]
          apply(simp add: exp_ind_def)
          done
        from this obtain stpa lna rna where stp1:  
          "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev (<Suc a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa =
          (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast
        moreover have 
          "\<exists>stp ln rn.
          steps (Suc 0,  Bk # Bk\<^bsup>lna\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp =
          (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<a>) + (2*rs + 2)  * 2 ^ a\<^esup> @ Bk\<^bsup>rn\<^esup>)"
          using ind2[of lna ires "2*rs + 2" rna] by simp   
        from this obtain stpb lnb rnb where stp2:  
          "steps (Suc 0,  Bk # Bk\<^bsup>lna\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb =
          (0, Bk # ires, Bk # Oc # Bk\<^bsup>lnb\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<a>) + (2*rs + 2)  * 2 ^ a\<^esup> @ Bk\<^bsup>rnb\<^esup>)"
          by blast
        from stp1 and stp2 show "?thesis"
          apply(rule_tac x = "stpa + stpb" in exI,
            rule_tac x = lnb in exI, rule_tac x = rnb in exI, simp)
          apply(simp add: steps_add bl_bin_nat_Suc exponent_def)
          done
      qed
    qed
  next
    fix aa list
    assume g: "Suc x = length args" "args \<noteq> []" "lm = <args>" "args = xs @ [a::nat]" "xs = (aa::nat) # list"
    thus "\<exists>stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev lm @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
      (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin lm + rs * 2 ^ (length lm - 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
    proof(induct a arbitrary: m n rs args lm, simp_all add: tape_of_nl_rev, 
        simp only: tape_of_nl_cons_app1, simp)
      fix m n rs args lm
      have "\<exists>stp ln rn.
        steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # rev (<(aa::nat) # list>) @ Bk # Bk # ires,
        Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
        (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ rev (<aa # list>) @ Bk # Bk # ires, 
        Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
        proof(simp add: tape_of_nl_rev)
          have "\<exists> xs. (<rev list @ [aa]>) = Oc # xs" by auto           
          from this obtain xs where "(<rev list @ [aa]>) = Oc # xs" ..
          thus "\<exists>stp ln rn.
            steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # <rev list @ [aa]> @ Bk # Bk # ires,
            Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
            (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ <rev list @ [aa]> @ Bk # Bk # ires, Bk # Oc\<^bsup>5 + 4 * rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
            apply(simp)
            using wcode_fourtimes_case[of m "xs @ Bk # Bk # ires" rs n]
            apply(simp)
            done
        qed
      from this obtain stpa lna rna where stp1:
        "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # rev (<aa # list>) @ Bk # Bk # ires,
        Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa =
        (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<aa # list>) @ Bk # Bk # ires, 
        Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast
      from g have 
        "\<exists> stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, 
        Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp = (0, Bk # ires, 
        Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<aa#list>)+ (4*rs + 4) * 2^(length (<aa#list>) - 1) \<^esup> @ Bk\<^bsup>rn\<^esup>)"
         apply(rule_tac args = "(aa::nat)#list" in ind, simp_all)
         done
       from this obtain stpb lnb rnb where stp2:
         "steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, 
         Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb = (0, Bk # ires, 
         Bk # Oc # Bk\<^bsup>lnb\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<aa#list>)+ (4*rs + 4) * 2^(length (<aa#list>) - 1) \<^esup> @ Bk\<^bsup>rnb\<^esup>)"
         by blast
       from stp1 and stp2 and h
       show "\<exists>stp ln rn.
         steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # <rev list @ [aa]> @ Bk # Bk # ires,
         Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
         (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk #
         Bk # Oc\<^bsup>bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # <list @ [0]>) + rs * (2 * 2 ^ (aa + length (<list @ [0]>)))\<^esup> @ Bk\<^bsup>rn\<^esup>)"
         apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI,
           rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_rev)
         done
     next
       fix ab m n rs args lm
       assume ind2:
         "\<And> m n rs args lm.
         \<lbrakk>lm = <aa # list @ [ab]>; args = aa # list @ [ab]\<rbrakk>
         \<Longrightarrow> \<exists>stp ln rn.
         steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ <ab # rev list @ [aa]> @ Bk # Bk # ires,
         Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
         (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk #
         Bk # Oc\<^bsup>bl_bin (<aa # list @ [ab]>) + rs * 2 ^ (length (<aa # list @ [ab]>) - Suc 0)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
         and k: "args = aa # list @ [Suc ab]" "lm = <aa # list @ [Suc ab]>"
       show "\<exists>stp ln rn.
         steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ <Suc ab # rev list @ [aa]> @ Bk # Bk # ires,
         Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
         (0, Bk # ires,Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # 
         Bk # Oc\<^bsup>bl_bin (<aa # list @ [Suc ab]>) + rs * 2 ^ (length (<aa # list @ [Suc ab]>) - Suc 0)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
       proof(simp add: tape_of_nl_cons_app1)
         have "\<exists>stp ln rn.
           steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc\<^bsup>Suc (Suc ab)\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, 
           Bk # Oc # Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp
           = (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc\<^bsup>Suc ab\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
           Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
           using wcode_double_case[of m "Oc\<^bsup>ab\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires"
                                      rs n]
           apply(simp add: exp_ind_def)
           done
         from this obtain stpa lna rna where stp1:
           "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc\<^bsup>Suc (Suc ab)\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, 
           Bk # Oc # Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa
           = (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ Oc\<^bsup>Suc ab\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
           Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast
         from k have 
           "\<exists> stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ <ab # rev list @ [aa]> @ Bk # Bk # ires,
           Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp
           = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk #
           Bk # Oc\<^bsup>bl_bin (<aa # list @ [ab]> ) +  (2*rs + 2)* 2^(length (<aa # list @ [ab]>) - Suc 0)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
           apply(rule_tac ind2, simp_all)
           done
         from this obtain stpb lnb rnb where stp2: 
           "steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @  <ab # rev list @ [aa]> @ Bk # Bk # ires,
           Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb
           = (0, Bk # ires, Bk # Oc # Bk\<^bsup>lnb\<^esup> @ Bk #
           Bk # Oc\<^bsup>bl_bin (<aa # list @ [ab]> ) +  (2*rs + 2)* 2^(length (<aa # list @ [ab]>) - Suc 0)\<^esup> @ Bk\<^bsup>rnb\<^esup>)" 
           by blast
         from stp1 and stp2 show 
           "\<exists>stp ln rn.
           steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc\<^bsup>Suc (Suc ab)\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
           Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
           (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # 
           Oc\<^bsup>bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # <list @ [Suc ab]>) + rs * (2 * 2 ^ (aa + length (<list @ [Suc ab]>)))\<^esup> 
           @ Bk\<^bsup>rn\<^esup>)"
           apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI,
             rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_cons_app1 exp_ind_def)
           done
       qed
     qed
   qed
 qed


         
(* turing_shift can be used*)
term t_wcode_main
definition t_wcode_prepare :: "tprog"
  where
  "t_wcode_prepare \<equiv> 
         [(W1, 2), (L, 1), (L, 3), (R, 2), (R, 4), (W0, 3),
          (R, 4), (R, 5), (R, 6), (R, 5), (R, 7), (R, 5),
          (W1, 7), (L, 0)]"

fun wprepare_add_one :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wprepare_add_one m lm (l, r) = 
      (\<exists> rn. l = [] \<and>
               (r = <m # lm> @ Bk\<^bsup>rn\<^esup> \<or> 
                r = Bk # <m # lm> @ Bk\<^bsup>rn\<^esup>))"

fun wprepare_goto_first_end :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wprepare_goto_first_end m lm (l, r) = 
      (\<exists> ml mr rn. l = Oc\<^bsup>ml\<^esup> \<and>
                      r = Oc\<^bsup>mr\<^esup> @ Bk # <lm> @ Bk\<^bsup>rn\<^esup> \<and>
                      ml + mr = Suc (Suc m))"

fun wprepare_erase :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow>  bool"
  where
  "wprepare_erase m lm (l, r) = 
     (\<exists> rn. l = Oc\<^bsup>Suc m\<^esup> \<and> 
               tl r = Bk # <lm> @ Bk\<^bsup>rn\<^esup>)"

fun wprepare_goto_start_pos_B :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wprepare_goto_start_pos_B m lm (l, r) = 
     (\<exists> rn. l = Bk # Oc\<^bsup>Suc m\<^esup> \<and>
               r = Bk # <lm> @ Bk\<^bsup>rn\<^esup>)"

fun wprepare_goto_start_pos_O :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wprepare_goto_start_pos_O m lm (l, r) = 
     (\<exists> rn. l = Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and>
               r = <lm> @ Bk\<^bsup>rn\<^esup>)"

fun wprepare_goto_start_pos :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wprepare_goto_start_pos m lm (l, r) = 
       (wprepare_goto_start_pos_B m lm (l, r) \<or>
        wprepare_goto_start_pos_O m lm (l, r))"

fun wprepare_loop_start_on_rightmost :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wprepare_loop_start_on_rightmost m lm (l, r) = 
     (\<exists> rn mr. rev l @ r = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm> @ Bk\<^bsup>rn\<^esup> \<and> l \<noteq> [] \<and>
                       r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup>)"

fun wprepare_loop_start_in_middle :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wprepare_loop_start_in_middle m lm (l, r) =
     (\<exists> rn (mr:: nat) (lm1::nat list). 
  rev l @ r = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm> @ Bk\<^bsup>rn\<^esup> \<and> l \<noteq> [] \<and>
  r = Oc\<^bsup>mr\<^esup> @ Bk # <lm1> @ Bk\<^bsup>rn\<^esup> \<and> lm1 \<noteq> [])"

fun wprepare_loop_start :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wprepare_loop_start m lm (l, r) = (wprepare_loop_start_on_rightmost m lm (l, r) \<or> 
                                      wprepare_loop_start_in_middle m lm (l, r))"

fun wprepare_loop_goon_on_rightmost :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wprepare_loop_goon_on_rightmost m lm (l, r) = 
     (\<exists> rn. l = Bk # <rev lm> @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and>
               r = Bk\<^bsup>rn\<^esup>)"

fun wprepare_loop_goon_in_middle :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wprepare_loop_goon_in_middle m lm (l, r) = 
     (\<exists> rn (mr:: nat) (lm1::nat list). 
  rev l @ r = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm> @ Bk\<^bsup>rn\<^esup> \<and> l \<noteq> [] \<and>
                     (if lm1 = [] then r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> 
                     else r = Oc\<^bsup>mr\<^esup> @ Bk # <lm1> @ Bk\<^bsup>rn\<^esup>) \<and> mr > 0)"

fun wprepare_loop_goon :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wprepare_loop_goon m lm (l, r) = 
              (wprepare_loop_goon_in_middle m lm (l, r) \<or> 
               wprepare_loop_goon_on_rightmost m lm (l, r))"

fun wprepare_add_one2 :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wprepare_add_one2 m lm (l, r) =
          (\<exists> rn. l = Bk # Bk # <rev lm> @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and>
               (r = [] \<or> tl r = Bk\<^bsup>rn\<^esup>))"

fun wprepare_stop :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wprepare_stop m lm (l, r) = 
         (\<exists> rn. l = Bk # <rev lm> @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and>
               r = Bk # Oc # Bk\<^bsup>rn\<^esup>)"

fun wprepare_inv :: "nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wprepare_inv st m lm (l, r) = 
        (if st = 0 then wprepare_stop m lm (l, r) 
         else if st = Suc 0 then wprepare_add_one m lm (l, r)
         else if st = Suc (Suc 0) then wprepare_goto_first_end m lm (l, r)
         else if st = Suc (Suc (Suc 0)) then wprepare_erase m lm (l, r)
         else if st = 4 then wprepare_goto_start_pos m lm (l, r)
         else if st = 5 then wprepare_loop_start m lm (l, r)
         else if st = 6 then wprepare_loop_goon m lm (l, r)
         else if st = 7 then wprepare_add_one2 m lm (l, r)
         else False)"

fun wprepare_stage :: "t_conf \<Rightarrow> nat"
  where
  "wprepare_stage (st, l, r) = 
      (if st \<ge> 1 \<and> st \<le> 4 then 3
       else if st = 5 \<or> st = 6 then 2
       else 1)"

fun wprepare_state :: "t_conf \<Rightarrow> nat"
  where
  "wprepare_state (st, l, r) = 
       (if st = 1 then 4
        else if st = Suc (Suc 0) then 3
        else if st = Suc (Suc (Suc 0)) then 2
        else if st = 4 then 1
        else if st = 7 then 2
        else 0)"

fun wprepare_step :: "t_conf \<Rightarrow> nat"
  where
  "wprepare_step (st, l, r) = 
      (if st = 1 then (if hd r = Oc then Suc (length l)
                       else 0)
       else if st = Suc (Suc 0) then length r
       else if st = Suc (Suc (Suc 0)) then (if hd r = Oc then 1
                            else 0)
       else if st = 4 then length r
       else if st = 5 then Suc (length r)
       else if st = 6 then (if r = [] then 0 else Suc (length r))
       else if st = 7 then (if (r \<noteq> [] \<and> hd r = Oc) then 0
                            else 1)
       else 0)"

fun wcode_prepare_measure :: "t_conf \<Rightarrow> nat \<times> nat \<times> nat"
  where
  "wcode_prepare_measure (st, l, r) = 
     (wprepare_stage (st, l, r), 
      wprepare_state (st, l, r), 
      wprepare_step (st, l, r))"

definition wcode_prepare_le :: "(t_conf \<times> t_conf) set"
  where "wcode_prepare_le \<equiv> (inv_image lex_triple wcode_prepare_measure)"

lemma [intro]: "wf lex_pair"
by(auto intro:wf_lex_prod simp:lex_pair_def)

lemma wf_wcode_prepare_le[intro]: "wf wcode_prepare_le"
by(auto intro:wf_inv_image simp: wcode_prepare_le_def 
           recursive.lex_triple_def)

declare wprepare_add_one.simps[simp del] wprepare_goto_first_end.simps[simp del]
        wprepare_erase.simps[simp del] wprepare_goto_start_pos.simps[simp del]
        wprepare_loop_start.simps[simp del] wprepare_loop_goon.simps[simp del]
        wprepare_add_one2.simps[simp del]

lemmas wprepare_invs = wprepare_add_one.simps wprepare_goto_first_end.simps
        wprepare_erase.simps wprepare_goto_start_pos.simps
        wprepare_loop_start.simps wprepare_loop_goon.simps
        wprepare_add_one2.simps

declare wprepare_inv.simps[simp del]
lemma [simp]: "fetch t_wcode_prepare (Suc 0) Bk = (W1, 2)"
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_prepare (Suc 0) Oc = (L, 1)"
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_prepare (Suc (Suc 0)) Bk = (L, 3)"
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_prepare (Suc (Suc 0)) Oc = (R, 2)"
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_prepare (Suc (Suc (Suc 0))) Bk = (R, 4)"
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_prepare (Suc (Suc (Suc 0))) Oc = (W0, 3)"
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_prepare 4 Bk = (R, 4)"
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_prepare 4 Oc = (R, 5)"
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_prepare 5 Oc = (R, 5)"
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_prepare 5 Bk = (R, 6)"
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_prepare 6 Oc = (R, 5)"
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_prepare 6 Bk = (R, 7)"
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_prepare 7 Oc = (L, 0)"
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_prepare 7 Bk = (W1, 7)"
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
done

lemma tape_of_nl_not_null: "lm \<noteq> [] \<Longrightarrow> <lm::nat list> \<noteq> []"
apply(case_tac lm, auto)
apply(case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
done

lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_add_one m lm (b, []) = False"
apply(simp add: wprepare_invs)
apply(simp add: tape_of_nl_not_null)
done

lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_goto_first_end m lm (b, []) = False"
apply(simp add: wprepare_invs)
done

lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_erase m lm (b, []) = False"
apply(simp add: wprepare_invs)
done



lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_goto_start_pos m lm (b, []) = False"
apply(simp add: wprepare_invs tape_of_nl_not_null)
done

lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [])\<rbrakk> \<Longrightarrow> b \<noteq> []"
apply(simp add: wprepare_invs tape_of_nl_not_null, auto)
done

lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [])\<rbrakk> \<Longrightarrow> 
                                  wprepare_loop_goon m lm (Bk # b, [])"
apply(simp only: wprepare_invs tape_of_nl_not_null)
apply(erule_tac disjE)
apply(rule_tac disjI2)
apply(simp add: wprepare_loop_start_on_rightmost.simps
                wprepare_loop_goon_on_rightmost.simps, auto)
apply(rule_tac rev_eq, simp add: tape_of_nl_rev)
done

lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, [])\<rbrakk> \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs tape_of_nl_not_null, auto)
done

lemma [simp]:"\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, [])\<rbrakk> \<Longrightarrow> 
  wprepare_add_one2 m lm (Bk # b, [])"
apply(simp only: wprepare_invs tape_of_nl_not_null, auto split: if_splits)
apply(case_tac mr, simp, simp add: exp_ind_def)
done

lemma [simp]: "wprepare_add_one2 m lm (b, []) \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs tape_of_nl_not_null, auto)
done

lemma [simp]: "wprepare_add_one2 m lm (b, []) \<Longrightarrow> wprepare_add_one2 m lm (b, [Oc])"
apply(simp only: wprepare_invs tape_of_nl_not_null, auto)
done

lemma [simp]: "Bk # list = <(m::nat) # lm> @ ys = False"
apply(case_tac lm, auto simp: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
done

lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_add_one m lm (b, Bk # list)\<rbrakk>
       \<Longrightarrow> (b = [] \<longrightarrow> wprepare_goto_first_end m lm ([], Oc # list)) \<and> 
           (b \<noteq> [] \<longrightarrow> wprepare_goto_first_end m lm (b, Oc # list))"
apply(simp only: wprepare_invs, auto)
apply(rule_tac x = 0 in exI, simp add: exp_ind_def)
apply(case_tac lm, simp, simp add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
apply(rule_tac x = rn in exI, simp)
done

lemma [simp]: "wprepare_goto_first_end m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs tape_of_nl_not_null, auto)
apply(case_tac mr, simp_all add: exp_ind_def)
done

lemma [simp]: "wprepare_goto_first_end m lm (b, Bk # list) \<Longrightarrow>
                          wprepare_erase m lm (tl b, hd b # Bk # list)"
apply(simp only: wprepare_invs tape_of_nl_not_null, auto)
apply(case_tac mr, simp_all add: exp_ind_def)
apply(case_tac mr, auto simp: exp_ind_def)
done

lemma [simp]: "wprepare_erase m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs exp_ind_def, auto)
done

lemma [simp]: "wprepare_erase m lm (b, Bk # list) \<Longrightarrow> 
                           wprepare_goto_start_pos m lm (Bk # b, list)"
apply(simp only: wprepare_invs, auto)
done

lemma [simp]: "\<lbrakk>wprepare_add_one m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
apply(simp only: wprepare_invs)
apply(case_tac lm, simp_all add: tape_of_nl_abv 
                         tape_of_nat_list.simps exp_ind_def, auto)
done
    
lemma [simp]: "\<lbrakk>lm \<noteq> [];  wprepare_goto_first_end m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
apply(simp only: wprepare_invs, auto)
apply(case_tac mr, simp_all add: exp_ind_def)
apply(simp add: tape_of_nl_not_null)
done
     
lemma [simp]: "\<lbrakk>lm \<noteq> [];  wprepare_goto_first_end m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs, auto)
apply(case_tac mr, simp_all add: exp_ind_def tape_of_nl_not_null)
done

lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_erase m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
apply(simp only: wprepare_invs, auto)
done

lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_erase m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs, auto simp: exp_ind_def)
done

lemma [simp]: "\<lbrakk>lm \<noteq> [];  wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
apply(simp only: wprepare_invs, auto)
apply(simp add: tape_of_nl_not_null)
apply(case_tac lm, simp, case_tac list)
apply(simp_all add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
done

lemma [simp]: "\<lbrakk>lm \<noteq> [];  wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs)
apply(auto)
done

lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs, auto)
done

lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> 
  (list = [] \<longrightarrow> wprepare_add_one2 m lm (Bk # b, [])) \<and> 
  (list \<noteq> [] \<longrightarrow> wprepare_add_one2 m lm (Bk # b, list))"
apply(simp only: wprepare_invs, simp)
apply(case_tac list, simp_all split: if_splits, auto)
apply(case_tac [1-3] mr, simp_all add: exp_ind_def)
apply(case_tac mr, simp_all add: exp_ind_def tape_of_nl_not_null)
apply(case_tac [1-2] mr, simp_all add: exp_ind_def)
apply(case_tac rn, simp, case_tac nat, auto simp: exp_ind_def)
done

lemma [simp]: "wprepare_add_one2 m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs, simp)
done

lemma [simp]: "wprepare_add_one2 m lm (b, Bk # list) \<Longrightarrow> 
      (list = [] \<longrightarrow> wprepare_add_one2 m lm (b, [Oc])) \<and> 
      (list \<noteq> [] \<longrightarrow> wprepare_add_one2 m lm (b, Oc # list))"
apply(simp only:  wprepare_invs, auto)
done

lemma [simp]: "wprepare_goto_first_end m lm (b, Oc # list)
       \<Longrightarrow> (b = [] \<longrightarrow> wprepare_goto_first_end m lm ([Oc], list)) \<and> 
           (b \<noteq> [] \<longrightarrow> wprepare_goto_first_end m lm (Oc # b, list))"
apply(simp only:  wprepare_invs, auto)
apply(rule_tac x = 1 in exI, auto)
apply(case_tac mr, simp_all add: exp_ind_def)
apply(case_tac ml, simp_all add: exp_ind_def)
apply(rule_tac x = rn in exI, simp)
apply(rule_tac x = "Suc ml" in exI, simp_all add: exp_ind_def)
apply(rule_tac x = "mr - 1" in exI, simp)
apply(case_tac mr, simp_all add: exp_ind_def, auto)
done

lemma [simp]: "wprepare_erase m lm (b, Oc # list) \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs, auto simp: exp_ind_def)
done

lemma [simp]: "wprepare_erase m lm (b, Oc # list)
  \<Longrightarrow> wprepare_erase m lm (b, Bk # list)"
apply(simp  only:wprepare_invs, auto simp: exp_ind_def)
done

lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk>
       \<Longrightarrow> wprepare_goto_start_pos m lm (Bk # b, list)"
apply(simp only:wprepare_invs, auto)
apply(case_tac [!] lm, simp, simp_all)
done

lemma [simp]: "wprepare_loop_start m lm (b, aa) \<Longrightarrow> b \<noteq> []"
apply(simp only:wprepare_invs, auto)
done
lemma [elim]: "Bk # list = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup>  \<Longrightarrow> \<exists>rn. list = Bk\<^bsup>rn\<^esup>"
apply(case_tac mr, simp_all)
apply(case_tac rn, simp_all add: exp_ind_def, auto)
done

lemma rev_equal_iff: "x = y \<Longrightarrow> rev x = rev y"
by simp

lemma tape_of_nl_false1:
  "lm \<noteq> [] \<Longrightarrow> rev b @ [Bk] \<noteq> Bk\<^bsup>ln\<^esup> @ Oc # Oc\<^bsup>m\<^esup> @ Bk # Bk # <lm::nat list>"
apply(auto)
apply(drule_tac rev_equal_iff, simp add: tape_of_nl_rev)
apply(case_tac "rev lm")
apply(case_tac [2] list, auto simp: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
done

lemma [simp]: "wprepare_loop_start_in_middle m lm (b, [Bk]) = False"
apply(simp add: wprepare_loop_start_in_middle.simps, auto)
apply(case_tac mr, simp_all add: exp_ind_def)
apply(case_tac lm1, simp, simp add: tape_of_nl_not_null)
done

declare wprepare_loop_start_in_middle.simps[simp del]

declare wprepare_loop_start_on_rightmost.simps[simp del] 
        wprepare_loop_goon_in_middle.simps[simp del]
        wprepare_loop_goon_on_rightmost.simps[simp del]

lemma [simp]: "wprepare_loop_goon_in_middle m lm (Bk # b, []) = False"
apply(simp add: wprepare_loop_goon_in_middle.simps, auto)
done

lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [Bk])\<rbrakk> \<Longrightarrow>
  wprepare_loop_goon m lm (Bk # b, [])"
apply(simp only: wprepare_invs, simp)
apply(simp add: wprepare_loop_goon_on_rightmost.simps 
  wprepare_loop_start_on_rightmost.simps, auto)
apply(case_tac mr, simp_all add: exp_ind_def)
apply(rule_tac rev_eq)
apply(simp add: tape_of_nl_rev)
apply(simp add: exp_ind_def[THEN sym] exp_ind)
done

lemma [simp]: "wprepare_loop_start_on_rightmost m lm (b, Bk # a # lista)
 \<Longrightarrow> wprepare_loop_goon_in_middle m lm (Bk # b, a # lista) = False"
apply(auto simp: wprepare_loop_start_on_rightmost.simps
                 wprepare_loop_goon_in_middle.simps)
apply(case_tac [!] mr, simp_all add: exp_ind_def)
done

lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start_on_rightmost m lm (b, Bk # a # lista)\<rbrakk>
    \<Longrightarrow> wprepare_loop_goon_on_rightmost m lm (Bk # b, a # lista)"
apply(simp only: wprepare_loop_start_on_rightmost.simps
                 wprepare_loop_goon_on_rightmost.simps, auto)
apply(case_tac mr, simp_all add: exp_ind_def)
apply(simp add: tape_of_nl_rev)
apply(simp add: exp_ind_def[THEN sym] exp_ind)
done

lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start_in_middle m lm (b, Bk # a # lista)\<rbrakk>
  \<Longrightarrow> wprepare_loop_goon_on_rightmost m lm (Bk # b, a # lista) = False"
apply(simp add: wprepare_loop_start_in_middle.simps
                wprepare_loop_goon_on_rightmost.simps, auto)
apply(case_tac mr, simp_all add: exp_ind_def)
apply(case_tac  "lm1::nat list", simp_all, case_tac  list, simp)
apply(simp add: tape_of_nl_abv tape_of_nat_list.simps tape_of_nat_abv exp_ind_def)
apply(case_tac [!] rna, simp_all add: exp_ind_def)
apply(case_tac mr, simp_all add: exp_ind_def)
apply(case_tac lm1, simp, case_tac list, simp)
apply(simp_all add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def tape_of_nat_abv)
done

lemma [simp]: 
  "\<lbrakk>lm \<noteq> []; wprepare_loop_start_in_middle m lm (b, Bk # a # lista)\<rbrakk> 
  \<Longrightarrow> wprepare_loop_goon_in_middle m lm (Bk # b, a # lista)"
apply(simp add: wprepare_loop_start_in_middle.simps
               wprepare_loop_goon_in_middle.simps, auto)
apply(rule_tac x = rn in exI, simp)
apply(case_tac mr, simp_all add: exp_ind_def)
apply(case_tac lm1, simp)
apply(rule_tac x = "Suc aa" in exI, simp)
apply(rule_tac x = list in exI)
apply(case_tac list, simp_all add: tape_of_nl_abv tape_of_nat_list.simps)
done

lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, Bk # a # lista)\<rbrakk> \<Longrightarrow> 
  wprepare_loop_goon m lm (Bk # b, a # lista)"
apply(simp add: wprepare_loop_start.simps 
                wprepare_loop_goon.simps)
apply(erule_tac disjE, simp, auto)
done

lemma start_2_goon:
  "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, Bk # list)\<rbrakk> \<Longrightarrow>
   (list = [] \<longrightarrow> wprepare_loop_goon m lm (Bk # b, [])) \<and>
  (list \<noteq> [] \<longrightarrow> wprepare_loop_goon m lm (Bk # b, list))"
apply(case_tac list, auto)
done

lemma add_one_2_add_one: "wprepare_add_one m lm (b, Oc # list)
  \<Longrightarrow> (hd b = Oc \<longrightarrow> (b = [] \<longrightarrow> wprepare_add_one m lm ([], Bk # Oc # list)) \<and>
                     (b \<noteq> [] \<longrightarrow> wprepare_add_one m lm (tl b, Oc # Oc # list))) \<and>
  (hd b \<noteq> Oc \<longrightarrow> (b = [] \<longrightarrow> wprepare_add_one m lm ([], Bk # Oc # list)) \<and>
                 (b \<noteq> [] \<longrightarrow> wprepare_add_one m lm (tl b, hd b # Oc # list)))"
apply(simp only: wprepare_add_one.simps, auto)
done

lemma [simp]: "wprepare_loop_start m lm (b, Oc # list) \<Longrightarrow> b \<noteq> []"
apply(simp)
done

lemma [simp]: "wprepare_loop_start_on_rightmost m lm (b, Oc # list) \<Longrightarrow> 
  wprepare_loop_start_on_rightmost m lm (Oc # b, list)"
apply(simp add: wprepare_loop_start_on_rightmost.simps, auto)
apply(rule_tac x = rn in exI, auto)
apply(case_tac mr, simp_all add: exp_ind_def)
apply(case_tac rn, auto simp: exp_ind_def)
done

lemma [simp]: "wprepare_loop_start_in_middle m lm (b, Oc # list) \<Longrightarrow> 
                wprepare_loop_start_in_middle m lm (Oc # b, list)"
apply(simp add: wprepare_loop_start_in_middle.simps, auto)
apply(rule_tac x = rn in exI, auto)
apply(case_tac mr, simp, simp add: exp_ind_def)
apply(rule_tac x = nat in exI, simp)
apply(rule_tac x = lm1 in exI, simp)
done

lemma start_2_start: "wprepare_loop_start m lm (b, Oc # list) \<Longrightarrow> 
       wprepare_loop_start m lm (Oc # b, list)"
apply(simp add: wprepare_loop_start.simps)
apply(erule_tac disjE, simp_all )
done

lemma [simp]: "wprepare_loop_goon m lm (b, Oc # list) \<Longrightarrow> b \<noteq> []"
apply(simp add: wprepare_loop_goon.simps     
                wprepare_loop_goon_in_middle.simps 
                wprepare_loop_goon_on_rightmost.simps)
apply(auto)
done

lemma [simp]: "wprepare_goto_start_pos m lm (b, Oc # list) \<Longrightarrow> b \<noteq> []"
apply(simp add: wprepare_goto_start_pos.simps)
done

lemma [simp]: "wprepare_loop_goon_on_rightmost m lm (b, Oc # list) = False"
apply(simp add: wprepare_loop_goon_on_rightmost.simps)
done
lemma wprepare_loop1: "\<lbrakk>rev b @ Oc\<^bsup>mr\<^esup> =  Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm>; 
         b \<noteq> []; 0 < mr; Oc # list = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup>\<rbrakk>
       \<Longrightarrow> wprepare_loop_start_on_rightmost m lm (Oc # b, list)"
apply(simp add: wprepare_loop_start_on_rightmost.simps)
apply(rule_tac x = rn in exI, simp)
apply(case_tac mr, simp, simp add: exp_ind_def, auto)
done

lemma wprepare_loop2: "\<lbrakk>rev b @ Oc\<^bsup>mr\<^esup> @ Bk # <a # lista> = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm>;
                b \<noteq> []; Oc # list = Oc\<^bsup>mr\<^esup> @ Bk # <(a::nat) # lista> @ Bk\<^bsup>rn\<^esup>\<rbrakk>
       \<Longrightarrow>  wprepare_loop_start_in_middle m lm (Oc # b, list)"
apply(simp add: wprepare_loop_start_in_middle.simps)
apply(rule_tac x = rn in exI, simp)
apply(case_tac mr, simp_all add: exp_ind_def)
apply(rule_tac x = nat in exI, simp)
apply(rule_tac x = "a#lista" in exI, simp)
done

lemma [simp]: "wprepare_loop_goon_in_middle m lm (b, Oc # list) \<Longrightarrow>
                wprepare_loop_start_on_rightmost m lm (Oc # b, list) \<or>
                wprepare_loop_start_in_middle m lm (Oc # b, list)"
apply(simp add: wprepare_loop_goon_in_middle.simps split: if_splits)
apply(case_tac lm1, simp_all add: wprepare_loop1 wprepare_loop2)
done

lemma [simp]: "wprepare_loop_goon m lm (b, Oc # list)
  \<Longrightarrow>  wprepare_loop_start m lm (Oc # b, list)"
apply(simp add: wprepare_loop_goon.simps
                wprepare_loop_start.simps)
done

lemma [simp]: "wprepare_add_one m lm (b, Oc # list)
       \<Longrightarrow> b = [] \<longrightarrow> wprepare_add_one m lm ([], Bk # Oc # list)"
apply(auto)
apply(simp add: wprepare_add_one.simps)
done

lemma [simp]: "wprepare_goto_start_pos m [a] (b, Oc # list)
              \<Longrightarrow> wprepare_loop_start_on_rightmost m [a] (Oc # b, list) "
apply(auto simp: wprepare_goto_start_pos.simps 
                 wprepare_loop_start_on_rightmost.simps)
apply(rule_tac x = rn in exI, simp)
apply(simp add: tape_of_nat_abv tape_of_nat_list.simps exp_ind_def, auto)
done

lemma [simp]:  "wprepare_goto_start_pos m (a # aa # listaa) (b, Oc # list)
       \<Longrightarrow>wprepare_loop_start_in_middle m (a # aa # listaa) (Oc # b, list)"
apply(auto simp: wprepare_goto_start_pos.simps
                 wprepare_loop_start_in_middle.simps)
apply(rule_tac x = rn in exI, simp)
apply(simp add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
apply(rule_tac x = a in exI, rule_tac x = "aa#listaa" in exI, simp)
done

lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_start_pos m lm (b, Oc # list)\<rbrakk>
       \<Longrightarrow> wprepare_loop_start m lm (Oc # b, list)"
apply(case_tac lm, simp_all)
apply(case_tac lista, simp_all add: wprepare_loop_start.simps)
done

lemma [simp]: "wprepare_add_one2 m lm (b, Oc # list) \<Longrightarrow> b \<noteq> []"
apply(auto simp: wprepare_add_one2.simps)
done

lemma add_one_2_stop:
  "wprepare_add_one2 m lm (b, Oc # list)      
  \<Longrightarrow>  wprepare_stop m lm (tl b, hd b # Oc # list)"
apply(simp add: wprepare_stop.simps wprepare_add_one2.simps)
done

declare wprepare_stop.simps[simp del]

lemma wprepare_correctness:
  assumes h: "lm \<noteq> []"
  shows "let P = (\<lambda> (st, l, r). st = 0) in 
  let Q = (\<lambda> (st, l, r). wprepare_inv st m lm (l, r)) in 
  let f = (\<lambda> stp. steps (Suc 0, [], (<m # lm>)) t_wcode_prepare stp) in
    \<exists> n .P (f n) \<and> Q (f n)"
proof -
  let ?P = "(\<lambda> (st, l, r). st = 0)"
  let ?Q = "(\<lambda> (st, l, r). wprepare_inv st m lm (l, r))"
  let ?f = "(\<lambda> stp. steps (Suc 0, [], (<m # lm>)) t_wcode_prepare stp)"
  have "\<exists> n. ?P (?f n) \<and> ?Q (?f n)"
  proof(rule_tac halt_lemma2)
    show "wf wcode_prepare_le" by auto
  next
    show "\<forall> n. \<not> ?P (?f n) \<and> ?Q (?f n) \<longrightarrow> 
                 ?Q (?f (Suc n)) \<and> (?f (Suc n), ?f n) \<in> wcode_prepare_le"
      using h
      apply(rule_tac allI, rule_tac impI, case_tac "?f n", 
            simp add: tstep_red tstep.simps)
      apply(case_tac c, simp, case_tac [2] aa)
      apply(simp_all add: wprepare_inv.simps wcode_prepare_le_def new_tape.simps
                          lex_triple_def lex_pair_def

                 split: if_splits)
      apply(simp_all add: start_2_goon  start_2_start
                           add_one_2_add_one add_one_2_stop)
      apply(auto simp: wprepare_add_one2.simps)
      done   
  next
    show "?Q (?f 0)"
      apply(simp add: steps.simps wprepare_inv.simps wprepare_invs)
      done
  next
    show "\<not> ?P (?f 0)"
      apply(simp add: steps.simps)
      done
  qed
  thus "?thesis"
    apply(auto)
    done
qed

lemma [intro]: "t_correct t_wcode_prepare"
apply(simp add: t_correct.simps t_wcode_prepare_def iseven_def)
apply(rule_tac x = 7 in exI, simp)
done
    
lemma twice_len_even: "length (tm_of abc_twice) mod 2 = 0"
apply(simp add: tm_even)
done

lemma fourtimes_len_even: "length (tm_of abc_fourtimes) mod 2 = 0"
apply(simp add: tm_even)
done

lemma t_correct_termi: "t_correct tp \<Longrightarrow> 
      list_all (\<lambda>(acn, st). (st \<le> Suc (length tp div 2))) (change_termi_state tp)"
apply(auto simp: t_correct.simps List.list_all_length)
apply(erule_tac x = n in allE, simp)
apply(case_tac "tp!n", auto simp: change_termi_state.simps split: if_splits)
done


lemma t_correct_shift:
         "list_all (\<lambda>(acn, st). (st \<le> y)) tp \<Longrightarrow>
          list_all (\<lambda>(acn, st). (st \<le> y + off)) (tshift tp off) "
apply(auto simp: t_correct.simps List.list_all_length)
apply(erule_tac x = n in allE, simp add: shift_length)
apply(case_tac "tp!n", auto simp: tshift.simps)
done

lemma [intro]: 
  "t_correct (tm_of abc_twice @ tMp (Suc 0) 
        (start_of twice_ly (length abc_twice) - Suc 0))"
apply(rule_tac t_compiled_correct, simp_all)
apply(simp add: twice_ly_def)
done

lemma [intro]: "t_correct (tm_of abc_fourtimes @ tMp (Suc 0) 
   (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))"
apply(rule_tac t_compiled_correct, simp_all)
apply(simp add: fourtimes_ly_def)
done


lemma [intro]: "t_correct t_wcode_main"
apply(auto simp: t_wcode_main_def t_correct.simps shift_length 
                 t_twice_def t_fourtimes_def)
proof -
  show "iseven (60 + (length (tm_of abc_twice) +
                 length (tm_of abc_fourtimes)))"
    using twice_len_even fourtimes_len_even
    apply(auto simp: iseven_def)
    apply(rule_tac x = "30 + q + qa" in exI, simp)
    done
next
  show " list_all (\<lambda>(acn, s). s \<le> (60 + (length (tm_of abc_twice) + 
           length (tm_of abc_fourtimes))) div 2) t_wcode_main_first_part"
    apply(auto simp: t_wcode_main_first_part_def shift_length t_twice_def)
    done
next
  have "list_all (\<lambda>(acn, s). s \<le> Suc (length (tm_of abc_twice @ tMp (Suc 0)
    (start_of twice_ly (length abc_twice) - Suc 0)) div 2))
    (change_termi_state (tm_of abc_twice @ tMp (Suc 0) 
    (start_of twice_ly (length abc_twice) - Suc 0)))"
    apply(rule_tac t_correct_termi, auto)
    done
  hence "list_all (\<lambda>(acn, s). s \<le>  Suc (length (tm_of abc_twice @ tMp (Suc 0)
    (start_of twice_ly (length abc_twice) - Suc 0)) div 2) + 12)
     (abacus.tshift (change_termi_state (tm_of abc_twice @ tMp (Suc 0) 
           (start_of twice_ly (length abc_twice) - Suc 0))) 12)"
    apply(rule_tac t_correct_shift, simp)
    done
  thus  "list_all (\<lambda>(acn, s). s \<le> 
           (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2)
     (abacus.tshift (change_termi_state (tm_of abc_twice @ tMp (Suc 0)
                 (start_of twice_ly (length abc_twice) - Suc 0))) 12)"
    apply(simp)
    apply(simp add: list_all_length, auto)
    done
next
  have "list_all (\<lambda>(acn, s). s \<le> Suc (length (tm_of abc_fourtimes @ tMp (Suc 0) 
    (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) div 2))
      (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) 
    (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) "
    apply(rule_tac t_correct_termi, auto)
    done
  hence "list_all (\<lambda>(acn, s). s \<le> Suc (length (tm_of abc_fourtimes @ tMp (Suc 0) 
    (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) div 2) + (t_twice_len + 13))
    (abacus.tshift (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) 
    (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) (t_twice_len + 13))"
    apply(rule_tac t_correct_shift, simp)
    done
  thus "list_all (\<lambda>(acn, s). s \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2)
    (abacus.tshift (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0)
    (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) (t_twice_len + 13))"
    apply(simp add: t_twice_len_def t_twice_def)
    using twice_len_even fourtimes_len_even
    apply(auto simp: list_all_length)
    done
qed

lemma [intro]: "t_correct (t_wcode_prepare |+| t_wcode_main)"
apply(auto intro: t_correct_add)
done

lemma prepare_mainpart_lemma:
  "args \<noteq> [] \<Longrightarrow> 
  \<exists> stp ln rn. steps (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) stp
              = (0,  Bk # Oc\<^bsup>Suc m\<^esup>, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
proof -
  let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <m # args>"
  let ?Q1 = "\<lambda> (l, r). wprepare_stop m args (l, r)"
  let ?P2 = ?Q1
  let ?Q2 = "\<lambda> (l, r). (\<exists> ln rn. l = Bk # Oc\<^bsup>Suc m\<^esup> \<and>
                           r =  Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  let ?P3 = "\<lambda> tp. False"
  assume h: "args \<noteq> []"
  have "?P1 \<turnstile>-> \<lambda> tp. (\<exists> stp tp'. steps (Suc 0, tp)
                      (t_wcode_prepare |+| t_wcode_main) stp = (0, tp') \<and> ?Q2 tp')"
  proof(rule_tac turing_merge.t_merge_halt[of t_wcode_prepare t_wcode_main ?P1 ?P2 ?P3 ?P3 ?Q1 ?Q2], 
        auto simp: turing_merge_def)
    show "\<exists>stp. case steps (Suc 0, [], <m # args>) t_wcode_prepare stp of (st, tp')
                  \<Rightarrow> st = 0 \<and> wprepare_stop m args tp'"
      using wprepare_correctness[of args m] h
      apply(simp, auto)
      apply(rule_tac x = n in exI, simp add: wprepare_inv.simps)
      done
  next
    fix a b
    assume "wprepare_stop m args (a, b)"
    thus "\<exists>stp. case steps (Suc 0, a, b) t_wcode_main stp of
      (st, tp') \<Rightarrow> (st = 0) \<and> (case tp' of (l, r) \<Rightarrow> l = Bk # Oc\<^bsup>Suc m\<^esup> \<and> 
      (\<exists>ln rn. r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>))"
      proof(simp only: wprepare_stop.simps, erule_tac exE)
        fix rn
        assume "a = Bk # <rev args> @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and> 
                   b = Bk # Oc # Bk\<^bsup>rn\<^esup>"
        thus "?thesis"
          using t_wcode_main_lemma_pre[of "args" "<args>" 0 "Oc\<^bsup>Suc m\<^esup>" 0 rn] h
          apply(simp)
          apply(erule_tac exE)+
          apply(rule_tac x = stp in exI, simp add: tape_of_nl_rev, auto)
          done
      qed
  next
    show "wprepare_stop m args \<turnstile>-> wprepare_stop m args"
      by(simp add: t_imply_def)
  qed
  thus "\<exists> stp ln rn. steps (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) stp
              = (0,  Bk # Oc\<^bsup>Suc m\<^esup>, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
    apply(simp add: t_imply_def)
    apply(erule_tac exE)+
    apply(auto)
    done
qed
      

lemma [simp]:  "tinres r r' \<Longrightarrow> 
  fetch t ss (case r of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) = 
  fetch t ss (case r' of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)"
apply(simp add: fetch.simps, auto split: if_splits simp: tinres_def)
apply(case_tac [!] r', simp_all)
apply(case_tac [!] n, simp_all add: exp_ind_def)
apply(case_tac [!] r, simp_all)
done

lemma [intro]: "\<exists> n. (a::block)\<^bsup>n\<^esup> = []"
by auto

lemma [simp]: "\<lbrakk>tinres r r'; r \<noteq> []; r' \<noteq> []\<rbrakk> \<Longrightarrow> hd r = hd r'"
apply(auto simp: tinres_def)
done

lemma [intro]: "hd (Bk\<^bsup>Suc n\<^esup>) = Bk"
apply(simp add: exp_ind_def)
done

lemma [simp]: "\<lbrakk>tinres r []; r \<noteq> []\<rbrakk> \<Longrightarrow> hd r = Bk"
apply(auto simp: tinres_def)
apply(case_tac n, auto)
done

lemma [simp]: "\<lbrakk>tinres [] r'; r' \<noteq> []\<rbrakk> \<Longrightarrow> hd r' = Bk"
apply(auto simp: tinres_def)
done

lemma [intro]: "\<exists>na. tl r = tl (r @ Bk\<^bsup>n\<^esup>) @ Bk\<^bsup>na\<^esup> \<or> tl (r @ Bk\<^bsup>n\<^esup>) = tl r @ Bk\<^bsup>na\<^esup>"
apply(case_tac r, simp)
apply(case_tac n, simp)
apply(rule_tac x = 0 in exI, simp)
apply(rule_tac x = nat in exI, simp add: exp_ind_def)
apply(simp)
apply(rule_tac x = n in exI, simp)
done

lemma [simp]: "tinres r r' \<Longrightarrow> tinres (tl r) (tl r')"
apply(auto simp: tinres_def)
apply(case_tac r', simp_all)
apply(case_tac n, simp_all add: exp_ind_def)
apply(rule_tac x = 0 in exI, simp)
apply(rule_tac x = nat in exI, simp_all)
apply(rule_tac x = n in exI, simp)
done

lemma [simp]: "\<lbrakk>tinres r [];  r \<noteq> []\<rbrakk> \<Longrightarrow> tinres (tl r) []"
apply(case_tac r, auto simp: tinres_def)
apply(case_tac n, simp_all add: exp_ind_def)
apply(rule_tac x = nat in exI, simp)
done

lemma [simp]: "\<lbrakk>tinres [] r'\<rbrakk> \<Longrightarrow> tinres [] (tl r')"
apply(case_tac r', auto simp: tinres_def)
apply(case_tac n, simp_all add: exp_ind_def)
apply(rule_tac x = nat in exI, simp)
done

lemma [simp]: "tinres r r' \<Longrightarrow> tinres (b # r) (b # r')"
apply(auto simp: tinres_def)
done

lemma tinres_step2: 
  "\<lbrakk>tinres r r'; tstep (ss, l, r) t = (sa, la, ra); tstep (ss, l, r') t = (sb, lb, rb)\<rbrakk>
    \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb"
apply(case_tac "ss = 0", simp add: tstep_0)
apply(simp add: tstep.simps [simp del])
apply(case_tac "fetch t ss (case r of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp)
apply(auto simp: new_tape.simps)
apply(simp_all split: taction.splits if_splits)
apply(auto)
done


lemma tinres_steps2: 
  "\<lbrakk>tinres r r'; steps (ss, l, r) t stp = (sa, la, ra); steps (ss, l, r') t stp = (sb, lb, rb)\<rbrakk>
    \<Longrightarrow> la = lb \<and> tinres 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> la = lb \<and> tinres ra rb \<and> sa = sb"
  and h: " tinres r r'" "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 "b = ba \<and> tinres c ca \<and> a = aa"
    apply(rule_tac ind, simp_all add: h)
    done
  thus "la = lb \<and> tinres ra rb \<and> sa = sb"
    apply(rule_tac l = b  and r = c  and ss = a and r' = ca   
            and t = t in tinres_step2)
    using h
    apply(simp, simp, simp)
    done
qed
 
definition t_wcode_adjust :: "tprog"
  where
  "t_wcode_adjust = [(W1, 1), (R, 2), (Nop, 2), (R, 3), (R, 3), (R, 4), 
                   (L, 8), (L, 5), (L, 6), (W0, 5), (L, 6), (R, 7), 
                   (W1, 2), (Nop, 7), (L, 9), (W0, 8), (L, 9), (L, 10), 
                    (L, 11), (L, 10), (R, 0), (L, 11)]"
                 
lemma [simp]: "fetch t_wcode_adjust (Suc 0) Bk = (W1, 1)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust (Suc 0) Oc = (R, 2)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust (Suc (Suc 0)) Oc = (R, 3)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust (Suc (Suc (Suc 0))) Oc = (R, 4)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust  (Suc (Suc (Suc 0))) Bk = (R, 3)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done
   
lemma [simp]: "fetch t_wcode_adjust 4 Bk = (L, 8)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust 4 Oc = (L, 5)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust 5 Oc = (W0, 5)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust 5 Bk = (L, 6)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust 6 Oc = (R, 7)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust 6 Bk = (L, 6)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust 7 Bk = (W1, 2)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust 8 Bk = (L, 9)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust 8 Oc = (W0, 8)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust 9 Oc = (L, 10)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust 9 Bk = (L, 9)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust 10 Bk = (L, 11)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust 10 Oc = (L, 10)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust 11 Oc = (L, 11)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

lemma [simp]: "fetch t_wcode_adjust 11 Bk = (R, 0)"
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
done

fun wadjust_start :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_start m rs (l, r) = 
         (\<exists> ln rn. l = Bk # Oc\<^bsup>Suc m\<^esup> \<and>
                   tl r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"

fun wadjust_loop_start :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_loop_start m rs (l, r) = 
          (\<exists> ln rn ml mr. l = Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup>  \<and>
                          r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
                          ml + mr = Suc (Suc rs) \<and> mr > 0)"

fun wadjust_loop_right_move :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_loop_right_move m rs (l, r) = 
   (\<exists> ml mr nl nr rn. l = Bk\<^bsup>nl\<^esup> @ Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
                      r = Bk\<^bsup>nr\<^esup> @ Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
                      ml + mr = Suc (Suc rs) \<and> mr > 0 \<and>
                      nl + nr > 0)"

fun wadjust_loop_check :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_loop_check m rs (l, r) = 
  (\<exists> ml mr ln rn. l = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
                  r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> ml + mr = (Suc rs))"

fun wadjust_loop_erase :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_loop_erase m rs (l, r) = 
    (\<exists> ml mr ln rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
                    tl r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> ml + mr = (Suc rs) \<and> mr > 0)"

fun wadjust_loop_on_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_loop_on_left_moving_O m rs (l, r) = 
      (\<exists> ml mr ln rn. l = Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m \<^esup>\<and>
                      r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
                      ml + mr = Suc rs \<and> mr > 0)"

fun wadjust_loop_on_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_loop_on_left_moving_B m rs (l, r) = 
      (\<exists> ml mr nl nr rn. l = Bk\<^bsup>nl\<^esup> @ Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
                         r = Bk\<^bsup>nr\<^esup> @ Bk # Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
                         ml + mr = Suc rs \<and> mr > 0)"

fun wadjust_loop_on_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_loop_on_left_moving m rs (l, r) = 
       (wadjust_loop_on_left_moving_O m rs (l, r) \<or>
       wadjust_loop_on_left_moving_B m rs (l, r))"

fun wadjust_loop_right_move2 :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_loop_right_move2 m rs (l, r) = 
        (\<exists> ml mr ln rn. l = Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
                        r = Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
                        ml + mr = Suc rs \<and> mr > 0)"

fun wadjust_erase2 :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_erase2 m rs (l, r) = 
     (\<exists> ln rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Oc # Oc\<^bsup>Suc rs\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
                     tl r = Bk\<^bsup>rn\<^esup>)"

fun wadjust_on_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_on_left_moving_O m rs (l, r) = 
        (\<exists> rn. l = Oc\<^bsup>Suc rs\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
                  r = Oc # Bk\<^bsup>rn\<^esup>)"

fun wadjust_on_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_on_left_moving_B m rs (l, r) = 
         (\<exists> ln rn. l = Bk\<^bsup>ln\<^esup> @ Oc # Oc\<^bsup>Suc rs\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
                   r = Bk\<^bsup>rn\<^esup>)"

fun wadjust_on_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_on_left_moving m rs (l, r) = 
      (wadjust_on_left_moving_O m rs (l, r) \<or>
       wadjust_on_left_moving_B m rs (l, r))"

fun wadjust_goon_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where 
  "wadjust_goon_left_moving_B m rs (l, r) = 
        (\<exists> rn. l = Oc\<^bsup>Suc m\<^esup> \<and> 
               r = Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"

fun wadjust_goon_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_goon_left_moving_O m rs (l, r) = 
      (\<exists> ml mr rn. l = Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
                      r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
                      ml + mr = Suc (Suc rs) \<and> mr > 0)"

fun wadjust_goon_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_goon_left_moving m rs (l, r) = 
            (wadjust_goon_left_moving_B m rs (l, r) \<or>
             wadjust_goon_left_moving_O m rs (l, r))"

fun wadjust_backto_standard_pos_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_backto_standard_pos_B m rs (l, r) =
        (\<exists> rn. l = [] \<and> 
               r = Bk # Oc\<^bsup>Suc m \<^esup>@ Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"

fun wadjust_backto_standard_pos_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_backto_standard_pos_O m rs (l, r) = 
      (\<exists> ml mr rn. l = Oc\<^bsup>ml\<^esup> \<and>
                      r = Oc\<^bsup>mr\<^esup> @ Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
                      ml + mr = Suc m \<and> mr > 0)"

fun wadjust_backto_standard_pos :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_backto_standard_pos m rs (l, r) = 
        (wadjust_backto_standard_pos_B m rs (l, r) \<or> 
        wadjust_backto_standard_pos_O m rs (l, r))"

fun wadjust_stop :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
  "wadjust_stop m rs (l, r) =
        (\<exists> rn. l = [Bk] \<and> 
               r = Oc\<^bsup>Suc m \<^esup>@ Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"

declare wadjust_start.simps[simp del]  wadjust_loop_start.simps[simp del]
        wadjust_loop_right_move.simps[simp del]  wadjust_loop_check.simps[simp del]
        wadjust_loop_erase.simps[simp del] wadjust_loop_on_left_moving.simps[simp del]
        wadjust_loop_right_move2.simps[simp del] wadjust_erase2.simps[simp del]
        wadjust_on_left_moving_O.simps[simp del] wadjust_on_left_moving_B.simps[simp del]
        wadjust_on_left_moving.simps[simp del] wadjust_goon_left_moving_B.simps[simp del]
        wadjust_goon_left_moving_O.simps[simp del] wadjust_goon_left_moving.simps[simp del]
        wadjust_backto_standard_pos.simps[simp del] wadjust_backto_standard_pos_B.simps[simp del]
        wadjust_backto_standard_pos_O.simps[simp del] wadjust_stop.simps[simp del]

fun wadjust_inv :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  where
  "wadjust_inv st m rs (l, r) = 
       (if st = Suc 0 then wadjust_start m rs (l, r) 
        else if st = Suc (Suc 0) then wadjust_loop_start m rs (l, r)
        else if st = Suc (Suc (Suc 0)) then wadjust_loop_right_move m rs (l, r)
        else if st = 4 then wadjust_loop_check m rs (l, r)
        else if st = 5 then wadjust_loop_erase m rs (l, r)
        else if st = 6 then wadjust_loop_on_left_moving m rs (l, r)
        else if st = 7 then wadjust_loop_right_move2 m rs (l, r)
        else if st = 8 then wadjust_erase2 m rs (l, r)
        else if st = 9 then wadjust_on_left_moving m rs (l, r)
        else if st = 10 then wadjust_goon_left_moving m rs (l, r)
        else if st = 11 then wadjust_backto_standard_pos m rs (l, r)
        else if st = 0 then wadjust_stop m rs (l, r)
        else False
)"

declare wadjust_inv.simps[simp del]

fun wadjust_phase :: "nat \<Rightarrow> t_conf \<Rightarrow> nat"
  where
  "wadjust_phase rs (st, l, r) = 
         (if st = 1 then 3 
          else if st \<ge> 2 \<and> st \<le> 7 then 2
          else if st \<ge> 8 \<and> st \<le> 11 then 1
          else 0)"

thm dropWhile.simps

fun wadjust_stage :: "nat \<Rightarrow> t_conf \<Rightarrow> nat"
  where
  "wadjust_stage rs (st, l, r) = 
           (if st \<ge> 2 \<and> st \<le> 7 then 
                  rs - length (takeWhile (\<lambda> a. a = Oc) 
                          (tl (dropWhile (\<lambda> a. a = Oc) (rev l @ r))))
            else 0)"

fun wadjust_state :: "nat \<Rightarrow> t_conf \<Rightarrow> nat"
  where
  "wadjust_state rs (st, l, r) = 
       (if st \<ge> 2 \<and> st \<le> 7 then 8 - st
        else if st \<ge> 8 \<and> st \<le> 11 then 12 - st
        else 0)"

fun wadjust_step :: "nat \<Rightarrow> t_conf \<Rightarrow> nat"
  where
  "wadjust_step rs (st, l, r) = 
       (if st = 1 then (if hd r = Bk then 1
                        else 0) 
        else if st = 3 then length r
        else if st = 5 then (if hd r = Oc then 1
                             else 0)
        else if st = 6 then length l
        else if st = 8 then (if hd r = Oc then 1
                             else 0)
        else if st = 9 then length l
        else if st = 10 then length l
        else if st = 11 then (if hd r = Bk then 0
                              else Suc (length l))
        else 0)"

fun wadjust_measure :: "(nat \<times> t_conf) \<Rightarrow> nat \<times> nat \<times> nat \<times> nat"
  where
  "wadjust_measure (rs, (st, l, r)) = 
     (wadjust_phase rs (st, l, r), 
      wadjust_stage rs (st, l, r),
      wadjust_state rs (st, l, r), 
      wadjust_step rs (st, l, r))"

definition wadjust_le :: "((nat \<times> t_conf) \<times> nat \<times> t_conf) set"
  where "wadjust_le \<equiv> (inv_image lex_square wadjust_measure)"

lemma [intro]: "wf lex_square"
by(auto intro:wf_lex_prod simp: abacus.lex_pair_def lex_square_def 
  abacus.lex_triple_def)

lemma wf_wadjust_le[intro]: "wf wadjust_le"
by(auto intro:wf_inv_image simp: wadjust_le_def
           abacus.lex_triple_def abacus.lex_pair_def)

lemma [simp]: "wadjust_start m rs (c, []) = False"
apply(auto simp: wadjust_start.simps)
done

lemma [simp]: "wadjust_loop_right_move m rs (c, []) \<Longrightarrow> c \<noteq> []"
apply(auto simp: wadjust_loop_right_move.simps)
done

lemma [simp]: "wadjust_loop_right_move m rs (c, [])
        \<Longrightarrow>  wadjust_loop_check m rs (Bk # c, [])"
apply(simp only: wadjust_loop_right_move.simps wadjust_loop_check.simps)
apply(auto)
apply(case_tac [!] mr, simp_all add: exp_ind_def)
done

lemma [simp]: "wadjust_loop_check m rs (c, []) \<Longrightarrow> c \<noteq> []"
apply(simp only: wadjust_loop_check.simps, auto)
done
 
lemma [simp]: "wadjust_loop_start m rs (c, []) = False"
apply(simp add: wadjust_loop_start.simps)
done

lemma [simp]: "wadjust_loop_right_move m rs (c, []) \<Longrightarrow> 
  wadjust_loop_right_move m rs (Bk # c, [])"
apply(simp only: wadjust_loop_right_move.simps)
apply(erule_tac exE)+
apply(auto)
apply(case_tac [!] mr, simp_all add: exp_ind_def)
done

lemma [simp]: "wadjust_loop_check m rs (c, []) \<Longrightarrow> wadjust_erase2 m rs (tl c, [hd c])"
apply(simp only: wadjust_loop_check.simps wadjust_erase2.simps, auto)
apply(case_tac mr, simp_all add: exp_ind_def, auto)
done

lemma [simp]: " wadjust_loop_erase m rs (c, [])
    \<Longrightarrow> (c = [] \<longrightarrow> wadjust_loop_on_left_moving m rs ([], [Bk])) \<and>
        (c \<noteq> [] \<longrightarrow> wadjust_loop_on_left_moving m rs (tl c, [hd c]))"
apply(simp add: wadjust_loop_erase.simps, auto)
apply(case_tac [!] mr, simp_all add: exp_ind_def)
done

lemma [simp]: "wadjust_loop_on_left_moving m rs (c, []) = False"
apply(auto simp: wadjust_loop_on_left_moving.simps)
done


lemma [simp]: "wadjust_loop_right_move2 m rs (c, []) = False"
apply(auto simp: wadjust_loop_right_move2.simps)
done
   
lemma [simp]: "wadjust_erase2 m rs ([], []) = False"
apply(auto simp: wadjust_erase2.simps)
done

lemma [simp]: "wadjust_on_left_moving_B m rs 
                 (Oc # Oc # Oc\<^bsup>rs\<^esup> @ Bk # Oc # Oc\<^bsup>m\<^esup>, [Bk])"
apply(simp add: wadjust_on_left_moving_B.simps, auto)
apply(rule_tac x = 0 in exI, simp add: exp_ind_def)
done

lemma [simp]: "wadjust_on_left_moving_B m rs 
                 (Bk\<^bsup>n\<^esup> @ Bk # Oc # Oc # Oc\<^bsup>rs\<^esup> @ Bk # Oc # Oc\<^bsup>m\<^esup>, [Bk])"
apply(simp add: wadjust_on_left_moving_B.simps exp_ind_def, auto)
apply(rule_tac x = "Suc n" in exI, simp add: exp_ind)
done

lemma [simp]: "\<lbrakk>wadjust_erase2 m rs (c, []); c \<noteq> []\<rbrakk> \<Longrightarrow>
            wadjust_on_left_moving m rs (tl c, [hd c])"
apply(simp only: wadjust_erase2.simps)
apply(erule_tac exE)+
apply(case_tac ln, simp_all add: exp_ind_def wadjust_on_left_moving.simps)
done

lemma [simp]: "wadjust_erase2 m rs (c, [])
    \<Longrightarrow> (c = [] \<longrightarrow> wadjust_on_left_moving m rs ([], [Bk])) \<and> 
       (c \<noteq> [] \<longrightarrow> wadjust_on_left_moving m rs (tl c, [hd c]))"
apply(auto)
done

lemma [simp]: "wadjust_on_left_moving m rs ([], []) = False"
apply(simp add: wadjust_on_left_moving.simps 
  wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps)
done

lemma [simp]: "wadjust_on_left_moving_O m rs (c, []) = False"
apply(simp add: wadjust_on_left_moving_O.simps)
done

lemma [simp]: " \<lbrakk>wadjust_on_left_moving_B m rs (c, []); c \<noteq> []; hd c = Bk\<rbrakk> \<Longrightarrow>
                                      wadjust_on_left_moving_B m rs (tl c, [Bk])"
apply(simp add: wadjust_on_left_moving_B.simps, auto)
apply(case_tac [!] ln, simp_all add: exp_ind_def, auto)
done

lemma [simp]: "\<lbrakk>wadjust_on_left_moving_B m rs (c, []); c \<noteq> []; hd c = Oc\<rbrakk> \<Longrightarrow>
                                  wadjust_on_left_moving_O m rs (tl c, [Oc])"
apply(simp add: wadjust_on_left_moving_B.simps wadjust_on_left_moving_O.simps, auto)
apply(case_tac [!] ln, simp_all add: exp_ind_def)
done

lemma [simp]: "\<lbrakk>wadjust_on_left_moving m rs (c, []); c \<noteq> []\<rbrakk> \<Longrightarrow> 
  wadjust_on_left_moving m rs (tl c, [hd c])"
apply(simp add: wadjust_on_left_moving.simps)
apply(case_tac "hd c", simp_all)
done

lemma [simp]: "wadjust_on_left_moving m rs (c, [])
    \<Longrightarrow> (c = [] \<longrightarrow> wadjust_on_left_moving m rs ([], [Bk])) \<and> 
       (c \<noteq> [] \<longrightarrow> wadjust_on_left_moving m rs (tl c, [hd c]))"
apply(auto)
done

lemma [simp]: "wadjust_goon_left_moving m rs (c, []) = False"
apply(auto simp: wadjust_goon_left_moving.simps wadjust_goon_left_moving_B.simps
                 wadjust_goon_left_moving_O.simps)
done

lemma [simp]: "wadjust_backto_standard_pos m rs (c, []) = False"
apply(auto simp: wadjust_backto_standard_pos.simps
 wadjust_backto_standard_pos_B.simps wadjust_backto_standard_pos_O.simps)
done

lemma [simp]:
  "wadjust_start m rs (c, Bk # list) \<Longrightarrow> 
  (c = [] \<longrightarrow> wadjust_start m rs ([], Oc # list)) \<and> 
  (c \<noteq> [] \<longrightarrow> wadjust_start m rs (c, Oc # list))"
apply(auto simp: wadjust_start.simps)
done

lemma [simp]: "wadjust_loop_start m rs (c, Bk # list) = False"
apply(auto simp: wadjust_loop_start.simps)
done

lemma [simp]: "wadjust_loop_right_move m rs (c, b) \<Longrightarrow> c \<noteq> []"
apply(simp only: wadjust_loop_right_move.simps, auto)
done

lemma [simp]: "wadjust_loop_right_move m rs (c, Bk # list)
    \<Longrightarrow> wadjust_loop_right_move m rs (Bk # c, list)"
apply(simp only: wadjust_loop_right_move.simps)
apply(erule_tac exE)+
apply(rule_tac x = ml in exI, simp)
apply(rule_tac x = mr in exI, simp)
apply(rule_tac x = "Suc nl" in exI, simp add: exp_ind_def)
apply(case_tac nr, simp, case_tac mr, simp_all add: exp_ind_def)
apply(rule_tac x = nat in exI, auto)
done

lemma [simp]: "wadjust_loop_check m rs (c, b) \<Longrightarrow> c \<noteq> []"
apply(simp only: wadjust_loop_check.simps, auto)
done

lemma [simp]: "wadjust_loop_check m rs (c, Bk # list)
              \<Longrightarrow>  wadjust_erase2 m rs (tl c, hd c # Bk # list)"
apply(auto simp: wadjust_loop_check.simps wadjust_erase2.simps)
apply(case_tac [!] mr, simp_all add: exp_ind_def, auto)
done

lemma [simp]: "wadjust_loop_erase m rs (c, b) \<Longrightarrow> c \<noteq> []"
apply(simp only: wadjust_loop_erase.simps, auto)
done

declare wadjust_loop_on_left_moving_O.simps[simp del]
        wadjust_loop_on_left_moving_B.simps[simp del]

lemma [simp]: "\<lbrakk>wadjust_loop_erase m rs (c, Bk # list); hd c = Bk\<rbrakk>
    \<Longrightarrow> wadjust_loop_on_left_moving_B m rs (tl c, Bk # Bk # list)"
apply(simp only: wadjust_loop_erase.simps 
  wadjust_loop_on_left_moving_B.simps)
apply(erule_tac exE)+
apply(rule_tac x = ml in exI, rule_tac x = mr in exI, 
      rule_tac x = ln in exI, rule_tac x = 0 in exI, simp)
apply(case_tac ln, simp_all add: exp_ind_def, auto)
apply(simp add: exp_ind exp_ind_def[THEN sym])
done

lemma [simp]: "\<lbrakk>wadjust_loop_erase m rs (c, Bk # list); c \<noteq> []; hd c = Oc\<rbrakk> \<Longrightarrow>
             wadjust_loop_on_left_moving_O m rs (tl c, Oc # Bk # list)"
apply(simp only: wadjust_loop_erase.simps wadjust_loop_on_left_moving_O.simps,
       auto)
apply(case_tac [!] ln, simp_all add: exp_ind_def)
done

lemma [simp]: "\<lbrakk>wadjust_loop_erase m rs (c, Bk # list); c \<noteq> []\<rbrakk> \<Longrightarrow> 
                wadjust_loop_on_left_moving m rs (tl c, hd c # Bk # list)"
apply(case_tac "hd c", simp_all add:wadjust_loop_on_left_moving.simps)
done

lemma [simp]: "wadjust_loop_on_left_moving m rs (c, b) \<Longrightarrow> c \<noteq> []"
apply(simp add: wadjust_loop_on_left_moving.simps 
wadjust_loop_on_left_moving_O.simps wadjust_loop_on_left_moving_B.simps, auto)
done

lemma [simp]: "wadjust_loop_on_left_moving_O m rs (c, Bk # list) = False"
apply(simp add: wadjust_loop_on_left_moving_O.simps)
done

lemma [simp]: "\<lbrakk>wadjust_loop_on_left_moving_B m rs (c, Bk # list); hd c = Bk\<rbrakk>
    \<Longrightarrow>  wadjust_loop_on_left_moving_B m rs (tl c, Bk # Bk # list)"
apply(simp only: wadjust_loop_on_left_moving_B.simps)
apply(erule_tac exE)+
apply(rule_tac x = ml in exI, rule_tac x = mr in exI)
apply(case_tac nl, simp_all add: exp_ind_def, auto)
apply(rule_tac x = "Suc nr" in exI, auto simp: exp_ind_def)
done

lemma [simp]: "\<lbrakk>wadjust_loop_on_left_moving_B m rs (c, Bk # list); hd c = Oc\<rbrakk>
    \<Longrightarrow> wadjust_loop_on_left_moving_O m rs (tl c, Oc # Bk # list)"
apply(simp only: wadjust_loop_on_left_moving_O.simps 
                 wadjust_loop_on_left_moving_B.simps)
apply(erule_tac exE)+
apply(rule_tac x = ml in exI, rule_tac x = mr in exI)
apply(case_tac nl, simp_all add: exp_ind_def, auto)
done

lemma [simp]: "wadjust_loop_on_left_moving m rs (c, Bk # list)
            \<Longrightarrow> wadjust_loop_on_left_moving m rs (tl c, hd c # Bk # list)"
apply(simp add: wadjust_loop_on_left_moving.simps)
apply(case_tac "hd c", simp_all)
done

lemma [simp]: "wadjust_loop_right_move2 m rs (c, b) \<Longrightarrow> c \<noteq> []"
apply(simp only: wadjust_loop_right_move2.simps, auto)
done

lemma [simp]: "wadjust_loop_right_move2 m rs (c, Bk # list) \<Longrightarrow>  wadjust_loop_start m rs (c, Oc # list)"
apply(auto simp: wadjust_loop_right_move2.simps wadjust_loop_start.simps)
apply(case_tac ln, simp_all add: exp_ind_def)
apply(rule_tac x = 0 in exI, simp)
apply(rule_tac x = rn in exI, simp)
apply(rule_tac x = "Suc ml" in exI, simp add: exp_ind_def, auto)
apply(rule_tac x = "Suc nat" in exI, simp add: exp_ind)
apply(rule_tac x = rn in exI, auto)
apply(rule_tac x = "Suc ml" in exI, auto simp: exp_ind_def)
done

lemma [simp]: "wadjust_erase2 m rs (c, Bk # list) \<Longrightarrow> c \<noteq> []"
apply(auto simp:wadjust_erase2.simps )
done

lemma [simp]: "wadjust_erase2 m rs (c, Bk # list) \<Longrightarrow> 
                 wadjust_on_left_moving m rs (tl c, hd c # Bk # list)"
apply(auto simp: wadjust_erase2.simps)
apply(case_tac ln, simp_all add: exp_ind_def wadjust_on_left_moving.simps 
        wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps)
apply(auto)
apply(rule_tac x = "(Suc (Suc rn))" in exI, simp add: exp_ind_def)
apply(rule_tac x = "Suc nat" in exI, simp add: exp_ind)
apply(rule_tac x = "(Suc (Suc rn))" in exI, simp add: exp_ind_def)
done

lemma [simp]: "wadjust_on_left_moving m rs (c,b) \<Longrightarrow> c \<noteq> []"
apply(simp only:wadjust_on_left_moving.simps
                wadjust_on_left_moving_O.simps
                wadjust_on_left_moving_B.simps
             , auto)
done

lemma [simp]: "wadjust_on_left_moving_O m rs (c, Bk # list) = False"
apply(simp add: wadjust_on_left_moving_O.simps)
done

lemma [simp]: "\<lbrakk>wadjust_on_left_moving_B m rs (c, Bk # list); hd c = Bk\<rbrakk>
    \<Longrightarrow> wadjust_on_left_moving_B m rs (tl c, Bk # Bk # list)"
apply(auto simp: wadjust_on_left_moving_B.simps)
apply(case_tac ln, simp_all add: exp_ind_def, auto)
done

lemma [simp]: "\<lbrakk>wadjust_on_left_moving_B m rs (c, Bk # list); hd c = Oc\<rbrakk>
    \<Longrightarrow> wadjust_on_left_moving_O m rs (tl c, Oc # Bk # list)"
apply(auto simp: wadjust_on_left_moving_O.simps
                 wadjust_on_left_moving_B.simps)
apply(case_tac ln, simp_all add: exp_ind_def)
done

lemma [simp]: "wadjust_on_left_moving  m rs (c, Bk # list) \<Longrightarrow>  
                  wadjust_on_left_moving m rs (tl c, hd c # Bk # list)"
apply(simp add: wadjust_on_left_moving.simps)
apply(case_tac "hd c", simp_all)
done

lemma [simp]: "wadjust_goon_left_moving m rs (c, b) \<Longrightarrow> c \<noteq> []"
apply(simp add: wadjust_goon_left_moving.simps
                wadjust_goon_left_moving_B.simps
                wadjust_goon_left_moving_O.simps exp_ind_def, auto)
done

lemma [simp]: "wadjust_goon_left_moving_O m rs (c, Bk # list) = False"
apply(simp add: wadjust_goon_left_moving_O.simps, auto)
apply(case_tac mr, simp_all add: exp_ind_def)
done

lemma [simp]: "\<lbrakk>wadjust_goon_left_moving_B m rs (c, Bk # list); hd c = Bk\<rbrakk>
    \<Longrightarrow> wadjust_backto_standard_pos_B m rs (tl c, Bk # Bk # list)"
apply(auto simp: wadjust_goon_left_moving_B.simps 
                 wadjust_backto_standard_pos_B.simps exp_ind_def)
done

lemma [simp]: "\<lbrakk>wadjust_goon_left_moving_B m rs (c, Bk # list); hd c = Oc\<rbrakk>
    \<Longrightarrow> wadjust_backto_standard_pos_O m rs (tl c, Oc # Bk # list)"
apply(auto simp: wadjust_goon_left_moving_B.simps 
                 wadjust_backto_standard_pos_O.simps exp_ind_def)
apply(rule_tac x = m in exI, simp, auto)
done

lemma [simp]: "wadjust_goon_left_moving m rs (c, Bk # list) \<Longrightarrow>
  wadjust_backto_standard_pos m rs (tl c, hd c # Bk # list)"
apply(case_tac "hd c", simp_all add: wadjust_backto_standard_pos.simps 
                                     wadjust_goon_left_moving.simps)
done

lemma [simp]: "wadjust_backto_standard_pos m rs (c, Bk # list) \<Longrightarrow>
  (c = [] \<longrightarrow> wadjust_stop m rs ([Bk], list)) \<and> (c \<noteq> [] \<longrightarrow> wadjust_stop m rs (Bk # c, list))"
apply(auto simp: wadjust_backto_standard_pos.simps 
                 wadjust_backto_standard_pos_B.simps
                 wadjust_backto_standard_pos_O.simps wadjust_stop.simps)
apply(case_tac [!] mr, simp_all add: exp_ind_def)
done

lemma [simp]: "wadjust_start m rs (c, Oc # list)
              \<Longrightarrow> (c = [] \<longrightarrow> wadjust_loop_start m rs ([Oc], list)) \<and>
                (c \<noteq> [] \<longrightarrow> wadjust_loop_start m rs (Oc # c, list))"
apply(auto simp:wadjust_loop_start.simps wadjust_start.simps )
apply(rule_tac x = ln in exI, rule_tac x = rn in exI,
      rule_tac x = "Suc 0" in exI, simp)
done

lemma [simp]: "wadjust_loop_start m rs (c, b) \<Longrightarrow> c \<noteq> []"
apply(simp add: wadjust_loop_start.simps, auto)
done

lemma [simp]: "wadjust_loop_start m rs (c, Oc # list)
              \<Longrightarrow> wadjust_loop_right_move m rs (Oc # c, list)"
apply(simp add: wadjust_loop_start.simps wadjust_loop_right_move.simps, auto)
apply(rule_tac x = ml in exI, rule_tac x = mr in exI, 
      rule_tac x = 0 in exI, simp)
apply(rule_tac x = "Suc ln" in exI, simp add: exp_ind, auto)
done

lemma [simp]: "wadjust_loop_right_move m rs (c, Oc # list) \<Longrightarrow> 
                       wadjust_loop_check m rs (Oc # c, list)"
apply(simp add: wadjust_loop_right_move.simps  
                 wadjust_loop_check.simps, auto)
apply(rule_tac [!] x = ml in exI, simp_all, auto)
apply(case_tac nl, auto simp: exp_ind_def)
apply(rule_tac x = "mr - 1" in exI, case_tac mr, simp_all add: exp_ind_def)
apply(case_tac [!] nr, simp_all add: exp_ind_def, auto)
done

lemma [simp]: "wadjust_loop_check m rs (c, Oc # list) \<Longrightarrow> 
               wadjust_loop_erase m rs (tl c, hd c # Oc # list)"
apply(simp only: wadjust_loop_check.simps wadjust_loop_erase.simps)
apply(erule_tac exE)+
apply(rule_tac x = ml in exI, rule_tac x = mr in exI, auto)
apply(case_tac mr, simp_all add: exp_ind_def)
apply(case_tac rn, simp_all add: exp_ind_def)
done

lemma [simp]: "wadjust_loop_erase m rs (c, Oc # list) \<Longrightarrow> 
                wadjust_loop_erase m rs (c, Bk # list)"
apply(auto simp: wadjust_loop_erase.simps)
done

lemma [simp]: "wadjust_loop_on_left_moving_B m rs (c, Oc # list) = False"
apply(auto simp: wadjust_loop_on_left_moving_B.simps)
apply(case_tac nr, simp_all add: exp_ind_def)
done

lemma [simp]: "wadjust_loop_on_left_moving m rs (c, Oc # list)
           \<Longrightarrow> wadjust_loop_right_move2 m rs (Oc # c, list)"
apply(simp add:wadjust_loop_on_left_moving.simps)
apply(auto simp: wadjust_loop_on_left_moving_O.simps
                 wadjust_loop_right_move2.simps)
done

lemma [simp]: "wadjust_loop_right_move2 m rs (c, Oc # list) = False"
apply(auto simp: wadjust_loop_right_move2.simps )
apply(case_tac ln, simp_all add: exp_ind_def)
done

lemma [simp]: "wadjust_erase2 m rs (c, Oc # list)
              \<Longrightarrow> (c = [] \<longrightarrow> wadjust_erase2 m rs ([], Bk # list))
               \<and> (c \<noteq> [] \<longrightarrow> wadjust_erase2 m rs (c, Bk # list))"
apply(auto simp: wadjust_erase2.simps )
done

lemma [simp]: "wadjust_on_left_moving_B m rs (c, Oc # list) = False"
apply(auto simp: wadjust_on_left_moving_B.simps)
done

lemma [simp]: "\<lbrakk>wadjust_on_left_moving_O m rs (c, Oc # list); hd c = Bk\<rbrakk> \<Longrightarrow> 
         wadjust_goon_left_moving_B m rs (tl c, Bk # Oc # list)"
apply(auto simp: wadjust_on_left_moving_O.simps 
     wadjust_goon_left_moving_B.simps exp_ind_def)
done

lemma [simp]: "\<lbrakk>wadjust_on_left_moving_O m rs (c, Oc # list); hd c = Oc\<rbrakk>
    \<Longrightarrow> wadjust_goon_left_moving_O m rs (tl c, Oc # Oc # list)"
apply(auto simp: wadjust_on_left_moving_O.simps 
                 wadjust_goon_left_moving_O.simps exp_ind_def)
apply(rule_tac x = rs in exI, simp)
apply(auto simp: exp_ind_def numeral_2_eq_2)
done


lemma [simp]: "wadjust_on_left_moving m rs (c, Oc # list) \<Longrightarrow> 
              wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)"
apply(simp add: wadjust_on_left_moving.simps   
                 wadjust_goon_left_moving.simps)
apply(case_tac "hd c", simp_all)
done

lemma [simp]: "wadjust_on_left_moving m rs (c, Oc # list) \<Longrightarrow> 
  wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)"
apply(simp add: wadjust_on_left_moving.simps 
  wadjust_goon_left_moving.simps)
apply(case_tac "hd c", simp_all)
done

lemma [simp]: "wadjust_goon_left_moving_B m rs (c, Oc # list) = False"
apply(auto simp: wadjust_goon_left_moving_B.simps)
done

lemma [simp]: "\<lbrakk>wadjust_goon_left_moving_O m rs (c, Oc # list); hd c = Bk\<rbrakk> 
               \<Longrightarrow> wadjust_goon_left_moving_B m rs (tl c, Bk # Oc # list)"
apply(auto simp: wadjust_goon_left_moving_O.simps wadjust_goon_left_moving_B.simps)
apply(case_tac [!] ml, auto simp: exp_ind_def)
done

lemma  [simp]: "\<lbrakk>wadjust_goon_left_moving_O m rs (c, Oc # list); hd c = Oc\<rbrakk> \<Longrightarrow> 
  wadjust_goon_left_moving_O m rs (tl c, Oc # Oc # list)"
apply(auto simp: wadjust_goon_left_moving_O.simps wadjust_goon_left_moving_B.simps)
apply(rule_tac x = "ml - 1" in exI, simp)
apply(case_tac ml, simp_all add: exp_ind_def)
apply(rule_tac x = "Suc mr" in exI, auto simp: exp_ind_def)
done

lemma [simp]: "wadjust_goon_left_moving m rs (c, Oc # list) \<Longrightarrow> 
  wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)"
apply(simp add: wadjust_goon_left_moving.simps)
apply(case_tac "hd c", simp_all)
done

lemma [simp]: "wadjust_backto_standard_pos_B m rs (c, Oc # list) = False"
apply(simp add: wadjust_backto_standard_pos_B.simps)
done

lemma [simp]: "wadjust_backto_standard_pos_O m rs (c, Bk # xs) = False"
apply(simp add: wadjust_backto_standard_pos_O.simps, auto)
apply(case_tac mr, simp_all add: exp_ind_def)
done



lemma [simp]: "wadjust_backto_standard_pos_O m rs ([], Oc # list) \<Longrightarrow> 
  wadjust_backto_standard_pos_B m rs ([], Bk # Oc # list)"
apply(auto simp: wadjust_backto_standard_pos_O.simps
                 wadjust_backto_standard_pos_B.simps)
apply(rule_tac x = rn in exI, simp)
apply(case_tac ml, simp_all add: exp_ind_def)
done


lemma [simp]: 
  "\<lbrakk>wadjust_backto_standard_pos_O m rs (c, Oc # list); c \<noteq> []; hd c = Bk\<rbrakk>
  \<Longrightarrow> wadjust_backto_standard_pos_B m rs (tl c, Bk # Oc # list)"
apply(simp add:wadjust_backto_standard_pos_O.simps 
        wadjust_backto_standard_pos_B.simps, auto)
apply(case_tac [!] ml, simp_all add: exp_ind_def)
done 

lemma [simp]: "\<lbrakk>wadjust_backto_standard_pos_O m rs (c, Oc # list); c \<noteq> []; hd c = Oc\<rbrakk>
          \<Longrightarrow>  wadjust_backto_standard_pos_O m rs (tl c, Oc # Oc # list)"
apply(simp add: wadjust_backto_standard_pos_O.simps, auto)
apply(case_tac ml, simp_all add: exp_ind_def, auto)
apply(rule_tac x = nat in exI, auto simp: exp_ind_def)
done

lemma [simp]: "wadjust_backto_standard_pos m rs (c, Oc # list)
  \<Longrightarrow> (c = [] \<longrightarrow> wadjust_backto_standard_pos m rs ([], Bk # Oc # list)) \<and> 
 (c \<noteq> [] \<longrightarrow> wadjust_backto_standard_pos m rs (tl c, hd c # Oc # list))"
apply(auto simp: wadjust_backto_standard_pos.simps)
apply(case_tac "hd c", simp_all)
done
thm wadjust_loop_right_move.simps

lemma [simp]: "wadjust_loop_right_move m rs (c, []) = False"
apply(simp only: wadjust_loop_right_move.simps)
apply(rule_tac iffI)
apply(erule_tac exE)+
apply(case_tac nr, simp_all add: exp_ind_def)
apply(case_tac mr, simp_all add: exp_ind_def)
done

lemma [simp]: "wadjust_loop_erase m rs (c, []) = False"
apply(simp only: wadjust_loop_erase.simps, auto)
apply(case_tac mr, simp_all add: exp_ind_def)
done

lemma [simp]: "\<lbrakk>Suc (Suc rs) = a;  wadjust_loop_erase m rs (c, Bk # list)\<rbrakk>
  \<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Bk # list))))
  < a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list)))) \<or>
  a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) =
  a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list))))"
apply(simp only: wadjust_loop_erase.simps)
apply(rule_tac disjI2)
apply(case_tac c, simp, simp)
done

lemma [simp]:
  "\<lbrakk>Suc (Suc rs) = a;  wadjust_loop_on_left_moving m rs (c, Bk # list)\<rbrakk>
  \<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Bk # list))))
  < a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list)))) \<or>
  a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) =
  a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list))))"
apply(subgoal_tac "c \<noteq> []")
apply(case_tac c, simp_all)
done

lemma dropWhile_exp1: "dropWhile (\<lambda>a. a = Oc) (Oc\<^bsup>n\<^esup> @ xs) = dropWhile (\<lambda>a. a = Oc) xs"
apply(induct n, simp_all add: exp_ind_def)
done
lemma takeWhile_exp1: "takeWhile (\<lambda>a. a = Oc) (Oc\<^bsup>n\<^esup> @ xs) = Oc\<^bsup>n\<^esup> @ takeWhile (\<lambda>a. a = Oc) xs"
apply(induct n, simp_all add: exp_ind_def)
done

lemma [simp]: "\<lbrakk>Suc (Suc rs) = a;  wadjust_loop_right_move2 m rs (c, Bk # list)\<rbrakk>
              \<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list))))
                 < a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list))))"
apply(simp add: wadjust_loop_right_move2.simps, auto)
apply(simp add: dropWhile_exp1 takeWhile_exp1)
apply(case_tac ln, simp, simp add: exp_ind_def)
done

lemma [simp]: "wadjust_loop_check m rs ([], b) = False"
apply(simp add: wadjust_loop_check.simps)
done

lemma [simp]: "\<lbrakk>Suc (Suc rs) = a;  wadjust_loop_check m rs (c, Oc # list)\<rbrakk>
  \<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Oc # list))))
  < a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list)))) \<or>
  a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Oc # list)))) =
  a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list))))"
apply(case_tac "c", simp_all)
done

lemma [simp]: 
  "\<lbrakk>Suc (Suc rs) = a;  wadjust_loop_erase m rs (c, Oc # list)\<rbrakk>
  \<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list))))
  < a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list)))) \<or>
  a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list)))) =
  a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list))))"
apply(simp add: wadjust_loop_erase.simps)
apply(rule_tac disjI2)
apply(auto)
apply(simp add: dropWhile_exp1 takeWhile_exp1)
done

declare numeral_2_eq_2[simp del]

lemma wadjust_correctness:
  shows "let P = (\<lambda> (len, st, l, r). st = 0) in 
  let Q = (\<lambda> (len, st, l, r). wadjust_inv st m rs (l, r)) in 
  let f = (\<lambda> stp. (Suc (Suc rs),  steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, 
                Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk #  Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust stp)) in
    \<exists> n .P (f n) \<and> Q (f n)"
proof -
  let ?P = "(\<lambda> (len, st, l, r). st = 0)"
  let ?Q = "\<lambda> (len, st, l, r). wadjust_inv st m rs (l, r)"
  let ?f = "\<lambda> stp. (Suc (Suc rs),  steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, 
                Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust stp)"
  have "\<exists> n. ?P (?f n) \<and> ?Q (?f n)"
  proof(rule_tac halt_lemma2)
    show "wf wadjust_le" by auto
  next
    show "\<forall> n. \<not> ?P (?f n) \<and> ?Q (?f n) \<longrightarrow> 
                 ?Q (?f (Suc n)) \<and> (?f (Suc n), ?f n) \<in> wadjust_le"
    proof(rule_tac allI, rule_tac impI, case_tac "?f n", 
            simp add: tstep_red tstep.simps, rule_tac conjI, erule_tac conjE,
          erule_tac conjE)      
      fix n a b c d
      assume "0 < b" "wadjust_inv b m rs (c, d)" "Suc (Suc rs) = a"
      thus "case case fetch t_wcode_adjust b (case d of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)
        of (ac, ns) \<Rightarrow> (ns, new_tape ac (c, d)) of (st, x) \<Rightarrow> wadjust_inv st m rs x"
        apply(case_tac d, simp, case_tac [2] aa)
        apply(simp_all add: wadjust_inv.simps wadjust_le_def new_tape.simps
          abacus.lex_triple_def abacus.lex_pair_def lex_square_def
          split: if_splits)
        done
    next
      fix n a b c d
      assume "0 < b \<and> wadjust_inv b m rs (c, d)"
        "Suc (Suc rs) = a \<and> steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>,
         Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust n = (b, c, d)"
      thus "((a, case fetch t_wcode_adjust b (case d of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)
        of (ac, ns) \<Rightarrow> (ns, new_tape ac (c, d))), a, b, c, d) \<in> wadjust_le"
      proof(erule_tac conjE, erule_tac conjE, erule_tac conjE)
        assume "0 < b" "wadjust_inv b m rs (c, d)" "Suc (Suc rs) = a"
        thus "?thesis"
          apply(case_tac d, case_tac [2] aa)
          apply(simp_all add: wadjust_inv.simps wadjust_le_def new_tape.simps
            abacus.lex_triple_def abacus.lex_pair_def lex_square_def
            split: if_splits)
          done
      qed
    qed
  next
    show "?Q (?f 0)"
      apply(simp add: steps.simps wadjust_inv.simps wadjust_start.simps)
      apply(rule_tac x = ln in exI,auto)
      done
  next
    show "\<not> ?P (?f 0)"
      apply(simp add: steps.simps)
      done
  qed
  thus "?thesis"
    apply(auto)
    done
qed

lemma [intro]: "t_correct t_wcode_adjust"
apply(auto simp: t_wcode_adjust_def t_correct.simps iseven_def)
apply(rule_tac x = 11 in exI, simp)
done

lemma wcode_lemma_pre':
  "args \<noteq> [] \<Longrightarrow> 
  \<exists> stp rn. steps (Suc 0, [], <m # args>) 
              ((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) stp
  = (0,  [Bk],  Oc\<^bsup>Suc m\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>)" 
proof -
  let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <m # args>"
  let ?Q1 = "\<lambda>(l, r). l = Bk # Oc\<^bsup>Suc m\<^esup> \<and>
    (\<exists>ln rn. r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  let ?P2 = ?Q1
  let ?Q2 = "\<lambda> (l, r). (wadjust_stop m (bl_bin (<args>) - 1) (l, r))"
  let ?P3 = "\<lambda> tp. False"
  assume h: "args \<noteq> []"
  have "?P1 \<turnstile>-> \<lambda> tp. (\<exists> stp tp'. steps (Suc 0, tp)
                      ((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) stp = (0, tp') \<and> ?Q2 tp')"
  proof(rule_tac turing_merge.t_merge_halt[of "t_wcode_prepare |+| t_wcode_main" 
               t_wcode_adjust ?P1 ?P2 ?P3 ?P3 ?Q1 ?Q2], 
        auto simp: turing_merge_def)

    show "\<exists>stp. case steps (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) stp of
          (st, tp') \<Rightarrow> st = 0 \<and> (case tp' of (l, r) \<Rightarrow> l = Bk # Oc\<^bsup>Suc m\<^esup> \<and>
                (\<exists>ln rn. r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>))"
      using h prepare_mainpart_lemma[of args m]
      apply(auto)
      apply(rule_tac x = stp in exI, simp)
      apply(rule_tac x = ln in exI, auto)
      done
  next
    fix ln rn
    show "\<exists>stp. case steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # 
                               Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust stp of
      (st, tp') \<Rightarrow> st = 0 \<and> wadjust_stop m (bl_bin (<args>) - Suc 0) tp'"
      using wadjust_correctness[of m "bl_bin (<args>) - 1" "Suc ln" rn]
      apply(subgoal_tac "bl_bin (<args>) > 0", auto simp: wadjust_inv.simps)
      apply(rule_tac x = n in exI, simp add: exp_ind)
      using h
      apply(case_tac args, simp_all, case_tac list,
            simp_all add: tape_of_nl_abv  tape_of_nat_list.simps exp_ind_def
            bl_bin.simps)
      done     
  next
    show "?Q1 \<turnstile>-> ?P2"
      by(simp add: t_imply_def)
  qed
  thus "\<exists>stp rn. steps (Suc 0, [], <m # args>) ((t_wcode_prepare |+| t_wcode_main) |+| 
        t_wcode_adjust) stp = (0, [Bk], Oc\<^bsup>Suc m\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>)"
    apply(simp add: t_imply_def)
    apply(erule_tac exE)+
    apply(subgoal_tac "bl_bin (<args>) > 0", auto simp: wadjust_stop.simps)
    using h
    apply(case_tac args, simp_all, case_tac list,  
          simp_all add: tape_of_nl_abv  tape_of_nat_list.simps exp_ind_def
            bl_bin.simps)
    done
qed

text {*
  The initialization TM @{text "t_wcode"}.
  *}
definition t_wcode :: "tprog"
  where
  "t_wcode = (t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust"


text {*
  The correctness of @{text "t_wcode"}.
  *}
lemma wcode_lemma_1:
  "args \<noteq> [] \<Longrightarrow> 
  \<exists> stp ln rn. steps (Suc 0, [], <m # args>)  (t_wcode) stp = 
              (0,  [Bk],  Oc\<^bsup>Suc m\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>)"
apply(simp add: wcode_lemma_pre' t_wcode_def)
done

lemma wcode_lemma: 
  "args \<noteq> [] \<Longrightarrow> 
  \<exists> stp ln rn. steps (Suc 0, [], <m # args>)  (t_wcode) stp = 
              (0,  [Bk],  <[m ,bl_bin (<args>)]> @ Bk\<^bsup>rn\<^esup>)"
using wcode_lemma_1[of args m]
apply(simp add: t_wcode_def tape_of_nl_abv tape_of_nat_list.simps)
done

section {* The universal TM *}

text {*
  This section gives the explicit construction of {\em Universal Turing Machine}, defined as @{text "UTM"} and proves its 
  correctness. It is pretty easy by composing the partial results we have got so far.
  *}


definition UTM :: "tprog"
  where
  "UTM = (let (aprog, rs_pos, a_md) = rec_ci rec_F in 
          let abc_F = aprog [+] dummy_abc (Suc (Suc 0)) in 
          (t_wcode |+| (tm_of abc_F @ tMp (Suc (Suc 0)) (start_of (layout_of abc_F) 
                                                   (length abc_F) - Suc 0))))"

definition F_aprog :: "abc_prog"
  where
  "F_aprog \<equiv> (let (aprog, rs_pos, a_md) = rec_ci rec_F in 
                       aprog [+] dummy_abc (Suc (Suc 0)))"

definition F_tprog :: "tprog"
  where
  "F_tprog = tm_of (F_aprog)"

definition t_utm :: "tprog"
  where
  "t_utm \<equiv>
     (F_tprog) @ tMp (Suc (Suc 0)) (start_of (layout_of (F_aprog)) 
                                  (length (F_aprog)) - Suc 0)"

definition UTM_pre :: "tprog"
  where
  "UTM_pre = t_wcode |+| t_utm"

lemma F_abc_halt_eq:
  "\<lbrakk>turing_basic.t_correct tp; 
    length lm = k;
    steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>);
    rs > 0\<rbrakk>
    \<Longrightarrow> \<exists> stp m. abc_steps_l (0, [code tp, bl2wc (<lm>)]) (F_aprog) stp =
                       (length (F_aprog), code tp # bl2wc (<lm>) # (rs - 1) # 0\<^bsup>m\<^esup>)"
apply(drule_tac  F_t_halt_eq, simp, simp, simp)
apply(case_tac "rec_ci rec_F")
apply(frule_tac abc_append_dummy_complie, simp, simp, erule_tac exE,
      erule_tac exE)
apply(rule_tac x = stp in exI, rule_tac x = m in exI)
apply(simp add: F_aprog_def dummy_abc_def)
done

lemma F_abc_utm_halt_eq: 
  "\<lbrakk>rs > 0; 
  abc_steps_l (0, [code tp, bl2wc (<lm>)]) F_aprog stp =
        (length F_aprog, code tp #  bl2wc (<lm>) # (rs - 1) # 0\<^bsup>m\<^esup>)\<rbrakk>
  \<Longrightarrow> \<exists>stp m n.(steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp =
                                             (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>))"
  thm abacus_turing_eq_halt
  using abacus_turing_eq_halt
  [of "layout_of F_aprog" "F_aprog" "F_tprog" "length (F_aprog)" 
    "[code tp, bl2wc (<lm>)]" stp "code tp # bl2wc (<lm>) # (rs - 1) # 0\<^bsup>m\<^esup>" "Suc (Suc 0)"
    "start_of (layout_of (F_aprog)) (length (F_aprog))" "[]" 0]
apply(simp add: F_tprog_def t_utm_def abc_lm_v.simps nth_append)
apply(erule_tac exE)+
apply(rule_tac x = stpa in exI, rule_tac x = "Suc (Suc ma)" in exI, 
       rule_tac x = l in exI, simp add: exp_ind)
done

declare tape_of_nl_abv_cons[simp del]

lemma t_utm_halt_eq': 
  "\<lbrakk>turing_basic.t_correct tp;
   0 < rs;
  steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>)\<rbrakk>
  \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp = 
                                                (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
apply(drule_tac  l = l in F_abc_halt_eq, simp, simp, simp)
apply(erule_tac exE, erule_tac exE)
apply(rule_tac F_abc_utm_halt_eq, simp_all)
done

lemma [simp]: "tinres xs (xs @ Bk\<^bsup>i\<^esup>)"
apply(auto simp: tinres_def)
done

lemma [elim]: "\<lbrakk>rs > 0; Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup> = c @ Bk\<^bsup>n\<^esup>\<rbrakk>
        \<Longrightarrow> \<exists>n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>"
apply(case_tac "na > n")
apply(subgoal_tac "\<exists> d. na = d + n", auto simp: exp_add)
apply(rule_tac x = "na - n" in exI, simp)
apply(subgoal_tac "\<exists> d. n = d + na", auto simp: exp_add)
apply(case_tac rs, simp_all add: exp_ind, case_tac d, 
           simp_all add: exp_ind)
apply(rule_tac x = "n - na" in exI, simp)
done


lemma t_utm_halt_eq'': 
  "\<lbrakk>turing_basic.t_correct tp;
   0 < rs;
   steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>)\<rbrakk>
  \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stp = 
                                                (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
apply(drule_tac t_utm_halt_eq', simp_all)
apply(erule_tac exE)+
proof -
  fix stpa ma na
  assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stpa = (0, Bk\<^bsup>ma\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>)"
  and gr: "rs > 0"
  thus "\<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
    apply(rule_tac x = stpa in exI, rule_tac x = ma in exI,  simp)
  proof(case_tac "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa", simp)
    fix a b c
    assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stpa = (0, Bk\<^bsup>ma\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>)"
            "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (a, b, c)"
    thus " a = 0 \<and> b = Bk\<^bsup>ma\<^esup> \<and> (\<exists>n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
      using tinres_steps2[of "<[code tp, bl2wc (<lm>)]>" "<[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>" 
                           "Suc 0" " [Bk, Bk]" t_utm stpa 0 "Bk\<^bsup>ma\<^esup>" "Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>" a b c]
      apply(simp)
      using gr
      apply(simp only: tinres_def, auto)
      apply(rule_tac x = "na + n" in exI, simp add: exp_add)
      done
  qed
qed

lemma [simp]: "tinres [Bk, Bk] [Bk]"
apply(auto simp: tinres_def)
done

lemma [elim]: "Bk\<^bsup>ma\<^esup> = b @ Bk\<^bsup>n\<^esup>  \<Longrightarrow> \<exists>m. b = Bk\<^bsup>m\<^esup>"
apply(subgoal_tac "ma = length b + n")
apply(rule_tac x = "ma - n" in exI, simp add: exp_add)
apply(drule_tac length_equal)
apply(simp)
done

lemma t_utm_halt_eq: 
  "\<lbrakk>turing_basic.t_correct tp;
   0 < rs;
   steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>)\<rbrakk>
  \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stp = 
                                                (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
apply(drule_tac i = i in t_utm_halt_eq'', simp_all)
apply(erule_tac exE)+
proof -
  fix stpa ma na
  assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (0, Bk\<^bsup>ma\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>)"
  and gr: "rs > 0"
  thus "\<exists>stp m n. steps (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
    apply(rule_tac x = stpa in exI)
  proof(case_tac "steps (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa", simp)
    fix a b c
    assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (0, Bk\<^bsup>ma\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>)"
            "steps (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (a, b, c)"
    thus "a = 0 \<and> (\<exists>m. b = Bk\<^bsup>m\<^esup>) \<and> (\<exists>n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
      using tinres_steps[of "[Bk, Bk]" "[Bk]" "Suc 0" "<[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>" t_utm stpa 0
                             "Bk\<^bsup>ma\<^esup>" "Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>" a b c]
      apply(simp)
      apply(auto simp: tinres_def)
      apply(rule_tac x = "ma + n" in exI, simp add: exp_add)
      done
  qed
qed

lemma [intro]: "t_correct t_wcode"
apply(simp add: t_wcode_def)
apply(auto)
done
      
lemma [intro]: "t_correct t_utm"
apply(simp add: t_utm_def F_tprog_def)
apply(rule_tac t_compiled_correct, auto)
done   

lemma UTM_halt_lemma_pre: 
  "\<lbrakk>turing_basic.t_correct tp;
   0 < rs;
   args \<noteq> [];
   steps (Suc 0, Bk\<^bsup>i\<^esup>, <args::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>k\<^esup>)\<rbrakk>
  \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [], <code tp # args>) UTM_pre stp = 
                                                (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
proof -
  let ?Q2 = "\<lambda> (l, r). (\<exists> ln rn. l = Bk\<^bsup>ln\<^esup> \<and> r = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  term ?Q2
  let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <code tp # args>"
  let ?Q1 = "\<lambda> (l, r). (l = [Bk] \<and>
             (\<exists> rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>))"
  let ?P2 = ?Q1
  let ?P3 = "\<lambda> (l, r). False"
  assume h: "turing_basic.t_correct tp" "0 < rs"
            "args \<noteq> []" "steps (Suc 0, Bk\<^bsup>i\<^esup>, <args::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>k\<^esup>)"
  have "?P1 \<turnstile>-> \<lambda> tp. (\<exists> stp tp'. steps (Suc 0, tp)
                    (t_wcode |+| t_utm) stp = (0, tp') \<and> ?Q2 tp')"
  proof(rule_tac turing_merge.t_merge_halt [of "t_wcode" "t_utm"
          ?P1 ?P2 ?P3 ?P3 ?Q1 ?Q2], auto simp: turing_merge_def)
    show "\<exists>stp. case steps (Suc 0, [], <code tp # args>) t_wcode stp of (st, tp') \<Rightarrow> 
       st = 0 \<and> (case tp' of (l, r) \<Rightarrow> l = [Bk] \<and>
                   (\<exists>rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>))"
      using wcode_lemma_1[of args "code tp"] h
      apply(simp, auto)
      apply(rule_tac x = stpa in exI, auto)
      done      
  next
    fix rn 
    show "\<exists>stp. case steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @
      Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp of
      (st, tp') \<Rightarrow> st = 0 \<and> (case tp' of (l, r) \<Rightarrow>
      (\<exists>ln. l = Bk\<^bsup>ln\<^esup>) \<and> (\<exists>rn. r = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>rn\<^esup>))"
      using t_utm_halt_eq[of tp rs i args stp m k rn] h
      apply(auto)
      apply(rule_tac x = stpa in exI, simp add: bin_wc_eq 
        tape_of_nat_list.simps tape_of_nl_abv)
      apply(auto)
      done
  next
    show "?Q1 \<turnstile>-> ?P2"
      apply(simp add: t_imply_def)
      done
  qed
  thus "?thesis"
    apply(simp add: t_imply_def)
    apply(auto simp: UTM_pre_def)
    done
qed

text {*
  The correctness of @{text "UTM"}, the halt case.
*}
lemma UTM_halt_lemma: 
  "\<lbrakk>turing_basic.t_correct tp;
   0 < rs;
   args \<noteq> [];
   steps (Suc 0, Bk\<^bsup>i\<^esup>, <args::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>k\<^esup>)\<rbrakk>
  \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [], <code tp # args>) UTM stp = 
                                                (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
using UTM_halt_lemma_pre[of tp rs args i stp m k]
apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def)
apply(case_tac "rec_ci rec_F", simp)
done

definition TSTD:: "t_conf \<Rightarrow> bool"
  where
  "TSTD c = (let (st, l, r) = c in 
             st = 0 \<and> (\<exists> m. l = Bk\<^bsup>m\<^esup>) \<and> (\<exists> rs n. r = Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>))"

thm abacus_turing_eq_uhalt

lemma nstd_case1: "0 < a \<Longrightarrow> NSTD (trpl_code (a, b, c))"
apply(simp add: NSTD.simps trpl_code.simps)
done

lemma [simp]: "\<forall>m. b \<noteq> Bk\<^bsup>m\<^esup> \<Longrightarrow> 0 < bl2wc b"
apply(rule classical, simp)
apply(induct b, erule_tac x = 0 in allE, simp)
apply(simp add: bl2wc.simps, case_tac a, simp_all 
  add: bl2nat.simps bl2nat_double)
apply(case_tac "\<exists> m. b = Bk\<^bsup>m\<^esup>",  erule exE)
apply(erule_tac x = "Suc m" in allE, simp add: exp_ind_def, simp)
done
lemma nstd_case2: "\<forall>m. b \<noteq> Bk\<^bsup>m\<^esup> \<Longrightarrow> NSTD (trpl_code (a, b, c))"
apply(simp add: NSTD.simps trpl_code.simps)
done

thm lg.simps
thm lgR.simps

lemma [elim]: "Suc (2 * x) = 2 * y \<Longrightarrow> RR"
apply(induct x arbitrary: y, simp, simp)
apply(case_tac y, simp, simp)
done

lemma bl2nat_zero_eq[simp]: "(bl2nat c 0 = 0) = (\<exists>n. c = Bk\<^bsup>n\<^esup>)"
apply(auto)
apply(induct c, simp add: bl2nat.simps)
apply(rule_tac x = 0 in exI, simp)
apply(case_tac a, auto simp: bl2nat.simps bl2nat_double)
done

lemma bl2wc_exp_ex: 
  "\<lbrakk>Suc (bl2wc c) = 2 ^  m\<rbrakk> \<Longrightarrow> \<exists> rs n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>"
apply(induct c arbitrary: m, simp add: bl2wc.simps bl2nat.simps)
apply(case_tac a, auto)
apply(case_tac m, simp_all add: bl2wc.simps, auto)
apply(rule_tac x = 0 in exI, rule_tac x = "Suc n" in exI, 
  simp add: exp_ind_def)
apply(simp add: bl2wc.simps bl2nat.simps bl2nat_double)
apply(case_tac m, simp, simp)
proof -
  fix c m nat
  assume ind: 
    "\<And>m. Suc (bl2nat c 0) = 2 ^ m \<Longrightarrow> \<exists>rs n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>"
  and h: 
    "Suc (Suc (2 * bl2nat c 0)) = 2 * 2 ^ nat"
  have "\<exists>rs n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>"
    apply(rule_tac m = nat in ind)
    using h
    apply(simp)
    done
  from this obtain rs n where " c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>" by blast 
  thus "\<exists>rs n. Oc # c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>"
    apply(rule_tac x = "Suc rs" in exI, simp add: exp_ind_def)
    apply(rule_tac x = n in exI, simp)
    done
qed

lemma [elim]: 
  "\<lbrakk>\<forall>rs n. c \<noteq> Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>; 
  bl2wc c = 2 ^ lg (Suc (bl2wc c)) 2 - Suc 0\<rbrakk> \<Longrightarrow> bl2wc c = 0"
apply(subgoal_tac "\<exists> m. Suc (bl2wc c) = 2^m", erule_tac exE)
apply(drule_tac bl2wc_exp_ex, simp, erule_tac exE, erule_tac exE)
apply(case_tac rs, simp, simp, erule_tac x = nat in allE,
  erule_tac x = n in allE, simp)
using bl2wc_exp_ex[of c "lg (Suc (bl2wc c)) 2"]
apply(case_tac "(2::nat) ^ lg (Suc (bl2wc c)) 2", 
  simp, simp, erule_tac exE, erule_tac exE, simp)
apply(simp add: bl2wc.simps)
apply(rule_tac x = rs in exI)
apply(case_tac "(2::nat)^rs", simp, simp)
done

lemma nstd_case3: 
  "\<forall>rs n. c \<noteq> Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup> \<Longrightarrow>  NSTD (trpl_code (a, b, c))"
apply(simp add: NSTD.simps trpl_code.simps)
apply(rule_tac impI)
apply(rule_tac disjI2, rule_tac disjI2, auto)
done

lemma NSTD_1: "\<not> TSTD (a, b, c)
    \<Longrightarrow> rec_exec rec_NSTD [trpl_code (a, b, c)] = Suc 0"
  using NSTD_lemma1[of "trpl_code (a, b, c)"]
       NSTD_lemma2[of "trpl_code (a, b, c)"]
  apply(simp add: TSTD_def)
  apply(erule_tac disjE, erule_tac nstd_case1)
  apply(erule_tac disjE, erule_tac nstd_case2)
  apply(erule_tac nstd_case3)
  done
 
lemma nonstop_t_uhalt_eq:
      "\<lbrakk>turing_basic.t_correct tp;
        steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (a, b, c);
       \<not> TSTD (a, b, c)\<rbrakk>
       \<Longrightarrow> rec_exec rec_nonstop [code tp, bl2wc (<lm>), stp] = Suc 0"
apply(simp add: rec_nonstop_def rec_exec.simps)
apply(subgoal_tac 
  "rec_exec rec_conf [code tp, bl2wc (<lm>), stp] =
  trpl_code (a, b, c)", simp)
apply(erule_tac NSTD_1)
using rec_t_eq_steps[of tp l lm stp]
apply(simp)
done

lemma nonstop_true:
  "\<lbrakk>turing_basic.t_correct tp;
  \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp))\<rbrakk>
     \<Longrightarrow> \<forall>y. rec_calc_rel rec_nonstop 
                        ([code tp, bl2wc (<lm>), y]) (Suc 0)"
apply(rule_tac allI, erule_tac x = y in allE)
apply(case_tac "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp y", simp)
apply(rule_tac nonstop_t_uhalt_eq, simp_all)
done

(*
lemma [simp]: 
  "\<forall>j<Suc k. Ex (rec_calc_rel (get_fstn_args (Suc k) (Suc k) ! j)
                                                     (code tp # lm))"
apply(auto simp: get_fstn_args_nth)
apply(rule_tac x = "(code tp # lm) ! j" in exI)
apply(rule_tac calc_id, simp_all)
done
*)
declare ci_cn_para_eq[simp]

lemma F_aprog_uhalt: 
  "\<lbrakk>turing_basic.t_correct tp; 
    \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp)); 
    rec_ci rec_F = (F_ap, rs_pos, a_md)\<rbrakk>
  \<Longrightarrow> \<forall> stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)] @ 0\<^bsup>a_md - rs_pos \<^esup>
               @ suflm) (F_ap) stp of (ss, e) \<Rightarrow> ss < length (F_ap)"
apply(case_tac "rec_ci (Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) rec_conf 
               ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])])")
apply(simp only: rec_F_def, rule_tac i = 0  and ga = a and gb = b and 
  gc = c in cn_gi_uhalt, simp, simp, simp, simp, simp, simp, simp)
apply(simp add: ci_cn_para_eq)
apply(case_tac "rec_ci (Cn (Suc (Suc 0)) rec_conf 
  ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt]))")
apply(rule_tac rf = "(Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) rec_conf
              ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])])" 
           and n = "Suc (Suc 0)" and f = rec_right and 
          gs = "[Cn (Suc (Suc 0)) rec_conf 
           ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])]"
           and i = 0 and ga = aa and gb = ba and gc = ca in 
          cn_gi_uhalt)
apply(simp, simp, simp, simp, simp, simp, simp, 
     simp add: ci_cn_para_eq)
apply(case_tac "rec_ci rec_halt")
apply(rule_tac rf = "(Cn (Suc (Suc 0)) rec_conf 
  ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt]))" 
  and n = "Suc (Suc 0)" and f = "rec_conf" and 
  gs = "([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])"  and 
  i = "Suc (Suc 0)" and gi = "rec_halt" and ga = ab and gb = bb and
  gc = cb in cn_gi_uhalt)
apply(simp, simp, simp, simp, simp add: nth_append, simp, 
  simp add: nth_append, simp add: rec_halt_def)
apply(simp only: rec_halt_def)
apply(case_tac [!] "rec_ci ((rec_nonstop))")
apply(rule_tac allI, rule_tac impI, simp)
apply(case_tac j, simp)
apply(rule_tac x = "code tp" in exI, rule_tac calc_id, simp, simp, simp, simp)
apply(rule_tac x = "bl2wc (<lm>)" in exI, rule_tac calc_id, simp, simp, simp)
apply(rule_tac rf = "Mn (Suc (Suc 0)) (rec_nonstop)"
  and f = "(rec_nonstop)" and n = "Suc (Suc 0)"
  and  aprog' = ac and rs_pos' =  bc and a_md' = cc in Mn_unhalt)
apply(simp, simp add: rec_halt_def , simp, simp)
apply(drule_tac  nonstop_true, simp_all)
apply(rule_tac allI)
apply(erule_tac x = y in allE)+
apply(simp)
done

thm abc_list_crsp_steps

lemma uabc_uhalt': 
  "\<lbrakk>turing_basic.t_correct tp;
  \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp));
  rec_ci rec_F = (ap, pos, md)\<rbrakk>
  \<Longrightarrow> \<forall> stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)]) ap stp of (ss, e)
           \<Rightarrow>  ss < length ap"
proof(frule_tac F_ap = ap and rs_pos = pos and a_md = md
    and suflm = "[]" in F_aprog_uhalt, auto)
  fix stp a b
  assume h: 
    "\<forall>stp. case abc_steps_l (0, code tp # bl2wc (<lm>) # 0\<^bsup>md - pos\<^esup>) ap stp of 
    (ss, e) \<Rightarrow> ss < length ap"
    "abc_steps_l (0, [code tp, bl2wc (<lm>)]) ap stp = (a, b)" 
    "turing_basic.t_correct tp" 
    "rec_ci rec_F = (ap, pos, md)"
  moreover have "ap \<noteq> []"
    using h apply(rule_tac rec_ci_not_null, simp)
    done
  ultimately show "a < length ap"
  proof(erule_tac x = stp in allE,
  case_tac "abc_steps_l (0, code tp # bl2wc (<lm>) # 0\<^bsup>md - pos\<^esup>) ap stp", simp)
    fix aa ba
    assume g: "aa < length ap" 
      "abc_steps_l (0, code tp # bl2wc (<lm>) # 0\<^bsup>md - pos\<^esup>) ap stp = (aa, ba)" 
      "ap \<noteq> []"
    thus "?thesis"
      using abc_list_crsp_steps[of "[code tp, bl2wc (<lm>)]"
                                   "md - pos" ap stp aa ba] h
      apply(simp)
      done
  qed
qed

lemma uabc_uhalt: 
  "\<lbrakk>turing_basic.t_correct tp; 
  \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp))\<rbrakk>
  \<Longrightarrow> \<forall> stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)]) F_aprog 
       stp of (ss, e) \<Rightarrow> ss < length F_aprog"
apply(case_tac "rec_ci rec_F", simp add: F_aprog_def)
thm uabc_uhalt'
apply(drule_tac ap = a and pos = b and md = c in uabc_uhalt', simp_all)
proof -
  fix a b c
  assume 
    "\<forall>stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)]) a stp of (ss, e) 
                                                   \<Rightarrow> ss < length a"
    "rec_ci rec_F = (a, b, c)"
  thus 
    "\<forall>stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)]) 
    (a [+] dummy_abc (Suc (Suc 0))) stp of (ss, e) \<Rightarrow> 
           ss < Suc (Suc (Suc (length a)))"
    using abc_append_uhalt1[of a "[code tp, bl2wc (<lm>)]" 
      "a [+] dummy_abc (Suc (Suc 0))" "[]" "dummy_abc (Suc (Suc 0))"]  
    apply(simp)
    done
qed

lemma tutm_uhalt': 
  "\<lbrakk>turing_basic.t_correct tp;
    \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp))\<rbrakk>
  \<Longrightarrow> \<forall> stp. \<not> isS0 (steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp)"
  using abacus_turing_eq_uhalt[of "layout_of (F_aprog)" 
               "F_aprog" "F_tprog" "[code tp, bl2wc (<lm>)]" 
               "start_of (layout_of (F_aprog )) (length (F_aprog))" 
               "Suc (Suc 0)"]
apply(simp add: F_tprog_def)
apply(subgoal_tac "\<forall>stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)])
  (F_aprog) stp of (as, am) \<Rightarrow> as < length (F_aprog)", simp)
thm abacus_turing_eq_uhalt
apply(simp add: t_utm_def F_tprog_def)
apply(rule_tac uabc_uhalt, simp_all)
done

lemma tinres_commute: "tinres r r' \<Longrightarrow> tinres r' r"
apply(auto simp: tinres_def)
done

lemma inres_tape:
  "\<lbrakk>steps (st, l, r) tp stp = (a, b, c); steps (st, l', r') tp stp = (a', b', c'); 
  tinres l l'; tinres r r'\<rbrakk>
  \<Longrightarrow> a = a' \<and> tinres b b' \<and> tinres c c'"
proof(case_tac "steps (st, l', r) tp stp")
  fix aa ba ca
  assume h: "steps (st, l, r) tp stp = (a, b, c)" 
            "steps (st, l', r') tp stp = (a', b', c')"
            "tinres l l'" "tinres r r'"
            "steps (st, l', r) tp stp = (aa, ba, ca)"
  have "tinres b ba \<and> c = ca \<and> a = aa"
    using h
    apply(rule_tac tinres_steps, auto)
    done

  thm tinres_steps2
  moreover have "b' = ba \<and> tinres c' ca \<and> a' =  aa"
    using h
    apply(rule_tac tinres_steps2, auto intro: tinres_commute)
    done
  ultimately show "?thesis"
    apply(auto intro: tinres_commute)
    done
qed

lemma tape_normalize: "\<forall> stp. \<not> isS0 (steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp)
      \<Longrightarrow> \<forall> stp. \<not> isS0 (steps (Suc 0, Bk\<^bsup>m\<^esup>, <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>n\<^esup>) t_utm stp)"
apply(rule_tac allI, case_tac "(steps (Suc 0, Bk\<^bsup>m\<^esup>, 
               <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>n\<^esup>) t_utm stp)", simp add: isS0_def)
apply(erule_tac x = stp in allE)
apply(case_tac "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp", simp)
apply(drule_tac inres_tape, auto)
apply(auto simp: tinres_def)
apply(case_tac "m > Suc (Suc 0)")
apply(rule_tac x = "m - Suc (Suc 0)" in exI) 
apply(case_tac m, simp_all add: exp_ind_def, case_tac nat, simp_all add: exp_ind_def)
apply(rule_tac x = "2 - m" in exI, simp add: exp_ind_def[THEN sym] exp_add[THEN sym])
apply(simp only: numeral_2_eq_2, simp add: exp_ind_def)
done

lemma tutm_uhalt: 
  "\<lbrakk>turing_basic.t_correct tp;
    \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <args>) tp stp))\<rbrakk>
  \<Longrightarrow> \<forall> stp. \<not> isS0 (steps (Suc 0, Bk\<^bsup>m\<^esup>, <[code tp, bl2wc (<args>)]> @ Bk\<^bsup>n\<^esup>) t_utm stp)"
apply(rule_tac tape_normalize)
apply(rule_tac tutm_uhalt', simp_all)
done

lemma UTM_uhalt_lemma_pre:
  "\<lbrakk>turing_basic.t_correct tp;
   \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <args>) tp stp));
   args \<noteq> []\<rbrakk>
  \<Longrightarrow>  \<forall> stp. \<not> isS0 (steps (Suc 0, [], <code tp # args>)  UTM_pre stp)"
proof -
  let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <code tp # args>"
  let ?Q1 = "\<lambda> (l, r). (l = [Bk] \<and>
             (\<exists> rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>))"
  let ?P4 = ?Q1
  let ?P3 = "\<lambda> (l, r). False"
  assume h: "turing_basic.t_correct tp" "\<forall>stp. \<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <args>) tp stp)"
            "args \<noteq> []"
  have "?P1 \<turnstile>-> \<lambda> tp. \<not> (\<exists> stp. isS0 (steps (Suc 0, tp) (t_wcode |+| t_utm) stp))"
  proof(rule_tac turing_merge.t_merge_uhalt [of "t_wcode" "t_utm"
          ?P1 ?P3 ?P3 ?P4 ?Q1 ?P3], auto simp: turing_merge_def)
    show "\<exists>stp. case steps (Suc 0, [], <code tp # args>) t_wcode stp of (st, tp') \<Rightarrow> 
       st = 0 \<and> (case tp' of (l, r) \<Rightarrow> l = [Bk] \<and>
                   (\<exists>rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>))"
      using wcode_lemma_1[of args "code tp"] h
      apply(simp, auto)
      apply(rule_tac x = stp in exI, auto)
      done      
  next
    fix rn  stp
    show " isS0 (steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp)
          \<Longrightarrow> False"
      using tutm_uhalt[of tp l args "Suc 0" rn] h
      apply(simp)
      apply(erule_tac x = stp in allE)
      apply(simp add: tape_of_nl_abv tape_of_nat_list.simps bin_wc_eq)
      done
  next
    fix rn stp
    show "isS0 (steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp) \<Longrightarrow>
      isS0 (steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp)"
      by simp
  next
    show "?Q1 \<turnstile>-> ?P4"
      apply(simp add: t_imply_def)
      done
  qed
  thus "?thesis"
    apply(simp add: t_imply_def UTM_pre_def)
    done
qed

text {*
  The correctness of @{text "UTM"}, the unhalt case.
  *}

lemma UTM_uhalt_lemma:
  "\<lbrakk>turing_basic.t_correct tp;
   \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <args>) tp stp));
   args \<noteq> []\<rbrakk>
  \<Longrightarrow>  \<forall> stp. \<not> isS0 (steps (Suc 0, [], <code tp # args>)  UTM stp)"
using UTM_uhalt_lemma_pre[of tp l args]
apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def)
apply(case_tac "rec_ci rec_F", simp)
done

end