utm/UTM.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Thu, 06 Dec 2012 16:32:03 +0000
changeset 373 0679a84b11ad
parent 371 48b231495281
permissions -rw-r--r--
added pip to a new repository
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
371
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
     1
theory UTM
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
     2
imports Main uncomputable recursive abacus UF GCD 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
     3
begin
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
     4
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
     5
section {* Wang coding of input arguments *}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
     6
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
     7
text {*
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
     8
  The direct compilation of the universal function @{text "rec_F"} can not give us UTM, because @{text "rec_F"} is of arity 2,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
     9
  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. 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    10
  (Notice, left number is always @{text "0"} at the very beginning). However, UTM needs to simulate the execution of any TM which may
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    11
  very well take many input arguments. Therefore, a initialization TM needs to run before the TM compiled from @{text "rec_F"}, and the sequential 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    12
  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 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    13
  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
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    14
  argument. 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    15
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    16
  However, this initialization TM (named @{text "t_wcode"}) can not be constructed by compiling from any resurve function, because every recursive 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    17
  function takes a fixed number of input arguments, while @{text "t_wcode"} needs to take varying number of arguments and tranform them into 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    18
  Wang's coding. Therefore, this section give a direct construction of @{text "t_wcode"} with just some parts being obtained from recursive functions.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    19
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    20
\newlength{\basewidth}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    21
\settowidth{\basewidth}{xxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxx}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    22
\newlength{\baseheight}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    23
\settoheight{\baseheight}{$B:R$}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    24
\newcommand{\vsep}{5\baseheight}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    25
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    26
The TM used to generate the Wang's code of input arguments is divided into three TMs
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    27
 executed sequentially, namely $prepare$, $mainwork$ and $adjust$¡£According to the
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    28
 convention, start state of ever TM is fixed to state $1$ while the final state is
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    29
 fixed to $0$.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    30
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    31
The input and output of $prepare$ are illustrated respectively by Figure
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    32
\ref{prepare_input} and \ref{prepare_output}.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    33
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    34
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    35
\begin{figure}[h!]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    36
\centering
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    37
\scalebox{1.2}{
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    38
\begin{tikzpicture}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    39
  [tbox/.style = {draw, thick, inner sep = 5pt}]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    40
  \node (0) {};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    41
  \node (1) [tbox, text height = 3.5pt, right = -0.9pt of 0] {\wuhao $m$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    42
  \node (2) [tbox, right = -0.9pt of 1] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    43
  \node (3) [tbox, right = -0.9pt of 2] {\wuhao $a_1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    44
  \node (4) [tbox, right = -0.9pt of 3] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    45
  \node (5) [tbox, right = -0.9pt of 4] {\wuhao $a_2$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    46
  \node (6) [right = -0.9pt of 5] {\ldots \ldots};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    47
  \node (7) [tbox, right = -0.9pt of 6] {\wuhao $a_n$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    48
  \draw [->, >=latex, thick] (1)+(0, -4\baseheight) -- (1);
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    49
\end{tikzpicture}}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    50
\caption{The input of TM $prepare$} \label{prepare_input}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    51
\end{figure}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    52
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    53
\begin{figure}[h!]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    54
\centering
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    55
\scalebox{1.2}{
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    56
\begin{tikzpicture}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    57
  \node (0) {};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    58
  \node (1) [draw, text height = 3.5pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    59
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    60
  \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    61
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $a_1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    62
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    63
  \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {\wuhao $a_2$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    64
  \node (7) [right = -0.9pt of 6] {\ldots \ldots};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    65
  \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {\wuhao $a_n$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    66
  \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    67
  \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    68
  \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {\wuhao $1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    69
  \draw [->, >=latex, thick] (10)+(0, -4\baseheight) -- (10);
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    70
\end{tikzpicture}}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    71
\caption{The output of TM $prepare$} \label{prepare_output}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    72
\end{figure}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    73
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    74
As shown in Figure \ref{prepare_input}, the input of $prepare$ is the same as the the input
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    75
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},
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    76
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,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    77
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}.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    78
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    79
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    80
\begin{figure}[h!]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    81
\centering
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    82
\scalebox{0.9}{
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    83
\begin{tikzpicture}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    84
     \node[circle,draw] (1) {$1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    85
     \node[circle,draw] (2) at ($(1)+(0.3\basewidth, 0)$) {$2$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    86
     \node[circle,draw] (3) at ($(2)+(0.3\basewidth, 0)$) {$3$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    87
     \node[circle,draw] (4) at ($(3)+(0.3\basewidth, 0)$) {$4$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    88
     \node[circle,draw] (5) at ($(4)+(0.3\basewidth, 0)$) {$5$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    89
     \node[circle,draw] (6) at ($(5)+(0.3\basewidth, 0)$) {$6$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    90
     \node[circle,draw] (7) at ($(6)+(0.3\basewidth, 0)$) {$7$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    91
     \node[circle,draw] (8) at ($(7)+(0.3\basewidth, 0)$) {$0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    92
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    93
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    94
     \draw [->, >=latex] (1) edge [loop above] node[above] {$S_1:L$} (1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    95
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    96
     \draw [->, >=latex] (1) -- node[above] {$S_0:S_1$} (2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    97
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    98
     \draw [->, >=latex] (2) edge [loop above] node[above] {$S_1:R$} (2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
    99
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   100
     \draw [->, >=latex] (2) -- node[above] {$S_0:L$} (3)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   101
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   102
     \draw [->, >=latex] (3) edge[loop above] node[above] {$S_1:S_0$} (3)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   103
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   104
     \draw [->, >=latex] (3) -- node[above] {$S_0:R$} (4)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   105
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   106
     \draw [->, >=latex] (4) edge[loop above] node[above] {$S_0:R$} (4)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   107
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   108
     \draw [->, >=latex] (4) -- node[above] {$S_0:R$} (5)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   109
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   110
     \draw [->, >=latex] (5) edge[loop above] node[above] {$S_1:R$} (5)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   111
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   112
     \draw [->, >=latex] (5) -- node[above] {$S_0:R$} (6)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   113
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   114
     \draw [->, >=latex] (6) edge[bend left = 50] node[below] {$S_1:R$} (5)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   115
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   116
     \draw [->, >=latex] (6) -- node[above] {$S_0:R$} (7)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   117
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   118
     \draw [->, >=latex] (7) edge[loop above] node[above] {$S_0:S_1$} (7)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   119
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   120
     \draw [->, >=latex] (7) -- node[above] {$S_1:L$} (8)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   121
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   122
 \end{tikzpicture}}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   123
\caption{The diagram of TM $prepare$} \label{prepare_diag}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   124
\end{figure}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   125
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   126
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.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   127
In order to detect the termination condition when the left most bit of $a_1$ is reached,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   128
TM $mainwork$ needs to look ahead and consider three different situations at the start of
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   129
every iteration:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   130
\begin{enumerate}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   131
    \item The TM configuration for the first situation is shown in Figure \ref{mainwork_case_one_input},
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   132
        where the accumulator is stored in $r$, both of the next two bits
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   133
        to be encoded are $1$. The configuration at the end of the iteration
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   134
        is shown in Figure \ref{mainwork_case_one_output}, where the first 1-bit has been
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   135
        encoded and cleared. Notice that the accumulator has been changed to
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   136
        $(r+1) \times 2$ to reflect the encoded bit.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   137
    \item The TM configuration for the second situation is shown in Figure
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   138
        \ref{mainwork_case_two_input},
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   139
        where the accumulator is stored in $r$, the next two bits
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   140
        to be encoded are $1$ and $0$. After the first
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   141
        $1$-bit was encoded and cleared, the second $0$-bit is difficult to detect
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   142
        and process. To solve this problem, these two consecutive bits are
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   143
        encoded in one iteration.  In this situation, only the first $1$-bit needs
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   144
        to be cleared since the second one is cleared by definition.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   145
        The configuration at the end of the iteration
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   146
        is shown in Figure \ref{mainwork_case_two_output}.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   147
        Notice that the accumulator has been changed to
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   148
        $(r+1) \times 4$ to reflect the two encoded bits.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   149
    \item The third situation corresponds to the case when the last bit of $a_1$ is reached.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   150
        The TM configurations at the start and end of the iteration are shown in
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   151
        Figure \ref{mainwork_case_three_input} and \ref{mainwork_case_three_output}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   152
        respectively. For this situation, only the read write head needs to be moved to
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   153
        the left to prepare a initial configuration for TM $adjust$ to start with.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   154
\end{enumerate}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   155
The diagram of $mainwork$ is given in Figure \ref{mainwork_diag}. The two rectangular nodes
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   156
labeled with $2 \times x$ and $4 \times x$ are two TMs compiling from recursive functions
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   157
so that we do not have to design and verify two quite complicated TMs.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   158
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   159
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   160
\begin{figure}[h!]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   161
\centering
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   162
\scalebox{1.2}{
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   163
\begin{tikzpicture}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   164
  \node (0) {};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   165
  \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   166
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   167
  \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   168
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $a_1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   169
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   170
  \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {\wuhao $a_2$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   171
  \node (7) [right = -0.9pt of 6] {\ldots \ldots};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   172
  \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {\wuhao $a_i$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   173
  \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {\wuhao $1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   174
  \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {\wuhao $1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   175
  \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   176
  \node (12) [right = -0.9pt of 11] {\ldots \ldots};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   177
  \node (13) [draw, right = -0.9pt of 12, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   178
  \node (14) [draw, text height = 3.9pt, right = -0.9pt of 13, thick, inner sep = 5pt] {\wuhao $r$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   179
  \draw [->, >=latex, thick] (13)+(0, -4\baseheight) -- (13);
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   180
\end{tikzpicture}}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   181
\caption{The first situation for TM $mainwork$ to consider} \label{mainwork_case_one_input}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   182
\end{figure}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   183
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   184
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   185
\begin{figure}[h!]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   186
\centering
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   187
\scalebox{1.2}{
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   188
\begin{tikzpicture}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   189
  \node (0) {};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   190
  \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   191
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   192
  \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   193
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $a_1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   194
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   195
  \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {\wuhao $a_2$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   196
  \node (7) [right = -0.9pt of 6] {\ldots \ldots};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   197
  \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {\wuhao $a_i$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   198
  \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {\wuhao $1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   199
  \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   200
  \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   201
  \node (12) [right = -0.9pt of 11] {\ldots \ldots};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   202
  \node (13) [draw, right = -0.9pt of 12, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   203
  \node (14) [draw, text height = 2.7pt, right = -0.9pt of 13, thick, inner sep = 5pt] {\wuhao $(r+1) \times 2$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   204
  \draw [->, >=latex, thick] (13)+(0, -4\baseheight) -- (13);
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   205
\end{tikzpicture}}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   206
\caption{The output for the first case of TM $mainwork$'s processing}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   207
\label{mainwork_case_one_output}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   208
\end{figure}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   209
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   210
\begin{figure}[h!]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   211
\centering
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   212
\scalebox{1.2}{
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   213
\begin{tikzpicture}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   214
  \node (0) {};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   215
  \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   216
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   217
  \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   218
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $a_1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   219
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   220
  \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {\wuhao $a_2$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   221
  \node (7) [right = -0.9pt of 6] {\ldots \ldots};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   222
  \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {\wuhao $a_i$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   223
  \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {\wuhao $1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   224
  \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   225
  \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {\wuhao $1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   226
  \node (12) [draw, right = -0.9pt of 11, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   227
  \node (13) [right = -0.9pt of 12] {\ldots \ldots};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   228
  \node (14) [draw, right = -0.9pt of 13, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   229
  \node (15) [draw, text height = 3.9pt, right = -0.9pt of 14, thick, inner sep = 5pt] {\wuhao $r$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   230
  \draw [->, >=latex, thick] (14)+(0, -4\baseheight) -- (14);
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   231
\end{tikzpicture}}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   232
\caption{The second situation for TM $mainwork$ to consider} \label{mainwork_case_two_input}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   233
\end{figure}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   234
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   235
\begin{figure}[h!]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   236
\centering
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   237
\scalebox{1.2}{
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   238
\begin{tikzpicture}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   239
  \node (0) {};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   240
  \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   241
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   242
  \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   243
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $a_1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   244
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   245
  \node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {\wuhao $a_2$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   246
  \node (7) [right = -0.9pt of 6] {\ldots \ldots};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   247
  \node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {\wuhao $a_i$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   248
  \node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {\wuhao $1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   249
  \node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   250
  \node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   251
  \node (12) [draw, right = -0.9pt of 11, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   252
  \node (13) [right = -0.9pt of 12] {\ldots \ldots};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   253
  \node (14) [draw, right = -0.9pt of 13, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   254
  \node (15) [draw, text height = 2.7pt, right = -0.9pt of 14, thick, inner sep = 5pt] {\wuhao $(r+1) \times 4$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   255
  \draw [->, >=latex, thick] (14)+(0, -4\baseheight) -- (14);
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   256
\end{tikzpicture}}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   257
\caption{The output for the second case of TM $mainwork$'s processing}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   258
\label{mainwork_case_two_output}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   259
\end{figure}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   260
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   261
\begin{figure}[h!]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   262
\centering
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   263
\scalebox{1.2}{
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   264
\begin{tikzpicture}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   265
  \node (0) {};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   266
  \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   267
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   268
  \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   269
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   270
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   271
  \node (6) [right = -0.9pt of 5] {\ldots \ldots};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   272
  \node (7) [draw, right = -0.9pt of 6, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   273
  \node (8) [draw, text height = 3.9pt, right = -0.9pt of 7, thick, inner sep = 5pt] {\wuhao $r$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   274
  \draw [->, >=latex, thick] (7)+(0, -4\baseheight) -- (7);
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   275
\end{tikzpicture}}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   276
\caption{The third situation for TM $mainwork$ to consider} \label{mainwork_case_three_input}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   277
\end{figure}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   278
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   279
\begin{figure}[h!]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   280
\centering
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   281
\scalebox{1.2}{
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   282
\begin{tikzpicture}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   283
  \node (0) {};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   284
  \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   285
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   286
  \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   287
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   288
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   289
  \node (6) [right = -0.9pt of 5] {\ldots \ldots};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   290
  \node (7) [draw, right = -0.9pt of 6, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   291
  \node (8) [draw, text height = 3.9pt, right = -0.9pt of 7, thick, inner sep = 5pt] {\wuhao $r$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   292
  \draw [->, >=latex, thick] (3)+(0, -4\baseheight) -- (3);
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   293
\end{tikzpicture}}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   294
\caption{The output for the third case of TM $mainwork$'s processing}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   295
\label{mainwork_case_three_output}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   296
\end{figure}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   297
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   298
\begin{figure}[h!]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   299
\centering
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   300
\scalebox{0.9}{
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   301
\begin{tikzpicture}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   302
     \node[circle,draw] (1) {$1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   303
     \node[circle,draw] (2) at ($(1)+(0.3\basewidth, 0)$) {$2$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   304
     \node[circle,draw] (3) at ($(2)+(0.3\basewidth, 0)$) {$3$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   305
     \node[circle,draw] (4) at ($(3)+(0.3\basewidth, 0)$) {$4$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   306
     \node[circle,draw] (5) at ($(4)+(0.3\basewidth, 0)$) {$5$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   307
     \node[circle,draw] (6) at ($(5)+(0.3\basewidth, 0)$) {$6$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   308
     \node[circle,draw] (7) at ($(2)+(0, -7\baseheight)$) {$7$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   309
     \node[circle,draw] (8) at ($(7)+(0, -7\baseheight)$) {$8$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   310
     \node[circle,draw] (9) at ($(8)+(0.3\basewidth, 0)$) {$9$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   311
     \node[circle,draw] (10) at ($(9)+(0.3\basewidth, 0)$) {$10$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   312
     \node[circle,draw] (11) at ($(10)+(0.3\basewidth, 0)$) {$11$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   313
     \node[circle,draw] (12) at ($(11)+(0.3\basewidth, 0)$) {$12$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   314
     \node[draw] (13) at ($(6)+(0.3\basewidth, 0)$) {$2 \times x$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   315
     \node[circle,draw] (14) at ($(13)+(0.3\basewidth, 0)$) {$j_1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   316
     \node[draw] (15) at ($(12)+(0.3\basewidth, 0)$) {$4 \times x$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   317
     \node[draw] (16) at ($(15)+(0.3\basewidth, 0)$) {$j_2$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   318
     \node[draw] (17) at ($(7)+(0.3\basewidth, 0)$) {$0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   319
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   320
     \draw [->, >=latex] (1) edge[loop left] node[above] {$S_0:L$} (1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   321
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   322
     \draw [->, >=latex] (1) -- node[above] {$S_1:L$} (2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   323
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   324
     \draw [->, >=latex] (2) -- node[above] {$S_1:R$} (3)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   325
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   326
     \draw [->, >=latex] (2) -- node[left] {$S_1:L$} (7)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   327
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   328
     \draw [->, >=latex] (3) edge[loop above] node[above] {$S_1:S_0$} (3)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   329
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   330
     \draw [->, >=latex] (3) -- node[above] {$S_0:R$} (4)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   331
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   332
     \draw [->, >=latex] (4) edge[loop above] node[above] {$S_0:R$} (4)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   333
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   334
     \draw [->, >=latex] (4) -- node[above] {$S_1:R$} (5)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   335
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   336
     \draw [->, >=latex] (5) edge[loop above] node[above] {$S_1:R$} (5)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   337
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   338
     \draw [->, >=latex] (5) -- node[above] {$S_0:S_1$} (6)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   339
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   340
     \draw [->, >=latex] (6) edge[loop above] node[above] {$S_1:L$} (6)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   341
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   342
     \draw [->, >=latex] (6) -- node[above] {$S_0:R$} (13)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   343
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   344
     \draw [->, >=latex] (13) -- (14)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   345
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   346
     \draw (14) -- ($(14)+(0, 6\baseheight)$) -- ($(1) + (0, 6\baseheight)$) node [above,midway] {$S_1:L$}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   347
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   348
     \draw [->, >=latex] ($(1) + (0, 6\baseheight)$) -- (1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   349
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   350
     \draw [->, >=latex] (7) -- node[above] {$S_0:R$} (17)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   351
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   352
     \draw [->, >=latex] (7) -- node[left] {$S_1:R$} (8)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   353
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   354
     \draw [->, >=latex] (8) -- node[above] {$S_0:R$} (9)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   355
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   356
     \draw [->, >=latex] (9) -- node[above] {$S_0:R$} (10)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   357
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   358
     \draw [->, >=latex] (10) -- node[above] {$S_1:R$} (11)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   359
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   360
     \draw [->, >=latex] (10) edge[loop above] node[above] {$S_0:R$} (10)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   361
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   362
     \draw [->, >=latex] (11) edge[loop above] node[above] {$S_1:R$} (11)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   363
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   364
     \draw [->, >=latex] (11) -- node[above] {$S_0:S_1$} (12)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   365
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   366
     \draw [->, >=latex] (12) -- node[above] {$S_0:R$} (15)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   367
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   368
     \draw [->, >=latex] (12) edge[loop above] node[above] {$S_1:L$} (12)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   369
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   370
     \draw [->, >=latex] (15) -- (16)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   371
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   372
     \draw (16) -- ($(16)+(0, -4\baseheight)$) -- ($(1) + (0, -18\baseheight)$) node [below,midway] {$S_1:L$}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   373
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   374
     \draw [->, >=latex] ($(1) + (0, -18\baseheight)$) -- (1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   375
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   376
 \end{tikzpicture}}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   377
\caption{The diagram of TM $mainwork$} \label{mainwork_diag}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   378
\end{figure}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   379
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   380
The purpose of TM $adjust$ is to encode the last bit of $a_1$. The initial and final configuration
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   381
of this TM are shown in Figure \ref{adjust_initial} and \ref{adjust_final} respectively.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   382
The diagram of TM $adjust$ is shown in Figure \ref{adjust_diag}.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   383
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   384
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   385
\begin{figure}[h!]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   386
\centering
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   387
\scalebox{1.2}{
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   388
\begin{tikzpicture}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   389
  \node (0) {};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   390
  \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   391
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   392
  \node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   393
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   394
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   395
  \node (6) [right = -0.9pt of 5] {\ldots \ldots};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   396
  \node (7) [draw, right = -0.9pt of 6, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   397
  \node (8) [draw, text height = 3.9pt, right = -0.9pt of 7, thick, inner sep = 5pt] {\wuhao $r$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   398
  \draw [->, >=latex, thick] (3)+(0, -4\baseheight) -- (3);
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   399
\end{tikzpicture}}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   400
\caption{Initial configuration of TM $adjust$} \label{adjust_initial}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   401
\end{figure}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   402
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   403
\begin{figure}[h!]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   404
\centering
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   405
\scalebox{1.2}{
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   406
\begin{tikzpicture}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   407
  \node (0) {};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   408
  \node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {\wuhao $m$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   409
  \node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   410
  \node (3) [draw, text height = 2.9pt, right = -0.9pt of 2, thick, inner sep = 5pt] {\wuhao $r+1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   411
  \node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   412
  \node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {\wuhao $0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   413
  \node (6) [right = -0.9pt of 5] {\ldots \ldots};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   414
  \draw [->, >=latex, thick] (1)+(0, -4\baseheight) -- (1);
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   415
\end{tikzpicture}}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   416
\caption{Final configuration of TM $adjust$} \label{adjust_final}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   417
\end{figure}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   418
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   419
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   420
\begin{figure}[h!]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   421
\centering
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   422
\scalebox{0.9}{
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   423
\begin{tikzpicture}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   424
     \node[circle,draw] (1) {$1$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   425
     \node[circle,draw] (2) at ($(1)+(0.3\basewidth, 0)$) {$2$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   426
     \node[circle,draw] (3) at ($(2)+(0.3\basewidth, 0)$) {$3$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   427
     \node[circle,draw] (4) at ($(3)+(0.3\basewidth, 0)$) {$4$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   428
     \node[circle,draw] (5) at ($(4)+(0.3\basewidth, 0)$) {$5$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   429
     \node[circle,draw] (6) at ($(5)+(0.3\basewidth, 0)$) {$6$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   430
     \node[circle,draw] (7) at ($(6)+(0.3\basewidth, 0)$) {$7$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   431
     \node[circle,draw] (8) at ($(4)+(0, -7\baseheight)$) {$8$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   432
     \node[circle,draw] (9) at ($(8)+(0.3\basewidth, 0)$) {$9$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   433
     \node[circle,draw] (10) at ($(9)+(0.3\basewidth, 0)$) {$10$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   434
     \node[circle,draw] (11) at ($(10)+(0.3\basewidth, 0)$) {$11$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   435
     \node[circle,draw] (12) at ($(11)+(0.3\basewidth, 0)$) {$0$};
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   436
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   437
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   438
     \draw [->, >=latex] (1) -- node[above] {$S_1:R$} (2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   439
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   440
     \draw [->, >=latex] (1) edge[loop above] node[above] {$S_0:S_1$} (1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   441
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   442
     \draw [->, >=latex] (2) -- node[above] {$S_1:R$} (3)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   443
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   444
     \draw [->, >=latex] (3) edge[loop above] node[above] {$S_0:R$} (3)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   445
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   446
     \draw [->, >=latex] (3) -- node[above] {$S_1:R$} (4)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   447
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   448
     \draw [->, >=latex] (4) -- node[above] {$S_1:L$} (5)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   449
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   450
     \draw [->, >=latex] (4) -- node[right] {$S_0:L$} (8)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   451
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   452
     \draw [->, >=latex] (5) -- node[above] {$S_0:L$} (6)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   453
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   454
     \draw [->, >=latex] (5) edge[loop above] node[above] {$S_1:S_0$} (5)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   455
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   456
     \draw [->, >=latex] (6) -- node[above] {$S_1:R$} (7)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   457
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   458
     \draw [->, >=latex] (6) edge[loop above] node[above] {$S_0:L$} (6)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   459
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   460
     \draw (7) -- ($(7)+(0, 6\baseheight)$) -- ($(2) + (0, 6\baseheight)$) node [above,midway] {$S_0:S_1$}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   461
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   462
     \draw [->, >=latex] ($(2) + (0, 6\baseheight)$) -- (2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   463
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   464
     \draw [->, >=latex] (8) edge[loop left] node[left] {$S_1:S_0$} (8)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   465
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   466
     \draw [->, >=latex] (8) -- node[above] {$S_0:L$} (9)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   467
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   468
     \draw [->, >=latex] (9) edge[loop above] node[above] {$S_0:L$} (9)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   469
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   470
     \draw [->, >=latex] (9) -- node[above] {$S_1:L$} (10)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   471
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   472
     \draw [->, >=latex] (10) edge[loop above] node[above] {$S_0:L$} (10)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   473
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   474
     \draw [->, >=latex] (10) -- node[above] {$S_0:L$} (11)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   475
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   476
     \draw [->, >=latex] (11) edge[loop above] node[above] {$S_1:L$} (11)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   477
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   478
     \draw [->, >=latex] (11) -- node[above] {$S_0:R$} (12)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   479
     ;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   480
 \end{tikzpicture}}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   481
\caption{Diagram of TM $adjust$} \label{adjust_diag}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   482
\end{figure}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   483
*}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   484
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   485
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   486
definition rec_twice :: "recf"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   487
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   488
  "rec_twice = Cn 1 rec_mult [id 1 0, constn 2]"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   489
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   490
definition rec_fourtimes  :: "recf"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   491
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   492
  "rec_fourtimes = Cn 1 rec_mult [id 1 0, constn 4]"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   493
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   494
definition abc_twice :: "abc_prog"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   495
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   496
  "abc_twice = (let (aprog, ary, fp) = rec_ci rec_twice in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   497
                       aprog [+] dummy_abc ((Suc 0)))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   498
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   499
definition abc_fourtimes :: "abc_prog"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   500
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   501
  "abc_fourtimes = (let (aprog, ary, fp) = rec_ci rec_fourtimes in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   502
                       aprog [+] dummy_abc ((Suc 0)))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   503
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   504
definition twice_ly :: "nat list"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   505
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   506
  "twice_ly = layout_of abc_twice"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   507
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   508
definition fourtimes_ly :: "nat list"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   509
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   510
  "fourtimes_ly = layout_of abc_fourtimes"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   511
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   512
definition t_twice :: "tprog"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   513
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   514
  "t_twice = change_termi_state (tm_of (abc_twice) @ (tMp 1 (start_of twice_ly (length abc_twice) - Suc 0)))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   515
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   516
definition t_fourtimes :: "tprog"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   517
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   518
  "t_fourtimes = change_termi_state (tm_of (abc_fourtimes) @ 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   519
             (tMp 1 (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   520
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   521
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   522
definition t_twice_len :: "nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   523
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   524
  "t_twice_len = length t_twice div 2"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   525
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   526
definition t_wcode_main_first_part:: "tprog"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   527
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   528
  "t_wcode_main_first_part \<equiv> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   529
                   [(L, 1), (L, 2), (L, 7), (R, 3),
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   530
                    (R, 4), (W0, 3), (R, 4), (R, 5),
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   531
                    (W1, 6), (R, 5), (R, 13), (L, 6),
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   532
                    (R, 0), (R, 8), (R, 9), (Nop, 8),
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   533
                    (R, 10), (W0, 9), (R, 10), (R, 11), 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   534
                    (W1, 12), (R, 11), (R, t_twice_len + 14), (L, 12)]"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   535
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   536
definition t_wcode_main :: "tprog"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   537
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   538
  "t_wcode_main = (t_wcode_main_first_part @ tshift t_twice 12 @ [(L, 1), (L, 1)]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   539
                    @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   540
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   541
fun bl_bin :: "block list \<Rightarrow> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   542
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   543
  "bl_bin [] = 0" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   544
| "bl_bin (Bk # xs) = 2 * bl_bin xs"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   545
| "bl_bin (Oc # xs) = Suc (2 * bl_bin xs)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   546
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   547
declare bl_bin.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   548
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   549
type_synonym bin_inv_t = "block list \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   550
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   551
fun wcode_before_double :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   552
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   553
  "wcode_before_double ires rs (l, r) =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   554
     (\<exists> ln rn. l = Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   555
               r = Oc\<^bsup>(Suc (Suc rs))\<^esup> @ Bk\<^bsup>rn \<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   556
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   557
declare wcode_before_double.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   558
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   559
fun wcode_after_double :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   560
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   561
  "wcode_after_double ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   562
     (\<exists> ln rn. l = Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   563
         r = Oc\<^bsup>Suc (Suc (Suc 2*rs))\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   564
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   565
declare wcode_after_double.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   566
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   567
fun wcode_on_left_moving_1_B :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   568
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   569
  "wcode_on_left_moving_1_B ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   570
     (\<exists> ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # Oc # ires \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   571
               r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   572
               ml + mr > Suc 0 \<and> mr > 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   573
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   574
declare wcode_on_left_moving_1_B.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   575
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   576
fun wcode_on_left_moving_1_O :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   577
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   578
  "wcode_on_left_moving_1_O ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   579
     (\<exists> ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   580
               l = Oc # ires \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   581
               r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   582
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   583
declare wcode_on_left_moving_1_O.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   584
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   585
fun wcode_on_left_moving_1 :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   586
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   587
  "wcode_on_left_moving_1 ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   588
          (wcode_on_left_moving_1_B ires rs (l, r) \<or> wcode_on_left_moving_1_O ires rs (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   589
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   590
declare wcode_on_left_moving_1.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   591
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   592
fun wcode_on_checking_1 :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   593
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   594
   "wcode_on_checking_1 ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   595
    (\<exists> ln rn. l = ires \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   596
              r = Oc # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   597
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   598
fun wcode_erase1 :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   599
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   600
"wcode_erase1 ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   601
       (\<exists> ln rn. l = Oc # ires \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   602
                 tl r = Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   603
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   604
declare wcode_erase1.simps [simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   605
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   606
fun wcode_on_right_moving_1 :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   607
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   608
  "wcode_on_right_moving_1 ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   609
       (\<exists> ml mr rn.        
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   610
             l = Bk\<^bsup>ml\<^esup> @ Oc # ires \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   611
             r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   612
             ml + mr > Suc 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   613
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   614
declare wcode_on_right_moving_1.simps [simp del] 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   615
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   616
declare wcode_on_right_moving_1.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   617
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   618
fun wcode_goon_right_moving_1 :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   619
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   620
  "wcode_goon_right_moving_1 ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   621
      (\<exists> ml mr ln rn. 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   622
            l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   623
            r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   624
            ml + mr = Suc rs)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   625
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   626
declare wcode_goon_right_moving_1.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   627
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   628
fun wcode_backto_standard_pos_B :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   629
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   630
  "wcode_backto_standard_pos_B ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   631
          (\<exists> ln rn. l =  Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   632
               r =  Bk # Oc\<^bsup>(Suc (Suc rs))\<^esup> @ Bk\<^bsup>rn \<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   633
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   634
declare wcode_backto_standard_pos_B.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   635
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   636
fun wcode_backto_standard_pos_O :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   637
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   638
   "wcode_backto_standard_pos_O ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   639
        (\<exists> ml mr ln rn. 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   640
            l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   641
            r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   642
            ml + mr = Suc (Suc rs) \<and> mr > 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   643
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   644
declare wcode_backto_standard_pos_O.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   645
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   646
fun wcode_backto_standard_pos :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   647
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   648
  "wcode_backto_standard_pos ires rs (l, r) = (wcode_backto_standard_pos_B ires rs (l, r) \<or>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   649
                                            wcode_backto_standard_pos_O ires rs (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   650
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   651
declare wcode_backto_standard_pos.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   652
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   653
lemma [simp]: "<0::nat> = [Oc]"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   654
apply(simp add: tape_of_nat_abv exponent_def tape_of_nat_list.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   655
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   656
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   657
lemma tape_of_Suc_nat: "<Suc (a ::nat)> = replicate a Oc @ [Oc, Oc]"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   658
apply(simp add: tape_of_nat_abv exp_ind tape_of_nat_list.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   659
apply(simp only: exp_ind_def[THEN sym])
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   660
apply(simp only: exp_ind, simp, simp add: exponent_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   661
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   662
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   663
lemma [simp]: "length (<a::nat>) = Suc a"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   664
apply(simp add: tape_of_nat_abv tape_of_nat_list.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   665
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   666
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   667
lemma [simp]: "<[a::nat]> = <a>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   668
apply(simp add: tape_of_nat_abv tape_of_nl_abv exponent_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   669
                tape_of_nat_list.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   670
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   671
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   672
lemma bin_wc_eq: "bl_bin xs = bl2wc xs"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   673
proof(induct xs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   674
  show " bl_bin [] = bl2wc []" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   675
    apply(simp add: bl_bin.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   676
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   677
next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   678
  fix a xs
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   679
  assume "bl_bin xs = bl2wc xs"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   680
  thus " bl_bin (a # xs) = bl2wc (a # xs)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   681
    apply(case_tac a, simp_all add: bl_bin.simps bl2wc.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   682
    apply(simp_all add: bl2nat.simps bl2nat_double)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   683
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   684
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   685
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   686
declare exp_def[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   687
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   688
lemma bl_bin_nat_Suc:  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   689
  "bl_bin (<Suc a>) = bl_bin (<a>) + 2^(Suc a)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   690
apply(simp add: tape_of_nat_abv bin_wc_eq)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   691
apply(simp add: bl2wc.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   692
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   693
lemma [simp]: " rev (a\<^bsup>aa\<^esup>) = a\<^bsup>aa\<^esup>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   694
apply(simp add: exponent_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   695
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   696
 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   697
declare tape_of_nl_abv_cons[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   698
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   699
lemma tape_of_nl_rev: "rev (<lm::nat list>) = (<rev lm>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   700
apply(induct lm rule: list_tl_induct, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   701
apply(case_tac "list = []", simp add: tape_of_nl_abv tape_of_nat_list.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   702
apply(simp add: tape_of_nat_list_butlast_last tape_of_nl_abv_cons)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   703
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   704
lemma [simp]: "a\<^bsup>Suc 0\<^esup> = [a]" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   705
by(simp add: exp_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   706
lemma tape_of_nl_cons_app1: "(<a # xs @ [b]>) = (Oc\<^bsup>Suc a\<^esup> @ Bk # (<xs@ [b]>))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   707
apply(case_tac xs, simp add: tape_of_nl_abv tape_of_nat_list.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   708
apply(simp add: tape_of_nl_abv  tape_of_nat_list.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   709
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   710
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   711
lemma bl_bin_bk_oc[simp]:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   712
  "bl_bin (xs @ [Bk, Oc]) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   713
  bl_bin xs + 2*2^(length xs)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   714
apply(simp add: bin_wc_eq)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   715
using bl2nat_cons_oc[of "xs @ [Bk]"]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   716
apply(simp add: bl2nat_cons_bk bl2wc.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   717
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   718
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   719
lemma tape_of_nat[simp]: "(<a::nat>) = Oc\<^bsup>Suc a\<^esup>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   720
apply(simp add: tape_of_nat_abv)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   721
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   722
lemma tape_of_nl_cons_app2: "(<c # xs @ [b]>) = (<c # xs> @ Bk # Oc\<^bsup>Suc b\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   723
proof(induct "length xs" arbitrary: xs c,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   724
  simp add: tape_of_nl_abv  tape_of_nat_list.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   725
  fix x xs c
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   726
  assume ind: "\<And>xs c. x = length xs \<Longrightarrow> <c # xs @ [b]> = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   727
    <c # xs> @ Bk # Oc\<^bsup>Suc b\<^esup>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   728
    and h: "Suc x = length (xs::nat list)" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   729
  show "<c # xs @ [b]> = <c # xs> @ Bk # Oc\<^bsup>Suc b\<^esup>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   730
  proof(case_tac xs, simp add: tape_of_nl_abv  tape_of_nat_list.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   731
    fix a list
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   732
    assume g: "xs = a # list"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   733
    hence k: "<a # list @ [b]> =  <a # list> @ Bk # Oc\<^bsup>Suc b\<^esup>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   734
      apply(rule_tac ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   735
      using h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   736
      apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   737
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   738
    from g and k show "<c # xs @ [b]> = <c # xs> @ Bk # Oc\<^bsup>Suc b\<^esup>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   739
      apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   740
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   741
  qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   742
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   743
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   744
lemma [simp]: "length (<aa # a # list>) = Suc (Suc aa) + length (<a # list>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   745
apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   746
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   747
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   748
lemma [simp]: "bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # tape_of_nat_list (a # lista) @ [Bk, Oc]) =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   749
              bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # tape_of_nat_list (a # lista)) + 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   750
              2* 2^(length (Oc\<^bsup>Suc aa\<^esup> @ Bk # tape_of_nat_list (a # lista)))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   751
using bl_bin_bk_oc[of "Oc\<^bsup>Suc aa\<^esup> @ Bk # tape_of_nat_list (a # lista)"]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   752
apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   753
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   754
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   755
lemma [simp]: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   756
  "bl_bin (<aa # list>) + (4 * rs + 4) * 2 ^ (length (<aa # list>) - Suc 0)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   757
  = bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # <list @ [0]>) + rs * (2 * 2 ^ (aa + length (<list @ [0]>)))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   758
apply(case_tac "list", simp add: add_mult_distrib, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   759
apply(simp add: tape_of_nl_cons_app2 add_mult_distrib)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   760
apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   761
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   762
  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   763
lemma tape_of_nl_app_Suc: "((<list @ [Suc ab]>)) = (<list @ [ab]>) @ [Oc]"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   764
apply(induct list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   765
apply(simp add: tape_of_nl_abv tape_of_nat_list.simps exp_ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   766
apply(case_tac list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   767
apply(simp_all add:tape_of_nl_abv tape_of_nat_list.simps exp_ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   768
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   769
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   770
lemma [simp]: "bl_bin (Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [ab]> @ [Oc])
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   771
              = bl_bin (Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [ab]>) +
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   772
              2^(length (Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [ab]>))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   773
apply(simp add: bin_wc_eq)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   774
apply(simp add: bl2nat_cons_oc bl2wc.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   775
using bl2nat_cons_oc[of "Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [ab]>"]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   776
apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   777
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   778
lemma [simp]: "bl_bin (Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [ab]>) + (4 * 2 ^ (aa + length (<list @ [ab]>)) +
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   779
         4 * (rs * 2 ^ (aa + length (<list @ [ab]>)))) =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   780
       bl_bin (Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [Suc ab]>) +
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   781
         rs * (2 * 2 ^ (aa + length (<list @ [Suc ab]>)))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   782
apply(simp add: tape_of_nl_app_Suc)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   783
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   784
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   785
declare tape_of_nat[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   786
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   787
fun wcode_double_case_inv :: "nat \<Rightarrow> bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   788
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   789
  "wcode_double_case_inv st ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   790
          (if st = Suc 0 then wcode_on_left_moving_1 ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   791
          else if st = Suc (Suc 0) then wcode_on_checking_1 ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   792
          else if st = 3 then wcode_erase1 ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   793
          else if st = 4 then wcode_on_right_moving_1 ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   794
          else if st = 5 then wcode_goon_right_moving_1 ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   795
          else if st = 6 then wcode_backto_standard_pos ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   796
          else if st = 13 then wcode_before_double ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   797
          else False)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   798
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   799
declare wcode_double_case_inv.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   800
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   801
fun wcode_double_case_state :: "t_conf \<Rightarrow> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   802
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   803
  "wcode_double_case_state (st, l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   804
   13 - st"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   805
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   806
fun wcode_double_case_step :: "t_conf \<Rightarrow> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   807
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   808
  "wcode_double_case_step (st, l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   809
      (if st = Suc 0 then (length l)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   810
      else if st = Suc (Suc 0) then (length r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   811
      else if st = 3 then 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   812
                 if hd r = Oc then 1 else 0
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   813
      else if st = 4 then (length r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   814
      else if st = 5 then (length r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   815
      else if st = 6 then (length l)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   816
      else 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   817
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   818
fun wcode_double_case_measure :: "t_conf \<Rightarrow> nat \<times> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   819
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   820
  "wcode_double_case_measure (st, l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   821
     (wcode_double_case_state (st, l, r), 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   822
      wcode_double_case_step (st, l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   823
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   824
definition wcode_double_case_le :: "(t_conf \<times> t_conf) set"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   825
  where "wcode_double_case_le \<equiv> (inv_image lex_pair wcode_double_case_measure)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   826
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   827
lemma [intro]: "wf lex_pair"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   828
by(auto intro:wf_lex_prod simp:lex_pair_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   829
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   830
lemma wf_wcode_double_case_le[intro]: "wf wcode_double_case_le"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   831
by(auto intro:wf_inv_image simp: wcode_double_case_le_def )
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   832
term fetch
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   833
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   834
lemma [simp]: "fetch t_wcode_main (Suc 0) Bk = (L, Suc 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   835
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   836
                fetch.simps nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   837
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   838
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   839
lemma [simp]: "fetch t_wcode_main (Suc 0) Oc = (L, Suc (Suc 0))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   840
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   841
                fetch.simps nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   842
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   843
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   844
lemma [simp]: "fetch t_wcode_main (Suc (Suc 0)) Oc = (R, 3)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   845
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   846
                fetch.simps nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   847
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   848
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   849
lemma [simp]: "fetch t_wcode_main (Suc (Suc (Suc 0))) Bk = (R, 4)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   850
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   851
                fetch.simps nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   852
done 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   853
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   854
lemma [simp]: "fetch t_wcode_main (Suc (Suc (Suc 0))) Oc = (W0, 3)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   855
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   856
                fetch.simps nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   857
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   858
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   859
lemma [simp]: "fetch t_wcode_main 4 Bk = (R, 4)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   860
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   861
                fetch.simps nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   862
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   863
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   864
lemma [simp]: "fetch t_wcode_main 4 Oc = (R, 5)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   865
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   866
                fetch.simps nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   867
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   868
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   869
lemma [simp]: "fetch t_wcode_main 5 Oc = (R, 5)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   870
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   871
                fetch.simps nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   872
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   873
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   874
lemma [simp]: "fetch t_wcode_main 5 Bk = (W1, 6)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   875
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   876
                fetch.simps nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   877
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   878
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   879
lemma [simp]: "fetch t_wcode_main 6 Bk = (R, 13)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   880
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   881
                fetch.simps nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   882
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   883
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   884
lemma [simp]: "fetch t_wcode_main 6 Oc = (L, 6)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   885
apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   886
                fetch.simps nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   887
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   888
lemma [elim]: "Bk\<^bsup>mr\<^esup> = [] \<Longrightarrow> mr = 0"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   889
apply(case_tac mr, auto simp: exponent_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   890
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   891
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   892
lemma [simp]: "wcode_on_left_moving_1 ires rs (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   893
apply(simp add: wcode_on_left_moving_1.simps wcode_on_left_moving_1_B.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   894
                wcode_on_left_moving_1_O.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   895
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   896
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   897
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   898
declare wcode_on_checking_1.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   899
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   900
lemmas wcode_double_case_inv_simps = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   901
  wcode_on_left_moving_1.simps wcode_on_left_moving_1_O.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   902
  wcode_on_left_moving_1_B.simps wcode_on_checking_1.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   903
  wcode_erase1.simps wcode_on_right_moving_1.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   904
  wcode_goon_right_moving_1.simps wcode_backto_standard_pos.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   905
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   906
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   907
lemma [simp]: "wcode_on_left_moving_1 ires rs (b, r) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   908
apply(simp add: wcode_double_case_inv_simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   909
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   910
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   911
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   912
lemma [elim]: "\<lbrakk>wcode_on_left_moving_1 ires rs (b, Bk # list);
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   913
                tl b = aa \<and> hd b # Bk # list = ba\<rbrakk> \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   914
               wcode_on_left_moving_1 ires rs (aa, ba)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   915
apply(simp only: wcode_on_left_moving_1.simps wcode_on_left_moving_1_O.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   916
                wcode_on_left_moving_1_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   917
apply(erule_tac disjE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   918
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   919
apply(case_tac ml, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   920
apply(rule_tac x = "mr - Suc (Suc 0)" in exI, rule_tac x = rn in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   921
apply(case_tac mr, simp, case_tac nat, simp, simp add: exp_ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   922
apply(rule_tac disjI1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   923
apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   924
      simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   925
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   926
apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   927
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   928
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   929
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   930
lemma [elim]: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   931
  "\<lbrakk>wcode_on_left_moving_1 ires rs (b, Oc # list); tl b = aa \<and> hd b # Oc # list = ba\<rbrakk> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   932
    \<Longrightarrow> wcode_on_checking_1 ires rs (aa, ba)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   933
apply(simp only: wcode_double_case_inv_simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   934
apply(erule_tac disjE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   935
apply(erule_tac [!] exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   936
apply(case_tac mr, simp, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   937
apply(rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   938
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   939
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   940
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   941
lemma [simp]: "wcode_on_checking_1 ires rs (b, []) = False" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   942
apply(auto simp: wcode_double_case_inv_simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   943
done         
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   944
 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   945
lemma [simp]: "wcode_on_checking_1 ires rs (b, Bk # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   946
apply(auto simp: wcode_double_case_inv_simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   947
done         
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   948
  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   949
lemma [elim]: "\<lbrakk>wcode_on_checking_1 ires rs (b, Oc # ba);Oc # b = aa \<and> list = ba\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   950
  \<Longrightarrow> wcode_erase1 ires rs (aa, ba)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   951
apply(simp only: wcode_double_case_inv_simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   952
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   953
apply(rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   954
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   955
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   956
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   957
lemma [simp]: "wcode_on_checking_1 ires rs (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   958
apply(simp add: wcode_double_case_inv_simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   959
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   960
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   961
lemma [simp]: "wcode_on_checking_1 ires rs ([], Bk # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   962
apply(simp add: wcode_double_case_inv_simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   963
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   964
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   965
lemma [simp]: "wcode_erase1 ires rs (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   966
apply(simp add: wcode_double_case_inv_simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   967
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   968
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   969
lemma [simp]: "wcode_on_right_moving_1 ires rs (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   970
apply(simp add: wcode_double_case_inv_simps exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   971
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   972
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   973
lemma [simp]: "wcode_on_right_moving_1 ires rs (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   974
apply(simp add: wcode_double_case_inv_simps exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   975
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   976
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   977
lemma [elim]: "\<lbrakk>wcode_on_right_moving_1 ires rs (b, Bk # ba);  Bk # b = aa \<and> list = b\<rbrakk> \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   978
  wcode_on_right_moving_1 ires rs (aa, ba)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   979
apply(simp only: wcode_double_case_inv_simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   980
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   981
apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - Suc 0" in exI,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   982
      rule_tac x = rn in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   983
apply(simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   984
apply(case_tac mr, simp, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   985
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   986
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   987
lemma [elim]: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   988
  "\<lbrakk>wcode_on_right_moving_1 ires rs (b, Oc # ba); Oc # b = aa \<and> list = ba\<rbrakk> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   989
  \<Longrightarrow> wcode_goon_right_moving_1 ires rs (aa, ba)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   990
apply(simp only: wcode_double_case_inv_simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   991
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   992
apply(rule_tac x = "Suc 0" in exI, rule_tac x = "rs" in exI,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   993
      rule_tac x = "ml - Suc (Suc 0)" in exI, rule_tac x = rn in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   994
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   995
apply(case_tac ml, simp, case_tac nat, simp, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   996
apply(simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   997
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   998
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
   999
lemma [simp]: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1000
  "wcode_on_right_moving_1 ires rs (b, []) \<Longrightarrow> False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1001
apply(simp add: wcode_double_case_inv_simps exponent_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1002
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1003
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1004
lemma [elim]: "\<lbrakk>wcode_erase1 ires rs (b, Bk # ba); Bk # b = aa \<and> list = ba; c = Bk # ba\<rbrakk> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1005
  \<Longrightarrow> wcode_on_right_moving_1 ires rs (aa, ba)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1006
apply(simp only: wcode_double_case_inv_simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1007
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1008
apply(rule_tac x = "Suc 0" in exI, rule_tac x = "Suc (Suc ln)" in exI, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1009
      rule_tac x = rn in exI, simp add: exp_ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1010
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1011
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1012
lemma [elim]: "\<lbrakk>wcode_erase1 ires rs (aa, Oc # list);  b = aa \<and> Bk # list = ba\<rbrakk> \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1013
  wcode_erase1 ires rs (aa, ba)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1014
apply(simp only: wcode_double_case_inv_simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1015
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1016
apply(rule_tac x = ln in exI, rule_tac x = rn in exI, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1017
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1018
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1019
lemma [elim]: "\<lbrakk>wcode_goon_right_moving_1 ires rs (aa, []); b = aa \<and> [Oc] = ba\<rbrakk> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1020
              \<Longrightarrow> wcode_backto_standard_pos ires rs (aa, ba)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1021
apply(simp only: wcode_double_case_inv_simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1022
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1023
apply(rule_tac disjI2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1024
apply(simp only:wcode_backto_standard_pos_O.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1025
apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, rule_tac x = ln in exI,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1026
      rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1027
apply(case_tac mr, simp_all add: exponent_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1028
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1029
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1030
lemma [elim]: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1031
  "\<lbrakk>wcode_goon_right_moving_1 ires rs (aa, Bk # list);  b = aa \<and> Oc # list = ba\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1032
  \<Longrightarrow> wcode_backto_standard_pos ires rs (aa, ba)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1033
apply(simp only: wcode_double_case_inv_simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1034
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1035
apply(rule_tac disjI2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1036
apply(simp only:wcode_backto_standard_pos_O.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1037
apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, rule_tac x = ln in exI,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1038
      rule_tac x = "rn - Suc 0" in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1039
apply(case_tac mr, simp, case_tac rn, simp, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1040
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1041
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1042
lemma [elim]: "\<lbrakk>wcode_goon_right_moving_1 ires rs (b, Oc # ba);  Oc # b = aa \<and> list = ba\<rbrakk> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1043
  \<Longrightarrow> wcode_goon_right_moving_1 ires rs (aa, ba)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1044
apply(simp only: wcode_double_case_inv_simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1045
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1046
apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - Suc 0" in exI, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1047
      rule_tac x = ln in exI, rule_tac x = rn in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1048
apply(simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1049
apply(case_tac mr, simp, case_tac rn, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1050
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1051
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1052
lemma [elim]: "\<lbrakk>wcode_backto_standard_pos ires rs (b, []);  Bk # b = aa\<rbrakk> \<Longrightarrow> False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1053
apply(auto simp: wcode_double_case_inv_simps wcode_backto_standard_pos_O.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1054
                 wcode_backto_standard_pos_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1055
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1056
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1057
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1058
lemma [elim]: "\<lbrakk>wcode_backto_standard_pos ires rs (b, Bk # ba); Bk # b = aa \<and> list = ba\<rbrakk> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1059
  \<Longrightarrow> wcode_before_double ires rs (aa, ba)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1060
apply(simp only: wcode_double_case_inv_simps wcode_backto_standard_pos_B.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1061
                 wcode_backto_standard_pos_O.simps wcode_before_double.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1062
apply(erule_tac disjE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1063
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1064
apply(rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1065
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1066
apply(case_tac [!] mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1067
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1068
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1069
lemma [simp]: "wcode_backto_standard_pos ires rs ([], Oc # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1070
apply(auto simp: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1071
                 wcode_backto_standard_pos_O.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1072
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1073
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1074
lemma [simp]: "wcode_backto_standard_pos ires rs (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1075
apply(auto simp: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1076
                 wcode_backto_standard_pos_O.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1077
apply(case_tac mr, simp, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1078
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1079
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1080
lemma [elim]: "\<lbrakk>wcode_backto_standard_pos ires rs (b, Oc # list); tl b = aa; hd b # Oc # list =  ba\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1081
       \<Longrightarrow> wcode_backto_standard_pos ires rs (aa, ba)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1082
apply(simp only:  wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1083
                 wcode_backto_standard_pos_O.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1084
apply(erule_tac disjE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1085
apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1086
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1087
apply(case_tac ml, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1088
apply(rule_tac disjI1, rule_tac conjI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1089
apply(rule_tac x = ln  in exI, simp, rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1090
apply(rule_tac disjI2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1091
apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = ln in exI, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1092
      rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1093
apply(simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1094
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1095
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1096
declare new_tape.simps[simp del] nth_of.simps[simp del] fetch.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1097
lemma wcode_double_case_first_correctness:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1098
  "let P = (\<lambda> (st, l, r). st = 13) in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1099
       let Q = (\<lambda> (st, l, r). wcode_double_case_inv st ires rs (l, r)) in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1100
       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
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1101
       \<exists> n .P (f n) \<and> Q (f (n::nat))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1102
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1103
  let ?P = "(\<lambda> (st, l, r). st = 13)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1104
  let ?Q = "(\<lambda> (st, l, r). wcode_double_case_inv st ires rs (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1105
  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)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1106
  have "\<exists> n. ?P (?f n) \<and> ?Q (?f (n::nat))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1107
  proof(rule_tac halt_lemma2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1108
    show "wf wcode_double_case_le"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1109
      by auto
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1110
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1111
    show "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1112
                   ?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_double_case_le"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1113
    proof(rule_tac allI, case_tac "?f na", simp add: tstep_red)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1114
      fix na a b c
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1115
      show "a \<noteq> 13 \<and> wcode_double_case_inv a ires rs (b, c) \<longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1116
               (case tstep (a, b, c) t_wcode_main of (st, x) \<Rightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1117
                   wcode_double_case_inv st ires rs x) \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1118
                (tstep (a, b, c) t_wcode_main, a, b, c) \<in> wcode_double_case_le"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1119
        apply(rule_tac impI, simp add: wcode_double_case_inv.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1120
        apply(auto split: if_splits simp: tstep.simps, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1121
              case_tac [!] c, simp_all, case_tac [!] "(c::block list)!0")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1122
        apply(simp_all add: new_tape.simps wcode_double_case_inv.simps wcode_double_case_le_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1123
                                        lex_pair_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1124
        apply(auto split: if_splits)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1125
        done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1126
    qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1127
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1128
    show "?Q (?f 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1129
      apply(simp add: steps.simps wcode_double_case_inv.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1130
                                  wcode_on_left_moving_1.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1131
                                  wcode_on_left_moving_1_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1132
      apply(rule_tac disjI1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1133
      apply(rule_tac x = "Suc m" in exI, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1134
      apply(rule_tac x = "Suc 0" in exI, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1135
      apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1136
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1137
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1138
    show "\<not> ?P (?f 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1139
      apply(simp add: steps.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1140
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1141
  qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1142
  thus "let P = \<lambda>(st, l, r). st = 13;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1143
    Q = \<lambda>(st, l, r). wcode_double_case_inv st ires rs (l, r);
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1144
    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
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1145
    in \<exists>n. P (f n) \<and> Q (f n)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1146
    apply(simp add: Let_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1147
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1148
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1149
    
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1150
lemma [elim]: "t_ncorrect tp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1151
    \<Longrightarrow> t_ncorrect (abacus.tshift tp a)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1152
apply(simp add: t_ncorrect.simps shift_length)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1153
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1154
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1155
lemma tshift_fetch: "\<lbrakk> fetch tp a b = (aa, st'); 0 < st'\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1156
       \<Longrightarrow> fetch (abacus.tshift tp (length tp1 div 2)) a b 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1157
          = (aa, st' + length tp1 div 2)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1158
apply(subgoal_tac "a > 0")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1159
apply(auto simp: fetch.simps nth_of.simps shift_length nth_map
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1160
                 tshift.simps split: block.splits if_splits)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1161
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1162
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1163
lemma t_steps_steps_eq: "\<lbrakk>steps (st, l, r) tp stp = (st', l', r');
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1164
         0 < st';  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1165
         0 < st \<and> st \<le> length tp div 2; 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1166
         t_ncorrect tp1;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1167
          t_ncorrect tp\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1168
    \<Longrightarrow> t_steps (st + length tp1 div 2, l, r) (tshift tp (length tp1 div 2), 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1169
                                                      length tp1 div 2) stp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1170
       = (st' + length tp1 div 2, l', r')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1171
apply(induct stp arbitrary: st' l' r', simp add: steps.simps t_steps.simps,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1172
      simp add: tstep_red stepn)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1173
apply(case_tac "(steps (st, l, r) tp stp)", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1174
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1175
  fix stp st' l' r' a b c
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1176
  assume ind: "\<And>st' l' r'.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1177
    \<lbrakk>a = st' \<and> b = l' \<and> c = r'; 0 < st'\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1178
    \<Longrightarrow> t_steps (st + length tp1 div 2, l, r) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1179
    (abacus.tshift tp (length tp1 div 2), length tp1 div 2) stp = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1180
     (st' + length tp1 div 2, l', r')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1181
  and h: "tstep (a, b, c) tp = (st', l', r')" "0 < st'" "t_ncorrect tp1"  "t_ncorrect tp"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1182
  have k: "t_steps (st + length tp1 div 2, l, r) (abacus.tshift tp (length tp1 div 2),
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1183
         length tp1 div 2) stp = (a + length tp1 div 2, b, c)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1184
    apply(rule_tac ind, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1185
    using h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1186
    apply(case_tac a, simp_all add: tstep.simps fetch.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1187
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1188
  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)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1189
           (abacus.tshift tp (length tp1 div 2), length tp1 div 2) =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1190
          (st' + length tp1 div 2, l', r')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1191
    apply(simp add: k)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1192
    apply(simp add: tstep.simps t_step.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1193
    apply(case_tac "fetch tp a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1194
    apply(subgoal_tac "fetch (abacus.tshift tp (length tp1 div 2)) a
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1195
                       (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) = (aa, st' + length tp1 div 2)", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1196
    apply(simp add: tshift_fetch)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1197
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1198
qed 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1199
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1200
lemma t_tshift_lemma: "\<lbrakk> steps (st, l, r) tp stp = (st', l', r'); 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1201
                         st' \<noteq> 0; 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1202
                         stp > 0;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1203
                         0 < st \<and> st \<le> length tp div 2;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1204
                         t_ncorrect tp1;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1205
                         t_ncorrect tp;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1206
                         t_ncorrect tp2
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1207
                         \<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1208
         \<Longrightarrow> \<exists> stp>0. steps (st + length tp1 div 2, l, r) (tp1 @ tshift tp (length tp1 div 2) @ tp2) stp 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1209
                  = (st' + length tp1 div 2, l', r')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1210
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1211
  assume h: "steps (st, l, r) tp stp = (st', l', r')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1212
    "st' \<noteq> 0" "stp > 0"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1213
    "0 < st \<and> st \<le> length tp div 2"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1214
    "t_ncorrect tp1"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1215
    "t_ncorrect tp"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1216
    "t_ncorrect tp2"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1217
  from h have 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1218
    "\<exists>stp>0. t_steps (st + length tp1 div 2, l, r) (tp1 @ abacus.tshift tp (length tp1 div 2) @ tp2, 0) stp = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1219
                            (st' + length tp1 div 2, l', r')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1220
    apply(rule_tac stp = stp in turing_shift, simp_all add: shift_length)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1221
    apply(simp add: t_steps_steps_eq)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1222
    apply(simp add: t_ncorrect.simps shift_length)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1223
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1224
  thus "\<exists> stp>0. steps (st + length tp1 div 2, l, r) (tp1 @ tshift tp (length tp1 div 2) @ tp2) stp 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1225
                  = (st' + length tp1 div 2, l', r')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1226
    apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1227
    apply(rule_tac x = stp in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1228
    apply(subgoal_tac "length (tp1 @ abacus.tshift tp (length tp1 div 2) @ tp2) mod 2 = 0")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1229
    apply(simp only: steps_eq)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1230
    using h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1231
    apply(auto simp: t_ncorrect.simps shift_length)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1232
    apply arith
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1233
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1234
qed  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1235
  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1236
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1237
lemma t_twice_len_ge: "Suc 0 \<le> length t_twice div 2"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1238
apply(simp add: t_twice_def tMp.simps shift_length)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1239
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1240
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1241
lemma [intro]: "rec_calc_rel (recf.id (Suc 0) 0) [rs] rs"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1242
  apply(rule_tac calc_id, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1243
  done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1244
  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1245
lemma [intro]: "rec_calc_rel (constn 2) [rs] 2"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1246
using prime_rel_exec_eq[of "constn 2" "[rs]" 2]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1247
apply(subgoal_tac "primerec (constn 2) 1", auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1248
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1249
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1250
lemma  [intro]: "rec_calc_rel rec_mult [rs, 2] (2 * rs)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1251
using prime_rel_exec_eq[of "rec_mult" "[rs, 2]"  "2*rs"]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1252
apply(subgoal_tac "primerec rec_mult (Suc (Suc 0))", auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1253
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1254
lemma t_twice_correct: "\<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1255
            (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0)) stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1256
       (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1257
proof(case_tac "rec_ci rec_twice")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1258
  fix a b c
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1259
  assume h: "rec_ci rec_twice = (a, b, c)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1260
  have "\<exists>stp m l. steps (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\<^bsup>n\<^esup>) (tm_of abc_twice @ tMp (Suc 0) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1261
    (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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1262
  proof(rule_tac t_compiled_by_rec)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1263
    show "rec_ci rec_twice = (a, b, c)" by (simp add: h)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1264
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1265
    show "rec_calc_rel rec_twice [rs] (2 * rs)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1266
      apply(simp add: rec_twice_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1267
      apply(rule_tac rs =  "[rs, 2]" in calc_cn, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1268
      apply(rule_tac allI, case_tac k, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1269
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1270
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1271
    show "length [rs] = Suc 0" by simp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1272
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1273
    show "layout_of (a [+] dummy_abc (Suc 0)) = layout_of (a [+] dummy_abc (Suc 0))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1274
      by simp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1275
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1276
    show "start_of twice_ly (length abc_twice) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1277
      start_of (layout_of (a [+] dummy_abc (Suc 0))) (length (a [+] dummy_abc (Suc 0)))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1278
      using h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1279
      apply(simp add: twice_ly_def abc_twice_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1280
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1281
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1282
    show "tm_of abc_twice = tm_of (a [+] dummy_abc (Suc 0))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1283
      using h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1284
      apply(simp add: abc_twice_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1285
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1286
  qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1287
  thus "\<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1288
            (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0)) stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1289
       (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1290
    apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1291
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1292
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1293
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1294
lemma change_termi_state_fetch: "\<lbrakk>fetch ap a b = (aa, st); st > 0\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1295
       \<Longrightarrow> fetch (change_termi_state ap) a b = (aa, st)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1296
apply(case_tac b, auto simp: fetch.simps nth_of.simps change_termi_state.simps nth_map
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1297
                       split: if_splits block.splits)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1298
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1299
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1300
lemma change_termi_state_exec_in_range:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1301
     "\<lbrakk>steps (st, l, r) ap stp = (st', l', r'); st' \<noteq> 0\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1302
    \<Longrightarrow> steps (st, l, r) (change_termi_state ap) stp = (st', l', r')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1303
proof(induct stp arbitrary: st l r st' l' r', simp add: steps.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1304
  fix stp st l r st' l' r'
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1305
  assume ind: "\<And>st l r st' l' r'. 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1306
    \<lbrakk>steps (st, l, r) ap stp = (st', l', r'); st' \<noteq> 0\<rbrakk> \<Longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1307
    steps (st, l, r) (change_termi_state ap) stp = (st', l', r')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1308
  and h: "steps (st, l, r) ap (Suc stp) = (st', l', r')" "st' \<noteq> 0"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1309
  from h show "steps (st, l, r) (change_termi_state ap) (Suc stp) = (st', l', r')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1310
  proof(simp add: tstep_red, case_tac "steps (st, l, r) ap stp", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1311
    fix a b c
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1312
    assume g: "steps (st, l, r) ap stp = (a, b, c)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1313
              "tstep (a, b, c) ap = (st', l', r')" "0 < st'"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1314
    hence "steps (st, l, r) (change_termi_state ap) stp = (a, b, c)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1315
      apply(rule_tac ind, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1316
      apply(case_tac a, simp_all add: tstep_0)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1317
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1318
    from g and this show "tstep (steps (st, l, r) (change_termi_state ap) stp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1319
      (change_termi_state ap) = (st', l', r')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1320
      apply(simp add: tstep.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1321
      apply(case_tac "fetch ap a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1322
      apply(subgoal_tac "fetch (change_termi_state ap) a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1323
                   = (aa, st')", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1324
      apply(simp add: change_termi_state_fetch)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1325
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1326
  qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1327
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1328
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1329
lemma change_termi_state_fetch0: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1330
  "\<lbrakk>0 < a; a \<le> length ap div 2; t_correct ap; fetch ap a b = (aa, 0)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1331
  \<Longrightarrow> fetch (change_termi_state ap) a b = (aa, Suc (length ap div 2))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1332
apply(case_tac b, auto simp: fetch.simps nth_of.simps change_termi_state.simps nth_map
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1333
                       split: if_splits block.splits)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1334
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1335
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1336
lemma turing_change_termi_state: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1337
  "\<lbrakk>steps (Suc 0, l, r) ap stp = (0, l', r'); t_correct ap\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1338
     \<Longrightarrow> \<exists> stp. steps (Suc 0, l, r) (change_termi_state ap) stp = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1339
        (Suc (length ap div 2), l', r')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1340
apply(drule first_halt_point)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1341
apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1342
apply(rule_tac x = "Suc stp" in exI, simp add: tstep_red)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1343
apply(case_tac "steps (Suc 0, l, r) ap stp")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1344
apply(simp add: isS0_def change_termi_state_exec_in_range)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1345
apply(subgoal_tac "steps (Suc 0, l, r) (change_termi_state ap) stp = (a, b, c)", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1346
apply(simp add: tstep.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1347
apply(case_tac "fetch ap a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1348
apply(subgoal_tac "fetch (change_termi_state ap) a 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1349
  (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) = (aa, Suc (length ap div 2))", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1350
apply(rule_tac ap = ap in change_termi_state_fetch0, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1351
apply(rule_tac tp = "(l, r)" and l = b and r = c  and stp = stp and A = ap in s_keep, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1352
apply(simp add: change_termi_state_exec_in_range)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1353
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1354
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1355
lemma t_twice_change_term_state:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1356
  "\<exists> stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1357
     = (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1358
using t_twice_correct[of ires rs n]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1359
apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1360
apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1361
apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1362
proof(drule_tac turing_change_termi_state)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1363
  fix stp ln rn
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1364
  show "t_correct (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1365
    apply(rule_tac t_compiled_correct, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1366
    apply(simp add: twice_ly_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1367
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1368
next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1369
  fix stp ln rn
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1370
  show "\<exists>stp. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1371
    (change_termi_state (tm_of abc_twice @ tMp (Suc 0) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1372
    (start_of twice_ly (length abc_twice) - Suc 0))) stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1373
    (Suc (length (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0)) div 2),
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1374
    Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>) \<Longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1375
    \<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1376
    (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1377
    apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1378
    apply(simp add: t_twice_len_def t_twice_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1379
    apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1380
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1381
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1382
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1383
lemma t_twice_append_pre:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1384
  "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1385
  = (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1386
   \<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>)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1387
     (t_wcode_main_first_part @ tshift t_twice (length t_wcode_main_first_part div 2) @
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1388
      ([(L, 1), (L, 1)] @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1389
    = (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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1390
proof(rule_tac t_tshift_lemma, simp_all add: t_twice_len_ge)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1391
  assume "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1392
    (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1393
  thus "0 < stp"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1394
    apply(case_tac stp, simp add: steps.simps t_twice_len_ge t_twice_len_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1395
    using t_twice_len_ge
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1396
    apply(simp, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1397
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1398
next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1399
  show "t_ncorrect t_wcode_main_first_part"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1400
    apply(simp add: t_ncorrect.simps t_wcode_main_first_part_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1401
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1402
next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1403
  show "t_ncorrect t_twice"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1404
    using length_tm_even[of abc_twice]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1405
    apply(auto simp: t_ncorrect.simps t_twice_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1406
    apply(arith)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1407
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1408
next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1409
  show "t_ncorrect ((L, Suc 0) # (L, Suc 0) #
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1410
       abacus.tshift t_fourtimes (t_twice_len + 13) @ [(L, Suc 0), (L, Suc 0)])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1411
    using length_tm_even[of abc_fourtimes]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1412
    apply(simp add: t_ncorrect.simps shift_length t_fourtimes_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1413
    apply arith
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1414
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1415
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1416
  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1417
lemma t_twice_append:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1418
  "\<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>)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1419
     (t_wcode_main_first_part @ tshift t_twice (length t_wcode_main_first_part div 2) @
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1420
      ([(L, 1), (L, 1)] @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1421
    = (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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1422
  using t_twice_change_term_state[of ires rs n]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1423
  apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1424
  apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1425
  apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1426
  apply(drule_tac t_twice_append_pre)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1427
  apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1428
  apply(rule_tac x = stpa in exI, rule_tac x = ln in exI, rule_tac x = rn in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1429
  apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1430
  done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1431
  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1432
lemma [simp]: "fetch t_wcode_main (Suc (t_twice_len + length t_wcode_main_first_part div 2)) Oc
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1433
     = (L, Suc 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1434
apply(subgoal_tac "length (t_twice) mod 2 = 0")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1435
apply(simp add: t_wcode_main_def nth_append fetch.simps t_wcode_main_first_part_def 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1436
  nth_of.simps shift_length t_twice_len_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1437
apply(simp add: t_twice_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1438
apply(subgoal_tac "length (tm_of abc_twice) mod 2 = 0")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1439
apply arith
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1440
apply(rule_tac tm_even)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1441
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1442
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1443
lemma wcode_jump1: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1444
  "\<exists> stp ln rn. steps (Suc (t_twice_len) + length t_wcode_main_first_part div 2,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1445
                       Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>n\<^esup>)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1446
     t_wcode_main stp 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1447
    = (Suc 0, Bk\<^bsup>ln\<^esup> @ Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1448
apply(rule_tac x = "Suc 0" in exI, rule_tac x = "m" in exI, rule_tac x = n in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1449
apply(simp add: steps.simps tstep.simps exp_ind_def new_tape.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1450
apply(case_tac m, simp, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1451
apply(simp add: exp_ind_def[THEN sym] exp_ind[THEN sym])
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1452
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1453
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1454
lemma wcode_main_first_part_len:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1455
  "length t_wcode_main_first_part = 24"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1456
  apply(simp add: t_wcode_main_first_part_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1457
  done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1458
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1459
lemma wcode_double_case: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1460
  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 =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1461
          (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1462
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1463
  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 =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1464
          (13,  Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1465
    using wcode_double_case_first_correctness[of ires rs m n]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1466
    apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1467
    apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1468
    apply(case_tac "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1469
           Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main na",
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1470
          auto simp: wcode_double_case_inv.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1471
                     wcode_before_double.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1472
    apply(rule_tac x = na in exI, rule_tac x = ln in exI, rule_tac x = rn in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1473
    apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1474
    done    
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1475
  from this obtain stpa lna rna where stp1: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1476
    "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 = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1477
    (13, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1478
  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 =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1479
    (13 + t_twice_len, Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1480
    using t_twice_append[of "Bk\<^bsup>lna\<^esup> @ Oc # ires" "Suc rs" rna]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1481
    apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1482
    apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1483
    apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1484
    apply(simp add: wcode_main_first_part_len)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1485
    apply(rule_tac x = stp in exI, rule_tac x = "ln + lna" in exI, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1486
          rule_tac x = rn in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1487
    apply(simp add: t_wcode_main_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1488
    apply(simp add: exp_ind_def[THEN sym] exp_add[THEN sym])
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1489
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1490
  from this obtain stpb lnb rnb where stp2: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1491
    "steps (13, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1492
    (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
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1493
  have "\<exists>stp ln rn. steps (13 + t_twice_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1494
    Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnb\<^esup>) t_wcode_main stp = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1495
       (Suc 0,  Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1496
    using wcode_jump1[of lnb "Oc # ires" "Suc rs" rnb]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1497
    apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1498
    apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1499
    apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1500
    apply(rule_tac x = stp in exI, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1501
          rule_tac x = ln in exI, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1502
          rule_tac x = rn in exI, simp add:wcode_main_first_part_len t_wcode_main_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1503
    apply(subgoal_tac "Bk\<^bsup>lnb\<^esup> @ Bk # Bk # Oc # ires = Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1504
    apply(simp add: exp_ind_def[THEN sym] exp_ind[THEN sym])
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1505
    apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1506
    apply(case_tac lnb, simp, simp add: exp_ind_def[THEN sym] exp_ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1507
    done               
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1508
  from this obtain stpc lnc rnc where stp3: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1509
    "steps (13 + t_twice_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1510
    Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnb\<^esup>) t_wcode_main stpc = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1511
       (Suc 0,  Bk # Bk\<^bsup>lnc\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnc\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1512
    by blast
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1513
  from stp1 stp2 stp3 show "?thesis"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1514
    apply(rule_tac x = "stpa + stpb + stpc" in exI, rule_tac x = lnc in exI,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1515
         rule_tac x = rnc in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1516
    apply(simp add: steps_add)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1517
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1518
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1519
    
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1520
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1521
(* Begin: fourtime_case*)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1522
fun wcode_on_left_moving_2_B :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1523
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1524
  "wcode_on_left_moving_2_B ires rs (l, r) =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1525
     (\<exists> ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # Bk # Oc # ires \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1526
                 r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1527
                 ml + mr > Suc 0 \<and> mr > 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1528
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1529
fun wcode_on_left_moving_2_O :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1530
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1531
  "wcode_on_left_moving_2_O ires rs (l, r) =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1532
     (\<exists> ln rn. l = Bk # Oc # ires \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1533
               r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1534
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1535
fun wcode_on_left_moving_2 :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1536
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1537
  "wcode_on_left_moving_2 ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1538
      (wcode_on_left_moving_2_B ires rs (l, r) \<or> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1539
      wcode_on_left_moving_2_O ires rs (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1540
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1541
fun wcode_on_checking_2 :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1542
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1543
  "wcode_on_checking_2 ires rs (l, r) =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1544
       (\<exists> ln rn. l = Oc#ires \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1545
                 r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1546
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1547
fun wcode_goon_checking :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1548
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1549
  "wcode_goon_checking ires rs (l, r) =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1550
       (\<exists> ln rn. l = ires \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1551
                 r = Oc # Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1552
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1553
fun wcode_right_move :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1554
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1555
  "wcode_right_move ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1556
     (\<exists> ln rn. l = Oc # ires \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1557
                 r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1558
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1559
fun wcode_erase2 :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1560
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1561
  "wcode_erase2 ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1562
        (\<exists> ln rn. l = Bk # Oc # ires \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1563
                 tl r = Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1564
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1565
fun wcode_on_right_moving_2 :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1566
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1567
  "wcode_on_right_moving_2 ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1568
        (\<exists> ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # ires \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1569
                     r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> ml + mr > Suc 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1570
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1571
fun wcode_goon_right_moving_2 :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1572
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1573
  "wcode_goon_right_moving_2 ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1574
        (\<exists> ml mr ln rn. l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1575
                        r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> ml + mr = Suc rs)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1576
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1577
fun wcode_backto_standard_pos_2_B :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1578
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1579
  "wcode_backto_standard_pos_2_B ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1580
           (\<exists> ln rn. l = Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1581
                     r = Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1582
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1583
fun wcode_backto_standard_pos_2_O :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1584
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1585
  "wcode_backto_standard_pos_2_O ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1586
          (\<exists> ml mr ln rn. l = Oc\<^bsup>ml \<^esup>@ Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1587
                          r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1588
                          ml + mr = (Suc (Suc rs)) \<and> mr > 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1589
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1590
fun wcode_backto_standard_pos_2 :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1591
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1592
  "wcode_backto_standard_pos_2 ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1593
           (wcode_backto_standard_pos_2_O ires rs (l, r) \<or> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1594
           wcode_backto_standard_pos_2_B ires rs (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1595
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1596
fun wcode_before_fourtimes :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1597
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1598
  "wcode_before_fourtimes ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1599
          (\<exists> ln rn. l = Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1600
                    r = Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1601
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1602
declare wcode_on_left_moving_2_B.simps[simp del] wcode_on_left_moving_2.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1603
        wcode_on_left_moving_2_O.simps[simp del] wcode_on_checking_2.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1604
        wcode_goon_checking.simps[simp del] wcode_right_move.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1605
        wcode_erase2.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1606
        wcode_on_right_moving_2.simps[simp del] wcode_goon_right_moving_2.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1607
        wcode_backto_standard_pos_2_B.simps[simp del] wcode_backto_standard_pos_2_O.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1608
        wcode_backto_standard_pos_2.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1609
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1610
lemmas wcode_fourtimes_invs = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1611
       wcode_on_left_moving_2_B.simps wcode_on_left_moving_2.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1612
        wcode_on_left_moving_2_O.simps wcode_on_checking_2.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1613
        wcode_goon_checking.simps wcode_right_move.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1614
        wcode_erase2.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1615
        wcode_on_right_moving_2.simps wcode_goon_right_moving_2.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1616
        wcode_backto_standard_pos_2_B.simps wcode_backto_standard_pos_2_O.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1617
        wcode_backto_standard_pos_2.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1618
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1619
fun wcode_fourtimes_case_inv :: "nat \<Rightarrow> bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1620
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1621
  "wcode_fourtimes_case_inv st ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1622
           (if st = Suc 0 then wcode_on_left_moving_2 ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1623
            else if st = Suc (Suc 0) then wcode_on_checking_2 ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1624
            else if st = 7 then wcode_goon_checking ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1625
            else if st = 8 then wcode_right_move ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1626
            else if st = 9 then wcode_erase2 ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1627
            else if st = 10 then wcode_on_right_moving_2 ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1628
            else if st = 11 then wcode_goon_right_moving_2 ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1629
            else if st = 12 then wcode_backto_standard_pos_2 ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1630
            else if st = t_twice_len + 14 then wcode_before_fourtimes ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1631
            else False)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1632
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1633
declare wcode_fourtimes_case_inv.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1634
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1635
fun wcode_fourtimes_case_state :: "t_conf \<Rightarrow> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1636
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1637
  "wcode_fourtimes_case_state (st, l, r) = 13 - st"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1638
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1639
fun wcode_fourtimes_case_step :: "t_conf \<Rightarrow> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1640
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1641
  "wcode_fourtimes_case_step (st, l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1642
         (if st = Suc 0 then length l
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1643
          else if st = 9 then 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1644
           (if hd r = Oc then 1
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1645
            else 0)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1646
          else if st = 10 then length r
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1647
          else if st = 11 then length r
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1648
          else if st = 12 then length l
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1649
          else 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1650
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1651
fun wcode_fourtimes_case_measure :: "t_conf \<Rightarrow> nat \<times> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1652
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1653
  "wcode_fourtimes_case_measure (st, l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1654
     (wcode_fourtimes_case_state (st, l, r), 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1655
      wcode_fourtimes_case_step (st, l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1656
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1657
definition wcode_fourtimes_case_le :: "(t_conf \<times> t_conf) set"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1658
  where "wcode_fourtimes_case_le \<equiv> (inv_image lex_pair wcode_fourtimes_case_measure)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1659
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1660
lemma wf_wcode_fourtimes_case_le[intro]: "wf wcode_fourtimes_case_le"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1661
by(auto intro:wf_inv_image simp: wcode_fourtimes_case_le_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1662
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1663
lemma [simp]: "fetch t_wcode_main (Suc (Suc 0)) Bk = (L, 7)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1664
apply(simp add: t_wcode_main_def fetch.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1665
  t_wcode_main_first_part_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1666
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1667
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1668
lemma [simp]: "fetch t_wcode_main 7 Oc = (R, 8)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1669
apply(simp add: t_wcode_main_def fetch.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1670
  t_wcode_main_first_part_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1671
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1672
 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1673
lemma [simp]: "fetch t_wcode_main 8 Bk = (R, 9)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1674
apply(simp add: t_wcode_main_def fetch.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1675
  t_wcode_main_first_part_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1676
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1677
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1678
lemma [simp]: "fetch t_wcode_main 9 Bk = (R, 10)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1679
apply(simp add: t_wcode_main_def fetch.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1680
  t_wcode_main_first_part_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1681
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1682
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1683
lemma [simp]: "fetch t_wcode_main 9 Oc = (W0, 9)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1684
apply(simp add: t_wcode_main_def fetch.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1685
  t_wcode_main_first_part_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1686
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1687
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1688
lemma [simp]: "fetch t_wcode_main 10 Bk = (R, 10)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1689
apply(simp add: t_wcode_main_def fetch.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1690
  t_wcode_main_first_part_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1691
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1692
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1693
lemma [simp]: "fetch t_wcode_main 10 Oc = (R, 11)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1694
apply(simp add: t_wcode_main_def fetch.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1695
  t_wcode_main_first_part_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1696
done 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1697
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1698
lemma [simp]: "fetch t_wcode_main 11 Bk = (W1, 12)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1699
apply(simp add: t_wcode_main_def fetch.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1700
  t_wcode_main_first_part_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1701
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1702
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1703
lemma [simp]: "fetch t_wcode_main 11 Oc = (R, 11)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1704
apply(simp add: t_wcode_main_def fetch.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1705
  t_wcode_main_first_part_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1706
done 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1707
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1708
lemma [simp]: "fetch t_wcode_main 12 Oc = (L, 12)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1709
apply(simp add: t_wcode_main_def fetch.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1710
  t_wcode_main_first_part_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1711
done 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1712
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1713
lemma [simp]: "fetch t_wcode_main 12 Bk = (R, t_twice_len + 14)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1714
apply(simp add: t_wcode_main_def fetch.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1715
  t_wcode_main_first_part_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1716
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1717
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1718
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1719
lemma [simp]: "wcode_on_left_moving_2 ires rs (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1720
apply(auto simp: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1721
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1722
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1723
lemma [simp]: "wcode_on_checking_2 ires rs (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1724
apply(auto simp: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1725
done          
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1726
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1727
lemma [simp]: "wcode_goon_checking ires rs (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1728
apply(auto simp: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1729
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1730
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1731
lemma [simp]: "wcode_right_move ires rs (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1732
apply(auto simp: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1733
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1734
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1735
lemma [simp]: "wcode_erase2 ires rs (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1736
apply(auto simp: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1737
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1738
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1739
lemma [simp]: "wcode_on_right_moving_2 ires rs (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1740
apply(auto simp: wcode_fourtimes_invs exponent_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1741
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1742
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1743
lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1744
apply(auto simp: wcode_fourtimes_invs exponent_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1745
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1746
    
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1747
lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1748
apply(simp add: wcode_fourtimes_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1749
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1750
        
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1751
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)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1752
apply(simp only: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1753
apply(erule_tac disjE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1754
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1755
apply(case_tac ml, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1756
apply(rule_tac x = "mr - (Suc (Suc 0))" in exI, rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1757
apply(case_tac mr, simp, case_tac nat, simp, simp add: exp_ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1758
apply(rule_tac disjI1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1759
apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1760
      simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1761
apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1762
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1763
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1764
lemma [simp]: "wcode_on_checking_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1765
apply(auto simp: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1766
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1767
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1768
lemma  [simp]: "wcode_on_checking_2 ires rs (b, Bk # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1769
       \<Longrightarrow>   wcode_goon_checking ires rs (tl b, hd b # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1770
apply(simp only: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1771
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1772
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1773
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1774
lemma [simp]: "wcode_goon_checking ires rs (b, Bk # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1775
apply(simp add: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1776
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1777
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1778
lemma [simp]: " wcode_right_move ires rs (b, Bk # list) \<Longrightarrow> b\<noteq> []" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1779
apply(simp add: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1780
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1781
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1782
lemma [simp]: "wcode_right_move ires rs (b, Bk # list) \<Longrightarrow>  wcode_erase2 ires rs (Bk # b, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1783
apply(auto simp:wcode_fourtimes_invs )
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1784
apply(rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1785
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1786
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1787
lemma [simp]: "wcode_erase2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1788
apply(auto simp: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1789
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1790
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1791
lemma [simp]: "wcode_erase2 ires rs (b, Bk # list) \<Longrightarrow> wcode_on_right_moving_2 ires rs (Bk # b, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1792
apply(auto simp:wcode_fourtimes_invs )
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1793
apply(rule_tac x = "Suc (Suc 0)" in exI, simp add: exp_ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1794
apply(rule_tac x =  "Suc (Suc ln)" in exI, simp add: exp_ind, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1795
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1796
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1797
lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1798
apply(auto simp:wcode_fourtimes_invs )
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1799
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1800
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1801
lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Bk # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1802
       \<Longrightarrow> wcode_on_right_moving_2 ires rs (Bk # b, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1803
apply(auto simp: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1804
apply(rule_tac x = "Suc ml" in exI, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1805
apply(rule_tac x = "mr - 1" in exI, case_tac mr, auto simp: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1806
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1807
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1808
lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1809
apply(auto simp: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1810
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1811
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1812
lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Bk # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1813
                 wcode_backto_standard_pos_2 ires rs (b, Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1814
apply(simp add: wcode_fourtimes_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1815
apply(rule_tac x = ml in exI, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1816
apply(rule_tac x = "Suc 0" in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1817
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1818
apply(rule_tac x = "rn - 1" in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1819
apply(case_tac rn, simp, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1820
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1821
   
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1822
lemma  [simp]: "wcode_backto_standard_pos_2 ires rs (b, Bk # list) \<Longrightarrow>  b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1823
apply(simp add: wcode_fourtimes_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1824
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1825
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1826
lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1827
apply(simp add: wcode_fourtimes_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1828
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1829
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1830
lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Oc # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1831
                     wcode_on_checking_2 ires rs (tl b, hd b # Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1832
apply(auto simp: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1833
apply(case_tac [!] mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1834
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1835
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1836
lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, []) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1837
apply(auto simp: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1838
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1839
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1840
lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, []) \<Longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1841
              wcode_backto_standard_pos_2 ires rs (b, [Oc])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1842
apply(simp only: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1843
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1844
apply(rule_tac disjI1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1845
apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1846
      rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1847
apply(case_tac mr, simp, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1848
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1849
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1850
lemma "wcode_backto_standard_pos_2 ires rs (b, Bk # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1851
       \<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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1852
apply(auto simp: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1853
apply(case_tac [!] mr, auto simp: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1854
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1855
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1856
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1857
lemma [simp]: "wcode_on_checking_2 ires rs (b, Oc # list) \<Longrightarrow> False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1858
apply(simp add: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1859
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1860
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1861
lemma [simp]: "wcode_goon_checking ires rs (b, Oc # list) \<Longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1862
  (b = [] \<longrightarrow> wcode_right_move ires rs ([Oc], list)) \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1863
  (b \<noteq> [] \<longrightarrow> wcode_right_move ires rs (Oc # b, list))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1864
apply(simp only: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1865
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1866
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1867
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1868
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1869
lemma [simp]: "wcode_right_move ires rs (b, Oc # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1870
apply(auto simp: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1871
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1872
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1873
lemma [simp]: " wcode_erase2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1874
apply(simp add: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1875
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1876
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1877
lemma [simp]: "wcode_erase2 ires rs (b, Oc # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1878
       \<Longrightarrow> wcode_erase2 ires rs (b, Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1879
apply(auto simp: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1880
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1881
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1882
lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1883
apply(simp only: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1884
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1885
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1886
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1887
lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Oc # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1888
       \<Longrightarrow> wcode_goon_right_moving_2 ires rs (Oc # b, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1889
apply(auto simp: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1890
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1891
apply(rule_tac x = "Suc 0" in exI, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1892
apply(rule_tac x = "ml - 2" in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1893
apply(case_tac ml, simp, case_tac nat, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1894
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1895
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1896
lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1897
apply(simp only:wcode_fourtimes_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1898
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1899
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1900
lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, Bk # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1901
       \<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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1902
apply(simp add: wcode_fourtimes_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1903
apply(case_tac [!] mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1904
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1905
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1906
lemma [simp]: "wcode_on_checking_2 ires rs (b, Oc # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1907
apply(simp add: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1908
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1909
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1910
lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Oc # list) \<Longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1911
       wcode_goon_right_moving_2 ires rs (Oc # b, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1912
apply(simp only:wcode_fourtimes_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1913
apply(rule_tac x = "Suc ml" in exI, auto simp: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1914
apply(rule_tac x = "mr - 1" in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1915
apply(case_tac mr, case_tac rn, auto simp: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1916
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1917
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1918
lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1919
apply(simp only: wcode_fourtimes_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1920
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1921
 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1922
lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, Oc # list)    
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1923
            \<Longrightarrow> wcode_backto_standard_pos_2 ires rs (tl b, hd b # Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1924
apply(simp only: wcode_fourtimes_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1925
apply(erule_tac disjE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1926
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1927
apply(case_tac ml, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1928
apply(rule_tac disjI2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1929
apply(rule_tac conjI, rule_tac x = ln in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1930
apply(rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1931
apply(rule_tac disjI1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1932
apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1933
      rule_tac x = ln in exI, rule_tac x = rn in exI, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1934
apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1935
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1936
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1937
lemma wcode_fourtimes_case_first_correctness:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1938
 shows "let P = (\<lambda> (st, l, r). st = t_twice_len + 14) in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1939
  let Q = (\<lambda> (st, l, r). wcode_fourtimes_case_inv st ires rs (l, r)) in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1940
  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
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1941
  \<exists> n .P (f n) \<and> Q (f (n::nat))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1942
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1943
  let ?P = "(\<lambda> (st, l, r). st = t_twice_len + 14)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1944
  let ?Q = "(\<lambda> (st, l, r). wcode_fourtimes_case_inv st ires rs (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1945
  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)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1946
  have "\<exists> n . ?P (?f n) \<and> ?Q (?f (n::nat))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1947
  proof(rule_tac halt_lemma2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1948
    show "wf wcode_fourtimes_case_le"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1949
      by auto
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1950
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1951
    show "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1952
                  ?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_fourtimes_case_le"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1953
    apply(rule_tac allI,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1954
     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,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1955
     rule_tac impI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1956
    apply(simp add: tstep_red tstep.simps, case_tac c, simp, case_tac [2] aa, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1957
    
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1958
    apply(simp_all add: wcode_fourtimes_case_inv.simps new_tape.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1959
                        wcode_fourtimes_case_le_def lex_pair_def split: if_splits)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1960
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1961
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1962
    show "?Q (?f 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1963
      apply(simp add: steps.simps wcode_fourtimes_case_inv.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1964
      apply(simp add: wcode_on_left_moving_2.simps wcode_on_left_moving_2_B.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1965
                      wcode_on_left_moving_2_O.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1966
      apply(rule_tac x = "Suc m" in exI, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1967
      apply(rule_tac x ="Suc 0" in exI, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1968
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1969
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1970
    show "\<not> ?P (?f 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1971
      apply(simp add: steps.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1972
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1973
  qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1974
  thus "?thesis"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1975
    apply(erule_tac exE, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1976
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1977
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1978
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1979
definition t_fourtimes_len :: "nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1980
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1981
  "t_fourtimes_len = (length t_fourtimes div 2)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1982
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1983
lemma t_fourtimes_len_gr:  "t_fourtimes_len > 0"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1984
apply(simp add: t_fourtimes_len_def t_fourtimes_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1985
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1986
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1987
lemma t_fourtimes_correct: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1988
  "\<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1989
    (tm_of abc_fourtimes @ tMp (Suc 0) (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1990
       (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1991
proof(case_tac "rec_ci rec_fourtimes")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1992
  fix a b c
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1993
  assume h: "rec_ci rec_fourtimes = (a, b, c)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1994
  have "\<exists>stp m l. steps (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\<^bsup>n\<^esup>) (tm_of abc_fourtimes @ tMp (Suc 0) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1995
    (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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1996
  proof(rule_tac t_compiled_by_rec)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1997
    show "rec_ci rec_fourtimes = (a, b, c)" by (simp add: h)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1998
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  1999
    show "rec_calc_rel rec_fourtimes [rs] (4 * rs)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2000
      using prime_rel_exec_eq [of rec_fourtimes "[rs]" "4 * rs"]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2001
      apply(subgoal_tac "primerec rec_fourtimes (length [rs])")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2002
      apply(simp_all add: rec_fourtimes_def rec_exec.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2003
      apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2004
      apply(simp only: Nat.One_nat_def[THEN sym], auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2005
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2006
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2007
    show "length [rs] = Suc 0" by simp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2008
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2009
    show "layout_of (a [+] dummy_abc (Suc 0)) = layout_of (a [+] dummy_abc (Suc 0))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2010
      by simp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2011
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2012
    show "start_of fourtimes_ly (length abc_fourtimes) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2013
      start_of (layout_of (a [+] dummy_abc (Suc 0))) (length (a [+] dummy_abc (Suc 0)))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2014
      using h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2015
      apply(simp add: fourtimes_ly_def abc_fourtimes_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2016
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2017
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2018
    show "tm_of abc_fourtimes = tm_of (a [+] dummy_abc (Suc 0))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2019
      using h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2020
      apply(simp add: abc_fourtimes_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2021
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2022
  qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2023
  thus "\<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2024
            (tm_of abc_fourtimes @ tMp (Suc 0) (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2025
       (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2026
    apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2027
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2028
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2029
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2030
lemma t_fourtimes_change_term_state:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2031
  "\<exists> stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2032
     = (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2033
using t_fourtimes_correct[of ires rs n]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2034
apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2035
apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2036
apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2037
proof(drule_tac turing_change_termi_state)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2038
  fix stp ln rn
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2039
  show "t_correct (tm_of abc_fourtimes @ tMp (Suc 0) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2040
    (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2041
    apply(rule_tac t_compiled_correct, auto simp: fourtimes_ly_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2042
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2043
next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2044
  fix stp ln rn
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2045
  show "\<exists>stp. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2046
    (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2047
        (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2048
    (Suc (length (tm_of abc_fourtimes @ tMp (Suc 0) (start_of fourtimes_ly 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2049
    (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>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2050
    \<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2051
    (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2052
    apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2053
    apply(simp add: t_fourtimes_len_def t_fourtimes_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2054
    apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2055
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2056
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2057
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2058
lemma t_fourtimes_append_pre:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2059
  "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2060
  = (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2061
   \<Longrightarrow> \<exists> stp>0. steps (Suc 0 + length (t_wcode_main_first_part @ 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2062
              tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2063
       Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2064
     ((t_wcode_main_first_part @ 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2065
  tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) @ 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2066
  tshift t_fourtimes (length (t_wcode_main_first_part @ 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2067
  tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2) @ ([(L, 1), (L, 1)])) stp 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2068
  = (Suc t_fourtimes_len + length (t_wcode_main_first_part @ 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2069
  tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2070
  Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2071
proof(rule_tac t_tshift_lemma, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2072
  assume "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2073
    (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2074
  thus "0 < stp"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2075
    using t_fourtimes_len_gr
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2076
    apply(case_tac stp, simp_all add: steps.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2077
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2078
next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2079
  show "Suc 0 \<le> length t_fourtimes div 2"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2080
    apply(simp add: t_fourtimes_def shift_length tMp.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2081
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2082
next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2083
  show "t_ncorrect (t_wcode_main_first_part @ 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2084
    abacus.tshift t_twice (length t_wcode_main_first_part div 2) @ 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2085
    [(L, Suc 0), (L, Suc 0)])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2086
    apply(simp add: t_ncorrect.simps t_wcode_main_first_part_def shift_length
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2087
                    t_twice_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2088
    using tm_even[of abc_twice]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2089
    by arith
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2090
next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2091
  show "t_ncorrect t_fourtimes"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2092
    apply(simp add: t_fourtimes_def steps.simps t_ncorrect.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2093
    using tm_even[of abc_fourtimes]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2094
    by arith
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2095
next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2096
  show "t_ncorrect [(L, Suc 0), (L, Suc 0)]"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2097
    apply(simp add: t_ncorrect.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2098
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2099
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2100
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2101
lemma [simp]: "length t_wcode_main_first_part = 24"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2102
apply(simp add: t_wcode_main_first_part_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2103
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2104
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2105
lemma [simp]: "(26 + length t_twice) div 2 = (length t_twice) div 2 + 13"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2106
using tm_even[of abc_twice]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2107
apply(simp add: t_twice_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2108
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2109
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2110
lemma [simp]: "((26 + length (abacus.tshift t_twice 12)) div 2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2111
             = (length (abacus.tshift t_twice 12) div 2 + 13)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2112
using tm_even[of abc_twice]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2113
apply(simp add: t_twice_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2114
done 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2115
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2116
lemma [simp]: "t_twice_len + 14 =  14 + length (abacus.tshift t_twice 12) div 2"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2117
using tm_even[of abc_twice]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2118
apply(simp add: t_twice_def t_twice_len_def shift_length)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2119
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2120
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2121
lemma t_fourtimes_append:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2122
  "\<exists> stp ln rn. 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2123
  steps (Suc 0 + length (t_wcode_main_first_part @ tshift t_twice
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2124
  (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2125
  Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2126
  ((t_wcode_main_first_part @ tshift t_twice (length t_wcode_main_first_part div 2) @
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2127
  [(L, 1), (L, 1)]) @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)]) stp 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2128
  = (Suc t_fourtimes_len + length (t_wcode_main_first_part @ tshift t_twice
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2129
  (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2130
                                                                 Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2131
  using t_fourtimes_change_term_state[of ires rs n]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2132
  apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2133
  apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2134
  apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2135
  apply(drule_tac t_fourtimes_append_pre)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2136
  apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2137
  apply(rule_tac x = stpa in exI, rule_tac x = ln in exI, rule_tac x = rn in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2138
  apply(simp add: t_twice_len_def shift_length)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2139
  done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2140
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2141
lemma t_wcode_main_len: "length t_wcode_main = length t_twice + length t_fourtimes + 28"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2142
apply(simp add: t_wcode_main_def shift_length)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2143
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2144
 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2145
lemma [simp]: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) b
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2146
             = (L, Suc 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2147
using tm_even[of "abc_twice"] tm_even[of "abc_fourtimes"]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2148
apply(case_tac b)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2149
apply(simp_all only: fetch.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2150
apply(auto simp: nth_of.simps t_wcode_main_len t_twice_len_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2151
                 t_fourtimes_def t_twice_def t_fourtimes_def t_fourtimes_len_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2152
apply(auto simp: t_wcode_main_def t_wcode_main_first_part_def shift_length t_twice_def nth_append 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2153
                    t_fourtimes_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2154
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2155
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2156
lemma wcode_jump2: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2157
  "\<exists> stp ln rn. steps (t_twice_len + 14 + t_fourtimes_len
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2158
  , Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, Oc\<^bsup>Suc (4 * rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>) t_wcode_main stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2159
  (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4 * rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2160
apply(rule_tac x = "Suc 0" in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2161
apply(simp add: steps.simps shift_length)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2162
apply(rule_tac x = lnb in exI, rule_tac x = rnb in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2163
apply(simp add: tstep.simps new_tape.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2164
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2165
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2166
lemma wcode_fourtimes_case:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2167
  shows "\<exists>stp ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2168
  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 =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2169
  (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2170
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2171
  have "\<exists>stp ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2172
  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 =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2173
  (t_twice_len + 14, Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2174
    using wcode_fourtimes_case_first_correctness[of ires rs m n]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2175
    apply(simp add: wcode_fourtimes_case_inv.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2176
    apply(rule_tac x = na in exI, rule_tac x = ln in exI,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2177
          rule_tac x = rn in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2178
    apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2179
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2180
  from this obtain stpa lna rna where stp1:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2181
    "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 =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2182
  (t_twice_len + 14, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2183
  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>)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2184
                     t_wcode_main stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2185
          (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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2186
    using t_fourtimes_append[of " Bk\<^bsup>lna\<^esup> @ Oc # ires" "rs + 1" rna]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2187
    apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2188
    apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2189
    apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2190
    apply(simp add: t_wcode_main_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2191
    apply(rule_tac x = stp in exI, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2192
          rule_tac x = "ln + lna" in exI,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2193
          rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2194
    apply(simp add: exp_ind_def[THEN sym] exp_add[THEN sym])
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2195
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2196
  from this obtain stpb lnb rnb where stp2:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2197
    "steps (t_twice_len + 14, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rna\<^esup>)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2198
                     t_wcode_main stpb =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2199
       (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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2200
    by blast
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2201
  have "\<exists>stp ln rn. steps (t_twice_len + 14 + t_fourtimes_len,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2202
    Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires,  Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2203
    t_wcode_main stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2204
    (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2205
    apply(rule wcode_jump2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2206
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2207
  from this obtain stpc lnc rnc where stp3: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2208
    "steps (t_twice_len + 14 + t_fourtimes_len,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2209
    Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires,  Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2210
    t_wcode_main stpc =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2211
    (Suc 0, Bk # Bk\<^bsup>lnc\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnc\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2212
    by blast
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2213
  from stp1 stp2 stp3 show "?thesis"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2214
    apply(rule_tac x = "stpa + stpb + stpc" in exI,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2215
          rule_tac x = lnc in exI, rule_tac x = rnc in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2216
    apply(simp add: steps_add)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2217
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2218
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2219
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2220
(**********************************************************)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2221
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2222
fun wcode_on_left_moving_3_B :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2223
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2224
  "wcode_on_left_moving_3_B ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2225
       (\<exists> ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # Bk # Bk # ires \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2226
                    r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2227
                    ml + mr > Suc 0 \<and> mr > 0 )"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2228
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2229
fun wcode_on_left_moving_3_O :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2230
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2231
  "wcode_on_left_moving_3_O ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2232
         (\<exists> ln rn. l = Bk # Bk # ires \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2233
                   r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2234
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2235
fun wcode_on_left_moving_3 :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2236
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2237
  "wcode_on_left_moving_3 ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2238
       (wcode_on_left_moving_3_B ires rs (l, r) \<or>  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2239
        wcode_on_left_moving_3_O ires rs (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2240
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2241
fun wcode_on_checking_3 :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2242
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2243
  "wcode_on_checking_3 ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2244
         (\<exists> ln rn. l = Bk # ires \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2245
             r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2246
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2247
fun wcode_goon_checking_3 :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2248
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2249
  "wcode_goon_checking_3 ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2250
         (\<exists> ln rn. l = ires \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2251
             r = Bk # Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2252
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2253
fun wcode_stop :: "bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2254
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2255
  "wcode_stop ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2256
          (\<exists> ln rn. l = Bk # ires \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2257
             r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2258
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2259
fun wcode_halt_case_inv :: "nat \<Rightarrow> bin_inv_t"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2260
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2261
  "wcode_halt_case_inv st ires rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2262
          (if st = 0 then wcode_stop ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2263
           else if st = Suc 0 then wcode_on_left_moving_3 ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2264
           else if st = Suc (Suc 0) then wcode_on_checking_3 ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2265
           else if st = 7 then wcode_goon_checking_3 ires rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2266
           else False)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2267
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2268
fun wcode_halt_case_state :: "t_conf \<Rightarrow> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2269
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2270
  "wcode_halt_case_state (st, l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2271
           (if st = 1 then 5
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2272
            else if st = Suc (Suc 0) then 4
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2273
            else if st = 7 then 3
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2274
            else 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2275
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2276
fun wcode_halt_case_step :: "t_conf \<Rightarrow> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2277
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2278
  "wcode_halt_case_step (st, l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2279
         (if st = 1 then length l
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2280
         else 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2281
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2282
fun wcode_halt_case_measure :: "t_conf \<Rightarrow> nat \<times> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2283
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2284
  "wcode_halt_case_measure (st, l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2285
     (wcode_halt_case_state (st, l, r), 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2286
      wcode_halt_case_step (st, l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2287
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2288
definition wcode_halt_case_le :: "(t_conf \<times> t_conf) set"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2289
  where "wcode_halt_case_le \<equiv> (inv_image lex_pair wcode_halt_case_measure)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2290
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2291
lemma wf_wcode_halt_case_le[intro]: "wf wcode_halt_case_le"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2292
by(auto intro:wf_inv_image simp: wcode_halt_case_le_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2293
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2294
declare wcode_on_left_moving_3_B.simps[simp del] wcode_on_left_moving_3_O.simps[simp del]  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2295
        wcode_on_checking_3.simps[simp del] wcode_goon_checking_3.simps[simp del] 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2296
        wcode_on_left_moving_3.simps[simp del] wcode_stop.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2297
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2298
lemmas wcode_halt_invs = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2299
  wcode_on_left_moving_3_B.simps wcode_on_left_moving_3_O.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2300
  wcode_on_checking_3.simps wcode_goon_checking_3.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2301
  wcode_on_left_moving_3.simps wcode_stop.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2302
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2303
lemma [simp]: "fetch t_wcode_main 7 Bk = (R, 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2304
apply(simp add: fetch.simps t_wcode_main_def nth_append nth_of.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2305
                t_wcode_main_first_part_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2306
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2307
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2308
lemma [simp]: "wcode_on_left_moving_3 ires rs (b, [])  = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2309
apply(simp only: wcode_halt_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2310
apply(simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2311
done    
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2312
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2313
lemma [simp]: "wcode_on_checking_3 ires rs (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2314
apply(simp add: wcode_halt_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2315
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2316
              
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2317
lemma [simp]: "wcode_goon_checking_3 ires rs (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2318
apply(simp add: wcode_halt_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2319
done 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2320
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2321
lemma [simp]: "wcode_on_left_moving_3 ires rs (b, Bk # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2322
 \<Longrightarrow> wcode_on_left_moving_3 ires rs (tl b, hd b # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2323
apply(simp only: wcode_halt_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2324
apply(erule_tac disjE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2325
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2326
apply(case_tac ml, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2327
apply(rule_tac x = "mr - 2" in exI, rule_tac x = rn in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2328
apply(case_tac mr, simp, simp add: exp_ind, simp add: exp_ind[THEN sym])
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2329
apply(rule_tac disjI1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2330
apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2331
      rule_tac x = rn in exI, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2332
apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2333
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2334
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2335
lemma [simp]: "wcode_goon_checking_3 ires rs (b, Bk # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2336
  (b = [] \<longrightarrow> wcode_stop ires rs ([Bk], list)) \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2337
  (b \<noteq> [] \<longrightarrow> wcode_stop ires rs (Bk # b, list))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2338
apply(auto simp: wcode_halt_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2339
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2340
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2341
lemma [simp]: "wcode_on_left_moving_3 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2342
apply(auto simp: wcode_halt_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2343
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2344
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2345
lemma [simp]: "wcode_on_left_moving_3 ires rs (b, Oc # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2346
               wcode_on_checking_3 ires rs (tl b, hd b # Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2347
apply(simp add:wcode_halt_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2348
apply(case_tac [!] mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2349
done     
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2350
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2351
lemma [simp]: "wcode_on_checking_3 ires rs (b, Oc # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2352
apply(auto simp: wcode_halt_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2353
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2354
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2355
lemma [simp]: "wcode_on_left_moving_3 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2356
apply(simp add: wcode_halt_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2357
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2358
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2359
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2360
lemma [simp]: "wcode_on_checking_3 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2361
apply(auto simp: wcode_halt_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2362
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2363
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2364
lemma [simp]: "wcode_on_checking_3 ires rs (b, Bk # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2365
  wcode_goon_checking_3 ires rs (tl b, hd b # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2366
apply(auto simp: wcode_halt_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2367
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2368
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2369
lemma [simp]: "wcode_goon_checking_3 ires rs (b, Oc # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2370
apply(simp add: wcode_goon_checking_3.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2371
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2372
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2373
lemma t_halt_case_correctness: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2374
shows "let P = (\<lambda> (st, l, r). st = 0) in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2375
       let Q = (\<lambda> (st, l, r). wcode_halt_case_inv st ires rs (l, r)) in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2376
       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
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2377
       \<exists> n .P (f n) \<and> Q (f (n::nat))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2378
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2379
  let ?P = "(\<lambda> (st, l, r). st = 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2380
  let ?Q = "(\<lambda> (st, l, r). wcode_halt_case_inv st ires rs (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2381
  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)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2382
  have "\<exists> n. ?P (?f n) \<and> ?Q (?f (n::nat))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2383
  proof(rule_tac halt_lemma2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2384
    show "wf wcode_halt_case_le" by auto
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2385
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2386
    show "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2387
                    ?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_halt_case_le"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2388
      apply(rule_tac allI, rule_tac impI, case_tac "?f na")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2389
      apply(simp add: tstep_red tstep.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2390
      apply(case_tac c, simp, case_tac [2] aa)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2391
      apply(simp_all split: if_splits add: new_tape.simps wcode_halt_case_le_def lex_pair_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2392
      done      
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2393
  next 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2394
    show "?Q (?f 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2395
      apply(simp add: steps.simps wcode_halt_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2396
      apply(rule_tac x = "Suc m" in exI, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2397
      apply(rule_tac x = "Suc 0" in exI, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2398
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2399
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2400
    show "\<not> ?P (?f 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2401
      apply(simp add: steps.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2402
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2403
  qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2404
  thus "?thesis"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2405
    apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2406
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2407
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2408
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2409
declare wcode_halt_case_inv.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2410
lemma [intro]: "\<exists> xs. (<rev list @ [aa::nat]> :: block list) = Oc # xs"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2411
apply(case_tac "rev list", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2412
apply(simp add: tape_of_nat_abv tape_of_nat_list.simps exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2413
apply(case_tac list, simp, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2414
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2415
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2416
lemma wcode_halt_case:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2417
  "\<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>)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2418
  t_wcode_main stp  = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2419
  using t_halt_case_correctness[of ires rs m n]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2420
apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2421
apply(erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2422
apply(case_tac "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2423
                Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main na")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2424
apply(auto simp: wcode_halt_case_inv.simps wcode_stop.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2425
apply(rule_tac x = na in exI, rule_tac x = ln in exI, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2426
      rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2427
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2428
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2429
lemma bl_bin_one: "bl_bin [Oc] =  Suc 0"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2430
apply(simp add: bl_bin.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2431
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2432
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2433
lemma t_wcode_main_lemma_pre:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2434
  "\<lbrakk>args \<noteq> []; lm = <args::nat list>\<rbrakk> \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2435
       \<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
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2436
                    stp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2437
      = (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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2438
proof(induct "length args" arbitrary: args lm rs m n, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2439
  fix x args lm rs m n
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2440
  assume ind:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2441
    "\<And>args lm rs m n.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2442
    \<lbrakk>x = length args; (args::nat list) \<noteq> []; lm = <args>\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2443
    \<Longrightarrow> \<exists>stp ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2444
    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 =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2445
    (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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2446
  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2447
    and h: "Suc x = length args" "(args::nat list) \<noteq> []" "lm = <args>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2448
  from h have "\<exists> (a::nat) xs. args = xs @ [a]"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2449
    apply(rule_tac x = "last args" in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2450
    apply(rule_tac x = "butlast args" in exI, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2451
    done    
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2452
  from this obtain a xs where "args = xs @ [a]" by blast
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2453
  from h and this show
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2454
    "\<exists>stp ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2455
    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 =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2456
    (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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2457
  proof(case_tac "xs::nat list", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2458
    show "\<exists>stp ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2459
      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 =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2460
      (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<a>) + rs * 2 ^ a\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2461
    proof(induct "a" arbitrary: m n rs ires, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2462
      fix m n rs ires
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2463
      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>)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2464
        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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2465
        apply(simp add: bl_bin_one)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2466
        apply(rule_tac wcode_halt_case)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2467
        done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2468
    next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2469
      fix a m n rs ires
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2470
      assume ind2: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2471
        "\<And>m n rs ires.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2472
        \<exists>stp ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2473
        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 =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2474
        (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<a>) + rs * 2 ^ a\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2475
      show "\<exists>stp ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2476
        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 =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2477
        (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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2478
      proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2479
        have "\<exists>stp ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2480
          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 =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2481
          (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2482
          apply(simp add: tape_of_nat)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2483
          using wcode_double_case[of m "Oc\<^bsup>a\<^esup> @ Bk # Bk # ires" rs n]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2484
          apply(simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2485
          done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2486
        from this obtain stpa lna rna where stp1:  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2487
          "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 =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2488
          (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
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2489
        moreover have 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2490
          "\<exists>stp ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2491
          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 =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2492
          (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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2493
          using ind2[of lna ires "2*rs + 2" rna] by simp   
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2494
        from this obtain stpb lnb rnb where stp2:  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2495
          "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 =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2496
          (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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2497
          by blast
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2498
        from stp1 and stp2 show "?thesis"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2499
          apply(rule_tac x = "stpa + stpb" in exI,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2500
            rule_tac x = lnb in exI, rule_tac x = rnb in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2501
          apply(simp add: steps_add bl_bin_nat_Suc exponent_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2502
          done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2503
      qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2504
    qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2505
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2506
    fix aa list
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2507
    assume g: "Suc x = length args" "args \<noteq> []" "lm = <args>" "args = xs @ [a::nat]" "xs = (aa::nat) # list"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2508
    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 =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2509
      (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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2510
    proof(induct a arbitrary: m n rs args lm, simp_all add: tape_of_nl_rev, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2511
        simp only: tape_of_nl_cons_app1, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2512
      fix m n rs args lm
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2513
      have "\<exists>stp ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2514
        steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # rev (<(aa::nat) # list>) @ Bk # Bk # ires,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2515
        Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2516
        (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ rev (<aa # list>) @ Bk # Bk # ires, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2517
        Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2518
        proof(simp add: tape_of_nl_rev)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2519
          have "\<exists> xs. (<rev list @ [aa]>) = Oc # xs" by auto           
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2520
          from this obtain xs where "(<rev list @ [aa]>) = Oc # xs" ..
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2521
          thus "\<exists>stp ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2522
            steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # <rev list @ [aa]> @ Bk # Bk # ires,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2523
            Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2524
            (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ <rev list @ [aa]> @ Bk # Bk # ires, Bk # Oc\<^bsup>5 + 4 * rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2525
            apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2526
            using wcode_fourtimes_case[of m "xs @ Bk # Bk # ires" rs n]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2527
            apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2528
            done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2529
        qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2530
      from this obtain stpa lna rna where stp1:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2531
        "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # rev (<aa # list>) @ Bk # Bk # ires,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2532
        Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2533
        (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<aa # list>) @ Bk # Bk # ires, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2534
        Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2535
      from g have 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2536
        "\<exists> stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2537
        Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp = (0, Bk # ires, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2538
        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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2539
         apply(rule_tac args = "(aa::nat)#list" in ind, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2540
         done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2541
       from this obtain stpb lnb rnb where stp2:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2542
         "steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2543
         Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb = (0, Bk # ires, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2544
         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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2545
         by blast
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2546
       from stp1 and stp2 and h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2547
       show "\<exists>stp ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2548
         steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # <rev list @ [aa]> @ Bk # Bk # ires,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2549
         Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2550
         (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk #
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2551
         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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2552
         apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2553
           rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_rev)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2554
         done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2555
     next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2556
       fix ab m n rs args lm
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2557
       assume ind2:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2558
         "\<And> m n rs args lm.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2559
         \<lbrakk>lm = <aa # list @ [ab]>; args = aa # list @ [ab]\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2560
         \<Longrightarrow> \<exists>stp ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2561
         steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ <ab # rev list @ [aa]> @ Bk # Bk # ires,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2562
         Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2563
         (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk #
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2564
         Bk # Oc\<^bsup>bl_bin (<aa # list @ [ab]>) + rs * 2 ^ (length (<aa # list @ [ab]>) - Suc 0)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2565
         and k: "args = aa # list @ [Suc ab]" "lm = <aa # list @ [Suc ab]>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2566
       show "\<exists>stp ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2567
         steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ <Suc ab # rev list @ [aa]> @ Bk # Bk # ires,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2568
         Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2569
         (0, Bk # ires,Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2570
         Bk # Oc\<^bsup>bl_bin (<aa # list @ [Suc ab]>) + rs * 2 ^ (length (<aa # list @ [Suc ab]>) - Suc 0)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2571
       proof(simp add: tape_of_nl_cons_app1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2572
         have "\<exists>stp ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2573
           steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc\<^bsup>Suc (Suc ab)\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2574
           Bk # Oc # Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2575
           = (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc\<^bsup>Suc ab\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2576
           Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2577
           using wcode_double_case[of m "Oc\<^bsup>ab\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2578
                                      rs n]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2579
           apply(simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2580
           done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2581
         from this obtain stpa lna rna where stp1:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2582
           "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc\<^bsup>Suc (Suc ab)\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2583
           Bk # Oc # Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2584
           = (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ Oc\<^bsup>Suc ab\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2585
           Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2586
         from k have 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2587
           "\<exists> stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ <ab # rev list @ [aa]> @ Bk # Bk # ires,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2588
           Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2589
           = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk #
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2590
           Bk # Oc\<^bsup>bl_bin (<aa # list @ [ab]> ) +  (2*rs + 2)* 2^(length (<aa # list @ [ab]>) - Suc 0)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2591
           apply(rule_tac ind2, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2592
           done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2593
         from this obtain stpb lnb rnb where stp2: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2594
           "steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @  <ab # rev list @ [aa]> @ Bk # Bk # ires,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2595
           Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2596
           = (0, Bk # ires, Bk # Oc # Bk\<^bsup>lnb\<^esup> @ Bk #
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2597
           Bk # Oc\<^bsup>bl_bin (<aa # list @ [ab]> ) +  (2*rs + 2)* 2^(length (<aa # list @ [ab]>) - Suc 0)\<^esup> @ Bk\<^bsup>rnb\<^esup>)" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2598
           by blast
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2599
         from stp1 and stp2 show 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2600
           "\<exists>stp ln rn.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2601
           steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc\<^bsup>Suc (Suc ab)\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2602
           Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2603
           (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2604
           Oc\<^bsup>bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # <list @ [Suc ab]>) + rs * (2 * 2 ^ (aa + length (<list @ [Suc ab]>)))\<^esup> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2605
           @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2606
           apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2607
             rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_cons_app1 exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2608
           done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2609
       qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2610
     qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2611
   qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2612
 qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2613
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2614
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2615
         
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2616
(* turing_shift can be used*)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2617
term t_wcode_main
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2618
definition t_wcode_prepare :: "tprog"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2619
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2620
  "t_wcode_prepare \<equiv> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2621
         [(W1, 2), (L, 1), (L, 3), (R, 2), (R, 4), (W0, 3),
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2622
          (R, 4), (R, 5), (R, 6), (R, 5), (R, 7), (R, 5),
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2623
          (W1, 7), (L, 0)]"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2624
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2625
fun wprepare_add_one :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2626
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2627
  "wprepare_add_one m lm (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2628
      (\<exists> rn. l = [] \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2629
               (r = <m # lm> @ Bk\<^bsup>rn\<^esup> \<or> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2630
                r = Bk # <m # lm> @ Bk\<^bsup>rn\<^esup>))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2631
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2632
fun wprepare_goto_first_end :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2633
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2634
  "wprepare_goto_first_end m lm (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2635
      (\<exists> ml mr rn. l = Oc\<^bsup>ml\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2636
                      r = Oc\<^bsup>mr\<^esup> @ Bk # <lm> @ Bk\<^bsup>rn\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2637
                      ml + mr = Suc (Suc m))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2638
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2639
fun wprepare_erase :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow>  bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2640
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2641
  "wprepare_erase m lm (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2642
     (\<exists> rn. l = Oc\<^bsup>Suc m\<^esup> \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2643
               tl r = Bk # <lm> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2644
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2645
fun wprepare_goto_start_pos_B :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2646
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2647
  "wprepare_goto_start_pos_B m lm (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2648
     (\<exists> rn. l = Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2649
               r = Bk # <lm> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2650
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2651
fun wprepare_goto_start_pos_O :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2652
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2653
  "wprepare_goto_start_pos_O m lm (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2654
     (\<exists> rn. l = Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2655
               r = <lm> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2656
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2657
fun wprepare_goto_start_pos :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2658
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2659
  "wprepare_goto_start_pos m lm (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2660
       (wprepare_goto_start_pos_B m lm (l, r) \<or>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2661
        wprepare_goto_start_pos_O m lm (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2662
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2663
fun wprepare_loop_start_on_rightmost :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2664
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2665
  "wprepare_loop_start_on_rightmost m lm (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2666
     (\<exists> rn mr. rev l @ r = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm> @ Bk\<^bsup>rn\<^esup> \<and> l \<noteq> [] \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2667
                       r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2668
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2669
fun wprepare_loop_start_in_middle :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2670
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2671
  "wprepare_loop_start_in_middle m lm (l, r) =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2672
     (\<exists> rn (mr:: nat) (lm1::nat list). 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2673
  rev l @ r = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm> @ Bk\<^bsup>rn\<^esup> \<and> l \<noteq> [] \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2674
  r = Oc\<^bsup>mr\<^esup> @ Bk # <lm1> @ Bk\<^bsup>rn\<^esup> \<and> lm1 \<noteq> [])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2675
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2676
fun wprepare_loop_start :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2677
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2678
  "wprepare_loop_start m lm (l, r) = (wprepare_loop_start_on_rightmost m lm (l, r) \<or> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2679
                                      wprepare_loop_start_in_middle m lm (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2680
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2681
fun wprepare_loop_goon_on_rightmost :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2682
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2683
  "wprepare_loop_goon_on_rightmost m lm (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2684
     (\<exists> rn. l = Bk # <rev lm> @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2685
               r = Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2686
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2687
fun wprepare_loop_goon_in_middle :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2688
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2689
  "wprepare_loop_goon_in_middle m lm (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2690
     (\<exists> rn (mr:: nat) (lm1::nat list). 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2691
  rev l @ r = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm> @ Bk\<^bsup>rn\<^esup> \<and> l \<noteq> [] \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2692
                     (if lm1 = [] then r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2693
                     else r = Oc\<^bsup>mr\<^esup> @ Bk # <lm1> @ Bk\<^bsup>rn\<^esup>) \<and> mr > 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2694
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2695
fun wprepare_loop_goon :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2696
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2697
  "wprepare_loop_goon m lm (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2698
              (wprepare_loop_goon_in_middle m lm (l, r) \<or> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2699
               wprepare_loop_goon_on_rightmost m lm (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2700
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2701
fun wprepare_add_one2 :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2702
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2703
  "wprepare_add_one2 m lm (l, r) =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2704
          (\<exists> rn. l = Bk # Bk # <rev lm> @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2705
               (r = [] \<or> tl r = Bk\<^bsup>rn\<^esup>))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2706
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2707
fun wprepare_stop :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2708
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2709
  "wprepare_stop m lm (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2710
         (\<exists> rn. l = Bk # <rev lm> @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2711
               r = Bk # Oc # Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2712
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2713
fun wprepare_inv :: "nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2714
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2715
  "wprepare_inv st m lm (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2716
        (if st = 0 then wprepare_stop m lm (l, r) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2717
         else if st = Suc 0 then wprepare_add_one m lm (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2718
         else if st = Suc (Suc 0) then wprepare_goto_first_end m lm (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2719
         else if st = Suc (Suc (Suc 0)) then wprepare_erase m lm (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2720
         else if st = 4 then wprepare_goto_start_pos m lm (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2721
         else if st = 5 then wprepare_loop_start m lm (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2722
         else if st = 6 then wprepare_loop_goon m lm (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2723
         else if st = 7 then wprepare_add_one2 m lm (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2724
         else False)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2725
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2726
fun wprepare_stage :: "t_conf \<Rightarrow> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2727
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2728
  "wprepare_stage (st, l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2729
      (if st \<ge> 1 \<and> st \<le> 4 then 3
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2730
       else if st = 5 \<or> st = 6 then 2
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2731
       else 1)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2732
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2733
fun wprepare_state :: "t_conf \<Rightarrow> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2734
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2735
  "wprepare_state (st, l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2736
       (if st = 1 then 4
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2737
        else if st = Suc (Suc 0) then 3
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2738
        else if st = Suc (Suc (Suc 0)) then 2
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2739
        else if st = 4 then 1
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2740
        else if st = 7 then 2
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2741
        else 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2742
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2743
fun wprepare_step :: "t_conf \<Rightarrow> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2744
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2745
  "wprepare_step (st, l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2746
      (if st = 1 then (if hd r = Oc then Suc (length l)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2747
                       else 0)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2748
       else if st = Suc (Suc 0) then length r
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2749
       else if st = Suc (Suc (Suc 0)) then (if hd r = Oc then 1
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2750
                            else 0)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2751
       else if st = 4 then length r
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2752
       else if st = 5 then Suc (length r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2753
       else if st = 6 then (if r = [] then 0 else Suc (length r))
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2754
       else if st = 7 then (if (r \<noteq> [] \<and> hd r = Oc) then 0
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2755
                            else 1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2756
       else 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2757
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2758
fun wcode_prepare_measure :: "t_conf \<Rightarrow> nat \<times> nat \<times> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2759
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2760
  "wcode_prepare_measure (st, l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2761
     (wprepare_stage (st, l, r), 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2762
      wprepare_state (st, l, r), 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2763
      wprepare_step (st, l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2764
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2765
definition wcode_prepare_le :: "(t_conf \<times> t_conf) set"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2766
  where "wcode_prepare_le \<equiv> (inv_image lex_triple wcode_prepare_measure)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2767
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2768
lemma [intro]: "wf lex_pair"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2769
by(auto intro:wf_lex_prod simp:lex_pair_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2770
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2771
lemma wf_wcode_prepare_le[intro]: "wf wcode_prepare_le"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2772
by(auto intro:wf_inv_image simp: wcode_prepare_le_def 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2773
           recursive.lex_triple_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2774
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2775
declare wprepare_add_one.simps[simp del] wprepare_goto_first_end.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2776
        wprepare_erase.simps[simp del] wprepare_goto_start_pos.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2777
        wprepare_loop_start.simps[simp del] wprepare_loop_goon.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2778
        wprepare_add_one2.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2779
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2780
lemmas wprepare_invs = wprepare_add_one.simps wprepare_goto_first_end.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2781
        wprepare_erase.simps wprepare_goto_start_pos.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2782
        wprepare_loop_start.simps wprepare_loop_goon.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2783
        wprepare_add_one2.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2784
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2785
declare wprepare_inv.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2786
lemma [simp]: "fetch t_wcode_prepare (Suc 0) Bk = (W1, 2)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2787
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2788
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2789
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2790
lemma [simp]: "fetch t_wcode_prepare (Suc 0) Oc = (L, 1)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2791
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2792
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2793
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2794
lemma [simp]: "fetch t_wcode_prepare (Suc (Suc 0)) Bk = (L, 3)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2795
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2796
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2797
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2798
lemma [simp]: "fetch t_wcode_prepare (Suc (Suc 0)) Oc = (R, 2)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2799
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2800
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2801
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2802
lemma [simp]: "fetch t_wcode_prepare (Suc (Suc (Suc 0))) Bk = (R, 4)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2803
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2804
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2805
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2806
lemma [simp]: "fetch t_wcode_prepare (Suc (Suc (Suc 0))) Oc = (W0, 3)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2807
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2808
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2809
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2810
lemma [simp]: "fetch t_wcode_prepare 4 Bk = (R, 4)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2811
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2812
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2813
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2814
lemma [simp]: "fetch t_wcode_prepare 4 Oc = (R, 5)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2815
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2816
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2817
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2818
lemma [simp]: "fetch t_wcode_prepare 5 Oc = (R, 5)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2819
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2820
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2821
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2822
lemma [simp]: "fetch t_wcode_prepare 5 Bk = (R, 6)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2823
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2824
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2825
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2826
lemma [simp]: "fetch t_wcode_prepare 6 Oc = (R, 5)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2827
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2828
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2829
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2830
lemma [simp]: "fetch t_wcode_prepare 6 Bk = (R, 7)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2831
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2832
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2833
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2834
lemma [simp]: "fetch t_wcode_prepare 7 Oc = (L, 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2835
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2836
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2837
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2838
lemma [simp]: "fetch t_wcode_prepare 7 Bk = (W1, 7)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2839
apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2840
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2841
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2842
lemma tape_of_nl_not_null: "lm \<noteq> [] \<Longrightarrow> <lm::nat list> \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2843
apply(case_tac lm, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2844
apply(case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2845
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2846
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2847
lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_add_one m lm (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2848
apply(simp add: wprepare_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2849
apply(simp add: tape_of_nl_not_null)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2850
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2851
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2852
lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_goto_first_end m lm (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2853
apply(simp add: wprepare_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2854
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2855
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2856
lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_erase m lm (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2857
apply(simp add: wprepare_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2858
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2859
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2860
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2861
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2862
lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_goto_start_pos m lm (b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2863
apply(simp add: wprepare_invs tape_of_nl_not_null)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2864
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2865
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2866
lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [])\<rbrakk> \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2867
apply(simp add: wprepare_invs tape_of_nl_not_null, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2868
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2869
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2870
lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [])\<rbrakk> \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2871
                                  wprepare_loop_goon m lm (Bk # b, [])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2872
apply(simp only: wprepare_invs tape_of_nl_not_null)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2873
apply(erule_tac disjE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2874
apply(rule_tac disjI2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2875
apply(simp add: wprepare_loop_start_on_rightmost.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2876
                wprepare_loop_goon_on_rightmost.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2877
apply(rule_tac rev_eq, simp add: tape_of_nl_rev)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2878
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2879
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2880
lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, [])\<rbrakk> \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2881
apply(simp only: wprepare_invs tape_of_nl_not_null, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2882
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2883
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2884
lemma [simp]:"\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, [])\<rbrakk> \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2885
  wprepare_add_one2 m lm (Bk # b, [])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2886
apply(simp only: wprepare_invs tape_of_nl_not_null, auto split: if_splits)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2887
apply(case_tac mr, simp, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2888
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2889
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2890
lemma [simp]: "wprepare_add_one2 m lm (b, []) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2891
apply(simp only: wprepare_invs tape_of_nl_not_null, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2892
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2893
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2894
lemma [simp]: "wprepare_add_one2 m lm (b, []) \<Longrightarrow> wprepare_add_one2 m lm (b, [Oc])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2895
apply(simp only: wprepare_invs tape_of_nl_not_null, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2896
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2897
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2898
lemma [simp]: "Bk # list = <(m::nat) # lm> @ ys = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2899
apply(case_tac lm, auto simp: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2900
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2901
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2902
lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_add_one m lm (b, Bk # list)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2903
       \<Longrightarrow> (b = [] \<longrightarrow> wprepare_goto_first_end m lm ([], Oc # list)) \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2904
           (b \<noteq> [] \<longrightarrow> wprepare_goto_first_end m lm (b, Oc # list))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2905
apply(simp only: wprepare_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2906
apply(rule_tac x = 0 in exI, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2907
apply(case_tac lm, simp, simp add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2908
apply(rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2909
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2910
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2911
lemma [simp]: "wprepare_goto_first_end m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2912
apply(simp only: wprepare_invs tape_of_nl_not_null, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2913
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2914
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2915
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2916
lemma [simp]: "wprepare_goto_first_end m lm (b, Bk # list) \<Longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2917
                          wprepare_erase m lm (tl b, hd b # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2918
apply(simp only: wprepare_invs tape_of_nl_not_null, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2919
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2920
apply(case_tac mr, auto simp: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2921
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2922
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2923
lemma [simp]: "wprepare_erase m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2924
apply(simp only: wprepare_invs exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2925
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2926
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2927
lemma [simp]: "wprepare_erase m lm (b, Bk # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2928
                           wprepare_goto_start_pos m lm (Bk # b, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2929
apply(simp only: wprepare_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2930
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2931
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2932
lemma [simp]: "\<lbrakk>wprepare_add_one m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2933
apply(simp only: wprepare_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2934
apply(case_tac lm, simp_all add: tape_of_nl_abv 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2935
                         tape_of_nat_list.simps exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2936
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2937
    
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2938
lemma [simp]: "\<lbrakk>lm \<noteq> [];  wprepare_goto_first_end m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2939
apply(simp only: wprepare_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2940
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2941
apply(simp add: tape_of_nl_not_null)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2942
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2943
     
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2944
lemma [simp]: "\<lbrakk>lm \<noteq> [];  wprepare_goto_first_end m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2945
apply(simp only: wprepare_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2946
apply(case_tac mr, simp_all add: exp_ind_def tape_of_nl_not_null)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2947
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2948
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2949
lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_erase m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2950
apply(simp only: wprepare_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2951
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2952
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2953
lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_erase m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2954
apply(simp only: wprepare_invs, auto simp: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2955
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2956
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2957
lemma [simp]: "\<lbrakk>lm \<noteq> [];  wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2958
apply(simp only: wprepare_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2959
apply(simp add: tape_of_nl_not_null)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2960
apply(case_tac lm, simp, case_tac list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2961
apply(simp_all add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2962
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2963
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2964
lemma [simp]: "\<lbrakk>lm \<noteq> [];  wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2965
apply(simp only: wprepare_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2966
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2967
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2968
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2969
lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2970
apply(simp only: wprepare_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2971
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2972
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2973
lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2974
  (list = [] \<longrightarrow> wprepare_add_one2 m lm (Bk # b, [])) \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2975
  (list \<noteq> [] \<longrightarrow> wprepare_add_one2 m lm (Bk # b, list))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2976
apply(simp only: wprepare_invs, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2977
apply(case_tac list, simp_all split: if_splits, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2978
apply(case_tac [1-3] mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2979
apply(case_tac mr, simp_all add: exp_ind_def tape_of_nl_not_null)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2980
apply(case_tac [1-2] mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2981
apply(case_tac rn, simp, case_tac nat, auto simp: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2982
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2983
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2984
lemma [simp]: "wprepare_add_one2 m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2985
apply(simp only: wprepare_invs, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2986
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2987
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2988
lemma [simp]: "wprepare_add_one2 m lm (b, Bk # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2989
      (list = [] \<longrightarrow> wprepare_add_one2 m lm (b, [Oc])) \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2990
      (list \<noteq> [] \<longrightarrow> wprepare_add_one2 m lm (b, Oc # list))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2991
apply(simp only:  wprepare_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2992
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2993
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2994
lemma [simp]: "wprepare_goto_first_end m lm (b, Oc # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2995
       \<Longrightarrow> (b = [] \<longrightarrow> wprepare_goto_first_end m lm ([Oc], list)) \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2996
           (b \<noteq> [] \<longrightarrow> wprepare_goto_first_end m lm (Oc # b, list))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2997
apply(simp only:  wprepare_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2998
apply(rule_tac x = 1 in exI, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  2999
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3000
apply(case_tac ml, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3001
apply(rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3002
apply(rule_tac x = "Suc ml" in exI, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3003
apply(rule_tac x = "mr - 1" in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3004
apply(case_tac mr, simp_all add: exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3005
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3006
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3007
lemma [simp]: "wprepare_erase m lm (b, Oc # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3008
apply(simp only: wprepare_invs, auto simp: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3009
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3010
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3011
lemma [simp]: "wprepare_erase m lm (b, Oc # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3012
  \<Longrightarrow> wprepare_erase m lm (b, Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3013
apply(simp  only:wprepare_invs, auto simp: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3014
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3015
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3016
lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3017
       \<Longrightarrow> wprepare_goto_start_pos m lm (Bk # b, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3018
apply(simp only:wprepare_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3019
apply(case_tac [!] lm, simp, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3020
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3021
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3022
lemma [simp]: "wprepare_loop_start m lm (b, aa) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3023
apply(simp only:wprepare_invs, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3024
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3025
lemma [elim]: "Bk # list = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup>  \<Longrightarrow> \<exists>rn. list = Bk\<^bsup>rn\<^esup>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3026
apply(case_tac mr, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3027
apply(case_tac rn, simp_all add: exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3028
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3029
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3030
lemma rev_equal_iff: "x = y \<Longrightarrow> rev x = rev y"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3031
by simp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3032
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3033
lemma tape_of_nl_false1:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3034
  "lm \<noteq> [] \<Longrightarrow> rev b @ [Bk] \<noteq> Bk\<^bsup>ln\<^esup> @ Oc # Oc\<^bsup>m\<^esup> @ Bk # Bk # <lm::nat list>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3035
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3036
apply(drule_tac rev_equal_iff, simp add: tape_of_nl_rev)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3037
apply(case_tac "rev lm")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3038
apply(case_tac [2] list, auto simp: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3039
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3040
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3041
lemma [simp]: "wprepare_loop_start_in_middle m lm (b, [Bk]) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3042
apply(simp add: wprepare_loop_start_in_middle.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3043
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3044
apply(case_tac lm1, simp, simp add: tape_of_nl_not_null)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3045
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3046
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3047
declare wprepare_loop_start_in_middle.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3048
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3049
declare wprepare_loop_start_on_rightmost.simps[simp del] 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3050
        wprepare_loop_goon_in_middle.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3051
        wprepare_loop_goon_on_rightmost.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3052
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3053
lemma [simp]: "wprepare_loop_goon_in_middle m lm (Bk # b, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3054
apply(simp add: wprepare_loop_goon_in_middle.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3055
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3056
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3057
lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [Bk])\<rbrakk> \<Longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3058
  wprepare_loop_goon m lm (Bk # b, [])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3059
apply(simp only: wprepare_invs, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3060
apply(simp add: wprepare_loop_goon_on_rightmost.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3061
  wprepare_loop_start_on_rightmost.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3062
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3063
apply(rule_tac rev_eq)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3064
apply(simp add: tape_of_nl_rev)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3065
apply(simp add: exp_ind_def[THEN sym] exp_ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3066
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3067
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3068
lemma [simp]: "wprepare_loop_start_on_rightmost m lm (b, Bk # a # lista)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3069
 \<Longrightarrow> wprepare_loop_goon_in_middle m lm (Bk # b, a # lista) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3070
apply(auto simp: wprepare_loop_start_on_rightmost.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3071
                 wprepare_loop_goon_in_middle.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3072
apply(case_tac [!] mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3073
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3074
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3075
lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start_on_rightmost m lm (b, Bk # a # lista)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3076
    \<Longrightarrow> wprepare_loop_goon_on_rightmost m lm (Bk # b, a # lista)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3077
apply(simp only: wprepare_loop_start_on_rightmost.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3078
                 wprepare_loop_goon_on_rightmost.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3079
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3080
apply(simp add: tape_of_nl_rev)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3081
apply(simp add: exp_ind_def[THEN sym] exp_ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3082
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3083
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3084
lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start_in_middle m lm (b, Bk # a # lista)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3085
  \<Longrightarrow> wprepare_loop_goon_on_rightmost m lm (Bk # b, a # lista) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3086
apply(simp add: wprepare_loop_start_in_middle.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3087
                wprepare_loop_goon_on_rightmost.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3088
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3089
apply(case_tac  "lm1::nat list", simp_all, case_tac  list, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3090
apply(simp add: tape_of_nl_abv tape_of_nat_list.simps tape_of_nat_abv exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3091
apply(case_tac [!] rna, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3092
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3093
apply(case_tac lm1, simp, case_tac list, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3094
apply(simp_all add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def tape_of_nat_abv)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3095
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3096
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3097
lemma [simp]: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3098
  "\<lbrakk>lm \<noteq> []; wprepare_loop_start_in_middle m lm (b, Bk # a # lista)\<rbrakk> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3099
  \<Longrightarrow> wprepare_loop_goon_in_middle m lm (Bk # b, a # lista)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3100
apply(simp add: wprepare_loop_start_in_middle.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3101
               wprepare_loop_goon_in_middle.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3102
apply(rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3103
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3104
apply(case_tac lm1, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3105
apply(rule_tac x = "Suc aa" in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3106
apply(rule_tac x = list in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3107
apply(case_tac list, simp_all add: tape_of_nl_abv tape_of_nat_list.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3108
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3109
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3110
lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, Bk # a # lista)\<rbrakk> \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3111
  wprepare_loop_goon m lm (Bk # b, a # lista)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3112
apply(simp add: wprepare_loop_start.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3113
                wprepare_loop_goon.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3114
apply(erule_tac disjE, simp, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3115
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3116
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3117
lemma start_2_goon:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3118
  "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, Bk # list)\<rbrakk> \<Longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3119
   (list = [] \<longrightarrow> wprepare_loop_goon m lm (Bk # b, [])) \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3120
  (list \<noteq> [] \<longrightarrow> wprepare_loop_goon m lm (Bk # b, list))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3121
apply(case_tac list, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3122
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3123
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3124
lemma add_one_2_add_one: "wprepare_add_one m lm (b, Oc # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3125
  \<Longrightarrow> (hd b = Oc \<longrightarrow> (b = [] \<longrightarrow> wprepare_add_one m lm ([], Bk # Oc # list)) \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3126
                     (b \<noteq> [] \<longrightarrow> wprepare_add_one m lm (tl b, Oc # Oc # list))) \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3127
  (hd b \<noteq> Oc \<longrightarrow> (b = [] \<longrightarrow> wprepare_add_one m lm ([], Bk # Oc # list)) \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3128
                 (b \<noteq> [] \<longrightarrow> wprepare_add_one m lm (tl b, hd b # Oc # list)))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3129
apply(simp only: wprepare_add_one.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3130
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3131
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3132
lemma [simp]: "wprepare_loop_start m lm (b, Oc # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3133
apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3134
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3135
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3136
lemma [simp]: "wprepare_loop_start_on_rightmost m lm (b, Oc # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3137
  wprepare_loop_start_on_rightmost m lm (Oc # b, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3138
apply(simp add: wprepare_loop_start_on_rightmost.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3139
apply(rule_tac x = rn in exI, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3140
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3141
apply(case_tac rn, auto simp: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3142
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3143
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3144
lemma [simp]: "wprepare_loop_start_in_middle m lm (b, Oc # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3145
                wprepare_loop_start_in_middle m lm (Oc # b, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3146
apply(simp add: wprepare_loop_start_in_middle.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3147
apply(rule_tac x = rn in exI, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3148
apply(case_tac mr, simp, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3149
apply(rule_tac x = nat in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3150
apply(rule_tac x = lm1 in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3151
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3152
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3153
lemma start_2_start: "wprepare_loop_start m lm (b, Oc # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3154
       wprepare_loop_start m lm (Oc # b, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3155
apply(simp add: wprepare_loop_start.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3156
apply(erule_tac disjE, simp_all )
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3157
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3158
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3159
lemma [simp]: "wprepare_loop_goon m lm (b, Oc # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3160
apply(simp add: wprepare_loop_goon.simps     
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3161
                wprepare_loop_goon_in_middle.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3162
                wprepare_loop_goon_on_rightmost.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3163
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3164
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3165
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3166
lemma [simp]: "wprepare_goto_start_pos m lm (b, Oc # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3167
apply(simp add: wprepare_goto_start_pos.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3168
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3169
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3170
lemma [simp]: "wprepare_loop_goon_on_rightmost m lm (b, Oc # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3171
apply(simp add: wprepare_loop_goon_on_rightmost.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3172
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3173
lemma wprepare_loop1: "\<lbrakk>rev b @ Oc\<^bsup>mr\<^esup> =  Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm>; 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3174
         b \<noteq> []; 0 < mr; Oc # list = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup>\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3175
       \<Longrightarrow> wprepare_loop_start_on_rightmost m lm (Oc # b, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3176
apply(simp add: wprepare_loop_start_on_rightmost.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3177
apply(rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3178
apply(case_tac mr, simp, simp add: exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3179
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3180
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3181
lemma wprepare_loop2: "\<lbrakk>rev b @ Oc\<^bsup>mr\<^esup> @ Bk # <a # lista> = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm>;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3182
                b \<noteq> []; Oc # list = Oc\<^bsup>mr\<^esup> @ Bk # <(a::nat) # lista> @ Bk\<^bsup>rn\<^esup>\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3183
       \<Longrightarrow>  wprepare_loop_start_in_middle m lm (Oc # b, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3184
apply(simp add: wprepare_loop_start_in_middle.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3185
apply(rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3186
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3187
apply(rule_tac x = nat in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3188
apply(rule_tac x = "a#lista" in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3189
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3190
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3191
lemma [simp]: "wprepare_loop_goon_in_middle m lm (b, Oc # list) \<Longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3192
                wprepare_loop_start_on_rightmost m lm (Oc # b, list) \<or>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3193
                wprepare_loop_start_in_middle m lm (Oc # b, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3194
apply(simp add: wprepare_loop_goon_in_middle.simps split: if_splits)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3195
apply(case_tac lm1, simp_all add: wprepare_loop1 wprepare_loop2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3196
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3197
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3198
lemma [simp]: "wprepare_loop_goon m lm (b, Oc # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3199
  \<Longrightarrow>  wprepare_loop_start m lm (Oc # b, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3200
apply(simp add: wprepare_loop_goon.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3201
                wprepare_loop_start.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3202
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3203
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3204
lemma [simp]: "wprepare_add_one m lm (b, Oc # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3205
       \<Longrightarrow> b = [] \<longrightarrow> wprepare_add_one m lm ([], Bk # Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3206
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3207
apply(simp add: wprepare_add_one.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3208
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3209
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3210
lemma [simp]: "wprepare_goto_start_pos m [a] (b, Oc # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3211
              \<Longrightarrow> wprepare_loop_start_on_rightmost m [a] (Oc # b, list) "
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3212
apply(auto simp: wprepare_goto_start_pos.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3213
                 wprepare_loop_start_on_rightmost.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3214
apply(rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3215
apply(simp add: tape_of_nat_abv tape_of_nat_list.simps exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3216
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3217
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3218
lemma [simp]:  "wprepare_goto_start_pos m (a # aa # listaa) (b, Oc # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3219
       \<Longrightarrow>wprepare_loop_start_in_middle m (a # aa # listaa) (Oc # b, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3220
apply(auto simp: wprepare_goto_start_pos.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3221
                 wprepare_loop_start_in_middle.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3222
apply(rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3223
apply(simp add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3224
apply(rule_tac x = a in exI, rule_tac x = "aa#listaa" in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3225
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3226
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3227
lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_start_pos m lm (b, Oc # list)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3228
       \<Longrightarrow> wprepare_loop_start m lm (Oc # b, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3229
apply(case_tac lm, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3230
apply(case_tac lista, simp_all add: wprepare_loop_start.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3231
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3232
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3233
lemma [simp]: "wprepare_add_one2 m lm (b, Oc # list) \<Longrightarrow> b \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3234
apply(auto simp: wprepare_add_one2.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3235
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3236
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3237
lemma add_one_2_stop:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3238
  "wprepare_add_one2 m lm (b, Oc # list)      
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3239
  \<Longrightarrow>  wprepare_stop m lm (tl b, hd b # Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3240
apply(simp add: wprepare_stop.simps wprepare_add_one2.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3241
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3242
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3243
declare wprepare_stop.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3244
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3245
lemma wprepare_correctness:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3246
  assumes h: "lm \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3247
  shows "let P = (\<lambda> (st, l, r). st = 0) in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3248
  let Q = (\<lambda> (st, l, r). wprepare_inv st m lm (l, r)) in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3249
  let f = (\<lambda> stp. steps (Suc 0, [], (<m # lm>)) t_wcode_prepare stp) in
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3250
    \<exists> n .P (f n) \<and> Q (f n)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3251
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3252
  let ?P = "(\<lambda> (st, l, r). st = 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3253
  let ?Q = "(\<lambda> (st, l, r). wprepare_inv st m lm (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3254
  let ?f = "(\<lambda> stp. steps (Suc 0, [], (<m # lm>)) t_wcode_prepare stp)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3255
  have "\<exists> n. ?P (?f n) \<and> ?Q (?f n)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3256
  proof(rule_tac halt_lemma2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3257
    show "wf wcode_prepare_le" by auto
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3258
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3259
    show "\<forall> n. \<not> ?P (?f n) \<and> ?Q (?f n) \<longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3260
                 ?Q (?f (Suc n)) \<and> (?f (Suc n), ?f n) \<in> wcode_prepare_le"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3261
      using h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3262
      apply(rule_tac allI, rule_tac impI, case_tac "?f n", 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3263
            simp add: tstep_red tstep.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3264
      apply(case_tac c, simp, case_tac [2] aa)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3265
      apply(simp_all add: wprepare_inv.simps wcode_prepare_le_def new_tape.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3266
                          lex_triple_def lex_pair_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3267
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3268
                 split: if_splits)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3269
      apply(simp_all add: start_2_goon  start_2_start
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3270
                           add_one_2_add_one add_one_2_stop)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3271
      apply(auto simp: wprepare_add_one2.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3272
      done   
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3273
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3274
    show "?Q (?f 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3275
      apply(simp add: steps.simps wprepare_inv.simps wprepare_invs)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3276
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3277
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3278
    show "\<not> ?P (?f 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3279
      apply(simp add: steps.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3280
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3281
  qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3282
  thus "?thesis"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3283
    apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3284
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3285
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3286
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3287
lemma [intro]: "t_correct t_wcode_prepare"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3288
apply(simp add: t_correct.simps t_wcode_prepare_def iseven_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3289
apply(rule_tac x = 7 in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3290
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3291
    
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3292
lemma twice_len_even: "length (tm_of abc_twice) mod 2 = 0"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3293
apply(simp add: tm_even)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3294
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3295
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3296
lemma fourtimes_len_even: "length (tm_of abc_fourtimes) mod 2 = 0"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3297
apply(simp add: tm_even)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3298
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3299
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3300
lemma t_correct_termi: "t_correct tp \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3301
      list_all (\<lambda>(acn, st). (st \<le> Suc (length tp div 2))) (change_termi_state tp)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3302
apply(auto simp: t_correct.simps List.list_all_length)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3303
apply(erule_tac x = n in allE, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3304
apply(case_tac "tp!n", auto simp: change_termi_state.simps split: if_splits)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3305
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3306
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3307
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3308
lemma t_correct_shift:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3309
         "list_all (\<lambda>(acn, st). (st \<le> y)) tp \<Longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3310
          list_all (\<lambda>(acn, st). (st \<le> y + off)) (tshift tp off) "
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3311
apply(auto simp: t_correct.simps List.list_all_length)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3312
apply(erule_tac x = n in allE, simp add: shift_length)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3313
apply(case_tac "tp!n", auto simp: tshift.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3314
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3315
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3316
lemma [intro]: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3317
  "t_correct (tm_of abc_twice @ tMp (Suc 0) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3318
        (start_of twice_ly (length abc_twice) - Suc 0))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3319
apply(rule_tac t_compiled_correct, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3320
apply(simp add: twice_ly_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3321
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3322
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3323
lemma [intro]: "t_correct (tm_of abc_fourtimes @ tMp (Suc 0) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3324
   (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3325
apply(rule_tac t_compiled_correct, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3326
apply(simp add: fourtimes_ly_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3327
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3328
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3329
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3330
lemma [intro]: "t_correct t_wcode_main"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3331
apply(auto simp: t_wcode_main_def t_correct.simps shift_length 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3332
                 t_twice_def t_fourtimes_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3333
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3334
  show "iseven (60 + (length (tm_of abc_twice) +
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3335
                 length (tm_of abc_fourtimes)))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3336
    using twice_len_even fourtimes_len_even
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3337
    apply(auto simp: iseven_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3338
    apply(rule_tac x = "30 + q + qa" in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3339
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3340
next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3341
  show " list_all (\<lambda>(acn, s). s \<le> (60 + (length (tm_of abc_twice) + 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3342
           length (tm_of abc_fourtimes))) div 2) t_wcode_main_first_part"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3343
    apply(auto simp: t_wcode_main_first_part_def shift_length t_twice_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3344
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3345
next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3346
  have "list_all (\<lambda>(acn, s). s \<le> Suc (length (tm_of abc_twice @ tMp (Suc 0)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3347
    (start_of twice_ly (length abc_twice) - Suc 0)) div 2))
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3348
    (change_termi_state (tm_of abc_twice @ tMp (Suc 0) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3349
    (start_of twice_ly (length abc_twice) - Suc 0)))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3350
    apply(rule_tac t_correct_termi, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3351
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3352
  hence "list_all (\<lambda>(acn, s). s \<le>  Suc (length (tm_of abc_twice @ tMp (Suc 0)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3353
    (start_of twice_ly (length abc_twice) - Suc 0)) div 2) + 12)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3354
     (abacus.tshift (change_termi_state (tm_of abc_twice @ tMp (Suc 0) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3355
           (start_of twice_ly (length abc_twice) - Suc 0))) 12)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3356
    apply(rule_tac t_correct_shift, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3357
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3358
  thus  "list_all (\<lambda>(acn, s). s \<le> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3359
           (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3360
     (abacus.tshift (change_termi_state (tm_of abc_twice @ tMp (Suc 0)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3361
                 (start_of twice_ly (length abc_twice) - Suc 0))) 12)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3362
    apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3363
    apply(simp add: list_all_length, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3364
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3365
next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3366
  have "list_all (\<lambda>(acn, s). s \<le> Suc (length (tm_of abc_fourtimes @ tMp (Suc 0) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3367
    (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) div 2))
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3368
      (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3369
    (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) "
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3370
    apply(rule_tac t_correct_termi, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3371
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3372
  hence "list_all (\<lambda>(acn, s). s \<le> Suc (length (tm_of abc_fourtimes @ tMp (Suc 0) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3373
    (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) div 2) + (t_twice_len + 13))
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3374
    (abacus.tshift (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3375
    (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) (t_twice_len + 13))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3376
    apply(rule_tac t_correct_shift, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3377
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3378
  thus "list_all (\<lambda>(acn, s). s \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3379
    (abacus.tshift (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3380
    (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) (t_twice_len + 13))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3381
    apply(simp add: t_twice_len_def t_twice_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3382
    using twice_len_even fourtimes_len_even
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3383
    apply(auto simp: list_all_length)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3384
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3385
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3386
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3387
lemma [intro]: "t_correct (t_wcode_prepare |+| t_wcode_main)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3388
apply(auto intro: t_correct_add)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3389
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3390
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3391
lemma prepare_mainpart_lemma:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3392
  "args \<noteq> [] \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3393
  \<exists> stp ln rn. steps (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) stp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3394
              = (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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3395
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3396
  let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <m # args>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3397
  let ?Q1 = "\<lambda> (l, r). wprepare_stop m args (l, r)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3398
  let ?P2 = ?Q1
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3399
  let ?Q2 = "\<lambda> (l, r). (\<exists> ln rn. l = Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3400
                           r =  Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3401
  let ?P3 = "\<lambda> tp. False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3402
  assume h: "args \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3403
  have "?P1 \<turnstile>-> \<lambda> tp. (\<exists> stp tp'. steps (Suc 0, tp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3404
                      (t_wcode_prepare |+| t_wcode_main) stp = (0, tp') \<and> ?Q2 tp')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3405
  proof(rule_tac turing_merge.t_merge_halt[of t_wcode_prepare t_wcode_main ?P1 ?P2 ?P3 ?P3 ?Q1 ?Q2], 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3406
        auto simp: turing_merge_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3407
    show "\<exists>stp. case steps (Suc 0, [], <m # args>) t_wcode_prepare stp of (st, tp')
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3408
                  \<Rightarrow> st = 0 \<and> wprepare_stop m args tp'"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3409
      using wprepare_correctness[of args m] h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3410
      apply(simp, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3411
      apply(rule_tac x = n in exI, simp add: wprepare_inv.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3412
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3413
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3414
    fix a b
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3415
    assume "wprepare_stop m args (a, b)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3416
    thus "\<exists>stp. case steps (Suc 0, a, b) t_wcode_main stp of
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3417
      (st, tp') \<Rightarrow> (st = 0) \<and> (case tp' of (l, r) \<Rightarrow> l = Bk # Oc\<^bsup>Suc m\<^esup> \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3418
      (\<exists>ln rn. r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3419
      proof(simp only: wprepare_stop.simps, erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3420
        fix rn
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3421
        assume "a = Bk # <rev args> @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3422
                   b = Bk # Oc # Bk\<^bsup>rn\<^esup>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3423
        thus "?thesis"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3424
          using t_wcode_main_lemma_pre[of "args" "<args>" 0 "Oc\<^bsup>Suc m\<^esup>" 0 rn] h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3425
          apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3426
          apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3427
          apply(rule_tac x = stp in exI, simp add: tape_of_nl_rev, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3428
          done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3429
      qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3430
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3431
    show "wprepare_stop m args \<turnstile>-> wprepare_stop m args"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3432
      by(simp add: t_imply_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3433
  qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3434
  thus "\<exists> stp ln rn. steps (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) stp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3435
              = (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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3436
    apply(simp add: t_imply_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3437
    apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3438
    apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3439
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3440
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3441
      
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3442
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3443
lemma [simp]:  "tinres r r' \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3444
  fetch t ss (case r of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3445
  fetch t ss (case r' of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3446
apply(simp add: fetch.simps, auto split: if_splits simp: tinres_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3447
apply(case_tac [!] r', simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3448
apply(case_tac [!] n, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3449
apply(case_tac [!] r, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3450
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3451
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3452
lemma [intro]: "\<exists> n. (a::block)\<^bsup>n\<^esup> = []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3453
by auto
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3454
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3455
lemma [simp]: "\<lbrakk>tinres r r'; r \<noteq> []; r' \<noteq> []\<rbrakk> \<Longrightarrow> hd r = hd r'"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3456
apply(auto simp: tinres_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3457
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3458
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3459
lemma [intro]: "hd (Bk\<^bsup>Suc n\<^esup>) = Bk"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3460
apply(simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3461
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3462
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3463
lemma [simp]: "\<lbrakk>tinres r []; r \<noteq> []\<rbrakk> \<Longrightarrow> hd r = Bk"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3464
apply(auto simp: tinres_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3465
apply(case_tac n, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3466
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3467
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3468
lemma [simp]: "\<lbrakk>tinres [] r'; r' \<noteq> []\<rbrakk> \<Longrightarrow> hd r' = Bk"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3469
apply(auto simp: tinres_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3470
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3471
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3472
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>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3473
apply(case_tac r, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3474
apply(case_tac n, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3475
apply(rule_tac x = 0 in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3476
apply(rule_tac x = nat in exI, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3477
apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3478
apply(rule_tac x = n in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3479
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3480
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3481
lemma [simp]: "tinres r r' \<Longrightarrow> tinres (tl r) (tl r')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3482
apply(auto simp: tinres_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3483
apply(case_tac r', simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3484
apply(case_tac n, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3485
apply(rule_tac x = 0 in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3486
apply(rule_tac x = nat in exI, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3487
apply(rule_tac x = n in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3488
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3489
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3490
lemma [simp]: "\<lbrakk>tinres r [];  r \<noteq> []\<rbrakk> \<Longrightarrow> tinres (tl r) []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3491
apply(case_tac r, auto simp: tinres_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3492
apply(case_tac n, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3493
apply(rule_tac x = nat in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3494
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3495
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3496
lemma [simp]: "\<lbrakk>tinres [] r'\<rbrakk> \<Longrightarrow> tinres [] (tl r')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3497
apply(case_tac r', auto simp: tinres_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3498
apply(case_tac n, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3499
apply(rule_tac x = nat in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3500
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3501
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3502
lemma [simp]: "tinres r r' \<Longrightarrow> tinres (b # r) (b # r')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3503
apply(auto simp: tinres_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3504
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3505
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3506
lemma tinres_step2: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3507
  "\<lbrakk>tinres r r'; tstep (ss, l, r) t = (sa, la, ra); tstep (ss, l, r') t = (sb, lb, rb)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3508
    \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3509
apply(case_tac "ss = 0", simp add: tstep_0)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3510
apply(simp add: tstep.simps [simp del])
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3511
apply(case_tac "fetch t ss (case r of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3512
apply(auto simp: new_tape.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3513
apply(simp_all split: taction.splits if_splits)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3514
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3515
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3516
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3517
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3518
lemma tinres_steps2: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3519
  "\<lbrakk>tinres r r'; steps (ss, l, r) t stp = (sa, la, ra); steps (ss, l, r') t stp = (sb, lb, rb)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3520
    \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3521
apply(induct stp arbitrary: sa la ra sb lb rb, simp add: steps.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3522
apply(simp add: tstep_red)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3523
apply(case_tac "(steps (ss, l, r) t stp)")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3524
apply(case_tac "(steps (ss, l, r') t stp)")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3525
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3526
  fix stp sa la ra sb lb rb a b c aa ba ca
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3527
  assume ind: "\<And>sa la ra sb lb rb. \<lbrakk>steps (ss, l, r) t stp = (sa, la, ra); 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3528
    steps (ss, l, r') t stp = (sb, lb, rb)\<rbrakk> \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3529
  and h: " tinres r r'" "tstep (steps (ss, l, r) t stp) t = (sa, la, ra)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3530
         "tstep (steps (ss, l, r') t stp) t = (sb, lb, rb)" "steps (ss, l, r) t stp = (a, b, c)" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3531
         "steps (ss, l, r') t stp = (aa, ba, ca)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3532
  have "b = ba \<and> tinres c ca \<and> a = aa"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3533
    apply(rule_tac ind, simp_all add: h)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3534
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3535
  thus "la = lb \<and> tinres ra rb \<and> sa = sb"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3536
    apply(rule_tac l = b  and r = c  and ss = a and r' = ca   
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3537
            and t = t in tinres_step2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3538
    using h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3539
    apply(simp, simp, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3540
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3541
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3542
 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3543
definition t_wcode_adjust :: "tprog"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3544
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3545
  "t_wcode_adjust = [(W1, 1), (R, 2), (Nop, 2), (R, 3), (R, 3), (R, 4), 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3546
                   (L, 8), (L, 5), (L, 6), (W0, 5), (L, 6), (R, 7), 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3547
                   (W1, 2), (Nop, 7), (L, 9), (W0, 8), (L, 9), (L, 10), 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3548
                    (L, 11), (L, 10), (R, 0), (L, 11)]"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3549
                 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3550
lemma [simp]: "fetch t_wcode_adjust (Suc 0) Bk = (W1, 1)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3551
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3552
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3553
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3554
lemma [simp]: "fetch t_wcode_adjust (Suc 0) Oc = (R, 2)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3555
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3556
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3557
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3558
lemma [simp]: "fetch t_wcode_adjust (Suc (Suc 0)) Oc = (R, 3)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3559
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3560
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3561
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3562
lemma [simp]: "fetch t_wcode_adjust (Suc (Suc (Suc 0))) Oc = (R, 4)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3563
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3564
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3565
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3566
lemma [simp]: "fetch t_wcode_adjust  (Suc (Suc (Suc 0))) Bk = (R, 3)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3567
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3568
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3569
   
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3570
lemma [simp]: "fetch t_wcode_adjust 4 Bk = (L, 8)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3571
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3572
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3573
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3574
lemma [simp]: "fetch t_wcode_adjust 4 Oc = (L, 5)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3575
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3576
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3577
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3578
lemma [simp]: "fetch t_wcode_adjust 5 Oc = (W0, 5)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3579
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3580
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3581
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3582
lemma [simp]: "fetch t_wcode_adjust 5 Bk = (L, 6)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3583
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3584
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3585
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3586
lemma [simp]: "fetch t_wcode_adjust 6 Oc = (R, 7)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3587
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3588
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3589
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3590
lemma [simp]: "fetch t_wcode_adjust 6 Bk = (L, 6)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3591
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3592
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3593
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3594
lemma [simp]: "fetch t_wcode_adjust 7 Bk = (W1, 2)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3595
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3596
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3597
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3598
lemma [simp]: "fetch t_wcode_adjust 8 Bk = (L, 9)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3599
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3600
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3601
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3602
lemma [simp]: "fetch t_wcode_adjust 8 Oc = (W0, 8)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3603
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3604
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3605
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3606
lemma [simp]: "fetch t_wcode_adjust 9 Oc = (L, 10)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3607
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3608
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3609
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3610
lemma [simp]: "fetch t_wcode_adjust 9 Bk = (L, 9)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3611
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3612
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3613
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3614
lemma [simp]: "fetch t_wcode_adjust 10 Bk = (L, 11)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3615
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3616
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3617
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3618
lemma [simp]: "fetch t_wcode_adjust 10 Oc = (L, 10)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3619
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3620
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3621
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3622
lemma [simp]: "fetch t_wcode_adjust 11 Oc = (L, 11)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3623
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3624
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3625
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3626
lemma [simp]: "fetch t_wcode_adjust 11 Bk = (R, 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3627
apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3628
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3629
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3630
fun wadjust_start :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3631
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3632
  "wadjust_start m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3633
         (\<exists> ln rn. l = Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3634
                   tl r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3635
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3636
fun wadjust_loop_start :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3637
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3638
  "wadjust_loop_start m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3639
          (\<exists> ln rn ml mr. l = Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup>  \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3640
                          r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3641
                          ml + mr = Suc (Suc rs) \<and> mr > 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3642
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3643
fun wadjust_loop_right_move :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3644
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3645
  "wadjust_loop_right_move m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3646
   (\<exists> ml mr nl nr rn. l = Bk\<^bsup>nl\<^esup> @ Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3647
                      r = Bk\<^bsup>nr\<^esup> @ Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3648
                      ml + mr = Suc (Suc rs) \<and> mr > 0 \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3649
                      nl + nr > 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3650
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3651
fun wadjust_loop_check :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3652
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3653
  "wadjust_loop_check m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3654
  (\<exists> ml mr ln rn. l = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3655
                  r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> ml + mr = (Suc rs))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3656
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3657
fun wadjust_loop_erase :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3658
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3659
  "wadjust_loop_erase m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3660
    (\<exists> ml mr ln rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3661
                    tl r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> ml + mr = (Suc rs) \<and> mr > 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3662
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3663
fun wadjust_loop_on_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3664
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3665
  "wadjust_loop_on_left_moving_O m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3666
      (\<exists> ml mr ln rn. l = Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m \<^esup>\<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3667
                      r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3668
                      ml + mr = Suc rs \<and> mr > 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3669
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3670
fun wadjust_loop_on_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3671
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3672
  "wadjust_loop_on_left_moving_B m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3673
      (\<exists> ml mr nl nr rn. l = Bk\<^bsup>nl\<^esup> @ Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3674
                         r = Bk\<^bsup>nr\<^esup> @ Bk # Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3675
                         ml + mr = Suc rs \<and> mr > 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3676
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3677
fun wadjust_loop_on_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3678
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3679
  "wadjust_loop_on_left_moving m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3680
       (wadjust_loop_on_left_moving_O m rs (l, r) \<or>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3681
       wadjust_loop_on_left_moving_B m rs (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3682
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3683
fun wadjust_loop_right_move2 :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3684
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3685
  "wadjust_loop_right_move2 m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3686
        (\<exists> ml mr ln rn. l = Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3687
                        r = Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3688
                        ml + mr = Suc rs \<and> mr > 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3689
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3690
fun wadjust_erase2 :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3691
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3692
  "wadjust_erase2 m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3693
     (\<exists> ln rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Oc # Oc\<^bsup>Suc rs\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3694
                     tl r = Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3695
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3696
fun wadjust_on_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3697
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3698
  "wadjust_on_left_moving_O m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3699
        (\<exists> rn. l = Oc\<^bsup>Suc rs\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3700
                  r = Oc # Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3701
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3702
fun wadjust_on_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3703
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3704
  "wadjust_on_left_moving_B m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3705
         (\<exists> ln rn. l = Bk\<^bsup>ln\<^esup> @ Oc # Oc\<^bsup>Suc rs\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3706
                   r = Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3707
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3708
fun wadjust_on_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3709
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3710
  "wadjust_on_left_moving m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3711
      (wadjust_on_left_moving_O m rs (l, r) \<or>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3712
       wadjust_on_left_moving_B m rs (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3713
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3714
fun wadjust_goon_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3715
  where 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3716
  "wadjust_goon_left_moving_B m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3717
        (\<exists> rn. l = Oc\<^bsup>Suc m\<^esup> \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3718
               r = Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3719
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3720
fun wadjust_goon_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3721
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3722
  "wadjust_goon_left_moving_O m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3723
      (\<exists> ml mr rn. l = Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3724
                      r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3725
                      ml + mr = Suc (Suc rs) \<and> mr > 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3726
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3727
fun wadjust_goon_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3728
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3729
  "wadjust_goon_left_moving m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3730
            (wadjust_goon_left_moving_B m rs (l, r) \<or>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3731
             wadjust_goon_left_moving_O m rs (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3732
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3733
fun wadjust_backto_standard_pos_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3734
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3735
  "wadjust_backto_standard_pos_B m rs (l, r) =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3736
        (\<exists> rn. l = [] \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3737
               r = Bk # Oc\<^bsup>Suc m \<^esup>@ Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3738
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3739
fun wadjust_backto_standard_pos_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3740
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3741
  "wadjust_backto_standard_pos_O m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3742
      (\<exists> ml mr rn. l = Oc\<^bsup>ml\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3743
                      r = Oc\<^bsup>mr\<^esup> @ Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3744
                      ml + mr = Suc m \<and> mr > 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3745
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3746
fun wadjust_backto_standard_pos :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3747
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3748
  "wadjust_backto_standard_pos m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3749
        (wadjust_backto_standard_pos_B m rs (l, r) \<or> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3750
        wadjust_backto_standard_pos_O m rs (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3751
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3752
fun wadjust_stop :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3753
where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3754
  "wadjust_stop m rs (l, r) =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3755
        (\<exists> rn. l = [Bk] \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3756
               r = Oc\<^bsup>Suc m \<^esup>@ Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3757
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3758
declare wadjust_start.simps[simp del]  wadjust_loop_start.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3759
        wadjust_loop_right_move.simps[simp del]  wadjust_loop_check.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3760
        wadjust_loop_erase.simps[simp del] wadjust_loop_on_left_moving.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3761
        wadjust_loop_right_move2.simps[simp del] wadjust_erase2.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3762
        wadjust_on_left_moving_O.simps[simp del] wadjust_on_left_moving_B.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3763
        wadjust_on_left_moving.simps[simp del] wadjust_goon_left_moving_B.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3764
        wadjust_goon_left_moving_O.simps[simp del] wadjust_goon_left_moving.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3765
        wadjust_backto_standard_pos.simps[simp del] wadjust_backto_standard_pos_B.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3766
        wadjust_backto_standard_pos_O.simps[simp del] wadjust_stop.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3767
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3768
fun wadjust_inv :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3769
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3770
  "wadjust_inv st m rs (l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3771
       (if st = Suc 0 then wadjust_start m rs (l, r) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3772
        else if st = Suc (Suc 0) then wadjust_loop_start m rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3773
        else if st = Suc (Suc (Suc 0)) then wadjust_loop_right_move m rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3774
        else if st = 4 then wadjust_loop_check m rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3775
        else if st = 5 then wadjust_loop_erase m rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3776
        else if st = 6 then wadjust_loop_on_left_moving m rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3777
        else if st = 7 then wadjust_loop_right_move2 m rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3778
        else if st = 8 then wadjust_erase2 m rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3779
        else if st = 9 then wadjust_on_left_moving m rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3780
        else if st = 10 then wadjust_goon_left_moving m rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3781
        else if st = 11 then wadjust_backto_standard_pos m rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3782
        else if st = 0 then wadjust_stop m rs (l, r)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3783
        else False
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3784
)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3785
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3786
declare wadjust_inv.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3787
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3788
fun wadjust_phase :: "nat \<Rightarrow> t_conf \<Rightarrow> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3789
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3790
  "wadjust_phase rs (st, l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3791
         (if st = 1 then 3 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3792
          else if st \<ge> 2 \<and> st \<le> 7 then 2
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3793
          else if st \<ge> 8 \<and> st \<le> 11 then 1
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3794
          else 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3795
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3796
thm dropWhile.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3797
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3798
fun wadjust_stage :: "nat \<Rightarrow> t_conf \<Rightarrow> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3799
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3800
  "wadjust_stage rs (st, l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3801
           (if st \<ge> 2 \<and> st \<le> 7 then 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3802
                  rs - length (takeWhile (\<lambda> a. a = Oc) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3803
                          (tl (dropWhile (\<lambda> a. a = Oc) (rev l @ r))))
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3804
            else 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3805
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3806
fun wadjust_state :: "nat \<Rightarrow> t_conf \<Rightarrow> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3807
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3808
  "wadjust_state rs (st, l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3809
       (if st \<ge> 2 \<and> st \<le> 7 then 8 - st
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3810
        else if st \<ge> 8 \<and> st \<le> 11 then 12 - st
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3811
        else 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3812
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3813
fun wadjust_step :: "nat \<Rightarrow> t_conf \<Rightarrow> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3814
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3815
  "wadjust_step rs (st, l, r) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3816
       (if st = 1 then (if hd r = Bk then 1
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3817
                        else 0) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3818
        else if st = 3 then length r
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3819
        else if st = 5 then (if hd r = Oc then 1
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3820
                             else 0)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3821
        else if st = 6 then length l
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3822
        else if st = 8 then (if hd r = Oc then 1
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3823
                             else 0)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3824
        else if st = 9 then length l
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3825
        else if st = 10 then length l
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3826
        else if st = 11 then (if hd r = Bk then 0
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3827
                              else Suc (length l))
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3828
        else 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3829
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3830
fun wadjust_measure :: "(nat \<times> t_conf) \<Rightarrow> nat \<times> nat \<times> nat \<times> nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3831
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3832
  "wadjust_measure (rs, (st, l, r)) = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3833
     (wadjust_phase rs (st, l, r), 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3834
      wadjust_stage rs (st, l, r),
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3835
      wadjust_state rs (st, l, r), 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3836
      wadjust_step rs (st, l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3837
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3838
definition wadjust_le :: "((nat \<times> t_conf) \<times> nat \<times> t_conf) set"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3839
  where "wadjust_le \<equiv> (inv_image lex_square wadjust_measure)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3840
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3841
lemma [intro]: "wf lex_square"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3842
by(auto intro:wf_lex_prod simp: abacus.lex_pair_def lex_square_def 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3843
  abacus.lex_triple_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3844
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3845
lemma wf_wadjust_le[intro]: "wf wadjust_le"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3846
by(auto intro:wf_inv_image simp: wadjust_le_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3847
           abacus.lex_triple_def abacus.lex_pair_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3848
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3849
lemma [simp]: "wadjust_start m rs (c, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3850
apply(auto simp: wadjust_start.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3851
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3852
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3853
lemma [simp]: "wadjust_loop_right_move m rs (c, []) \<Longrightarrow> c \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3854
apply(auto simp: wadjust_loop_right_move.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3855
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3856
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3857
lemma [simp]: "wadjust_loop_right_move m rs (c, [])
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3858
        \<Longrightarrow>  wadjust_loop_check m rs (Bk # c, [])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3859
apply(simp only: wadjust_loop_right_move.simps wadjust_loop_check.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3860
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3861
apply(case_tac [!] mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3862
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3863
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3864
lemma [simp]: "wadjust_loop_check m rs (c, []) \<Longrightarrow> c \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3865
apply(simp only: wadjust_loop_check.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3866
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3867
 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3868
lemma [simp]: "wadjust_loop_start m rs (c, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3869
apply(simp add: wadjust_loop_start.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3870
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3871
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3872
lemma [simp]: "wadjust_loop_right_move m rs (c, []) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3873
  wadjust_loop_right_move m rs (Bk # c, [])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3874
apply(simp only: wadjust_loop_right_move.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3875
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3876
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3877
apply(case_tac [!] mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3878
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3879
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3880
lemma [simp]: "wadjust_loop_check m rs (c, []) \<Longrightarrow> wadjust_erase2 m rs (tl c, [hd c])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3881
apply(simp only: wadjust_loop_check.simps wadjust_erase2.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3882
apply(case_tac mr, simp_all add: exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3883
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3884
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3885
lemma [simp]: " wadjust_loop_erase m rs (c, [])
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3886
    \<Longrightarrow> (c = [] \<longrightarrow> wadjust_loop_on_left_moving m rs ([], [Bk])) \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3887
        (c \<noteq> [] \<longrightarrow> wadjust_loop_on_left_moving m rs (tl c, [hd c]))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3888
apply(simp add: wadjust_loop_erase.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3889
apply(case_tac [!] mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3890
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3891
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3892
lemma [simp]: "wadjust_loop_on_left_moving m rs (c, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3893
apply(auto simp: wadjust_loop_on_left_moving.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3894
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3895
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3896
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3897
lemma [simp]: "wadjust_loop_right_move2 m rs (c, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3898
apply(auto simp: wadjust_loop_right_move2.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3899
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3900
   
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3901
lemma [simp]: "wadjust_erase2 m rs ([], []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3902
apply(auto simp: wadjust_erase2.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3903
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3904
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3905
lemma [simp]: "wadjust_on_left_moving_B m rs 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3906
                 (Oc # Oc # Oc\<^bsup>rs\<^esup> @ Bk # Oc # Oc\<^bsup>m\<^esup>, [Bk])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3907
apply(simp add: wadjust_on_left_moving_B.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3908
apply(rule_tac x = 0 in exI, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3909
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3910
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3911
lemma [simp]: "wadjust_on_left_moving_B m rs 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3912
                 (Bk\<^bsup>n\<^esup> @ Bk # Oc # Oc # Oc\<^bsup>rs\<^esup> @ Bk # Oc # Oc\<^bsup>m\<^esup>, [Bk])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3913
apply(simp add: wadjust_on_left_moving_B.simps exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3914
apply(rule_tac x = "Suc n" in exI, simp add: exp_ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3915
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3916
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3917
lemma [simp]: "\<lbrakk>wadjust_erase2 m rs (c, []); c \<noteq> []\<rbrakk> \<Longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3918
            wadjust_on_left_moving m rs (tl c, [hd c])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3919
apply(simp only: wadjust_erase2.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3920
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3921
apply(case_tac ln, simp_all add: exp_ind_def wadjust_on_left_moving.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3922
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3923
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3924
lemma [simp]: "wadjust_erase2 m rs (c, [])
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3925
    \<Longrightarrow> (c = [] \<longrightarrow> wadjust_on_left_moving m rs ([], [Bk])) \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3926
       (c \<noteq> [] \<longrightarrow> wadjust_on_left_moving m rs (tl c, [hd c]))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3927
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3928
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3929
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3930
lemma [simp]: "wadjust_on_left_moving m rs ([], []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3931
apply(simp add: wadjust_on_left_moving.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3932
  wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3933
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3934
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3935
lemma [simp]: "wadjust_on_left_moving_O m rs (c, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3936
apply(simp add: wadjust_on_left_moving_O.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3937
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3938
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3939
lemma [simp]: " \<lbrakk>wadjust_on_left_moving_B m rs (c, []); c \<noteq> []; hd c = Bk\<rbrakk> \<Longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3940
                                      wadjust_on_left_moving_B m rs (tl c, [Bk])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3941
apply(simp add: wadjust_on_left_moving_B.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3942
apply(case_tac [!] ln, simp_all add: exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3943
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3944
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3945
lemma [simp]: "\<lbrakk>wadjust_on_left_moving_B m rs (c, []); c \<noteq> []; hd c = Oc\<rbrakk> \<Longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3946
                                  wadjust_on_left_moving_O m rs (tl c, [Oc])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3947
apply(simp add: wadjust_on_left_moving_B.simps wadjust_on_left_moving_O.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3948
apply(case_tac [!] ln, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3949
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3950
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3951
lemma [simp]: "\<lbrakk>wadjust_on_left_moving m rs (c, []); c \<noteq> []\<rbrakk> \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3952
  wadjust_on_left_moving m rs (tl c, [hd c])"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3953
apply(simp add: wadjust_on_left_moving.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3954
apply(case_tac "hd c", simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3955
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3956
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3957
lemma [simp]: "wadjust_on_left_moving m rs (c, [])
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3958
    \<Longrightarrow> (c = [] \<longrightarrow> wadjust_on_left_moving m rs ([], [Bk])) \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3959
       (c \<noteq> [] \<longrightarrow> wadjust_on_left_moving m rs (tl c, [hd c]))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3960
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3961
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3962
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3963
lemma [simp]: "wadjust_goon_left_moving m rs (c, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3964
apply(auto simp: wadjust_goon_left_moving.simps wadjust_goon_left_moving_B.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3965
                 wadjust_goon_left_moving_O.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3966
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3967
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3968
lemma [simp]: "wadjust_backto_standard_pos m rs (c, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3969
apply(auto simp: wadjust_backto_standard_pos.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3970
 wadjust_backto_standard_pos_B.simps wadjust_backto_standard_pos_O.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3971
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3972
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3973
lemma [simp]:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3974
  "wadjust_start m rs (c, Bk # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3975
  (c = [] \<longrightarrow> wadjust_start m rs ([], Oc # list)) \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3976
  (c \<noteq> [] \<longrightarrow> wadjust_start m rs (c, Oc # list))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3977
apply(auto simp: wadjust_start.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3978
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3979
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3980
lemma [simp]: "wadjust_loop_start m rs (c, Bk # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3981
apply(auto simp: wadjust_loop_start.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3982
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3983
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3984
lemma [simp]: "wadjust_loop_right_move m rs (c, b) \<Longrightarrow> c \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3985
apply(simp only: wadjust_loop_right_move.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3986
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3987
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3988
lemma [simp]: "wadjust_loop_right_move m rs (c, Bk # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3989
    \<Longrightarrow> wadjust_loop_right_move m rs (Bk # c, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3990
apply(simp only: wadjust_loop_right_move.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3991
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3992
apply(rule_tac x = ml in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3993
apply(rule_tac x = mr in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3994
apply(rule_tac x = "Suc nl" in exI, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3995
apply(case_tac nr, simp, case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3996
apply(rule_tac x = nat in exI, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3997
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3998
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  3999
lemma [simp]: "wadjust_loop_check m rs (c, b) \<Longrightarrow> c \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4000
apply(simp only: wadjust_loop_check.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4001
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4002
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4003
lemma [simp]: "wadjust_loop_check m rs (c, Bk # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4004
              \<Longrightarrow>  wadjust_erase2 m rs (tl c, hd c # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4005
apply(auto simp: wadjust_loop_check.simps wadjust_erase2.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4006
apply(case_tac [!] mr, simp_all add: exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4007
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4008
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4009
lemma [simp]: "wadjust_loop_erase m rs (c, b) \<Longrightarrow> c \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4010
apply(simp only: wadjust_loop_erase.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4011
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4012
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4013
declare wadjust_loop_on_left_moving_O.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4014
        wadjust_loop_on_left_moving_B.simps[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4015
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4016
lemma [simp]: "\<lbrakk>wadjust_loop_erase m rs (c, Bk # list); hd c = Bk\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4017
    \<Longrightarrow> wadjust_loop_on_left_moving_B m rs (tl c, Bk # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4018
apply(simp only: wadjust_loop_erase.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4019
  wadjust_loop_on_left_moving_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4020
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4021
apply(rule_tac x = ml in exI, rule_tac x = mr in exI, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4022
      rule_tac x = ln in exI, rule_tac x = 0 in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4023
apply(case_tac ln, simp_all add: exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4024
apply(simp add: exp_ind exp_ind_def[THEN sym])
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4025
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4026
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4027
lemma [simp]: "\<lbrakk>wadjust_loop_erase m rs (c, Bk # list); c \<noteq> []; hd c = Oc\<rbrakk> \<Longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4028
             wadjust_loop_on_left_moving_O m rs (tl c, Oc # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4029
apply(simp only: wadjust_loop_erase.simps wadjust_loop_on_left_moving_O.simps,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4030
       auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4031
apply(case_tac [!] ln, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4032
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4033
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4034
lemma [simp]: "\<lbrakk>wadjust_loop_erase m rs (c, Bk # list); c \<noteq> []\<rbrakk> \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4035
                wadjust_loop_on_left_moving m rs (tl c, hd c # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4036
apply(case_tac "hd c", simp_all add:wadjust_loop_on_left_moving.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4037
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4038
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4039
lemma [simp]: "wadjust_loop_on_left_moving m rs (c, b) \<Longrightarrow> c \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4040
apply(simp add: wadjust_loop_on_left_moving.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4041
wadjust_loop_on_left_moving_O.simps wadjust_loop_on_left_moving_B.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4042
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4043
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4044
lemma [simp]: "wadjust_loop_on_left_moving_O m rs (c, Bk # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4045
apply(simp add: wadjust_loop_on_left_moving_O.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4046
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4047
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4048
lemma [simp]: "\<lbrakk>wadjust_loop_on_left_moving_B m rs (c, Bk # list); hd c = Bk\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4049
    \<Longrightarrow>  wadjust_loop_on_left_moving_B m rs (tl c, Bk # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4050
apply(simp only: wadjust_loop_on_left_moving_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4051
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4052
apply(rule_tac x = ml in exI, rule_tac x = mr in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4053
apply(case_tac nl, simp_all add: exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4054
apply(rule_tac x = "Suc nr" in exI, auto simp: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4055
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4056
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4057
lemma [simp]: "\<lbrakk>wadjust_loop_on_left_moving_B m rs (c, Bk # list); hd c = Oc\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4058
    \<Longrightarrow> wadjust_loop_on_left_moving_O m rs (tl c, Oc # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4059
apply(simp only: wadjust_loop_on_left_moving_O.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4060
                 wadjust_loop_on_left_moving_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4061
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4062
apply(rule_tac x = ml in exI, rule_tac x = mr in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4063
apply(case_tac nl, simp_all add: exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4064
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4065
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4066
lemma [simp]: "wadjust_loop_on_left_moving m rs (c, Bk # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4067
            \<Longrightarrow> wadjust_loop_on_left_moving m rs (tl c, hd c # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4068
apply(simp add: wadjust_loop_on_left_moving.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4069
apply(case_tac "hd c", simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4070
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4071
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4072
lemma [simp]: "wadjust_loop_right_move2 m rs (c, b) \<Longrightarrow> c \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4073
apply(simp only: wadjust_loop_right_move2.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4074
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4075
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4076
lemma [simp]: "wadjust_loop_right_move2 m rs (c, Bk # list) \<Longrightarrow>  wadjust_loop_start m rs (c, Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4077
apply(auto simp: wadjust_loop_right_move2.simps wadjust_loop_start.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4078
apply(case_tac ln, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4079
apply(rule_tac x = 0 in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4080
apply(rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4081
apply(rule_tac x = "Suc ml" in exI, simp add: exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4082
apply(rule_tac x = "Suc nat" in exI, simp add: exp_ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4083
apply(rule_tac x = rn in exI, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4084
apply(rule_tac x = "Suc ml" in exI, auto simp: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4085
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4086
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4087
lemma [simp]: "wadjust_erase2 m rs (c, Bk # list) \<Longrightarrow> c \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4088
apply(auto simp:wadjust_erase2.simps )
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4089
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4090
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4091
lemma [simp]: "wadjust_erase2 m rs (c, Bk # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4092
                 wadjust_on_left_moving m rs (tl c, hd c # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4093
apply(auto simp: wadjust_erase2.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4094
apply(case_tac ln, simp_all add: exp_ind_def wadjust_on_left_moving.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4095
        wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4096
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4097
apply(rule_tac x = "(Suc (Suc rn))" in exI, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4098
apply(rule_tac x = "Suc nat" in exI, simp add: exp_ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4099
apply(rule_tac x = "(Suc (Suc rn))" in exI, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4100
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4101
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4102
lemma [simp]: "wadjust_on_left_moving m rs (c,b) \<Longrightarrow> c \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4103
apply(simp only:wadjust_on_left_moving.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4104
                wadjust_on_left_moving_O.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4105
                wadjust_on_left_moving_B.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4106
             , auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4107
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4108
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4109
lemma [simp]: "wadjust_on_left_moving_O m rs (c, Bk # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4110
apply(simp add: wadjust_on_left_moving_O.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4111
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4112
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4113
lemma [simp]: "\<lbrakk>wadjust_on_left_moving_B m rs (c, Bk # list); hd c = Bk\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4114
    \<Longrightarrow> wadjust_on_left_moving_B m rs (tl c, Bk # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4115
apply(auto simp: wadjust_on_left_moving_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4116
apply(case_tac ln, simp_all add: exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4117
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4118
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4119
lemma [simp]: "\<lbrakk>wadjust_on_left_moving_B m rs (c, Bk # list); hd c = Oc\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4120
    \<Longrightarrow> wadjust_on_left_moving_O m rs (tl c, Oc # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4121
apply(auto simp: wadjust_on_left_moving_O.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4122
                 wadjust_on_left_moving_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4123
apply(case_tac ln, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4124
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4125
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4126
lemma [simp]: "wadjust_on_left_moving  m rs (c, Bk # list) \<Longrightarrow>  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4127
                  wadjust_on_left_moving m rs (tl c, hd c # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4128
apply(simp add: wadjust_on_left_moving.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4129
apply(case_tac "hd c", simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4130
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4131
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4132
lemma [simp]: "wadjust_goon_left_moving m rs (c, b) \<Longrightarrow> c \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4133
apply(simp add: wadjust_goon_left_moving.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4134
                wadjust_goon_left_moving_B.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4135
                wadjust_goon_left_moving_O.simps exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4136
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4137
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4138
lemma [simp]: "wadjust_goon_left_moving_O m rs (c, Bk # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4139
apply(simp add: wadjust_goon_left_moving_O.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4140
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4141
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4142
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4143
lemma [simp]: "\<lbrakk>wadjust_goon_left_moving_B m rs (c, Bk # list); hd c = Bk\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4144
    \<Longrightarrow> wadjust_backto_standard_pos_B m rs (tl c, Bk # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4145
apply(auto simp: wadjust_goon_left_moving_B.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4146
                 wadjust_backto_standard_pos_B.simps exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4147
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4148
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4149
lemma [simp]: "\<lbrakk>wadjust_goon_left_moving_B m rs (c, Bk # list); hd c = Oc\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4150
    \<Longrightarrow> wadjust_backto_standard_pos_O m rs (tl c, Oc # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4151
apply(auto simp: wadjust_goon_left_moving_B.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4152
                 wadjust_backto_standard_pos_O.simps exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4153
apply(rule_tac x = m in exI, simp, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4154
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4155
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4156
lemma [simp]: "wadjust_goon_left_moving m rs (c, Bk # list) \<Longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4157
  wadjust_backto_standard_pos m rs (tl c, hd c # Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4158
apply(case_tac "hd c", simp_all add: wadjust_backto_standard_pos.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4159
                                     wadjust_goon_left_moving.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4160
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4161
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4162
lemma [simp]: "wadjust_backto_standard_pos m rs (c, Bk # list) \<Longrightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4163
  (c = [] \<longrightarrow> wadjust_stop m rs ([Bk], list)) \<and> (c \<noteq> [] \<longrightarrow> wadjust_stop m rs (Bk # c, list))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4164
apply(auto simp: wadjust_backto_standard_pos.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4165
                 wadjust_backto_standard_pos_B.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4166
                 wadjust_backto_standard_pos_O.simps wadjust_stop.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4167
apply(case_tac [!] mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4168
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4169
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4170
lemma [simp]: "wadjust_start m rs (c, Oc # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4171
              \<Longrightarrow> (c = [] \<longrightarrow> wadjust_loop_start m rs ([Oc], list)) \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4172
                (c \<noteq> [] \<longrightarrow> wadjust_loop_start m rs (Oc # c, list))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4173
apply(auto simp:wadjust_loop_start.simps wadjust_start.simps )
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4174
apply(rule_tac x = ln in exI, rule_tac x = rn in exI,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4175
      rule_tac x = "Suc 0" in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4176
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4177
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4178
lemma [simp]: "wadjust_loop_start m rs (c, b) \<Longrightarrow> c \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4179
apply(simp add: wadjust_loop_start.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4180
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4181
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4182
lemma [simp]: "wadjust_loop_start m rs (c, Oc # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4183
              \<Longrightarrow> wadjust_loop_right_move m rs (Oc # c, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4184
apply(simp add: wadjust_loop_start.simps wadjust_loop_right_move.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4185
apply(rule_tac x = ml in exI, rule_tac x = mr in exI, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4186
      rule_tac x = 0 in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4187
apply(rule_tac x = "Suc ln" in exI, simp add: exp_ind, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4188
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4189
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4190
lemma [simp]: "wadjust_loop_right_move m rs (c, Oc # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4191
                       wadjust_loop_check m rs (Oc # c, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4192
apply(simp add: wadjust_loop_right_move.simps  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4193
                 wadjust_loop_check.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4194
apply(rule_tac [!] x = ml in exI, simp_all, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4195
apply(case_tac nl, auto simp: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4196
apply(rule_tac x = "mr - 1" in exI, case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4197
apply(case_tac [!] nr, simp_all add: exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4198
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4199
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4200
lemma [simp]: "wadjust_loop_check m rs (c, Oc # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4201
               wadjust_loop_erase m rs (tl c, hd c # Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4202
apply(simp only: wadjust_loop_check.simps wadjust_loop_erase.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4203
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4204
apply(rule_tac x = ml in exI, rule_tac x = mr in exI, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4205
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4206
apply(case_tac rn, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4207
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4208
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4209
lemma [simp]: "wadjust_loop_erase m rs (c, Oc # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4210
                wadjust_loop_erase m rs (c, Bk # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4211
apply(auto simp: wadjust_loop_erase.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4212
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4213
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4214
lemma [simp]: "wadjust_loop_on_left_moving_B m rs (c, Oc # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4215
apply(auto simp: wadjust_loop_on_left_moving_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4216
apply(case_tac nr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4217
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4218
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4219
lemma [simp]: "wadjust_loop_on_left_moving m rs (c, Oc # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4220
           \<Longrightarrow> wadjust_loop_right_move2 m rs (Oc # c, list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4221
apply(simp add:wadjust_loop_on_left_moving.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4222
apply(auto simp: wadjust_loop_on_left_moving_O.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4223
                 wadjust_loop_right_move2.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4224
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4225
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4226
lemma [simp]: "wadjust_loop_right_move2 m rs (c, Oc # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4227
apply(auto simp: wadjust_loop_right_move2.simps )
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4228
apply(case_tac ln, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4229
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4230
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4231
lemma [simp]: "wadjust_erase2 m rs (c, Oc # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4232
              \<Longrightarrow> (c = [] \<longrightarrow> wadjust_erase2 m rs ([], Bk # list))
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4233
               \<and> (c \<noteq> [] \<longrightarrow> wadjust_erase2 m rs (c, Bk # list))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4234
apply(auto simp: wadjust_erase2.simps )
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4235
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4236
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4237
lemma [simp]: "wadjust_on_left_moving_B m rs (c, Oc # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4238
apply(auto simp: wadjust_on_left_moving_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4239
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4240
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4241
lemma [simp]: "\<lbrakk>wadjust_on_left_moving_O m rs (c, Oc # list); hd c = Bk\<rbrakk> \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4242
         wadjust_goon_left_moving_B m rs (tl c, Bk # Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4243
apply(auto simp: wadjust_on_left_moving_O.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4244
     wadjust_goon_left_moving_B.simps exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4245
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4246
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4247
lemma [simp]: "\<lbrakk>wadjust_on_left_moving_O m rs (c, Oc # list); hd c = Oc\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4248
    \<Longrightarrow> wadjust_goon_left_moving_O m rs (tl c, Oc # Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4249
apply(auto simp: wadjust_on_left_moving_O.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4250
                 wadjust_goon_left_moving_O.simps exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4251
apply(rule_tac x = rs in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4252
apply(auto simp: exp_ind_def numeral_2_eq_2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4253
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4254
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4255
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4256
lemma [simp]: "wadjust_on_left_moving m rs (c, Oc # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4257
              wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4258
apply(simp add: wadjust_on_left_moving.simps   
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4259
                 wadjust_goon_left_moving.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4260
apply(case_tac "hd c", simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4261
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4262
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4263
lemma [simp]: "wadjust_on_left_moving m rs (c, Oc # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4264
  wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4265
apply(simp add: wadjust_on_left_moving.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4266
  wadjust_goon_left_moving.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4267
apply(case_tac "hd c", simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4268
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4269
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4270
lemma [simp]: "wadjust_goon_left_moving_B m rs (c, Oc # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4271
apply(auto simp: wadjust_goon_left_moving_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4272
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4273
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4274
lemma [simp]: "\<lbrakk>wadjust_goon_left_moving_O m rs (c, Oc # list); hd c = Bk\<rbrakk> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4275
               \<Longrightarrow> wadjust_goon_left_moving_B m rs (tl c, Bk # Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4276
apply(auto simp: wadjust_goon_left_moving_O.simps wadjust_goon_left_moving_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4277
apply(case_tac [!] ml, auto simp: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4278
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4279
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4280
lemma  [simp]: "\<lbrakk>wadjust_goon_left_moving_O m rs (c, Oc # list); hd c = Oc\<rbrakk> \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4281
  wadjust_goon_left_moving_O m rs (tl c, Oc # Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4282
apply(auto simp: wadjust_goon_left_moving_O.simps wadjust_goon_left_moving_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4283
apply(rule_tac x = "ml - 1" in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4284
apply(case_tac ml, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4285
apply(rule_tac x = "Suc mr" in exI, auto simp: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4286
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4287
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4288
lemma [simp]: "wadjust_goon_left_moving m rs (c, Oc # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4289
  wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4290
apply(simp add: wadjust_goon_left_moving.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4291
apply(case_tac "hd c", simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4292
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4293
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4294
lemma [simp]: "wadjust_backto_standard_pos_B m rs (c, Oc # list) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4295
apply(simp add: wadjust_backto_standard_pos_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4296
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4297
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4298
lemma [simp]: "wadjust_backto_standard_pos_O m rs (c, Bk # xs) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4299
apply(simp add: wadjust_backto_standard_pos_O.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4300
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4301
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4302
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4303
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4304
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4305
lemma [simp]: "wadjust_backto_standard_pos_O m rs ([], Oc # list) \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4306
  wadjust_backto_standard_pos_B m rs ([], Bk # Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4307
apply(auto simp: wadjust_backto_standard_pos_O.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4308
                 wadjust_backto_standard_pos_B.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4309
apply(rule_tac x = rn in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4310
apply(case_tac ml, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4311
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4312
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4313
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4314
lemma [simp]: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4315
  "\<lbrakk>wadjust_backto_standard_pos_O m rs (c, Oc # list); c \<noteq> []; hd c = Bk\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4316
  \<Longrightarrow> wadjust_backto_standard_pos_B m rs (tl c, Bk # Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4317
apply(simp add:wadjust_backto_standard_pos_O.simps 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4318
        wadjust_backto_standard_pos_B.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4319
apply(case_tac [!] ml, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4320
done 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4321
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4322
lemma [simp]: "\<lbrakk>wadjust_backto_standard_pos_O m rs (c, Oc # list); c \<noteq> []; hd c = Oc\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4323
          \<Longrightarrow>  wadjust_backto_standard_pos_O m rs (tl c, Oc # Oc # list)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4324
apply(simp add: wadjust_backto_standard_pos_O.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4325
apply(case_tac ml, simp_all add: exp_ind_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4326
apply(rule_tac x = nat in exI, auto simp: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4327
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4328
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4329
lemma [simp]: "wadjust_backto_standard_pos m rs (c, Oc # list)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4330
  \<Longrightarrow> (c = [] \<longrightarrow> wadjust_backto_standard_pos m rs ([], Bk # Oc # list)) \<and> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4331
 (c \<noteq> [] \<longrightarrow> wadjust_backto_standard_pos m rs (tl c, hd c # Oc # list))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4332
apply(auto simp: wadjust_backto_standard_pos.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4333
apply(case_tac "hd c", simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4334
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4335
thm wadjust_loop_right_move.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4336
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4337
lemma [simp]: "wadjust_loop_right_move m rs (c, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4338
apply(simp only: wadjust_loop_right_move.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4339
apply(rule_tac iffI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4340
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4341
apply(case_tac nr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4342
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4343
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4344
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4345
lemma [simp]: "wadjust_loop_erase m rs (c, []) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4346
apply(simp only: wadjust_loop_erase.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4347
apply(case_tac mr, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4348
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4349
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4350
lemma [simp]: "\<lbrakk>Suc (Suc rs) = a;  wadjust_loop_erase m rs (c, Bk # list)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4351
  \<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Bk # list))))
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4352
  < a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list)))) \<or>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4353
  a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4354
  a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list))))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4355
apply(simp only: wadjust_loop_erase.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4356
apply(rule_tac disjI2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4357
apply(case_tac c, simp, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4358
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4359
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4360
lemma [simp]:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4361
  "\<lbrakk>Suc (Suc rs) = a;  wadjust_loop_on_left_moving m rs (c, Bk # list)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4362
  \<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Bk # list))))
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4363
  < a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list)))) \<or>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4364
  a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4365
  a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list))))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4366
apply(subgoal_tac "c \<noteq> []")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4367
apply(case_tac c, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4368
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4369
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4370
lemma dropWhile_exp1: "dropWhile (\<lambda>a. a = Oc) (Oc\<^bsup>n\<^esup> @ xs) = dropWhile (\<lambda>a. a = Oc) xs"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4371
apply(induct n, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4372
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4373
lemma takeWhile_exp1: "takeWhile (\<lambda>a. a = Oc) (Oc\<^bsup>n\<^esup> @ xs) = Oc\<^bsup>n\<^esup> @ takeWhile (\<lambda>a. a = Oc) xs"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4374
apply(induct n, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4375
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4376
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4377
lemma [simp]: "\<lbrakk>Suc (Suc rs) = a;  wadjust_loop_right_move2 m rs (c, Bk # list)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4378
              \<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list))))
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4379
                 < a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list))))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4380
apply(simp add: wadjust_loop_right_move2.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4381
apply(simp add: dropWhile_exp1 takeWhile_exp1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4382
apply(case_tac ln, simp, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4383
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4384
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4385
lemma [simp]: "wadjust_loop_check m rs ([], b) = False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4386
apply(simp add: wadjust_loop_check.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4387
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4388
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4389
lemma [simp]: "\<lbrakk>Suc (Suc rs) = a;  wadjust_loop_check m rs (c, Oc # list)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4390
  \<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Oc # list))))
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4391
  < a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list)))) \<or>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4392
  a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Oc # list)))) =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4393
  a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list))))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4394
apply(case_tac "c", simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4395
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4396
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4397
lemma [simp]: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4398
  "\<lbrakk>Suc (Suc rs) = a;  wadjust_loop_erase m rs (c, Oc # list)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4399
  \<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list))))
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4400
  < a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list)))) \<or>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4401
  a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list)))) =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4402
  a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list))))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4403
apply(simp add: wadjust_loop_erase.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4404
apply(rule_tac disjI2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4405
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4406
apply(simp add: dropWhile_exp1 takeWhile_exp1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4407
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4408
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4409
declare numeral_2_eq_2[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4410
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4411
lemma wadjust_correctness:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4412
  shows "let P = (\<lambda> (len, st, l, r). st = 0) in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4413
  let Q = (\<lambda> (len, st, l, r). wadjust_inv st m rs (l, r)) in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4414
  let f = (\<lambda> stp. (Suc (Suc rs),  steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4415
                Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk #  Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust stp)) in
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4416
    \<exists> n .P (f n) \<and> Q (f n)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4417
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4418
  let ?P = "(\<lambda> (len, st, l, r). st = 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4419
  let ?Q = "\<lambda> (len, st, l, r). wadjust_inv st m rs (l, r)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4420
  let ?f = "\<lambda> stp. (Suc (Suc rs),  steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4421
                Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust stp)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4422
  have "\<exists> n. ?P (?f n) \<and> ?Q (?f n)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4423
  proof(rule_tac halt_lemma2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4424
    show "wf wadjust_le" by auto
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4425
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4426
    show "\<forall> n. \<not> ?P (?f n) \<and> ?Q (?f n) \<longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4427
                 ?Q (?f (Suc n)) \<and> (?f (Suc n), ?f n) \<in> wadjust_le"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4428
    proof(rule_tac allI, rule_tac impI, case_tac "?f n", 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4429
            simp add: tstep_red tstep.simps, rule_tac conjI, erule_tac conjE,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4430
          erule_tac conjE)      
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4431
      fix n a b c d
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4432
      assume "0 < b" "wadjust_inv b m rs (c, d)" "Suc (Suc rs) = a"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4433
      thus "case case fetch t_wcode_adjust b (case d of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4434
        of (ac, ns) \<Rightarrow> (ns, new_tape ac (c, d)) of (st, x) \<Rightarrow> wadjust_inv st m rs x"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4435
        apply(case_tac d, simp, case_tac [2] aa)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4436
        apply(simp_all add: wadjust_inv.simps wadjust_le_def new_tape.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4437
          abacus.lex_triple_def abacus.lex_pair_def lex_square_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4438
          split: if_splits)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4439
        done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4440
    next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4441
      fix n a b c d
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4442
      assume "0 < b \<and> wadjust_inv b m rs (c, d)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4443
        "Suc (Suc rs) = a \<and> steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4444
         Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust n = (b, c, d)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4445
      thus "((a, case fetch t_wcode_adjust b (case d of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4446
        of (ac, ns) \<Rightarrow> (ns, new_tape ac (c, d))), a, b, c, d) \<in> wadjust_le"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4447
      proof(erule_tac conjE, erule_tac conjE, erule_tac conjE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4448
        assume "0 < b" "wadjust_inv b m rs (c, d)" "Suc (Suc rs) = a"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4449
        thus "?thesis"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4450
          apply(case_tac d, case_tac [2] aa)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4451
          apply(simp_all add: wadjust_inv.simps wadjust_le_def new_tape.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4452
            abacus.lex_triple_def abacus.lex_pair_def lex_square_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4453
            split: if_splits)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4454
          done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4455
      qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4456
    qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4457
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4458
    show "?Q (?f 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4459
      apply(simp add: steps.simps wadjust_inv.simps wadjust_start.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4460
      apply(rule_tac x = ln in exI,auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4461
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4462
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4463
    show "\<not> ?P (?f 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4464
      apply(simp add: steps.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4465
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4466
  qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4467
  thus "?thesis"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4468
    apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4469
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4470
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4471
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4472
lemma [intro]: "t_correct t_wcode_adjust"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4473
apply(auto simp: t_wcode_adjust_def t_correct.simps iseven_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4474
apply(rule_tac x = 11 in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4475
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4476
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4477
lemma wcode_lemma_pre':
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4478
  "args \<noteq> [] \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4479
  \<exists> stp rn. steps (Suc 0, [], <m # args>) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4480
              ((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) stp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4481
  = (0,  [Bk],  Oc\<^bsup>Suc m\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>)" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4482
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4483
  let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <m # args>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4484
  let ?Q1 = "\<lambda>(l, r). l = Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4485
    (\<exists>ln rn. r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4486
  let ?P2 = ?Q1
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4487
  let ?Q2 = "\<lambda> (l, r). (wadjust_stop m (bl_bin (<args>) - 1) (l, r))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4488
  let ?P3 = "\<lambda> tp. False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4489
  assume h: "args \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4490
  have "?P1 \<turnstile>-> \<lambda> tp. (\<exists> stp tp'. steps (Suc 0, tp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4491
                      ((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) stp = (0, tp') \<and> ?Q2 tp')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4492
  proof(rule_tac turing_merge.t_merge_halt[of "t_wcode_prepare |+| t_wcode_main" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4493
               t_wcode_adjust ?P1 ?P2 ?P3 ?P3 ?Q1 ?Q2], 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4494
        auto simp: turing_merge_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4495
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4496
    show "\<exists>stp. case steps (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) stp of
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4497
          (st, tp') \<Rightarrow> st = 0 \<and> (case tp' of (l, r) \<Rightarrow> l = Bk # Oc\<^bsup>Suc m\<^esup> \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4498
                (\<exists>ln rn. r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4499
      using h prepare_mainpart_lemma[of args m]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4500
      apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4501
      apply(rule_tac x = stp in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4502
      apply(rule_tac x = ln in exI, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4503
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4504
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4505
    fix ln rn
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4506
    show "\<exists>stp. case steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4507
                               Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust stp of
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4508
      (st, tp') \<Rightarrow> st = 0 \<and> wadjust_stop m (bl_bin (<args>) - Suc 0) tp'"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4509
      using wadjust_correctness[of m "bl_bin (<args>) - 1" "Suc ln" rn]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4510
      apply(subgoal_tac "bl_bin (<args>) > 0", auto simp: wadjust_inv.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4511
      apply(rule_tac x = n in exI, simp add: exp_ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4512
      using h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4513
      apply(case_tac args, simp_all, case_tac list,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4514
            simp_all add: tape_of_nl_abv  tape_of_nat_list.simps exp_ind_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4515
            bl_bin.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4516
      done     
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4517
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4518
    show "?Q1 \<turnstile>-> ?P2"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4519
      by(simp add: t_imply_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4520
  qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4521
  thus "\<exists>stp rn. steps (Suc 0, [], <m # args>) ((t_wcode_prepare |+| t_wcode_main) |+| 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4522
        t_wcode_adjust) stp = (0, [Bk], Oc\<^bsup>Suc m\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4523
    apply(simp add: t_imply_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4524
    apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4525
    apply(subgoal_tac "bl_bin (<args>) > 0", auto simp: wadjust_stop.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4526
    using h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4527
    apply(case_tac args, simp_all, case_tac list,  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4528
          simp_all add: tape_of_nl_abv  tape_of_nat_list.simps exp_ind_def
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4529
            bl_bin.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4530
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4531
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4532
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4533
text {*
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4534
  The initialization TM @{text "t_wcode"}.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4535
  *}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4536
definition t_wcode :: "tprog"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4537
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4538
  "t_wcode = (t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4539
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4540
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4541
text {*
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4542
  The correctness of @{text "t_wcode"}.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4543
  *}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4544
lemma wcode_lemma_1:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4545
  "args \<noteq> [] \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4546
  \<exists> stp ln rn. steps (Suc 0, [], <m # args>)  (t_wcode) stp = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4547
              (0,  [Bk],  Oc\<^bsup>Suc m\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4548
apply(simp add: wcode_lemma_pre' t_wcode_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4549
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4550
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4551
lemma wcode_lemma: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4552
  "args \<noteq> [] \<Longrightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4553
  \<exists> stp ln rn. steps (Suc 0, [], <m # args>)  (t_wcode) stp = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4554
              (0,  [Bk],  <[m ,bl_bin (<args>)]> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4555
using wcode_lemma_1[of args m]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4556
apply(simp add: t_wcode_def tape_of_nl_abv tape_of_nat_list.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4557
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4558
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4559
section {* The universal TM *}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4560
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4561
text {*
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4562
  This section gives the explicit construction of {\em Universal Turing Machine}, defined as @{text "UTM"} and proves its 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4563
  correctness. It is pretty easy by composing the partial results we have got so far.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4564
  *}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4565
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4566
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4567
definition UTM :: "tprog"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4568
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4569
  "UTM = (let (aprog, rs_pos, a_md) = rec_ci rec_F in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4570
          let abc_F = aprog [+] dummy_abc (Suc (Suc 0)) in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4571
          (t_wcode |+| (tm_of abc_F @ tMp (Suc (Suc 0)) (start_of (layout_of abc_F) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4572
                                                   (length abc_F) - Suc 0))))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4573
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4574
definition F_aprog :: "abc_prog"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4575
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4576
  "F_aprog \<equiv> (let (aprog, rs_pos, a_md) = rec_ci rec_F in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4577
                       aprog [+] dummy_abc (Suc (Suc 0)))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4578
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4579
definition F_tprog :: "tprog"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4580
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4581
  "F_tprog = tm_of (F_aprog)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4582
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4583
definition t_utm :: "tprog"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4584
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4585
  "t_utm \<equiv>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4586
     (F_tprog) @ tMp (Suc (Suc 0)) (start_of (layout_of (F_aprog)) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4587
                                  (length (F_aprog)) - Suc 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4588
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4589
definition UTM_pre :: "tprog"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4590
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4591
  "UTM_pre = t_wcode |+| t_utm"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4592
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4593
lemma F_abc_halt_eq:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4594
  "\<lbrakk>turing_basic.t_correct tp; 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4595
    length lm = k;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4596
    steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>);
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4597
    rs > 0\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4598
    \<Longrightarrow> \<exists> stp m. abc_steps_l (0, [code tp, bl2wc (<lm>)]) (F_aprog) stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4599
                       (length (F_aprog), code tp # bl2wc (<lm>) # (rs - 1) # 0\<^bsup>m\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4600
apply(drule_tac  F_t_halt_eq, simp, simp, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4601
apply(case_tac "rec_ci rec_F")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4602
apply(frule_tac abc_append_dummy_complie, simp, simp, erule_tac exE,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4603
      erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4604
apply(rule_tac x = stp in exI, rule_tac x = m in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4605
apply(simp add: F_aprog_def dummy_abc_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4606
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4607
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4608
lemma F_abc_utm_halt_eq: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4609
  "\<lbrakk>rs > 0; 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4610
  abc_steps_l (0, [code tp, bl2wc (<lm>)]) F_aprog stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4611
        (length F_aprog, code tp #  bl2wc (<lm>) # (rs - 1) # 0\<^bsup>m\<^esup>)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4612
  \<Longrightarrow> \<exists>stp m n.(steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4613
                                             (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4614
  thm abacus_turing_eq_halt
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4615
  using abacus_turing_eq_halt
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4616
  [of "layout_of F_aprog" "F_aprog" "F_tprog" "length (F_aprog)" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4617
    "[code tp, bl2wc (<lm>)]" stp "code tp # bl2wc (<lm>) # (rs - 1) # 0\<^bsup>m\<^esup>" "Suc (Suc 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4618
    "start_of (layout_of (F_aprog)) (length (F_aprog))" "[]" 0]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4619
apply(simp add: F_tprog_def t_utm_def abc_lm_v.simps nth_append)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4620
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4621
apply(rule_tac x = stpa in exI, rule_tac x = "Suc (Suc ma)" in exI, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4622
       rule_tac x = l in exI, simp add: exp_ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4623
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4624
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4625
declare tape_of_nl_abv_cons[simp del]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4626
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4627
lemma t_utm_halt_eq': 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4628
  "\<lbrakk>turing_basic.t_correct tp;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4629
   0 < rs;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4630
  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>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4631
  \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4632
                                                (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4633
apply(drule_tac  l = l in F_abc_halt_eq, simp, simp, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4634
apply(erule_tac exE, erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4635
apply(rule_tac F_abc_utm_halt_eq, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4636
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4637
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4638
lemma [simp]: "tinres xs (xs @ Bk\<^bsup>i\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4639
apply(auto simp: tinres_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4640
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4641
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4642
lemma [elim]: "\<lbrakk>rs > 0; Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup> = c @ Bk\<^bsup>n\<^esup>\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4643
        \<Longrightarrow> \<exists>n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4644
apply(case_tac "na > n")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4645
apply(subgoal_tac "\<exists> d. na = d + n", auto simp: exp_add)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4646
apply(rule_tac x = "na - n" in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4647
apply(subgoal_tac "\<exists> d. n = d + na", auto simp: exp_add)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4648
apply(case_tac rs, simp_all add: exp_ind, case_tac d, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4649
           simp_all add: exp_ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4650
apply(rule_tac x = "n - na" in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4651
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4652
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4653
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4654
lemma t_utm_halt_eq'': 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4655
  "\<lbrakk>turing_basic.t_correct tp;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4656
   0 < rs;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4657
   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>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4658
  \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stp = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4659
                                                (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4660
apply(drule_tac t_utm_halt_eq', simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4661
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4662
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4663
  fix stpa ma na
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4664
  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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4665
  and gr: "rs > 0"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4666
  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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4667
    apply(rule_tac x = stpa in exI, rule_tac x = ma in exI,  simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4668
  proof(case_tac "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4669
    fix a b c
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4670
    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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4671
            "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (a, b, c)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4672
    thus " a = 0 \<and> b = Bk\<^bsup>ma\<^esup> \<and> (\<exists>n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4673
      using tinres_steps2[of "<[code tp, bl2wc (<lm>)]>" "<[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4674
                           "Suc 0" " [Bk, Bk]" t_utm stpa 0 "Bk\<^bsup>ma\<^esup>" "Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>" a b c]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4675
      apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4676
      using gr
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4677
      apply(simp only: tinres_def, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4678
      apply(rule_tac x = "na + n" in exI, simp add: exp_add)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4679
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4680
  qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4681
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4682
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4683
lemma [simp]: "tinres [Bk, Bk] [Bk]"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4684
apply(auto simp: tinres_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4685
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4686
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4687
lemma [elim]: "Bk\<^bsup>ma\<^esup> = b @ Bk\<^bsup>n\<^esup>  \<Longrightarrow> \<exists>m. b = Bk\<^bsup>m\<^esup>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4688
apply(subgoal_tac "ma = length b + n")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4689
apply(rule_tac x = "ma - n" in exI, simp add: exp_add)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4690
apply(drule_tac length_equal)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4691
apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4692
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4693
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4694
lemma t_utm_halt_eq: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4695
  "\<lbrakk>turing_basic.t_correct tp;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4696
   0 < rs;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4697
   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>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4698
  \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stp = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4699
                                                (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4700
apply(drule_tac i = i in t_utm_halt_eq'', simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4701
apply(erule_tac exE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4702
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4703
  fix stpa ma na
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4704
  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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4705
  and gr: "rs > 0"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4706
  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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4707
    apply(rule_tac x = stpa in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4708
  proof(case_tac "steps (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4709
    fix a b c
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4710
    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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4711
            "steps (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (a, b, c)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4712
    thus "a = 0 \<and> (\<exists>m. b = Bk\<^bsup>m\<^esup>) \<and> (\<exists>n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4713
      using tinres_steps[of "[Bk, Bk]" "[Bk]" "Suc 0" "<[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>" t_utm stpa 0
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4714
                             "Bk\<^bsup>ma\<^esup>" "Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>" a b c]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4715
      apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4716
      apply(auto simp: tinres_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4717
      apply(rule_tac x = "ma + n" in exI, simp add: exp_add)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4718
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4719
  qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4720
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4721
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4722
lemma [intro]: "t_correct t_wcode"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4723
apply(simp add: t_wcode_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4724
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4725
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4726
      
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4727
lemma [intro]: "t_correct t_utm"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4728
apply(simp add: t_utm_def F_tprog_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4729
apply(rule_tac t_compiled_correct, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4730
done   
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4731
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4732
lemma UTM_halt_lemma_pre: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4733
  "\<lbrakk>turing_basic.t_correct tp;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4734
   0 < rs;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4735
   args \<noteq> [];
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4736
   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>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4737
  \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [], <code tp # args>) UTM_pre stp = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4738
                                                (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4739
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4740
  let ?Q2 = "\<lambda> (l, r). (\<exists> ln rn. l = Bk\<^bsup>ln\<^esup> \<and> r = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4741
  term ?Q2
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4742
  let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <code tp # args>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4743
  let ?Q1 = "\<lambda> (l, r). (l = [Bk] \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4744
             (\<exists> rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4745
  let ?P2 = ?Q1
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4746
  let ?P3 = "\<lambda> (l, r). False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4747
  assume h: "turing_basic.t_correct tp" "0 < rs"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4748
            "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>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4749
  have "?P1 \<turnstile>-> \<lambda> tp. (\<exists> stp tp'. steps (Suc 0, tp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4750
                    (t_wcode |+| t_utm) stp = (0, tp') \<and> ?Q2 tp')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4751
  proof(rule_tac turing_merge.t_merge_halt [of "t_wcode" "t_utm"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4752
          ?P1 ?P2 ?P3 ?P3 ?Q1 ?Q2], auto simp: turing_merge_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4753
    show "\<exists>stp. case steps (Suc 0, [], <code tp # args>) t_wcode stp of (st, tp') \<Rightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4754
       st = 0 \<and> (case tp' of (l, r) \<Rightarrow> l = [Bk] \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4755
                   (\<exists>rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4756
      using wcode_lemma_1[of args "code tp"] h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4757
      apply(simp, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4758
      apply(rule_tac x = stpa in exI, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4759
      done      
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4760
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4761
    fix rn 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4762
    show "\<exists>stp. case steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4763
      Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp of
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4764
      (st, tp') \<Rightarrow> st = 0 \<and> (case tp' of (l, r) \<Rightarrow>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4765
      (\<exists>ln. l = Bk\<^bsup>ln\<^esup>) \<and> (\<exists>rn. r = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>rn\<^esup>))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4766
      using t_utm_halt_eq[of tp rs i args stp m k rn] h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4767
      apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4768
      apply(rule_tac x = stpa in exI, simp add: bin_wc_eq 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4769
        tape_of_nat_list.simps tape_of_nl_abv)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4770
      apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4771
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4772
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4773
    show "?Q1 \<turnstile>-> ?P2"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4774
      apply(simp add: t_imply_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4775
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4776
  qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4777
  thus "?thesis"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4778
    apply(simp add: t_imply_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4779
    apply(auto simp: UTM_pre_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4780
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4781
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4782
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4783
text {*
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4784
  The correctness of @{text "UTM"}, the halt case.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4785
*}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4786
lemma UTM_halt_lemma: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4787
  "\<lbrakk>turing_basic.t_correct tp;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4788
   0 < rs;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4789
   args \<noteq> [];
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4790
   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>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4791
  \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [], <code tp # args>) UTM stp = 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4792
                                                (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4793
using UTM_halt_lemma_pre[of tp rs args i stp m k]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4794
apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4795
apply(case_tac "rec_ci rec_F", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4796
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4797
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4798
definition TSTD:: "t_conf \<Rightarrow> bool"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4799
  where
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4800
  "TSTD c = (let (st, l, r) = c in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4801
             st = 0 \<and> (\<exists> m. l = Bk\<^bsup>m\<^esup>) \<and> (\<exists> rs n. r = Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4802
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4803
thm abacus_turing_eq_uhalt
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4804
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4805
lemma nstd_case1: "0 < a \<Longrightarrow> NSTD (trpl_code (a, b, c))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4806
apply(simp add: NSTD.simps trpl_code.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4807
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4808
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4809
lemma [simp]: "\<forall>m. b \<noteq> Bk\<^bsup>m\<^esup> \<Longrightarrow> 0 < bl2wc b"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4810
apply(rule classical, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4811
apply(induct b, erule_tac x = 0 in allE, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4812
apply(simp add: bl2wc.simps, case_tac a, simp_all 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4813
  add: bl2nat.simps bl2nat_double)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4814
apply(case_tac "\<exists> m. b = Bk\<^bsup>m\<^esup>",  erule exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4815
apply(erule_tac x = "Suc m" in allE, simp add: exp_ind_def, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4816
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4817
lemma nstd_case2: "\<forall>m. b \<noteq> Bk\<^bsup>m\<^esup> \<Longrightarrow> NSTD (trpl_code (a, b, c))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4818
apply(simp add: NSTD.simps trpl_code.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4819
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4820
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4821
thm lg.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4822
thm lgR.simps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4823
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4824
lemma [elim]: "Suc (2 * x) = 2 * y \<Longrightarrow> RR"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4825
apply(induct x arbitrary: y, simp, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4826
apply(case_tac y, simp, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4827
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4828
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4829
lemma bl2nat_zero_eq[simp]: "(bl2nat c 0 = 0) = (\<exists>n. c = Bk\<^bsup>n\<^esup>)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4830
apply(auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4831
apply(induct c, simp add: bl2nat.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4832
apply(rule_tac x = 0 in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4833
apply(case_tac a, auto simp: bl2nat.simps bl2nat_double)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4834
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4835
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4836
lemma bl2wc_exp_ex: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4837
  "\<lbrakk>Suc (bl2wc c) = 2 ^  m\<rbrakk> \<Longrightarrow> \<exists> rs n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4838
apply(induct c arbitrary: m, simp add: bl2wc.simps bl2nat.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4839
apply(case_tac a, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4840
apply(case_tac m, simp_all add: bl2wc.simps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4841
apply(rule_tac x = 0 in exI, rule_tac x = "Suc n" in exI, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4842
  simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4843
apply(simp add: bl2wc.simps bl2nat.simps bl2nat_double)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4844
apply(case_tac m, simp, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4845
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4846
  fix c m nat
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4847
  assume ind: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4848
    "\<And>m. Suc (bl2nat c 0) = 2 ^ m \<Longrightarrow> \<exists>rs n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4849
  and h: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4850
    "Suc (Suc (2 * bl2nat c 0)) = 2 * 2 ^ nat"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4851
  have "\<exists>rs n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4852
    apply(rule_tac m = nat in ind)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4853
    using h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4854
    apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4855
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4856
  from this obtain rs n where " c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>" by blast 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4857
  thus "\<exists>rs n. Oc # c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4858
    apply(rule_tac x = "Suc rs" in exI, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4859
    apply(rule_tac x = n in exI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4860
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4861
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4862
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4863
lemma [elim]: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4864
  "\<lbrakk>\<forall>rs n. c \<noteq> Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>; 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4865
  bl2wc c = 2 ^ lg (Suc (bl2wc c)) 2 - Suc 0\<rbrakk> \<Longrightarrow> bl2wc c = 0"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4866
apply(subgoal_tac "\<exists> m. Suc (bl2wc c) = 2^m", erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4867
apply(drule_tac bl2wc_exp_ex, simp, erule_tac exE, erule_tac exE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4868
apply(case_tac rs, simp, simp, erule_tac x = nat in allE,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4869
  erule_tac x = n in allE, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4870
using bl2wc_exp_ex[of c "lg (Suc (bl2wc c)) 2"]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4871
apply(case_tac "(2::nat) ^ lg (Suc (bl2wc c)) 2", 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4872
  simp, simp, erule_tac exE, erule_tac exE, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4873
apply(simp add: bl2wc.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4874
apply(rule_tac x = rs in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4875
apply(case_tac "(2::nat)^rs", simp, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4876
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4877
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4878
lemma nstd_case3: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4879
  "\<forall>rs n. c \<noteq> Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup> \<Longrightarrow>  NSTD (trpl_code (a, b, c))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4880
apply(simp add: NSTD.simps trpl_code.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4881
apply(rule_tac impI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4882
apply(rule_tac disjI2, rule_tac disjI2, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4883
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4884
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4885
lemma NSTD_1: "\<not> TSTD (a, b, c)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4886
    \<Longrightarrow> rec_exec rec_NSTD [trpl_code (a, b, c)] = Suc 0"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4887
  using NSTD_lemma1[of "trpl_code (a, b, c)"]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4888
       NSTD_lemma2[of "trpl_code (a, b, c)"]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4889
  apply(simp add: TSTD_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4890
  apply(erule_tac disjE, erule_tac nstd_case1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4891
  apply(erule_tac disjE, erule_tac nstd_case2)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4892
  apply(erule_tac nstd_case3)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4893
  done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4894
 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4895
lemma nonstop_t_uhalt_eq:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4896
      "\<lbrakk>turing_basic.t_correct tp;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4897
        steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (a, b, c);
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4898
       \<not> TSTD (a, b, c)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4899
       \<Longrightarrow> rec_exec rec_nonstop [code tp, bl2wc (<lm>), stp] = Suc 0"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4900
apply(simp add: rec_nonstop_def rec_exec.simps)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4901
apply(subgoal_tac 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4902
  "rec_exec rec_conf [code tp, bl2wc (<lm>), stp] =
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4903
  trpl_code (a, b, c)", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4904
apply(erule_tac NSTD_1)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4905
using rec_t_eq_steps[of tp l lm stp]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4906
apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4907
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4908
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4909
lemma nonstop_true:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4910
  "\<lbrakk>turing_basic.t_correct tp;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4911
  \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp))\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4912
     \<Longrightarrow> \<forall>y. rec_calc_rel rec_nonstop 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4913
                        ([code tp, bl2wc (<lm>), y]) (Suc 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4914
apply(rule_tac allI, erule_tac x = y in allE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4915
apply(case_tac "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp y", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4916
apply(rule_tac nonstop_t_uhalt_eq, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4917
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4918
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4919
(*
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4920
lemma [simp]: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4921
  "\<forall>j<Suc k. Ex (rec_calc_rel (get_fstn_args (Suc k) (Suc k) ! j)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4922
                                                     (code tp # lm))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4923
apply(auto simp: get_fstn_args_nth)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4924
apply(rule_tac x = "(code tp # lm) ! j" in exI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4925
apply(rule_tac calc_id, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4926
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4927
*)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4928
declare ci_cn_para_eq[simp]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4929
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4930
lemma F_aprog_uhalt: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4931
  "\<lbrakk>turing_basic.t_correct tp; 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4932
    \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp)); 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4933
    rec_ci rec_F = (F_ap, rs_pos, a_md)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4934
  \<Longrightarrow> \<forall> stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)] @ 0\<^bsup>a_md - rs_pos \<^esup>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4935
               @ suflm) (F_ap) stp of (ss, e) \<Rightarrow> ss < length (F_ap)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4936
apply(case_tac "rec_ci (Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) rec_conf 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4937
               ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])])")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4938
apply(simp only: rec_F_def, rule_tac i = 0  and ga = a and gb = b and 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4939
  gc = c in cn_gi_uhalt, simp, simp, simp, simp, simp, simp, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4940
apply(simp add: ci_cn_para_eq)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4941
apply(case_tac "rec_ci (Cn (Suc (Suc 0)) rec_conf 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4942
  ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt]))")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4943
apply(rule_tac rf = "(Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) rec_conf
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4944
              ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])])" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4945
           and n = "Suc (Suc 0)" and f = rec_right and 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4946
          gs = "[Cn (Suc (Suc 0)) rec_conf 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4947
           ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])]"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4948
           and i = 0 and ga = aa and gb = ba and gc = ca in 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4949
          cn_gi_uhalt)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4950
apply(simp, simp, simp, simp, simp, simp, simp, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4951
     simp add: ci_cn_para_eq)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4952
apply(case_tac "rec_ci rec_halt")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4953
apply(rule_tac rf = "(Cn (Suc (Suc 0)) rec_conf 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4954
  ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt]))" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4955
  and n = "Suc (Suc 0)" and f = "rec_conf" and 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4956
  gs = "([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])"  and 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4957
  i = "Suc (Suc 0)" and gi = "rec_halt" and ga = ab and gb = bb and
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4958
  gc = cb in cn_gi_uhalt)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4959
apply(simp, simp, simp, simp, simp add: nth_append, simp, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4960
  simp add: nth_append, simp add: rec_halt_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4961
apply(simp only: rec_halt_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4962
apply(case_tac [!] "rec_ci ((rec_nonstop))")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4963
apply(rule_tac allI, rule_tac impI, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4964
apply(case_tac j, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4965
apply(rule_tac x = "code tp" in exI, rule_tac calc_id, simp, simp, simp, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4966
apply(rule_tac x = "bl2wc (<lm>)" in exI, rule_tac calc_id, simp, simp, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4967
apply(rule_tac rf = "Mn (Suc (Suc 0)) (rec_nonstop)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4968
  and f = "(rec_nonstop)" and n = "Suc (Suc 0)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4969
  and  aprog' = ac and rs_pos' =  bc and a_md' = cc in Mn_unhalt)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4970
apply(simp, simp add: rec_halt_def , simp, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4971
apply(drule_tac  nonstop_true, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4972
apply(rule_tac allI)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4973
apply(erule_tac x = y in allE)+
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4974
apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4975
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4976
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4977
thm abc_list_crsp_steps
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4978
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4979
lemma uabc_uhalt': 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4980
  "\<lbrakk>turing_basic.t_correct tp;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4981
  \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp));
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4982
  rec_ci rec_F = (ap, pos, md)\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4983
  \<Longrightarrow> \<forall> stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)]) ap stp of (ss, e)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4984
           \<Rightarrow>  ss < length ap"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4985
proof(frule_tac F_ap = ap and rs_pos = pos and a_md = md
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4986
    and suflm = "[]" in F_aprog_uhalt, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4987
  fix stp a b
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4988
  assume h: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4989
    "\<forall>stp. case abc_steps_l (0, code tp # bl2wc (<lm>) # 0\<^bsup>md - pos\<^esup>) ap stp of 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4990
    (ss, e) \<Rightarrow> ss < length ap"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4991
    "abc_steps_l (0, [code tp, bl2wc (<lm>)]) ap stp = (a, b)" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4992
    "turing_basic.t_correct tp" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4993
    "rec_ci rec_F = (ap, pos, md)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4994
  moreover have "ap \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4995
    using h apply(rule_tac rec_ci_not_null, simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4996
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4997
  ultimately show "a < length ap"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4998
  proof(erule_tac x = stp in allE,
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  4999
  case_tac "abc_steps_l (0, code tp # bl2wc (<lm>) # 0\<^bsup>md - pos\<^esup>) ap stp", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5000
    fix aa ba
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5001
    assume g: "aa < length ap" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5002
      "abc_steps_l (0, code tp # bl2wc (<lm>) # 0\<^bsup>md - pos\<^esup>) ap stp = (aa, ba)" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5003
      "ap \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5004
    thus "?thesis"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5005
      using abc_list_crsp_steps[of "[code tp, bl2wc (<lm>)]"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5006
                                   "md - pos" ap stp aa ba] h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5007
      apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5008
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5009
  qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5010
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5011
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5012
lemma uabc_uhalt: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5013
  "\<lbrakk>turing_basic.t_correct tp; 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5014
  \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp))\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5015
  \<Longrightarrow> \<forall> stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)]) F_aprog 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5016
       stp of (ss, e) \<Rightarrow> ss < length F_aprog"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5017
apply(case_tac "rec_ci rec_F", simp add: F_aprog_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5018
thm uabc_uhalt'
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5019
apply(drule_tac ap = a and pos = b and md = c in uabc_uhalt', simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5020
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5021
  fix a b c
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5022
  assume 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5023
    "\<forall>stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)]) a stp of (ss, e) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5024
                                                   \<Rightarrow> ss < length a"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5025
    "rec_ci rec_F = (a, b, c)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5026
  thus 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5027
    "\<forall>stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)]) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5028
    (a [+] dummy_abc (Suc (Suc 0))) stp of (ss, e) \<Rightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5029
           ss < Suc (Suc (Suc (length a)))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5030
    using abc_append_uhalt1[of a "[code tp, bl2wc (<lm>)]" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5031
      "a [+] dummy_abc (Suc (Suc 0))" "[]" "dummy_abc (Suc (Suc 0))"]  
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5032
    apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5033
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5034
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5035
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5036
lemma tutm_uhalt': 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5037
  "\<lbrakk>turing_basic.t_correct tp;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5038
    \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp))\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5039
  \<Longrightarrow> \<forall> stp. \<not> isS0 (steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5040
  using abacus_turing_eq_uhalt[of "layout_of (F_aprog)" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5041
               "F_aprog" "F_tprog" "[code tp, bl2wc (<lm>)]" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5042
               "start_of (layout_of (F_aprog )) (length (F_aprog))" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5043
               "Suc (Suc 0)"]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5044
apply(simp add: F_tprog_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5045
apply(subgoal_tac "\<forall>stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)])
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5046
  (F_aprog) stp of (as, am) \<Rightarrow> as < length (F_aprog)", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5047
thm abacus_turing_eq_uhalt
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5048
apply(simp add: t_utm_def F_tprog_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5049
apply(rule_tac uabc_uhalt, simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5050
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5051
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5052
lemma tinres_commute: "tinres r r' \<Longrightarrow> tinres r' r"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5053
apply(auto simp: tinres_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5054
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5055
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5056
lemma inres_tape:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5057
  "\<lbrakk>steps (st, l, r) tp stp = (a, b, c); steps (st, l', r') tp stp = (a', b', c'); 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5058
  tinres l l'; tinres r r'\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5059
  \<Longrightarrow> a = a' \<and> tinres b b' \<and> tinres c c'"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5060
proof(case_tac "steps (st, l', r) tp stp")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5061
  fix aa ba ca
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5062
  assume h: "steps (st, l, r) tp stp = (a, b, c)" 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5063
            "steps (st, l', r') tp stp = (a', b', c')"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5064
            "tinres l l'" "tinres r r'"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5065
            "steps (st, l', r) tp stp = (aa, ba, ca)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5066
  have "tinres b ba \<and> c = ca \<and> a = aa"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5067
    using h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5068
    apply(rule_tac tinres_steps, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5069
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5070
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5071
  thm tinres_steps2
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5072
  moreover have "b' = ba \<and> tinres c' ca \<and> a' =  aa"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5073
    using h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5074
    apply(rule_tac tinres_steps2, auto intro: tinres_commute)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5075
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5076
  ultimately show "?thesis"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5077
    apply(auto intro: tinres_commute)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5078
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5079
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5080
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5081
lemma tape_normalize: "\<forall> stp. \<not> isS0 (steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5082
      \<Longrightarrow> \<forall> stp. \<not> isS0 (steps (Suc 0, Bk\<^bsup>m\<^esup>, <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>n\<^esup>) t_utm stp)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5083
apply(rule_tac allI, case_tac "(steps (Suc 0, Bk\<^bsup>m\<^esup>, 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5084
               <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>n\<^esup>) t_utm stp)", simp add: isS0_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5085
apply(erule_tac x = stp in allE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5086
apply(case_tac "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5087
apply(drule_tac inres_tape, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5088
apply(auto simp: tinres_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5089
apply(case_tac "m > Suc (Suc 0)")
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5090
apply(rule_tac x = "m - Suc (Suc 0)" in exI) 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5091
apply(case_tac m, simp_all add: exp_ind_def, case_tac nat, simp_all add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5092
apply(rule_tac x = "2 - m" in exI, simp add: exp_ind_def[THEN sym] exp_add[THEN sym])
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5093
apply(simp only: numeral_2_eq_2, simp add: exp_ind_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5094
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5095
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5096
lemma tutm_uhalt: 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5097
  "\<lbrakk>turing_basic.t_correct tp;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5098
    \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <args>) tp stp))\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5099
  \<Longrightarrow> \<forall> stp. \<not> isS0 (steps (Suc 0, Bk\<^bsup>m\<^esup>, <[code tp, bl2wc (<args>)]> @ Bk\<^bsup>n\<^esup>) t_utm stp)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5100
apply(rule_tac tape_normalize)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5101
apply(rule_tac tutm_uhalt', simp_all)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5102
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5103
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5104
lemma UTM_uhalt_lemma_pre:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5105
  "\<lbrakk>turing_basic.t_correct tp;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5106
   \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <args>) tp stp));
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5107
   args \<noteq> []\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5108
  \<Longrightarrow>  \<forall> stp. \<not> isS0 (steps (Suc 0, [], <code tp # args>)  UTM_pre stp)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5109
proof -
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5110
  let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <code tp # args>"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5111
  let ?Q1 = "\<lambda> (l, r). (l = [Bk] \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5112
             (\<exists> rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5113
  let ?P4 = ?Q1
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5114
  let ?P3 = "\<lambda> (l, r). False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5115
  assume h: "turing_basic.t_correct tp" "\<forall>stp. \<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <args>) tp stp)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5116
            "args \<noteq> []"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5117
  have "?P1 \<turnstile>-> \<lambda> tp. \<not> (\<exists> stp. isS0 (steps (Suc 0, tp) (t_wcode |+| t_utm) stp))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5118
  proof(rule_tac turing_merge.t_merge_uhalt [of "t_wcode" "t_utm"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5119
          ?P1 ?P3 ?P3 ?P4 ?Q1 ?P3], auto simp: turing_merge_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5120
    show "\<exists>stp. case steps (Suc 0, [], <code tp # args>) t_wcode stp of (st, tp') \<Rightarrow> 
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5121
       st = 0 \<and> (case tp' of (l, r) \<Rightarrow> l = [Bk] \<and>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5122
                   (\<exists>rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>))"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5123
      using wcode_lemma_1[of args "code tp"] h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5124
      apply(simp, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5125
      apply(rule_tac x = stp in exI, auto)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5126
      done      
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5127
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5128
    fix rn  stp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5129
    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)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5130
          \<Longrightarrow> False"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5131
      using tutm_uhalt[of tp l args "Suc 0" rn] h
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5132
      apply(simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5133
      apply(erule_tac x = stp in allE)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5134
      apply(simp add: tape_of_nl_abv tape_of_nat_list.simps bin_wc_eq)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5135
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5136
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5137
    fix rn stp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5138
    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>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5139
      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)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5140
      by simp
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5141
  next
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5142
    show "?Q1 \<turnstile>-> ?P4"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5143
      apply(simp add: t_imply_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5144
      done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5145
  qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5146
  thus "?thesis"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5147
    apply(simp add: t_imply_def UTM_pre_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5148
    done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5149
qed
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5150
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5151
text {*
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5152
  The correctness of @{text "UTM"}, the unhalt case.
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5153
  *}
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5154
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5155
lemma UTM_uhalt_lemma:
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5156
  "\<lbrakk>turing_basic.t_correct tp;
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5157
   \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <args>) tp stp));
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5158
   args \<noteq> []\<rbrakk>
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5159
  \<Longrightarrow>  \<forall> stp. \<not> isS0 (steps (Suc 0, [], <code tp # args>)  UTM stp)"
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5160
using UTM_uhalt_lemma_pre[of tp l args]
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5161
apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5162
apply(case_tac "rec_ci rec_F", simp)
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5163
done
48b231495281 Some illustration added together with more explanations.
zhang
parents: 370
diff changeset
  5164
370
1ce04eb1c8ad Initial upload of the formal construction of Universal Turing Machine.
zhang
parents:
diff changeset
  5165
end