utm/abacus.thy
changeset 370 1ce04eb1c8ad
child 371 48b231495281
equal deleted inserted replaced
369:cbb4ac6c8081 370:1ce04eb1c8ad
       
     1 header {* 
       
     2   {\em Abacus} (a kind of register machine) 
       
     3 *}
       
     4 
       
     5 theory abacus
       
     6 imports Main turing_basic
       
     7 begin
       
     8 
       
     9 text {*
       
    10   {\em Abacus} instructions:
       
    11 *}
       
    12 
       
    13 datatype abc_inst =
       
    14   -- {* @{text "Inc n"} increments the memory cell (or register) with address @{text "n"} by one.
       
    15      *}
       
    16      Inc nat
       
    17   -- {*
       
    18      @{text "Dec n label"} decrements the memory cell with address @{text "n"} by one. 
       
    19       If cell @{text "n"} is already zero, no decrements happens and the executio jumps to
       
    20       the instruction labeled by @{text "label"}.
       
    21      *}
       
    22    | Dec nat nat
       
    23   -- {*
       
    24   @{text "Goto label"} unconditionally jumps to the instruction labeled by @{text "label"}.
       
    25   *}
       
    26    | Goto nat
       
    27   
       
    28 
       
    29 text {*
       
    30   Abacus programs are defined as lists of Abacus instructions.
       
    31 *}
       
    32 type_synonym abc_prog = "abc_inst list"
       
    33 
       
    34 section {*
       
    35   Sample Abacus programs
       
    36   *}
       
    37 
       
    38 text {*
       
    39   Abacus for addition and clearance.
       
    40 *}
       
    41 fun plus_clear :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"  
       
    42   where
       
    43   "plus_clear m n e = [Dec m e, Inc n, Goto 0]"
       
    44 
       
    45 text {*
       
    46   Abacus for clearing memory untis.
       
    47 *}
       
    48 fun clear :: "nat \<Rightarrow> nat \<Rightarrow> abc_prog"
       
    49   where
       
    50   "clear n e = [Dec n e, Goto 0]"
       
    51 
       
    52 fun plus:: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
       
    53   where
       
    54   "plus m n p e = [Dec m 4, Inc n, Inc p,
       
    55                    Goto 0, Dec p e, Inc m, Goto 4]"
       
    56 
       
    57 fun mult :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
       
    58   where
       
    59   "mult m1 m2 n p e = [Dec m1 e]@ plus m1 m2 p 1"
       
    60 
       
    61 fun expo :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
       
    62   where
       
    63   "expo n m1 m2 p e = [Inc n, Dec m1 e] @ mult m2 n n p 2"
       
    64 
       
    65 
       
    66 text {*
       
    67   The state of Abacus machine.
       
    68   *}
       
    69 type_synonym abc_state = nat
       
    70 
       
    71 (* text {*
       
    72   The memory of Abacus machine is defined as a function from address to contents.
       
    73 *}
       
    74 type_synonym abc_mem = "nat \<Rightarrow> nat" *)
       
    75 
       
    76 text {*
       
    77   The memory of Abacus machine is defined as a list of contents, with 
       
    78   every units addressed by index into the list.
       
    79   *}
       
    80 type_synonym abc_lm = "nat list"
       
    81 
       
    82 text {*
       
    83   Fetching contents out of memory. Units not represented by list elements are considered
       
    84   as having content @{text "0"}.
       
    85 *}
       
    86 fun abc_lm_v :: "abc_lm \<Rightarrow> nat \<Rightarrow> nat"
       
    87   where 
       
    88     "abc_lm_v lm n = (if (n < length lm) then (lm!n) else 0)"         
       
    89 
       
    90 (*
       
    91 fun abc_l2m :: "abc_lm \<Rightarrow> abc_mem"
       
    92   where 
       
    93     "abc_l2m lm = abc_lm_v lm"
       
    94 *)
       
    95 
       
    96 text {*
       
    97   Set the content of memory unit @{text "n"} to value @{text "v"}.
       
    98   @{text "am"} is the Abacus memory before setting.
       
    99   If address @{text "n"} is outside to scope of @{text "am"}, @{text "am"} 
       
   100   is extended so that @{text "n"} becomes in scope.
       
   101 *}
       
   102 fun abc_lm_s :: "abc_lm \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_lm"
       
   103   where
       
   104     "abc_lm_s am n v = (if (n < length am) then (am[n:=v]) else 
       
   105                            am@ (replicate (n - length am) 0) @ [v])"
       
   106 
       
   107 
       
   108 text {*
       
   109   The configuration of Abaucs machines consists of its current state and its
       
   110   current memory:
       
   111 *}
       
   112 type_synonym abc_conf_l = "abc_state \<times> abc_lm"
       
   113 
       
   114 text {*
       
   115   Fetch instruction out of Abacus program:
       
   116 *}
       
   117 
       
   118 fun abc_fetch :: "nat \<Rightarrow> abc_prog \<Rightarrow> abc_inst option" 
       
   119   where
       
   120   "abc_fetch s p = (if (s < length p) then Some (p ! s)
       
   121                     else None)"
       
   122 
       
   123 text {*
       
   124   Single step execution of Abacus machine. If no instruction is feteched, 
       
   125   configuration does not change.
       
   126 *}
       
   127 fun abc_step_l :: "abc_conf_l \<Rightarrow> abc_inst option \<Rightarrow> abc_conf_l"
       
   128   where
       
   129   "abc_step_l (s, lm) a = (case a of 
       
   130                None \<Rightarrow> (s, lm) |
       
   131                Some (Inc n)  \<Rightarrow> (let nv = abc_lm_v lm n in
       
   132                        (s + 1, abc_lm_s lm n (nv + 1))) |
       
   133                Some (Dec n e) \<Rightarrow> (let nv = abc_lm_v lm n in
       
   134                        if (nv = 0) then (e, abc_lm_s lm n 0) 
       
   135                        else (s + 1,  abc_lm_s lm n (nv - 1))) |
       
   136                Some (Goto n) \<Rightarrow> (n, lm) 
       
   137                )"
       
   138 
       
   139 text {*
       
   140   Multi-step execution of Abacus machine.
       
   141 *}
       
   142 fun abc_steps_l :: "abc_conf_l \<Rightarrow> abc_prog \<Rightarrow> nat \<Rightarrow> abc_conf_l"
       
   143   where
       
   144   "abc_steps_l (s, lm) p 0 = (s, lm)" |
       
   145   "abc_steps_l (s, lm) p (Suc n) = abc_steps_l (abc_step_l (s, lm) (abc_fetch s p)) p n"
       
   146 
       
   147 section {*
       
   148   Compiling Abacus machines into Truing machines
       
   149 *}
       
   150 
       
   151 
       
   152 subsection {*
       
   153   Compiling functions
       
   154 *}
       
   155 
       
   156 text {*
       
   157   @{text "findnth n"} returns the TM which locates the represention of
       
   158   memory cell @{text "n"} on the tape and changes representation of zero
       
   159   on the way.
       
   160 *}
       
   161 
       
   162 fun findnth :: "nat \<Rightarrow> tprog"
       
   163   where
       
   164   "findnth 0 = []" |
       
   165   "findnth (Suc n) = (findnth n @ [(W1, 2 * n + 1), 
       
   166            (R, 2 * n + 2), (R, 2 * n + 3), (R, 2 * n + 2)])"
       
   167 
       
   168 text {*
       
   169   @{text "tinc_b"} returns the TM which increments the representation 
       
   170   of the memory cell under rw-head by one and move the representation 
       
   171   of cells afterwards to the right accordingly.
       
   172   *}
       
   173 
       
   174 definition tinc_b :: "tprog"
       
   175   where
       
   176   "tinc_b \<equiv> [(W1, 1), (R, 2), (W1, 3), (R, 2), (W1, 3), (R, 4), 
       
   177              (L, 7), (W0, 5), (R, 6), (W0, 5), (W1, 3), (R, 6),
       
   178              (L, 8), (L, 7), (R, 9), (L, 7), (R, 10), (W0, 9)]" 
       
   179 
       
   180 text {*
       
   181   @{text "tshift tm off"} shifts @{text "tm"} by offset @{text "off"}, leaving 
       
   182   instructions concerning state @{text "0"} unchanged, because state @{text "0"} 
       
   183   is the end state, which needs not be changed with shift operation.
       
   184   *}
       
   185 
       
   186 fun tshift :: "tprog \<Rightarrow> nat \<Rightarrow> tprog"
       
   187   where
       
   188   "tshift tp off = (map (\<lambda> (action, state). 
       
   189        (action, (if state = 0 then 0
       
   190                  else state + off))) tp)"
       
   191 
       
   192 text {*
       
   193   @{text "tinc ss n"} returns the TM which simulates the execution of 
       
   194   Abacus instruction @{text "Inc n"}, assuming that TM is located at
       
   195   location @{text "ss"} in the final TM complied from the whole
       
   196   Abacus program.
       
   197 *}
       
   198 
       
   199 fun tinc :: "nat \<Rightarrow> nat \<Rightarrow> tprog"
       
   200   where
       
   201   "tinc ss n = tshift (findnth n @ tshift tinc_b (2 * n)) (ss - 1)"
       
   202 
       
   203 text {*
       
   204   @{text "tinc_b"} returns the TM which decrements the representation 
       
   205   of the memory cell under rw-head by one and move the representation 
       
   206   of cells afterwards to the left accordingly.
       
   207   *}
       
   208 
       
   209 definition tdec_b :: "tprog"
       
   210   where
       
   211   "tdec_b \<equiv>  [(W1, 1), (R, 2), (L, 14), (R, 3), (L, 4), (R, 3),
       
   212               (R, 5), (W0, 4), (R, 6), (W0, 5), (L, 7), (L, 8),
       
   213               (L, 11), (W0, 7), (W1, 8), (R, 9), (L, 10), (R, 9),
       
   214               (R, 5), (W0, 10), (L, 12), (L, 11), (R, 13), (L, 11),
       
   215               (R, 17), (W0, 13), (L, 15), (L, 14), (R, 16), (L, 14),
       
   216               (R, 0), (W0, 16)]"
       
   217 
       
   218 text {*
       
   219   @{text "sete tp e"} attaches the termination edges (edges leading to state @{text "0"}) 
       
   220   of TM @{text "tp"} to the intruction labelled by @{text "e"}.
       
   221   *}
       
   222 
       
   223 fun sete :: "tprog \<Rightarrow> nat \<Rightarrow> tprog"
       
   224   where
       
   225   "sete tp e = map (\<lambda> (action, state). (action, if state = 0 then e else state)) tp"
       
   226 
       
   227 text {*
       
   228   @{text "tdec ss n label"} returns the TM which simulates the execution of 
       
   229   Abacus instruction @{text "Dec n label"}, assuming that TM is located at
       
   230   location @{text "ss"} in the final TM complied from the whole
       
   231   Abacus program.
       
   232 *}
       
   233 
       
   234 fun tdec :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> tprog"
       
   235   where
       
   236   "tdec ss n e = sete (tshift (findnth n @ tshift tdec_b (2 * n)) 
       
   237                  (ss - 1)) e"
       
   238  
       
   239 text {*
       
   240   @{text "tgoto f(label)"} returns the TM simulating the execution of Abacus instruction
       
   241   @{text "Goto label"}, where @{text "f(label)"} is the corresponding location of
       
   242   @{text "label"} in the final TM compiled from the overall Abacus program.
       
   243 *}
       
   244 
       
   245 fun tgoto :: "nat \<Rightarrow> tprog"
       
   246   where
       
   247   "tgoto n = [(Nop, n), (Nop, n)]"
       
   248 
       
   249 text {*
       
   250   The layout of the final TM compiled from an Abacus program is represented
       
   251   as a list of natural numbers, where the list element at index @{text "n"} represents the 
       
   252   starting state of the TM simulating the execution of @{text "n"}-th instruction
       
   253   in the Abacus program.
       
   254 *}
       
   255 
       
   256 type_synonym layout = "nat list"
       
   257 
       
   258 text {*
       
   259   @{text "length_of i"} is the length of the 
       
   260   TM simulating the Abacus instruction @{text "i"}.
       
   261 *}
       
   262 fun length_of :: "abc_inst \<Rightarrow> nat"
       
   263   where
       
   264   "length_of i = (case i of 
       
   265                     Inc n   \<Rightarrow> 2 * n + 9 |
       
   266                     Dec n e \<Rightarrow> 2 * n + 16 |
       
   267                     Goto n  \<Rightarrow> 1)"
       
   268 
       
   269 text {*
       
   270   @{text "layout_of ap"} returns the layout of Abacus program @{text "ap"}.
       
   271 *}
       
   272 fun layout_of :: "abc_prog \<Rightarrow> layout"
       
   273   where "layout_of ap = map length_of ap"
       
   274 
       
   275 
       
   276 text {*
       
   277   @{text "start_of layout n"} looks out the starting state of @{text "n"}-th
       
   278   TM in the finall TM.
       
   279 *}
       
   280 
       
   281 fun start_of :: "nat list \<Rightarrow> nat \<Rightarrow> nat"
       
   282   where
       
   283   "start_of ly 0 = Suc 0" |
       
   284   "start_of ly (Suc as) = 
       
   285         (if as < length ly then start_of ly as + (ly ! as)
       
   286          else start_of ly as)"
       
   287 
       
   288 text {*
       
   289   @{text "ci lo ss i"} complies Abacus instruction @{text "i"}
       
   290   assuming the TM of @{text "i"} starts from state @{text "ss"} 
       
   291   within the overal layout @{text "lo"}.
       
   292 *}
       
   293 
       
   294 fun ci :: "layout \<Rightarrow> nat \<Rightarrow> abc_inst \<Rightarrow> tprog"
       
   295   where
       
   296   "ci ly ss i = (case i of 
       
   297                     Inc n   \<Rightarrow> tinc ss n |
       
   298                     Dec n e \<Rightarrow> tdec ss n (start_of ly e) |
       
   299                     Goto n  \<Rightarrow> tgoto (start_of ly n))"
       
   300 
       
   301 text {*
       
   302   @{text "tpairs_of ap"} transfroms Abacus program @{text "ap"} pairing
       
   303   every instruction with its starting state.
       
   304 *}
       
   305 fun tpairs_of :: "abc_prog \<Rightarrow> (nat \<times> abc_inst) list"
       
   306   where "tpairs_of ap = (zip (map (start_of (layout_of ap)) 
       
   307                          [0..<(length ap)]) ap)"
       
   308 
       
   309 
       
   310 text {*
       
   311   @{text "tms_of ap"} returns the list of TMs, where every one of them simulates
       
   312   the corresponding Abacus intruction in @{text "ap"}.
       
   313 *}
       
   314 
       
   315 fun tms_of :: "abc_prog \<Rightarrow> tprog list"
       
   316   where "tms_of ap = map (\<lambda> (n, tm). ci (layout_of ap) n tm) 
       
   317                          (tpairs_of ap)"
       
   318 
       
   319 text {*
       
   320   @{text "tm_of ap"} returns the final TM machine compiled from Abacus program @{text "ap"}.
       
   321 *}
       
   322 fun tm_of :: "abc_prog \<Rightarrow> tprog"
       
   323   where "tm_of ap = concat (tms_of ap)"
       
   324 
       
   325 text {*
       
   326   The following two functions specify the well-formedness of complied TM.
       
   327 *}
       
   328 fun t_ncorrect :: "tprog \<Rightarrow> bool"
       
   329   where
       
   330   "t_ncorrect tp = (length tp mod 2 = 0)"
       
   331 
       
   332 fun abc2t_correct :: "abc_prog \<Rightarrow> bool"
       
   333   where 
       
   334   "abc2t_correct ap = list_all (\<lambda> (n, tm). 
       
   335              t_ncorrect (ci (layout_of ap) n tm)) (tpairs_of ap)"
       
   336 
       
   337 lemma findnth_length: "length (findnth n) div 2 = 2 * n"
       
   338 apply(induct n, simp, simp)
       
   339 done
       
   340 
       
   341 lemma ci_length : "length (ci ns n ai) div 2 = length_of ai"
       
   342 apply(auto simp: ci.simps tinc_b_def tdec_b_def findnth_length
       
   343                  split: abc_inst.splits)
       
   344 done
       
   345 
       
   346 subsection {*
       
   347   Representation of Abacus memory by TM tape
       
   348 *}
       
   349 
       
   350 consts tape_of :: "'a \<Rightarrow> block list" ("<_>" 100)
       
   351 
       
   352 text {*
       
   353   @{text "tape_of_nat_list am"} returns the TM tape representing
       
   354   Abacus memory @{text "am"}.
       
   355   *}
       
   356 
       
   357 fun tape_of_nat_list :: "nat list \<Rightarrow> block list"
       
   358   where 
       
   359   "tape_of_nat_list [] = []" |
       
   360   "tape_of_nat_list [n] = Oc\<^bsup>n+1\<^esup>" |
       
   361   "tape_of_nat_list (n#ns) = (Oc\<^bsup>n+1\<^esup>) @ [Bk] @ (tape_of_nat_list ns)"
       
   362 
       
   363 defs (overloaded)
       
   364   tape_of_nl_abv: "<am> \<equiv> tape_of_nat_list am"
       
   365   tape_of_nat_abv : "<(n::nat)> \<equiv> Oc\<^bsup>n+1\<^esup>"
       
   366 
       
   367 text {*
       
   368   @{text "crsp_l acf tcf"} meams the abacus configuration @{text "acf"}
       
   369   is corretly represented by the TM configuration @{text "tcf"}.
       
   370 *}
       
   371 
       
   372 fun crsp_l :: "layout \<Rightarrow> abc_conf_l \<Rightarrow> t_conf \<Rightarrow> block list \<Rightarrow> bool"
       
   373   where 
       
   374   "crsp_l ly (as, lm) (ts, (l, r)) inres = 
       
   375            (ts = start_of ly as \<and> (\<exists> rn. r = <lm> @ Bk\<^bsup>rn\<^esup>)
       
   376             \<and> l = Bk # Bk # inres)"
       
   377 
       
   378 declare crsp_l.simps[simp del]
       
   379 
       
   380 subsection {*
       
   381   A more general definition of TM execution. 
       
   382 *}
       
   383 
       
   384 (*
       
   385 fun nnth_of :: "(taction \<times> nat) list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> (taction \<times> nat)"
       
   386   where
       
   387   "nnth_of p s b = (if 2*s < length p 
       
   388                     then (p ! (2*s + b)) else (Nop, 0))"
       
   389 
       
   390 thm nth_of.simps
       
   391 
       
   392 fun nfetch :: "tprog \<Rightarrow> nat \<Rightarrow> block \<Rightarrow> taction \<times> nat"
       
   393   where
       
   394   "nfetch p 0 b = (Nop, 0)" |
       
   395   "nfetch p (Suc s) b = 
       
   396              (case b of 
       
   397                 Bk \<Rightarrow> nnth_of p s 0 |
       
   398                 Oc \<Rightarrow> nnth_of p s 1)"
       
   399 *)
       
   400 
       
   401 text {*
       
   402   @{text "t_step tcf (tp, ss)"} returns the result of one step exection of TM @{text "tp"}
       
   403   assuming @{text "tp"} starts from instial state @{text "ss"}.
       
   404 *}
       
   405 
       
   406 fun t_step :: "t_conf \<Rightarrow> (tprog \<times> nat) \<Rightarrow> t_conf"
       
   407   where 
       
   408   "t_step c (p, off) = 
       
   409            (let (state, leftn, rightn) = c in
       
   410             let (action, next_state) = fetch p (state-off)
       
   411                              (case rightn of 
       
   412                                 [] \<Rightarrow> Bk | 
       
   413                                 Bk # xs \<Rightarrow> Bk |
       
   414                                 Oc # xs \<Rightarrow> Oc
       
   415                              ) 
       
   416              in 
       
   417             (next_state, new_tape action (leftn, rightn)))"
       
   418 
       
   419 
       
   420 text {*
       
   421   @{text "t_steps tcf (tp, ss) n"} returns the result of @{text "n"}-step exection 
       
   422   of TM @{text "tp"} assuming @{text "tp"} starts from instial state @{text "ss"}.
       
   423 *}
       
   424 
       
   425 fun t_steps :: "t_conf \<Rightarrow> (tprog \<times> nat) \<Rightarrow> nat \<Rightarrow> t_conf"
       
   426   where
       
   427   "t_steps c (p, off) 0 = c" |
       
   428   "t_steps c (p, off) (Suc n) = t_steps 
       
   429                      (t_step c (p, off)) (p, off) n" 
       
   430 
       
   431 lemma stepn: "t_steps c (p, off) (Suc n) = 
       
   432               t_step (t_steps c (p, off) n) (p, off)"
       
   433 apply(induct n arbitrary: c, simp add: t_steps.simps)
       
   434 apply(simp add: t_steps.simps)
       
   435 done
       
   436 
       
   437 text {*
       
   438   The type of invarints expressing correspondence between 
       
   439   Abacus configuration and TM configuration.
       
   440 *}
       
   441 
       
   442 type_synonym inc_inv_t = "abc_conf_l \<Rightarrow> t_conf \<Rightarrow> block list \<Rightarrow> bool"
       
   443 
       
   444 declare tms_of.simps[simp del] tm_of.simps[simp del]
       
   445         layout_of.simps[simp del] abc_fetch.simps [simp del]  
       
   446         t_step.simps[simp del] t_steps.simps[simp del] 
       
   447         tpairs_of.simps[simp del] start_of.simps[simp del]
       
   448         fetch.simps [simp del] t_ncorrect.simps[simp del]
       
   449         new_tape.simps [simp del] ci.simps [simp del] length_of.simps[simp del] 
       
   450         layout_of.simps[simp del] crsp_l.simps[simp del]
       
   451         abc2t_correct.simps[simp del]
       
   452 
       
   453 lemma tct_div2: "t_ncorrect tp \<Longrightarrow> (length tp) mod 2 = 0"
       
   454 apply(simp add: t_ncorrect.simps)
       
   455 done
       
   456 
       
   457 lemma t_shift_fetch: 
       
   458     "\<lbrakk>t_ncorrect tp1; t_ncorrect tp; 
       
   459       length tp1 div 2 < a \<and> a \<le> length tp1 div 2 + length tp div 2\<rbrakk>
       
   460     \<Longrightarrow> fetch tp (a - length tp1 div 2) b = 
       
   461          fetch (tp1 @ tp @ tp2) a b"
       
   462 apply(subgoal_tac "\<exists> x. a = length tp1 div 2 + x", erule exE, simp)
       
   463 apply(case_tac x, simp)
       
   464 apply(subgoal_tac "length tp1 div 2 + Suc nat = 
       
   465              Suc (length tp1 div 2 + nat)")
       
   466 apply(simp only: fetch.simps nth_of.simps, auto)
       
   467 apply(case_tac b, simp)
       
   468 apply(subgoal_tac "2 * (length tp1 div 2) = length tp1", simp)
       
   469 apply(subgoal_tac "2 * nat < length tp", simp add: nth_append, simp)
       
   470 apply(simp add: t_ncorrect.simps, auto)
       
   471 apply(subgoal_tac "2 * (length tp1 div 2) = length tp1", simp)
       
   472 apply(subgoal_tac "2 * nat < length tp", simp add: nth_append, auto)
       
   473 apply(simp add: t_ncorrect.simps, auto)
       
   474 apply(rule_tac x = "a - length tp1 div 2" in exI, simp)
       
   475 done
       
   476 
       
   477 lemma t_shift_in_step:
       
   478       "\<lbrakk>t_step (a, aa, ba) (tp, length tp1 div 2) = (s, l, r);
       
   479         t_ncorrect tp1; t_ncorrect tp;
       
   480         length tp1 div 2 < a \<and> a \<le> length tp1 div 2 + length tp div 2\<rbrakk>
       
   481        \<Longrightarrow> t_step (a, aa, ba) (tp1 @ tp @ tp2, 0) = (s, l, r)"
       
   482 apply(simp add: t_step.simps)
       
   483 apply(subgoal_tac "fetch tp (a - length tp1 div 2) (case ba of [] \<Rightarrow> 
       
   484                    Bk | x # xs \<Rightarrow> x)
       
   485              = fetch (tp1 @ tp @ tp2) a (case ba of [] \<Rightarrow> Bk | x # xs
       
   486                    \<Rightarrow> x)")
       
   487 apply(case_tac "fetch tp (a - length tp1 div 2) (case ba of [] \<Rightarrow> Bk
       
   488                 | x # xs \<Rightarrow> x)")
       
   489 apply(auto intro: t_shift_fetch)
       
   490 apply(case_tac ba, simp, simp)
       
   491 apply(case_tac aaa, simp, simp)
       
   492 done
       
   493 
       
   494 declare add_Suc_right[simp del]
       
   495 lemma t_step_add: "t_steps c (p, off) (m + n) = 
       
   496           t_steps (t_steps c (p, off) m) (p, off) n"
       
   497 apply(induct m arbitrary: n,  simp add: t_steps.simps, simp)
       
   498 apply(subgoal_tac "t_steps c (p, off) (Suc (m + n)) = 
       
   499                          t_steps c (p, off) (m + Suc n)", simp)
       
   500 apply(subgoal_tac "t_steps (t_steps c (p, off) m) (p, off) (Suc n) =
       
   501                 t_steps (t_step (t_steps c (p, off) m) (p, off)) 
       
   502                          (p, off) n")
       
   503 apply(simp, simp add: stepn)
       
   504 apply(simp only: t_steps.simps)
       
   505 apply(simp only: add_Suc_right)
       
   506 done
       
   507 declare add_Suc_right[simp]
       
   508 
       
   509 lemma s_out_fetch: "\<lbrakk>t_ncorrect tp; 
       
   510         \<not> (length tp1 div 2 < a \<and> a \<le> length tp1 div 2 + 
       
   511          length tp div 2)\<rbrakk>
       
   512       \<Longrightarrow> fetch tp (a - length tp1 div 2) b = (Nop, 0)"
       
   513 apply(auto)
       
   514 apply(simp add: fetch.simps)
       
   515 apply(subgoal_tac "\<exists> x. a - length tp1 div 2 = length tp div 2 + x")
       
   516 apply(erule exE, simp)
       
   517 apply(case_tac x, simp)
       
   518 apply(auto simp add: fetch.simps)
       
   519 apply(subgoal_tac "2 * (length tp div 2) =  length tp")
       
   520 apply(auto simp: t_ncorrect.simps split: block.splits)
       
   521 apply(rule_tac x = "a - length tp1 div 2 - length tp div 2" in exI
       
   522      , simp)
       
   523 done 
       
   524 
       
   525 lemma conf_keep_step: 
       
   526       "\<lbrakk>t_ncorrect tp; 
       
   527         \<not> (length tp1 div 2 < a \<and> a \<le> length tp1 div 2 + 
       
   528        length tp div 2)\<rbrakk>
       
   529       \<Longrightarrow> t_step (a, aa, ba) (tp, length tp1 div 2) = (0, aa, ba)"
       
   530 apply(simp add: t_step.simps)
       
   531 apply(subgoal_tac "fetch tp (a - length tp1 div 2) (case ba of [] \<Rightarrow> 
       
   532   Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc) = (Nop, 0)")
       
   533 apply(simp add: new_tape.simps)
       
   534 apply(rule s_out_fetch, simp, simp)
       
   535 done
       
   536 
       
   537 lemma conf_keep: 
       
   538       "\<lbrakk>t_ncorrect tp; 
       
   539         \<not> (length tp1 div 2 < a \<and>
       
   540         a \<le> length tp1 div 2 + length tp div 2); n > 0\<rbrakk>
       
   541       \<Longrightarrow> t_steps (a, aa, ba) (tp, length tp1 div 2) n = (0, aa, ba)"
       
   542 apply(induct n, simp)
       
   543 apply(case_tac n, simp add: t_steps.simps)
       
   544 apply(rule_tac conf_keep_step, simp+)
       
   545 apply(subgoal_tac " t_steps (a, aa, ba) 
       
   546                (tp, length tp1 div 2) (Suc (Suc nat))
       
   547          = t_step (t_steps (a, aa, ba) 
       
   548             (tp, length tp1 div 2) (Suc nat)) (tp, length tp1 div 2)")
       
   549 apply(simp)
       
   550 apply(rule_tac conf_keep_step, simp, simp)
       
   551 apply(rule stepn)
       
   552 done
       
   553 
       
   554 lemma state_bef_inside: 
       
   555     "\<lbrakk>t_ncorrect tp1; t_ncorrect tp; 
       
   556       t_steps (s0, l0, r0) (tp, length tp1 div 2) stp = (s, l, r);
       
   557       length tp1 div 2 < s0 \<and> 
       
   558          s0 \<le> length tp1 div 2 + length tp div 2;
       
   559       length tp1 div 2 < s \<and> s \<le> length tp1 div 2 + length tp div 2; 
       
   560       n < stp; t_steps (s0, l0, r0) (tp, length tp1 div 2) n = 
       
   561       (a, aa, ba)\<rbrakk>
       
   562       \<Longrightarrow>  length tp1 div 2 < a \<and> 
       
   563          a \<le> length tp1 div 2 + length tp div 2"
       
   564 apply(subgoal_tac "\<exists> x. stp = n + x", erule exE)
       
   565 apply(simp only: t_step_add)
       
   566 apply(rule classical)
       
   567 apply(subgoal_tac "t_steps (a, aa, ba) 
       
   568           (tp, length tp1 div 2) x = (0, aa, ba)")
       
   569 apply(simp)
       
   570 apply(rule conf_keep, simp, simp, simp)
       
   571 apply(rule_tac x = "stp - n" in exI, simp)
       
   572 done
       
   573 
       
   574 lemma turing_shift_inside: 
       
   575        "\<lbrakk>t_steps (s0, l0, r0) (tp, length tp1 div 2) stp = (s, l, r);
       
   576          length tp1 div 2 < s0 \<and> 
       
   577          s0 \<le> length tp1 div 2 + length tp div 2; 
       
   578          t_ncorrect tp1; t_ncorrect tp;
       
   579          length tp1 div 2 < s \<and> 
       
   580          s \<le> length tp1 div 2 + length tp div 2\<rbrakk>
       
   581        \<Longrightarrow> t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) stp = (s, l, r)"
       
   582 apply(induct stp arbitrary: s l r)
       
   583 apply(simp add: t_steps.simps)
       
   584 apply(subgoal_tac " t_steps (s0, l0, r0) 
       
   585         (tp, length tp1 div 2) (Suc stp)
       
   586                   = t_step (t_steps (s0, l0, r0) 
       
   587         (tp, length tp1 div 2) stp) (tp, length tp1 div 2)")
       
   588 apply(case_tac "t_steps (s0, l0, r0) (tp, length tp1 div 2) stp")
       
   589 apply(subgoal_tac "length tp1 div 2 < a \<and> 
       
   590             a \<le> length tp1 div 2 + length tp div 2")
       
   591 apply(subgoal_tac "t_steps (s0, l0, r0) 
       
   592            (tp1 @ tp @ tp2, 0) stp = (a, b, c)")
       
   593 apply(simp only: stepn, simp)
       
   594 apply(rule_tac t_shift_in_step, simp+)
       
   595 defer
       
   596 apply(rule stepn)
       
   597 apply(rule_tac n = stp and stp = "Suc stp" and a = a 
       
   598                and aa = b and ba = c in state_bef_inside, simp+)
       
   599 done
       
   600 
       
   601 lemma take_Suc_last[elim]: "Suc as \<le> length xs \<Longrightarrow> 
       
   602             take (Suc as) xs = take as xs @ [xs ! as]"
       
   603 apply(induct xs arbitrary: as, simp, simp)
       
   604 apply(case_tac as, simp, simp)
       
   605 done
       
   606 
       
   607 lemma concat_suc: "Suc as \<le> length xs \<Longrightarrow> 
       
   608        concat (take (Suc as) xs) = concat (take as xs) @ xs! as"
       
   609 apply(subgoal_tac "take (Suc as) xs = take as xs @ [xs ! as]", simp)
       
   610 by auto
       
   611 
       
   612 lemma concat_take_suc_iff: "Suc n \<le> length tps \<Longrightarrow> 
       
   613        concat (take n tps) @ (tps ! n) = concat (take (Suc n) tps)"
       
   614 apply(drule_tac concat_suc, simp)
       
   615 done
       
   616 
       
   617 lemma concat_drop_suc_iff: 
       
   618    "Suc n < length tps \<Longrightarrow> concat (drop (Suc n) tps) = 
       
   619            tps ! Suc n @ concat (drop (Suc (Suc n)) tps)"
       
   620 apply(induct tps arbitrary: n, simp, simp)
       
   621 apply(case_tac tps, simp, simp)
       
   622 apply(case_tac n, simp, simp)
       
   623 done
       
   624 
       
   625 declare append_assoc[simp del]
       
   626 
       
   627 lemma  tm_append: "\<lbrakk>n < length tps; tp = tps ! n\<rbrakk> \<Longrightarrow> 
       
   628            \<exists> tp1 tp2. concat tps = tp1 @ tp @ tp2 \<and> tp1 = 
       
   629               concat (take n tps) \<and> tp2 = concat (drop (Suc n) tps)"
       
   630 apply(rule_tac x = "concat (take n tps)" in exI)
       
   631 apply(rule_tac x = "concat (drop (Suc n) tps)" in exI)
       
   632 apply(auto)
       
   633 apply(induct n, simp)
       
   634 apply(case_tac tps, simp, simp, simp)
       
   635 apply(subgoal_tac "concat (take n tps) @ (tps ! n) = 
       
   636                concat (take (Suc n) tps)")
       
   637 apply(simp only: append_assoc[THEN sym], simp only: append_assoc)
       
   638 apply(subgoal_tac " concat (drop (Suc n) tps) = tps ! Suc n @ 
       
   639                   concat (drop (Suc (Suc n)) tps)", simp)
       
   640 apply(rule_tac concat_drop_suc_iff, simp)
       
   641 apply(rule_tac concat_take_suc_iff, simp)
       
   642 done
       
   643 
       
   644 declare append_assoc[simp]
       
   645 
       
   646 lemma map_of:  "n < length xs \<Longrightarrow> (map f xs) ! n = f (xs ! n)"
       
   647 by(auto)
       
   648 
       
   649 lemma [simp]: "length (tms_of aprog) = length aprog"
       
   650 apply(auto simp: tms_of.simps tpairs_of.simps)
       
   651 done
       
   652 
       
   653 lemma ci_nth: "\<lbrakk>ly = layout_of aprog; as < length aprog; 
       
   654                 abc_fetch as aprog = Some ins\<rbrakk>
       
   655     \<Longrightarrow> ci ly (start_of ly as) ins = tms_of aprog ! as"
       
   656 apply(simp add: tms_of.simps tpairs_of.simps 
       
   657       abc_fetch.simps  map_of del: map_append)
       
   658 done
       
   659 
       
   660 lemma t_split:"\<lbrakk>
       
   661         ly = layout_of aprog;
       
   662         as < length aprog; abc_fetch as aprog = Some ins\<rbrakk>
       
   663       \<Longrightarrow> \<exists> tp1 tp2. concat (tms_of aprog) = 
       
   664             tp1 @ (ci ly (start_of ly as) ins) @ tp2
       
   665             \<and> tp1 = concat (take as (tms_of aprog)) \<and> 
       
   666               tp2 = concat (drop (Suc as) (tms_of aprog))"
       
   667 apply(insert tm_append[of "as" "tms_of aprog" 
       
   668                              "ci ly (start_of ly as) ins"], simp)
       
   669 apply(subgoal_tac "ci ly (start_of ly as) ins = (tms_of aprog) ! as")
       
   670 apply(subgoal_tac "length (tms_of aprog) = length aprog", simp, simp)
       
   671 apply(rule_tac ci_nth, auto)
       
   672 done
       
   673 
       
   674 lemma math_sub: "\<lbrakk>x >= Suc 0; x - 1 = z\<rbrakk> \<Longrightarrow> x + y - Suc 0 = z + y"
       
   675 by auto
       
   676 
       
   677 lemma start_more_one: "as \<noteq> 0 \<Longrightarrow> start_of ly as >= Suc 0"
       
   678 apply(induct as, simp add: start_of.simps)
       
   679 apply(case_tac as, auto simp: start_of.simps)
       
   680 done
       
   681 
       
   682 lemma tm_ct: "\<lbrakk>abc2t_correct aprog; tp \<in> set (tms_of aprog)\<rbrakk> \<Longrightarrow> 
       
   683                            t_ncorrect tp"
       
   684 apply(simp add: abc2t_correct.simps tms_of.simps)
       
   685 apply(auto)
       
   686 apply(simp add:list_all_iff, auto)
       
   687 done
       
   688 
       
   689 lemma div_apart: "\<lbrakk>x mod (2::nat) = 0; y mod 2 = 0\<rbrakk> 
       
   690           \<Longrightarrow> (x + y) div 2 = x div 2 + y div 2"
       
   691 apply(drule mod_eqD)+
       
   692 apply(auto)
       
   693 done
       
   694 
       
   695 lemma div_apart_iff: "\<lbrakk>x mod (2::nat) = 0; y mod 2 = 0\<rbrakk> \<Longrightarrow> 
       
   696            (x + y) mod 2 = 0"
       
   697 apply(auto)
       
   698 done
       
   699 
       
   700 lemma tms_ct: "\<lbrakk>abc2t_correct aprog; n < length aprog\<rbrakk> \<Longrightarrow> 
       
   701          t_ncorrect (concat (take n (tms_of aprog)))"
       
   702 apply(induct n, simp add: t_ncorrect.simps, simp)
       
   703 apply(subgoal_tac "concat (take (Suc n) (tms_of aprog)) = 
       
   704         concat (take n (tms_of aprog)) @ (tms_of aprog ! n)", simp)
       
   705 apply(simp add: t_ncorrect.simps)
       
   706 apply(rule_tac div_apart_iff, simp)
       
   707 apply(subgoal_tac "t_ncorrect (tms_of aprog ! n)", 
       
   708             simp add: t_ncorrect.simps)
       
   709 apply(rule_tac tm_ct, simp)
       
   710 apply(rule_tac nth_mem, simp add: tms_of.simps tpairs_of.simps)
       
   711 apply(rule_tac concat_suc, simp add: tms_of.simps tpairs_of.simps)
       
   712 done
       
   713 
       
   714 lemma tcorrect_div2: "\<lbrakk>abc2t_correct aprog; Suc as < length aprog\<rbrakk>
       
   715   \<Longrightarrow> (length (concat (take as (tms_of aprog))) + length (tms_of aprog
       
   716  ! as)) div 2 = length (concat (take as (tms_of aprog))) div 2 + 
       
   717                  length (tms_of aprog ! as) div 2"
       
   718 apply(subgoal_tac "t_ncorrect (tms_of aprog ! as)")
       
   719 apply(subgoal_tac "t_ncorrect (concat (take as (tms_of aprog)))")
       
   720 apply(rule_tac div_apart)
       
   721 apply(rule tct_div2, simp)+
       
   722 apply(erule_tac tms_ct, simp)
       
   723 apply(rule_tac tm_ct, simp)
       
   724 apply(rule_tac nth_mem)
       
   725 apply(simp add: tms_of.simps tpairs_of.simps)
       
   726 done
       
   727 
       
   728 lemma [simp]: "length (layout_of aprog) = length aprog"
       
   729 apply(auto simp: layout_of.simps)
       
   730 done
       
   731 
       
   732 lemma start_of_ind: "\<lbrakk>as < length aprog; ly = layout_of aprog\<rbrakk> \<Longrightarrow> 
       
   733        start_of ly (Suc as) = start_of ly as + 
       
   734                           length ((tms_of aprog) ! as) div 2"
       
   735 apply(simp only: start_of.simps, simp)
       
   736 apply(auto simp: start_of.simps tms_of.simps layout_of.simps 
       
   737                  tpairs_of.simps)
       
   738 apply(simp add: ci_length)
       
   739 done
       
   740 
       
   741 lemma concat_take_suc: "Suc n \<le> length xs \<Longrightarrow>
       
   742   concat (take (Suc n) xs) = concat (take n xs) @ (xs ! n)"
       
   743 apply(subgoal_tac "take (Suc n) xs =
       
   744                    take n xs @ [xs ! n]")
       
   745 apply(auto)
       
   746 done
       
   747 
       
   748 lemma ci_length_not0: "Suc 0 <= length (ci ly as i) div 2"
       
   749 apply(subgoal_tac "length (ci ly as i) div 2 = length_of i")
       
   750 apply(simp add: length_of.simps split: abc_inst.splits)
       
   751 apply(rule ci_length)
       
   752 done
       
   753  
       
   754 lemma findnth_length2: "length (findnth n) = 4 * n"
       
   755 apply(induct n, simp)
       
   756 apply(simp)
       
   757 done
       
   758 
       
   759 lemma ci_length2: "length (ci ly as i) = 2 * (length_of i)"
       
   760 apply(simp add: ci.simps length_of.simps tinc_b_def tdec_b_def
       
   761               split: abc_inst.splits, auto)
       
   762 apply(simp add: findnth_length2)+
       
   763 done
       
   764 
       
   765 lemma tm_mod2: "as < length aprog \<Longrightarrow> 
       
   766              length (tms_of aprog ! as) mod 2 = 0"
       
   767 apply(simp add: tms_of.simps)
       
   768 apply(subgoal_tac "map (\<lambda>(x, y). ci (layout_of aprog) x y) 
       
   769               (tpairs_of aprog) ! as
       
   770                 = (\<lambda>(x, y). ci (layout_of aprog) x y) 
       
   771               ((tpairs_of aprog) ! as)", simp)
       
   772 apply(case_tac "(tpairs_of aprog ! as)", simp)
       
   773 apply(subgoal_tac "length (ci (layout_of aprog) a b) =
       
   774                  2 * (length_of b)", simp)
       
   775 apply(rule ci_length2)
       
   776 apply(rule map_of, simp add: tms_of.simps tpairs_of.simps)
       
   777 done
       
   778 
       
   779 lemma tms_mod2: "as \<le> length aprog \<Longrightarrow> 
       
   780         length (concat (take as (tms_of aprog))) mod 2 = 0"
       
   781 apply(induct as, simp, simp)
       
   782 apply(subgoal_tac "concat (take (Suc as) (tms_of aprog))
       
   783                   = concat (take as (tms_of aprog)) @ 
       
   784                        (tms_of aprog ! as)", auto)
       
   785 apply(rule div_apart_iff, simp, rule tm_mod2, simp)
       
   786 apply(rule concat_take_suc, simp add: tms_of.simps tpairs_of.simps)
       
   787 done
       
   788 
       
   789 lemma [simp]: "\<lbrakk>as < length aprog; (abc_fetch as aprog) = Some ins\<rbrakk>
       
   790        \<Longrightarrow> ci (layout_of aprog) 
       
   791           (start_of (layout_of aprog) as) (ins) \<in> set (tms_of aprog)"
       
   792 apply(insert ci_nth[of "layout_of aprog" aprog as], simp)
       
   793 done
       
   794 
       
   795 lemma startof_not0: "start_of ly as > 0"
       
   796 apply(induct as, simp add: start_of.simps)
       
   797 apply(case_tac as, auto simp: start_of.simps)
       
   798 done
       
   799 
       
   800 declare abc_step_l.simps[simp del]
       
   801 lemma pre_lheq: "\<lbrakk>tp = concat (take as (tms_of aprog));
       
   802    abc2t_correct aprog; as \<le> length aprog\<rbrakk> \<Longrightarrow> 
       
   803          start_of (layout_of aprog) as - Suc 0 = length tp div 2"
       
   804 apply(induct as arbitrary: tp, simp add: start_of.simps, simp)
       
   805 proof - 
       
   806   fix as tp
       
   807   assume h1: "\<And>tp. tp = concat (take as (tms_of aprog)) \<Longrightarrow> 
       
   808      start_of (layout_of aprog) as - Suc 0 = 
       
   809             length (concat (take as (tms_of aprog))) div 2"
       
   810   and h2: " abc2t_correct aprog" "Suc as \<le> length aprog" 
       
   811   from h2 show "start_of (layout_of aprog) (Suc as) - Suc 0 = 
       
   812           length (concat (take (Suc as) (tms_of aprog))) div 2"
       
   813     apply(insert h1[of "concat (take as (tms_of aprog))"], simp)
       
   814     apply(insert start_of_ind[of as aprog "layout_of aprog"], simp)
       
   815     apply(subgoal_tac "(take (Suc as) (tms_of aprog)) = 
       
   816             take as (tms_of aprog) @ [(tms_of aprog) ! as]", simp)
       
   817     apply(subgoal_tac "(length (concat (take as (tms_of aprog))) + 
       
   818                        length (tms_of aprog ! as)) div 2
       
   819             = length (concat (take as (tms_of aprog))) div 2 + 
       
   820               length (tms_of aprog ! as) div 2", simp)
       
   821     apply(subgoal_tac "start_of (layout_of aprog) as = 
       
   822        length (concat (take as (tms_of aprog))) div 2 + Suc 0", simp)
       
   823     apply(subgoal_tac "start_of (layout_of aprog) as > 0", simp, 
       
   824            rule_tac startof_not0)
       
   825     apply(insert tm_mod2[of as aprog], simp)
       
   826     apply(insert tms_mod2[of as aprog], simp, arith)
       
   827     apply(rule take_Suc_last, simp)
       
   828     done
       
   829 qed
       
   830 
       
   831 lemma crsp2stateq: 
       
   832   "\<lbrakk>as < length aprog; abc2t_correct aprog;
       
   833        crsp_l (layout_of aprog) (as, am) (a, aa, ba) inres\<rbrakk> \<Longrightarrow> 
       
   834         a = length (concat (take as (tms_of aprog))) div 2 + 1"
       
   835 apply(simp add: crsp_l.simps)
       
   836 apply(insert pre_lheq[of "(concat (take as (tms_of aprog)))" as aprog]
       
   837 , simp)   
       
   838 apply(subgoal_tac "start_of (layout_of aprog) as > 0", 
       
   839       auto intro: startof_not0)
       
   840 done
       
   841 
       
   842 lemma turing_shift_outside: 
       
   843      "\<lbrakk>t_steps (s0, l0, r0) (tp, length tp1 div 2) stp = (s, l, r); 
       
   844        s \<noteq> 0; stp > 0;
       
   845        length tp1 div 2 < s0 \<and> 
       
   846        s0 \<le> length tp1 div 2 + length tp div 2; 
       
   847        t_ncorrect tp1; t_ncorrect tp;
       
   848        \<not> (length tp1 div 2 < s \<and> 
       
   849       s \<le> length tp1 div 2 + length tp div 2)\<rbrakk>
       
   850     \<Longrightarrow> \<exists>stp' > 0. t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) stp' 
       
   851                 = (s, l, r)"
       
   852 apply(rule_tac x = stp in exI)
       
   853 apply(case_tac stp, simp add: t_steps.simps)
       
   854 apply(simp only: stepn)
       
   855 apply(case_tac "t_steps (s0, l0, r0) (tp, length tp1 div 2) nat")
       
   856 apply(subgoal_tac "length tp1 div 2 < a \<and> 
       
   857                    a \<le> length tp1 div 2 + length tp div 2")
       
   858 apply(subgoal_tac "t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) nat 
       
   859                    = (a, b, c)", simp)
       
   860 apply(rule_tac t_shift_in_step, simp+)
       
   861 apply(rule_tac turing_shift_inside, simp+)
       
   862 apply(rule classical)
       
   863 apply(subgoal_tac "t_step (a,b,c) 
       
   864             (tp, length tp1 div 2) = (0, b, c)", simp)
       
   865 apply(rule_tac conf_keep_step, simp+)
       
   866 done
       
   867 
       
   868 lemma turing_shift: 
       
   869   "\<lbrakk>t_steps (s0, (l0, r0)) (tp, (length tp1 div 2)) stp
       
   870    = (s, (l, r)); s \<noteq> 0; stp > 0;
       
   871   (length tp1 div 2 < s0 \<and> s0 <= length tp1 div 2 + length tp div 2);
       
   872   t_ncorrect tp1; t_ncorrect tp\<rbrakk> \<Longrightarrow> 
       
   873          \<exists> stp' > 0. t_steps (s0, (l0, r0)) (tp1 @ tp @ tp2, 0) stp' =
       
   874                     (s, (l, r))"
       
   875 apply(case_tac "s > length tp1 div 2 \<and> 
       
   876               s <= length tp1 div 2 + length tp div 2")
       
   877 apply(subgoal_tac " t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) stp = 
       
   878                    (s, l, r)")
       
   879 apply(rule_tac x = stp in exI, simp)
       
   880 apply(rule_tac turing_shift_inside, simp+)
       
   881 apply(rule_tac turing_shift_outside, simp+)
       
   882 done
       
   883 
       
   884 lemma inc_startof_not0:  "start_of ly as \<ge> Suc 0"
       
   885 apply(induct as, simp add: start_of.simps)
       
   886 apply(simp add: start_of.simps)
       
   887 done
       
   888 
       
   889 lemma s_crsp:
       
   890   "\<lbrakk>as < length aprog; abc_fetch as aprog = Some ins;
       
   891   abc2t_correct aprog;
       
   892   crsp_l (layout_of aprog) (as, am) (a, aa, ba) inres\<rbrakk> \<Longrightarrow>  
       
   893   length (concat (take as (tms_of aprog))) div 2 < a 
       
   894       \<and> a \<le> length (concat (take as (tms_of aprog))) div 2 + 
       
   895          length (ci (layout_of aprog) (start_of (layout_of aprog) as)
       
   896          ins) div 2"
       
   897 apply(subgoal_tac "a = length (concat (take as (tms_of aprog))) div 
       
   898                    2 + 1", simp)
       
   899 apply(rule_tac ci_length_not0)
       
   900 apply(rule crsp2stateq, simp+)
       
   901 done
       
   902 
       
   903 lemma tms_out_ex:
       
   904   "\<lbrakk>ly = layout_of aprog; tprog = tm_of aprog;
       
   905   abc2t_correct aprog;
       
   906   crsp_l ly (as, am) tc inres; as < length aprog;
       
   907   abc_fetch as aprog = Some ins;
       
   908   t_steps tc (ci ly (start_of ly as) ins, 
       
   909   (start_of ly as) - 1) n = (s, l, r);
       
   910   n > 0; 
       
   911   abc_step_l (as, am) (abc_fetch as aprog) = (as', am');
       
   912   s = start_of ly as'
       
   913   \<rbrakk>
       
   914   \<Longrightarrow> \<exists> stp > 0. (t_steps tc (tprog, 0) stp = (s, (l, r)))"
       
   915 apply(simp only: tm_of.simps)
       
   916 apply(subgoal_tac "\<exists> tp1 tp2. concat (tms_of aprog) = 
       
   917       tp1 @ (ci ly (start_of ly as) ins) @ tp2
       
   918     \<and> tp1 = concat (take as (tms_of aprog)) \<and> 
       
   919       tp2 = concat (drop (Suc as) (tms_of aprog))")
       
   920 apply(erule exE, erule exE, erule conjE, erule conjE,
       
   921       case_tac tc, simp)
       
   922 apply(rule turing_shift)
       
   923 apply(subgoal_tac "start_of (layout_of aprog) as - Suc 0 
       
   924                 = length tp1 div 2", simp)
       
   925 apply(rule_tac pre_lheq, simp, simp, simp)
       
   926 apply(simp add: startof_not0, simp)
       
   927 apply(rule_tac s_crsp, simp, simp, simp, simp)
       
   928 apply(rule tms_ct, simp, simp)
       
   929 apply(rule tm_ct, simp)
       
   930 apply(subgoal_tac "ci (layout_of aprog) 
       
   931                  (start_of (layout_of aprog) as) ins
       
   932                 = (tms_of aprog ! as)", simp)
       
   933 apply(simp add: tms_of.simps tpairs_of.simps)
       
   934 apply(simp add: tms_of.simps tpairs_of.simps abc_fetch.simps)
       
   935 apply(erule_tac t_split, auto simp: tm_of.simps)
       
   936 done
       
   937 
       
   938 subsubsection {* The compilation of @{text "Inc n"} *}
       
   939 
       
   940 text {*
       
   941   The lemmas in this section lead to the correctness of 
       
   942   the compilation of @{text "Inc n"} instruction.
       
   943 *}
       
   944 
       
   945 (*****Begin: inc crsp*******)
       
   946 fun at_begin_fst_bwtn :: "inc_inv_t"
       
   947   where
       
   948   "at_begin_fst_bwtn (as, lm) (s, l, r) ires = 
       
   949       (\<exists> lm1 tn rn. lm1 = (lm @ (0\<^bsup>tn\<^esup>)) \<and> length lm1 = s \<and> 
       
   950           (if lm1 = [] then l = Bk # Bk # ires
       
   951            else l = [Bk]@<rev lm1>@Bk#Bk#ires) \<and> r = (Bk\<^bsup>rn\<^esup>))" 
       
   952 
       
   953 
       
   954 fun at_begin_fst_awtn :: "inc_inv_t"
       
   955   where
       
   956   "at_begin_fst_awtn (as, lm) (s, l, r) ires = 
       
   957       (\<exists> lm1 tn rn. lm1 = (lm @ (0\<^bsup>tn\<^esup>)) \<and> length lm1 = s \<and>
       
   958          (if lm1 = []  then l = Bk # Bk # ires
       
   959           else l = [Bk]@<rev lm1>@Bk#Bk#ires) \<and> r = [Oc]@Bk\<^bsup>rn\<^esup>
       
   960   )"
       
   961 
       
   962 fun at_begin_norm :: "inc_inv_t"
       
   963   where
       
   964   "at_begin_norm (as, lm) (s, l, r) ires= 
       
   965       (\<exists> lm1 lm2 rn. lm = lm1 @ lm2 \<and> length lm1 = s \<and> 
       
   966         (if lm1 = [] then l = Bk # Bk # ires
       
   967          else l = Bk # <rev lm1> @ Bk# Bk # ires ) \<and> r = <lm2> @ (Bk\<^bsup>rn\<^esup>))"
       
   968 
       
   969 fun in_middle :: "inc_inv_t"
       
   970   where
       
   971   "in_middle (as, lm) (s, l, r) ires = 
       
   972       (\<exists> lm1 lm2 tn m ml mr rn. lm @ 0\<^bsup>tn\<^esup> = lm1 @ [m] @ lm2
       
   973        \<and> length lm1 = s \<and> m + 1 = ml + mr \<and>  
       
   974          ml \<noteq> 0 \<and> tn = s + 1 - length lm \<and> 
       
   975        (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires 
       
   976         else l = (Oc\<^bsup>ml\<^esup>)@[Bk]@<rev lm1>@
       
   977                  Bk # Bk # ires) \<and> (r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ <lm2>@ (Bk\<^bsup>rn\<^esup>) \<or> 
       
   978       (lm2 = [] \<and> r = (Oc\<^bsup>mr\<^esup>)))
       
   979       )"
       
   980 
       
   981 fun inv_locate_a :: "inc_inv_t"
       
   982   where "inv_locate_a (as, lm) (s, l, r) ires = 
       
   983      (at_begin_norm (as, lm) (s, l, r) ires \<or>
       
   984       at_begin_fst_bwtn (as, lm) (s, l, r) ires \<or>
       
   985       at_begin_fst_awtn (as, lm) (s, l, r) ires
       
   986       )"
       
   987 
       
   988 fun inv_locate_b :: "inc_inv_t"
       
   989   where "inv_locate_b (as, lm) (s, l, r) ires = 
       
   990         (in_middle (as, lm) (s, l, r)) ires "
       
   991 
       
   992 fun inv_after_write :: "inc_inv_t"
       
   993   where "inv_after_write (as, lm) (s, l, r) ires = 
       
   994            (\<exists> rn m lm1 lm2. lm = lm1 @ m # lm2 \<and>
       
   995              (if lm1 = [] then l = Oc\<^bsup>m\<^esup> @ Bk # Bk # ires
       
   996               else Oc # l = Oc\<^bsup>Suc m \<^esup>@ Bk # <rev lm1> @ 
       
   997                       Bk # Bk # ires) \<and> r = [Oc] @ <lm2> @ (Bk\<^bsup>rn\<^esup>))"
       
   998 
       
   999 fun inv_after_move :: "inc_inv_t"
       
  1000   where "inv_after_move (as, lm) (s, l, r) ires = 
       
  1001       (\<exists> rn m lm1 lm2. lm = lm1 @ m # lm2 \<and>
       
  1002         (if lm1 = [] then l = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # ires
       
  1003          else l = Oc\<^bsup>Suc m\<^esup>@ Bk # <rev lm1> @ Bk # Bk # ires) \<and> 
       
  1004         r = <lm2> @ (Bk\<^bsup>rn\<^esup>))"
       
  1005 
       
  1006 fun inv_after_clear :: "inc_inv_t"
       
  1007   where "inv_after_clear (as, lm) (s, l, r) ires =
       
  1008        (\<exists> rn m lm1 lm2 r'. lm = lm1 @ m # lm2 \<and> 
       
  1009         (if lm1 = [] then l = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # ires
       
  1010          else l = Oc\<^bsup>Suc m\<^esup>@ Bk # <rev lm1> @ Bk # Bk # ires) \<and> 
       
  1011           r = Bk # r' \<and> Oc # r' = <lm2> @ (Bk\<^bsup>rn\<^esup>))"
       
  1012 
       
  1013 fun inv_on_right_moving :: "inc_inv_t"
       
  1014   where "inv_on_right_moving (as, lm) (s, l, r) ires = 
       
  1015        (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
       
  1016             ml + mr = m \<and> 
       
  1017           (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
       
  1018           else l = (Oc\<^bsup>ml\<^esup>) @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
       
  1019          ((r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ <lm2> @ (Bk\<^bsup>rn\<^esup>)) \<or> 
       
  1020           (r = (Oc\<^bsup>mr\<^esup>) \<and> lm2 = [])))"
       
  1021 
       
  1022 fun inv_on_left_moving_norm :: "inc_inv_t"
       
  1023   where "inv_on_left_moving_norm (as, lm) (s, l, r) ires =
       
  1024       (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and>  
       
  1025              ml + mr = Suc m \<and> mr > 0 \<and> (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
       
  1026                                          else l = (Oc\<^bsup>ml\<^esup>) @ Bk # <rev lm1> @ Bk # Bk # ires)
       
  1027         \<and> (r = (Oc\<^bsup>mr\<^esup>) @ Bk # <lm2> @ (Bk\<^bsup>rn\<^esup>) \<or> 
       
  1028            (lm2 = [] \<and> r = Oc\<^bsup>mr\<^esup>)))"
       
  1029 
       
  1030 fun inv_on_left_moving_in_middle_B:: "inc_inv_t"
       
  1031   where "inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires =
       
  1032                 (\<exists> lm1 lm2 rn. lm = lm1 @ lm2 \<and>  
       
  1033                      (if lm1 = [] then l = Bk # ires
       
  1034                       else l = <rev lm1> @ Bk # Bk # ires) \<and> 
       
  1035                       r = Bk # <lm2> @ (Bk\<^bsup>rn\<^esup>))"
       
  1036 
       
  1037 fun inv_on_left_moving :: "inc_inv_t"
       
  1038   where "inv_on_left_moving (as, lm) (s, l, r) ires = 
       
  1039        (inv_on_left_moving_norm  (as, lm) (s, l, r) ires \<or>
       
  1040         inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires)"
       
  1041 
       
  1042 
       
  1043 fun inv_check_left_moving_on_leftmost :: "inc_inv_t"
       
  1044   where "inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires = 
       
  1045                 (\<exists> rn. l = ires \<and> r = [Bk, Bk] @ <lm> @ (Bk\<^bsup>rn\<^esup>))"
       
  1046 
       
  1047 fun inv_check_left_moving_in_middle :: "inc_inv_t"
       
  1048   where "inv_check_left_moving_in_middle (as, lm) (s, l, r) ires = 
       
  1049 
       
  1050               (\<exists> lm1 lm2 r' rn. lm = lm1 @ lm2 \<and>
       
  1051                  (Oc # l = <rev lm1> @ Bk # Bk # ires) \<and> r = Oc # Bk # r' \<and> 
       
  1052                            r' = <lm2> @ (Bk\<^bsup>rn\<^esup>))"
       
  1053 
       
  1054 fun inv_check_left_moving :: "inc_inv_t"
       
  1055   where "inv_check_left_moving (as, lm) (s, l, r) ires = 
       
  1056              (inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires \<or>
       
  1057              inv_check_left_moving_in_middle (as, lm) (s, l, r) ires)"
       
  1058 
       
  1059 fun inv_after_left_moving :: "inc_inv_t"
       
  1060   where "inv_after_left_moving (as, lm) (s, l, r) ires= 
       
  1061               (\<exists> rn. l = Bk # ires \<and> r = Bk # <lm> @ (Bk\<^bsup>rn\<^esup>))"
       
  1062 
       
  1063 fun inv_stop :: "inc_inv_t"
       
  1064   where "inv_stop (as, lm) (s, l, r) ires= 
       
  1065               (\<exists> rn. l = Bk # Bk # ires \<and> r = <lm> @ (Bk\<^bsup>rn\<^esup>))"
       
  1066 
       
  1067 
       
  1068 fun inc_inv :: "layout \<Rightarrow> nat \<Rightarrow> inc_inv_t"
       
  1069   where
       
  1070   "inc_inv ly n (as, lm) (s, l, r) ires =
       
  1071               (let ss = start_of ly as in
       
  1072                let lm' = abc_lm_s lm n ((abc_lm_v lm n)+1) in
       
  1073                 if s = 0 then False
       
  1074                 else if s < ss then False
       
  1075                 else if s < ss + 2 * n then 
       
  1076                    if (s - ss) mod 2 = 0 then 
       
  1077                        inv_locate_a (as, lm) ((s - ss) div 2, l, r) ires
       
  1078                    else inv_locate_b (as, lm) ((s - ss) div 2, l, r) ires
       
  1079                 else if s = ss + 2 * n then 
       
  1080                         inv_locate_a (as, lm) (n, l, r) ires
       
  1081                 else if s = ss + 2 * n + 1 then 
       
  1082                    inv_locate_b (as, lm) (n, l, r) ires
       
  1083                 else if s = ss + 2 * n + 2 then 
       
  1084                    inv_after_write (as, lm') (s - ss, l, r) ires
       
  1085                 else if s = ss + 2 * n + 3 then 
       
  1086                    inv_after_move (as, lm') (s - ss, l, r) ires
       
  1087                 else if s = ss + 2 * n + 4 then 
       
  1088                    inv_after_clear (as, lm') (s - ss, l, r) ires
       
  1089                 else if s = ss + 2 * n + 5 then 
       
  1090                    inv_on_right_moving (as, lm') (s - ss, l, r) ires
       
  1091                 else if s = ss + 2 * n + 6 then 
       
  1092                    inv_on_left_moving (as, lm') (s - ss, l, r) ires
       
  1093                 else if s = ss + 2 * n + 7 then 
       
  1094                    inv_check_left_moving (as, lm') (s - ss, l, r) ires
       
  1095                 else if s = ss + 2 * n + 8 then 
       
  1096                    inv_after_left_moving (as, lm') (s - ss, l, r) ires
       
  1097                 else if s = ss + 2 * n + 9 then 
       
  1098                    inv_stop (as, lm') (s - ss, l, r) ires
       
  1099                 else False) "
       
  1100 
       
  1101 lemma fetch_intro: 
       
  1102   "\<lbrakk>\<And>xs.\<lbrakk>ba = Oc # xs\<rbrakk> \<Longrightarrow> P (fetch prog i Oc);
       
  1103    \<And>xs.\<lbrakk>ba = Bk # xs\<rbrakk> \<Longrightarrow> P (fetch prog i Bk);
       
  1104    ba = [] \<Longrightarrow> P (fetch prog i Bk)
       
  1105    \<rbrakk> \<Longrightarrow> P (fetch prog i 
       
  1106              (case ba of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))"
       
  1107 by (auto split:list.splits block.splits)
       
  1108 
       
  1109 lemma length_findnth[simp]: "length (findnth n) = 4 * n"
       
  1110 apply(induct n, simp)
       
  1111 apply(simp)
       
  1112 done
       
  1113 
       
  1114 declare tshift.simps[simp del]
       
  1115 declare findnth.simps[simp del]
       
  1116 
       
  1117 lemma findnth_nth: 
       
  1118  "\<lbrakk>n > q; x < 4\<rbrakk> \<Longrightarrow> 
       
  1119         (findnth n) ! (4 * q + x) = (findnth (Suc q) ! (4 * q + x))"
       
  1120 apply(induct n, simp)
       
  1121 apply(case_tac "q < n", simp add: findnth.simps, auto)
       
  1122 apply(simp add: nth_append)
       
  1123 apply(subgoal_tac "q = n", simp)
       
  1124 apply(arith)
       
  1125 done
       
  1126 
       
  1127 lemma Suc_pre[simp]: "\<not> a < start_of ly as \<Longrightarrow> 
       
  1128           (Suc a - start_of ly as) = Suc (a - start_of ly as)"
       
  1129 apply(arith)
       
  1130 done
       
  1131 
       
  1132 lemma fetch_locate_a_o: "
       
  1133 \<And>a  q xs.
       
  1134     \<lbrakk>\<not> a < start_of (layout_of aprog) as; 
       
  1135       a < start_of (layout_of aprog) as + 2 * n; 
       
  1136       a - start_of (layout_of aprog) as = 2 * q; 
       
  1137       start_of (layout_of aprog) as > 0\<rbrakk>
       
  1138     \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as)
       
  1139          (Inc n)) (Suc (2 * q)) Oc) = (R, a+1)"
       
  1140 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  1141                   nth_of.simps tshift.simps nth_append Suc_pre)
       
  1142 apply(subgoal_tac "(findnth n ! Suc (4 * q)) = 
       
  1143                  findnth (Suc q) ! (4 * q + 1)")
       
  1144 apply(simp add: findnth.simps nth_append)
       
  1145 apply(subgoal_tac " findnth n !(4 * q + 1) = 
       
  1146                  findnth (Suc q) ! (4 * q + 1)", simp)
       
  1147 apply(rule_tac findnth_nth, auto)
       
  1148 done
       
  1149 
       
  1150 lemma fetch_locate_a_b: "
       
  1151 \<And>a  q xs.
       
  1152     \<lbrakk>abc_fetch as aprog = Some (Inc n);  
       
  1153      \<not> a < start_of (layout_of aprog) as; 
       
  1154      a < start_of (layout_of aprog) as + 2 * n; 
       
  1155      a - start_of (layout_of aprog) as = 2 * q; 
       
  1156      start_of (layout_of aprog) as > 0\<rbrakk>
       
  1157     \<Longrightarrow> (fetch (ci (layout_of aprog) 
       
  1158       (start_of (layout_of aprog) as) (Inc n)) (Suc (2 * q)) Bk)
       
  1159        = (W1, a)"
       
  1160 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  1161                  tshift.simps nth_append)
       
  1162 apply(subgoal_tac "(findnth n ! (4 * q)) = 
       
  1163                            findnth (Suc q) ! (4 * q )")
       
  1164 apply(simp add: findnth.simps nth_append)
       
  1165 apply(subgoal_tac " findnth n !(4 * q + 0) =
       
  1166                             findnth (Suc q) ! (4 * q + 0)", simp)
       
  1167 apply(rule_tac findnth_nth, auto)
       
  1168 done
       
  1169 
       
  1170 lemma [intro]: "x mod 2 = Suc 0 \<Longrightarrow> \<exists> q. x = Suc (2 * q)"
       
  1171 apply(drule mod_eqD, auto)
       
  1172 done
       
  1173 
       
  1174 lemma  add3_Suc: "x + 3 = Suc (Suc (Suc x))"
       
  1175 apply(arith)
       
  1176 done
       
  1177 
       
  1178 declare start_of.simps[simp]
       
  1179 (*
       
  1180 lemma layout_not0: "start_of ly as > 0"
       
  1181 by(induct as, auto)
       
  1182 *)
       
  1183 lemma [simp]: 
       
  1184  "\<lbrakk>\<not> a < start_of (layout_of aprog) as; 
       
  1185    a - start_of (layout_of aprog) as = Suc (2 * q); 
       
  1186    abc_fetch as aprog = Some (Inc n); 
       
  1187    start_of (layout_of aprog) as > 0\<rbrakk>
       
  1188     \<Longrightarrow> Suc (Suc (2 * q + start_of (layout_of aprog) as - Suc 0)) = a"
       
  1189 apply(subgoal_tac 
       
  1190 "Suc (Suc (2 * q + start_of (layout_of aprog) as - Suc 0)) 
       
  1191               = 2 + 2 * q + start_of (layout_of aprog) as - Suc 0", 
       
  1192   simp, simp add: inc_startof_not0)
       
  1193 done
       
  1194 
       
  1195 lemma fetch_locate_b_o: "
       
  1196 \<And>a  xs.
       
  1197     \<lbrakk>0 < a; \<not> a < start_of (layout_of aprog) as; 
       
  1198   a < start_of (layout_of aprog) as + 2 * n; 
       
  1199  (a - start_of (layout_of aprog) as) mod 2 = Suc 0; 
       
  1200  start_of (layout_of aprog) as > 0\<rbrakk>
       
  1201     \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  1202       (Inc n)) (Suc (a - start_of (layout_of aprog) as)) Oc) = (R, a)"
       
  1203 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  1204                  nth_of.simps tshift.simps nth_append)
       
  1205 apply(subgoal_tac "\<exists> q. (a - start_of (layout_of aprog) as) = 
       
  1206                          2 * q + 1", auto)
       
  1207 apply(subgoal_tac "(findnth n ! Suc (Suc (Suc (4 * q)))) 
       
  1208                   = findnth (Suc q) ! (4 * q + 3)")
       
  1209 apply(simp add: findnth.simps nth_append)
       
  1210 apply(subgoal_tac " findnth n ! (4 * q + 3) = 
       
  1211                  findnth (Suc q) ! (4 * q + 3)", simp add: add3_Suc)
       
  1212 apply(rule_tac findnth_nth, auto)
       
  1213 done
       
  1214 
       
  1215 lemma fetch_locate_b_b: "
       
  1216 \<And>a  xs.
       
  1217     \<lbrakk>0 < a;  \<not> a < start_of (layout_of aprog) as; 
       
  1218      a < start_of (layout_of aprog) as + 2 * n; 
       
  1219      (a - start_of (layout_of aprog) as) mod 2 = Suc 0; 
       
  1220      start_of (layout_of aprog) as > 0\<rbrakk>
       
  1221     \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as)
       
  1222         (Inc n)) (Suc (a - start_of (layout_of aprog) as)) Bk) 
       
  1223         = (R, a + 1)"
       
  1224 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  1225                   nth_of.simps tshift.simps nth_append)
       
  1226 apply(subgoal_tac "\<exists> q. (a - start_of (layout_of aprog) as) = 
       
  1227                   2 * q + 1", auto)
       
  1228 apply(subgoal_tac "(findnth n ! Suc ((Suc (4 * q)))) = 
       
  1229                     findnth (Suc q) ! (4 * q + 2)")
       
  1230 apply(simp add: findnth.simps nth_append)
       
  1231 apply(subgoal_tac " findnth n ! (4 * q + 2) = 
       
  1232                     findnth (Suc q) ! (4 * q + 2)", simp)
       
  1233 apply(rule_tac findnth_nth, auto)
       
  1234 done
       
  1235 
       
  1236 lemma fetch_locate_n_a_o: 
       
  1237        "start_of (layout_of aprog) as > 0
       
  1238        \<Longrightarrow> (fetch (ci (layout_of aprog) 
       
  1239       (start_of (layout_of aprog) as) (Inc n)) (Suc (2 * n)) Oc) = 
       
  1240              (R, start_of (layout_of aprog) as + 2 * n + 1)"
       
  1241 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  1242                   nth_of.simps tshift.simps nth_append tinc_b_def)
       
  1243 done
       
  1244 
       
  1245 lemma fetch_locate_n_a_b: "
       
  1246        start_of (layout_of aprog) as > 0
       
  1247       \<Longrightarrow> (fetch (ci (layout_of aprog) 
       
  1248     (start_of (layout_of aprog) as) (Inc n)) (Suc (2 * n)) Bk) 
       
  1249    = (W1, start_of (layout_of aprog) as + 2 * n)"
       
  1250 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  1251                   nth_of.simps tshift.simps nth_append tinc_b_def)
       
  1252 done
       
  1253 
       
  1254 lemma fetch_locate_n_b_o: "
       
  1255     start_of (layout_of aprog) as > 0
       
  1256     \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  1257      (Inc n)) (Suc (Suc (2 * n))) Oc) = 
       
  1258                       (R, start_of (layout_of aprog) as + 2 * n + 1)"
       
  1259 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  1260                   nth_of.simps tshift.simps nth_append tinc_b_def)
       
  1261 done
       
  1262 
       
  1263 lemma fetch_locate_n_b_b: "
       
  1264     start_of (layout_of aprog) as > 0
       
  1265    \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  1266    (Inc n)) (Suc (Suc (2 * n))) Bk) = 
       
  1267        (W1, start_of (layout_of aprog) as + 2 * n + 2)"
       
  1268 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  1269                   nth_of.simps tshift.simps nth_append tinc_b_def)
       
  1270 done
       
  1271 
       
  1272 lemma fetch_after_write_o: "
       
  1273     start_of (layout_of aprog) as > 0
       
  1274     \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  1275             (Inc n)) (Suc (Suc (Suc (2 * n)))) Oc) = 
       
  1276         (R, start_of (layout_of aprog) as + 2*n + 3)"
       
  1277 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  1278                   nth_of.simps tshift.simps nth_append tinc_b_def)
       
  1279 done
       
  1280 
       
  1281 lemma fetch_after_move_o: "
       
  1282       start_of (layout_of aprog) as > 0
       
  1283       \<Longrightarrow> (fetch (ci (layout_of aprog) 
       
  1284               (start_of (layout_of aprog) as) (Inc n)) (4 + 2 * n) Oc)
       
  1285         = (W0, start_of (layout_of aprog) as + 2 * n + 4)"
       
  1286 apply(auto simp: ci.simps findnth.simps tshift.simps 
       
  1287                  tinc_b_def add3_Suc)
       
  1288 apply(subgoal_tac "4 + 2*n = Suc (2*n + 3)", simp only: fetch.simps)
       
  1289 apply(auto simp: nth_of.simps nth_append)
       
  1290 done
       
  1291 
       
  1292 lemma fetch_after_move_b: "
       
  1293       start_of (layout_of aprog) as > 0
       
  1294       \<Longrightarrow>(fetch (ci (layout_of aprog) 
       
  1295             (start_of (layout_of aprog) as) (Inc n)) (4 + 2 * n) Bk)
       
  1296        = (L, start_of (layout_of aprog) as + 2 * n + 6)"
       
  1297 apply(auto simp: ci.simps findnth.simps tshift.simps 
       
  1298                  tinc_b_def add3_Suc)
       
  1299 apply(subgoal_tac "4 + 2*n = Suc (2*n + 3)", simp only: fetch.simps)
       
  1300 apply(auto simp: nth_of.simps nth_append)
       
  1301 done
       
  1302 
       
  1303 lemma fetch_clear_b: "
       
  1304      start_of (layout_of aprog) as > 0
       
  1305       \<Longrightarrow> (fetch (ci (layout_of aprog) 
       
  1306               (start_of (layout_of aprog) as) (Inc n)) (5 + 2 * n) Bk)
       
  1307       = (R, start_of (layout_of aprog) as + 2 * n + 5)"
       
  1308 apply(auto simp: ci.simps findnth.simps 
       
  1309                      tshift.simps tinc_b_def add3_Suc)
       
  1310 apply(subgoal_tac "5 + 2*n = Suc (2*n + 4)", simp only: fetch.simps)
       
  1311 apply(auto simp: nth_of.simps nth_append)
       
  1312 done
       
  1313 
       
  1314 lemma fetch_right_move_o: "
       
  1315       start_of (layout_of aprog) as > 0
       
  1316       \<Longrightarrow> (fetch (ci (layout_of aprog) 
       
  1317                 (start_of (layout_of aprog) as) (Inc n)) (6 + 2*n) Oc)
       
  1318       = (R, start_of (layout_of aprog) as + 2 * n + 5)"
       
  1319 apply(auto simp: ci.simps findnth.simps tshift.simps 
       
  1320                  tinc_b_def add3_Suc)
       
  1321 apply(subgoal_tac "6 + 2*n = Suc (2*n + 5)", simp only: fetch.simps)
       
  1322 apply(auto simp: nth_of.simps nth_append)
       
  1323 done
       
  1324 
       
  1325 lemma fetch_right_move_b: "
       
  1326       start_of (layout_of aprog) as > 0
       
  1327       \<Longrightarrow> (fetch (ci (layout_of aprog) 
       
  1328                 (start_of (layout_of aprog) as) (Inc n)) (6 + 2*n) Bk)
       
  1329       = (W1, start_of (layout_of aprog) as + 2 * n + 2)"
       
  1330 apply(auto simp: ci.simps findnth.simps 
       
  1331                  tshift.simps tinc_b_def add3_Suc)
       
  1332 apply(subgoal_tac "6 + 2*n = Suc (2*n + 5)", simp only: fetch.simps)
       
  1333 apply(auto simp: nth_of.simps nth_append)
       
  1334 done
       
  1335 
       
  1336 lemma fetch_left_move_o: "
       
  1337       start_of (layout_of aprog) as > 0
       
  1338       \<Longrightarrow> (fetch (ci (layout_of aprog) 
       
  1339                (start_of (layout_of aprog) as) (Inc n)) (7 + 2*n) Oc)
       
  1340       = (L, start_of (layout_of aprog) as + 2 * n + 6)"
       
  1341 apply(auto simp: ci.simps findnth.simps tshift.simps 
       
  1342                  tinc_b_def add3_Suc)
       
  1343 apply(subgoal_tac "7 + 2*n = Suc (2*n + 6)", simp only: fetch.simps)
       
  1344 apply(auto simp: nth_of.simps nth_append)
       
  1345 done
       
  1346 
       
  1347 lemma fetch_left_move_b: "
       
  1348       start_of (layout_of aprog) as > 0
       
  1349       \<Longrightarrow> (fetch (ci (layout_of aprog) 
       
  1350                (start_of (layout_of aprog) as) (Inc n)) (7 + 2*n) Bk)
       
  1351       = (L, start_of (layout_of aprog) as + 2 * n + 7)"
       
  1352 apply(auto simp: ci.simps findnth.simps 
       
  1353                  tshift.simps tinc_b_def add3_Suc)
       
  1354 apply(subgoal_tac "7 + 2*n = Suc (2*n + 6)", simp only: fetch.simps)
       
  1355 apply(auto simp: nth_of.simps nth_append)
       
  1356 done
       
  1357 
       
  1358 lemma fetch_check_left_move_o: "
       
  1359       start_of (layout_of aprog) as > 0
       
  1360       \<Longrightarrow> (fetch (ci (layout_of aprog) 
       
  1361                (start_of (layout_of aprog) as) (Inc n)) (8 + 2*n) Oc)
       
  1362       = (L, start_of (layout_of aprog) as + 2 * n + 6)"
       
  1363 apply(auto simp: ci.simps findnth.simps tshift.simps tinc_b_def)
       
  1364 apply(subgoal_tac "8 + 2 * n = Suc (2 * n + 7)", 
       
  1365                                   simp only: fetch.simps)
       
  1366 apply(auto simp: nth_of.simps nth_append)
       
  1367 done
       
  1368 
       
  1369 lemma fetch_check_left_move_b: "
       
  1370       start_of (layout_of aprog) as > 0
       
  1371       \<Longrightarrow> (fetch (ci (layout_of aprog) 
       
  1372               (start_of (layout_of aprog) as) (Inc n)) (8 + 2*n) Bk)
       
  1373       = (R, start_of (layout_of aprog) as + 2 * n + 8)  "
       
  1374 apply(auto simp: ci.simps findnth.simps 
       
  1375                  tshift.simps tinc_b_def add3_Suc)
       
  1376 apply(subgoal_tac "8 + 2*n= Suc (2*n + 7)", simp only: fetch.simps)
       
  1377 apply(auto simp: nth_of.simps nth_append)
       
  1378 done
       
  1379 
       
  1380 lemma fetch_after_left_move: "
       
  1381       start_of (layout_of aprog) as > 0
       
  1382       \<Longrightarrow> (fetch (ci (layout_of aprog) 
       
  1383               (start_of (layout_of aprog) as) (Inc n)) (9 + 2*n) Bk)
       
  1384      = (R, start_of (layout_of aprog) as + 2 * n + 9)"
       
  1385 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  1386                   nth_of.simps tshift.simps nth_append tinc_b_def)
       
  1387 done
       
  1388 
       
  1389 lemma fetch_stop: "
       
  1390        start_of (layout_of aprog) as > 0
       
  1391       \<Longrightarrow> (fetch (ci (layout_of aprog) 
       
  1392              (start_of (layout_of aprog) as) (Inc n)) (10 + 2 *n)  b)
       
  1393      = (Nop, 0)"
       
  1394 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  1395                   nth_of.simps tshift.simps nth_append tinc_b_def
       
  1396             split: block.splits)
       
  1397 done
       
  1398 
       
  1399 lemma fetch_state0: "
       
  1400        (fetch (ci (layout_of aprog) 
       
  1401                (start_of (layout_of aprog) as) (Inc n)) 0 b)
       
  1402      = (Nop, 0)"
       
  1403 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  1404                   nth_of.simps tshift.simps nth_append tinc_b_def)
       
  1405 done
       
  1406 
       
  1407 lemmas fetch_simps = 
       
  1408   fetch_locate_a_o fetch_locate_a_b fetch_locate_b_o fetch_locate_b_b 
       
  1409   fetch_locate_n_a_b fetch_locate_n_a_o fetch_locate_n_b_o 
       
  1410   fetch_locate_n_b_b fetch_after_write_o fetch_after_move_o 
       
  1411   fetch_after_move_b fetch_clear_b fetch_right_move_o 
       
  1412   fetch_right_move_b fetch_left_move_o fetch_left_move_b
       
  1413   fetch_after_left_move fetch_check_left_move_o fetch_stop 
       
  1414   fetch_state0 fetch_check_left_move_b
       
  1415 
       
  1416 text {* *}
       
  1417 declare exponent_def[simp del] tape_of_nat_list.simps[simp del] 
       
  1418    at_begin_norm.simps[simp del] at_begin_fst_bwtn.simps[simp del] 
       
  1419    at_begin_fst_awtn.simps[simp del] in_middle.simps[simp del] 
       
  1420    abc_lm_s.simps[simp del] abc_lm_v.simps[simp del]  
       
  1421    ci.simps[simp del] t_step.simps[simp del]
       
  1422    inv_after_move.simps[simp del] 
       
  1423    inv_on_left_moving_norm.simps[simp del] 
       
  1424    inv_on_left_moving_in_middle_B.simps[simp del]
       
  1425    inv_after_clear.simps[simp del] 
       
  1426    inv_after_write.simps[simp del] inv_on_left_moving.simps[simp del]
       
  1427    inv_on_right_moving.simps[simp del] 
       
  1428    inv_check_left_moving.simps[simp del] 
       
  1429    inv_check_left_moving_in_middle.simps[simp del]
       
  1430    inv_check_left_moving_on_leftmost.simps[simp del] 
       
  1431    inv_after_left_moving.simps[simp del]
       
  1432    inv_stop.simps[simp del] inv_locate_a.simps[simp del] 
       
  1433    inv_locate_b.simps[simp del]
       
  1434 declare tms_of.simps[simp del] tm_of.simps[simp del]
       
  1435         layout_of.simps[simp del] abc_fetch.simps [simp del] 
       
  1436         t_step.simps[simp del] t_steps.simps[simp del] 
       
  1437         tpairs_of.simps[simp del] start_of.simps[simp del]
       
  1438         fetch.simps [simp del] new_tape.simps [simp del] 
       
  1439         nth_of.simps [simp del] ci.simps [simp del]
       
  1440         length_of.simps[simp del]
       
  1441 
       
  1442 (*! Start point *)
       
  1443 lemma [simp]: "Suc (2 * q) mod 2 = Suc 0"
       
  1444 by arith
       
  1445 
       
  1446 lemma [simp]: "Suc (2 * q) div 2 = q"
       
  1447 by arith
       
  1448 
       
  1449 lemma [simp]: "\<lbrakk> \<not> a < start_of ly as; 
       
  1450           a < start_of ly as + 2 * n; a - start_of ly as = 2 * q\<rbrakk>
       
  1451              \<Longrightarrow> Suc a < start_of ly as + 2 * n"
       
  1452 apply(arith)
       
  1453 done
       
  1454 
       
  1455 lemma [simp]: "x mod 2 = Suc 0 \<Longrightarrow> (Suc x) mod 2 = 0"
       
  1456 by arith
       
  1457 
       
  1458 lemma [simp]: "x mod 2 = Suc 0 \<Longrightarrow> (Suc x) div 2 = Suc (x div 2)"
       
  1459 by arith
       
  1460 lemma exp_def[simp]: "a\<^bsup>Suc n \<^esup>= a # a\<^bsup>n\<^esup>"
       
  1461 by(simp add: exponent_def)
       
  1462 lemma [intro]: "Bk # r = Oc\<^bsup>mr\<^esup> @ r' \<Longrightarrow> mr = 0"
       
  1463 by(case_tac mr, auto simp: exponent_def)
       
  1464 
       
  1465 lemma [intro]: "Bk # r = replicate mr Oc \<Longrightarrow> mr = 0"
       
  1466 by(case_tac mr, auto)
       
  1467 lemma tape_of_nl_abv_cons[simp]: "xs \<noteq> [] \<Longrightarrow> 
       
  1468                    <x # xs> = Oc\<^bsup>Suc x\<^esup>@ Bk # <xs>"
       
  1469 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
       
  1470 apply(case_tac xs, simp, simp add: tape_of_nat_list.simps)
       
  1471 done
       
  1472 
       
  1473 lemma [simp]: "<[]::nat list> = []"
       
  1474 by(auto simp: tape_of_nl_abv tape_of_nat_list.simps)
       
  1475 lemma [simp]: "Oc # r = <(lm::nat list)> @ Bk\<^bsup>rn\<^esup>\<Longrightarrow> lm \<noteq> []"
       
  1476 apply(auto simp: tape_of_nl_abv tape_of_nat_list.simps)
       
  1477 apply(case_tac rn, auto simp: exponent_def)
       
  1478 done
       
  1479 lemma BkCons_nil: "Bk # xs = <lm::nat list> @ Bk\<^bsup>rn\<^esup>\<Longrightarrow> lm = []"
       
  1480 apply(case_tac lm, simp)
       
  1481 apply(case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps)
       
  1482 done
       
  1483 lemma BkCons_nil': "Bk # xs = <lm::nat list> @ Bk\<^bsup>ln\<^esup>\<Longrightarrow> lm = []"
       
  1484 by(auto intro: BkCons_nil)
       
  1485 
       
  1486 lemma hd_tl_tape_of_nat_list:  
       
  1487    "tl (lm::nat list) \<noteq> [] \<Longrightarrow> <lm> = <hd lm> @ Bk # <tl lm>"
       
  1488 apply(frule tape_of_nl_abv_cons[of "tl lm" "hd lm"])
       
  1489 apply(simp add: tape_of_nat_abv Bk_def del: tape_of_nl_abv_cons)
       
  1490 apply(subgoal_tac "lm = hd lm # tl lm", auto)
       
  1491 apply(case_tac lm, auto)
       
  1492 done
       
  1493 lemma [simp]: "Oc # xs = Oc\<^bsup>mr\<^esup> @ Bk # <lm2> @ Bk\<^bsup>rn\<^esup>\<Longrightarrow> mr > 0"
       
  1494 apply(case_tac mr, auto simp: exponent_def)
       
  1495 done
       
  1496 
       
  1497 lemma tape_of_nat_list_cons: "xs \<noteq> [] \<Longrightarrow> tape_of_nat_list (x # xs) =
       
  1498               replicate (Suc x) Oc @ Bk # tape_of_nat_list xs"
       
  1499 apply(drule tape_of_nl_abv_cons[of xs x])
       
  1500 apply(auto simp: tape_of_nl_abv tape_of_nat_abv Oc_def Bk_def exponent_def)
       
  1501 done
       
  1502 
       
  1503 lemma rev_eq: "rev xs = rev ys \<Longrightarrow> xs = ys"
       
  1504 by simp
       
  1505 
       
  1506 lemma tape_of_nat_list_eq: " xs = ys \<Longrightarrow> 
       
  1507         tape_of_nat_list xs = tape_of_nat_list ys"
       
  1508 by simp
       
  1509 
       
  1510 lemma tape_of_nl_nil_eq: "<(lm::nat list)> = [] = (lm = [])"
       
  1511 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
       
  1512 apply(case_tac lm, simp add: tape_of_nat_list.simps)
       
  1513 apply(case_tac "list")
       
  1514 apply(auto simp: tape_of_nat_list.simps)
       
  1515 done
       
  1516 
       
  1517 lemma rep_ind: "replicate (Suc n) a = replicate n a @ [a]"
       
  1518 apply(induct n, simp, simp)
       
  1519 done
       
  1520 
       
  1521 lemma [simp]: "Oc # r = <lm::nat list> @ replicate rn Bk \<Longrightarrow> Suc 0 \<le> length lm"
       
  1522 apply(rule_tac classical, auto)
       
  1523 apply(case_tac lm, simp, case_tac rn, auto)
       
  1524 done
       
  1525 lemma Oc_Bk_Cons: "Oc # Bk # list = <lm::nat list> @ Bk\<^bsup>ln\<^esup> \<Longrightarrow> 
       
  1526                    lm \<noteq> [] \<and> hd lm = 0"
       
  1527 apply(case_tac lm, simp, case_tac ln, simp add: exponent_def, simp add: exponent_def, simp)
       
  1528 apply(case_tac lista, auto simp: tape_of_nl_abv tape_of_nat_list.simps)
       
  1529 done 
       
  1530 (*lemma Oc_Oc_Cons: "Oc # Oc # list = <lm::nat list> @ Bk\<^bsup>ln\<^esup> \<Longrightarrow> 
       
  1531                   lm \<noteq> [] \<and> hd lm > 0"
       
  1532 apply(case_tac lm, simp add: exponent_def, case_tac ln, simp, simp)
       
  1533 apply(case_tac lista, 
       
  1534         auto simp: tape_of_nl_abv tape_of_nat_list.simps exponent_def)
       
  1535 apply(case_tac [!] a, auto)
       
  1536 apply(case_tac ln, auto)
       
  1537 done
       
  1538 *)
       
  1539 lemma Oc_nil_zero[simp]: "[Oc] = <lm::nat list> @ Bk\<^bsup>ln\<^esup> 
       
  1540                  \<Longrightarrow> lm = [0] \<and> ln = 0"
       
  1541 apply(case_tac lm, simp)
       
  1542 apply(case_tac ln, auto simp: exponent_def)
       
  1543 apply(case_tac [!] list, 
       
  1544         auto simp: tape_of_nl_abv tape_of_nat_list.simps)
       
  1545 done
       
  1546 
       
  1547 lemma  [simp]: "Oc # r = <lm2> @ replicate rn Bk \<Longrightarrow> 
       
  1548        (\<exists>rn. r = replicate (hd lm2) Oc @ Bk # <tl lm2> @ 
       
  1549                       replicate rn Bk) \<or> 
       
  1550           tl lm2 = [] \<and> r = replicate (hd lm2) Oc"
       
  1551 apply(rule_tac disjCI, simp)
       
  1552 apply(case_tac "tl lm2 = []", simp)
       
  1553 apply(case_tac lm2, simp add: tape_of_nl_abv tape_of_nat_list.simps)
       
  1554 apply(case_tac rn, simp, simp, simp)
       
  1555 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps exponent_def)
       
  1556 apply(case_tac rn, simp, simp)
       
  1557 apply(rule_tac x = rn in exI)
       
  1558 apply(simp add: hd_tl_tape_of_nat_list)
       
  1559 apply(simp add: tape_of_nat_abv Oc_def exponent_def)
       
  1560 done
       
  1561 
       
  1562 (*inv: from locate_a to locate_b*)
       
  1563 lemma [simp]: 
       
  1564       "inv_locate_a (as, lm) (q, l, Oc # r) ires
       
  1565        \<Longrightarrow> inv_locate_b (as, lm) (q, Oc # l, r) ires"
       
  1566 apply(simp only: inv_locate_a.simps inv_locate_b.simps in_middle.simps
       
  1567           at_begin_norm.simps at_begin_fst_bwtn.simps
       
  1568           at_begin_fst_awtn.simps)
       
  1569 apply(erule disjE, erule exE, erule exE, erule exE)
       
  1570 apply(rule_tac x = lm1 in exI, rule_tac x = "tl lm2" in exI, simp)
       
  1571 apply(rule_tac x = "0" in exI, rule_tac x = "hd lm2" in exI, 
       
  1572                 auto simp: exponent_def)
       
  1573 apply(rule_tac x = "Suc 0" in exI, simp add:exponent_def)
       
  1574 apply(rule_tac x = "lm @ replicate tn 0" in exI, 
       
  1575       rule_tac x = "[]" in exI,    
       
  1576       rule_tac x = "Suc tn" in exI, rule_tac x = 0 in exI)
       
  1577 apply(simp only: rep_ind, simp)
       
  1578 apply(rule_tac x = "Suc 0" in exI, auto)
       
  1579 apply(case_tac [1-3] rn, simp_all )
       
  1580 apply(rule_tac x = "lm @ replicate tn 0" in exI, 
       
  1581       rule_tac x = "[]" in exI, 
       
  1582       rule_tac x = "Suc tn" in exI, 
       
  1583       rule_tac x = 0 in exI, simp add: rep_ind del: replicate_Suc split:if_splits)
       
  1584 apply(rule_tac x = "Suc 0" in exI, auto)
       
  1585 apply(case_tac rn, simp, simp)
       
  1586 apply(rule_tac [!] x = "Suc 0" in exI, auto)
       
  1587 apply(case_tac [!] rn, simp_all)
       
  1588 done
       
  1589 
       
  1590 (*inv: from locate_a to _locate_a*)
       
  1591 lemma locate_a_2_locate_a[simp]: "inv_locate_a (as, am) (q, aaa, Bk # xs) ires
       
  1592        \<Longrightarrow> inv_locate_a (as, am) (q, aaa, Oc # xs) ires"
       
  1593 apply(simp only: inv_locate_a.simps at_begin_norm.simps 
       
  1594                  at_begin_fst_bwtn.simps at_begin_fst_awtn.simps)
       
  1595 apply(erule_tac disjE, erule exE, erule exE, erule exE, 
       
  1596       rule disjI2, rule disjI2)
       
  1597 defer
       
  1598 apply(erule_tac disjE, erule exE, erule exE, 
       
  1599       erule exE, rule disjI2, rule disjI2)
       
  1600 prefer 2
       
  1601 apply(simp)
       
  1602 proof-
       
  1603   fix lm1 tn rn
       
  1604   assume k: "lm1 = am @ 0\<^bsup>tn\<^esup> \<and> length lm1 = q \<and> (if lm1 = [] then aaa = Bk # Bk # 
       
  1605     ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> Bk # xs = Bk\<^bsup>rn\<^esup>"
       
  1606   thus "\<exists>lm1 tn rn. lm1 = am @ 0\<^bsup>tn\<^esup> \<and> length lm1 = q \<and> (if lm1 = [] then 
       
  1607     aaa = Bk # Bk # ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> Oc # xs = [Oc] @ Bk\<^bsup>rn\<^esup>"
       
  1608     (is "\<exists>lm1 tn rn. ?P lm1 tn rn")
       
  1609   proof -
       
  1610     from k have "?P lm1 tn (rn - 1)"
       
  1611       apply(auto simp: Oc_def)
       
  1612       by(case_tac [!] "rn::nat", auto simp: exponent_def)
       
  1613     thus ?thesis by blast
       
  1614   qed
       
  1615 next
       
  1616   fix lm1 lm2 rn
       
  1617   assume h1: "am = lm1 @ lm2 \<and> length lm1 = q \<and> (if lm1 = [] 
       
  1618     then aaa = Bk # Bk # ires else aaa = Bk # <rev lm1> @ Bk # Bk # ires) \<and>
       
  1619     Bk # xs = <lm2> @ Bk\<^bsup>rn\<^esup>"
       
  1620   from h1 have h2: "lm2 = []"
       
  1621   proof(rule_tac xs = xs and rn = rn in BkCons_nil, simp)
       
  1622   qed
       
  1623   from h1 and h2 show "\<exists>lm1 tn rn. lm1 = am @ 0\<^bsup>tn\<^esup> \<and> length lm1 = q \<and> 
       
  1624     (if lm1 = [] then aaa = Bk # Bk # ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and>
       
  1625     Oc # xs = [Oc] @ Bk\<^bsup>rn\<^esup>" 
       
  1626     (is "\<exists>lm1 tn rn. ?P lm1 tn rn")
       
  1627   proof -
       
  1628     from h1 and h2  have "?P lm1 0 (rn - 1)"
       
  1629       apply(auto simp: Oc_def exponent_def 
       
  1630                       tape_of_nl_abv tape_of_nat_list.simps)
       
  1631       by(case_tac "rn::nat", simp, simp)
       
  1632     thus ?thesis by blast
       
  1633   qed
       
  1634 qed
       
  1635 
       
  1636 lemma [intro]: "\<exists>rn. [a] = a\<^bsup>rn\<^esup>"
       
  1637 by(rule_tac x = "Suc 0" in exI, simp add: exponent_def)
       
  1638 
       
  1639 lemma [intro]: "\<exists>tn. [] = a\<^bsup>tn\<^esup>"
       
  1640 apply(rule_tac x = 0 in exI, simp add: exponent_def)
       
  1641 done
       
  1642 
       
  1643 lemma [intro]:  "at_begin_norm (as, am) (q, aaa, []) ires
       
  1644              \<Longrightarrow> at_begin_norm (as, am) (q, aaa, [Bk]) ires"
       
  1645 apply(simp add: at_begin_norm.simps, erule_tac exE, erule_tac exE)
       
  1646 apply(rule_tac x = lm1 in exI, simp, auto)
       
  1647 done
       
  1648 
       
  1649 lemma [intro]: "at_begin_fst_bwtn (as, am) (q, aaa, []) ires 
       
  1650             \<Longrightarrow> at_begin_fst_bwtn (as, am) (q, aaa, [Bk]) ires"
       
  1651 apply(simp only: at_begin_fst_bwtn.simps, erule_tac exE, erule_tac exE, erule_tac exE)
       
  1652 apply(rule_tac x = "am @ 0\<^bsup>tn\<^esup>" in exI, auto)
       
  1653 done
       
  1654 
       
  1655 lemma [intro]: "at_begin_fst_awtn (as, am) (q, aaa, []) ires
       
  1656            \<Longrightarrow> at_begin_fst_awtn (as, am) (q, aaa, [Bk]) ires"
       
  1657 apply(auto simp: at_begin_fst_awtn.simps)
       
  1658 done 
       
  1659 
       
  1660 lemma [intro]: "inv_locate_a (as, am) (q, aaa, []) ires
       
  1661             \<Longrightarrow> inv_locate_a (as, am) (q, aaa, [Bk]) ires"
       
  1662 apply(simp only: inv_locate_a.simps)
       
  1663 apply(erule disj_forward)
       
  1664 defer
       
  1665 apply(erule disj_forward, auto)
       
  1666 done
       
  1667 
       
  1668 lemma [simp]: "inv_locate_a (as, am) (q, aaa, []) ires \<Longrightarrow> 
       
  1669                inv_locate_a (as, am) (q, aaa, [Oc]) ires"
       
  1670 apply(insert locate_a_2_locate_a [of as am q aaa "[]"])
       
  1671 apply(subgoal_tac "inv_locate_a (as, am) (q, aaa, [Bk]) ires", auto)
       
  1672 done
       
  1673 
       
  1674 (*inv: from locate_b to locate_b*)
       
  1675 lemma [simp]: "inv_locate_b (as, am) (q, aaa, Oc # xs) ires
       
  1676          \<Longrightarrow> inv_locate_b (as, am) (q, Oc # aaa, xs) ires"
       
  1677 apply(simp only: inv_locate_b.simps in_middle.simps)
       
  1678 apply(erule exE)+
       
  1679 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
       
  1680       rule_tac x = tn in exI, rule_tac x = m in exI)
       
  1681 apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - 1" in exI,
       
  1682       rule_tac x = rn in exI)
       
  1683 apply(case_tac mr, simp_all add: exponent_def, auto)
       
  1684 done
       
  1685 lemma zero_and_nil[intro]: "(Bk # Bk\<^bsup>n\<^esup> = Oc\<^bsup>mr\<^esup> @ Bk # <lm::nat list> @ 
       
  1686                              Bk\<^bsup>rn \<^esup>) \<or> (lm2 = [] \<and> Bk # Bk\<^bsup>n\<^esup> = Oc\<^bsup>mr\<^esup>)
       
  1687        \<Longrightarrow> mr = 0 \<and> lm = []"
       
  1688 apply(rule context_conjI)
       
  1689 apply(case_tac mr, auto simp:exponent_def)
       
  1690 apply(insert BkCons_nil[of "replicate (n - 1) Bk" lm rn])
       
  1691 apply(case_tac n, auto simp: exponent_def Bk_def  tape_of_nl_nil_eq)
       
  1692 done
       
  1693 
       
  1694 lemma tape_of_nat_def: "<[m::nat]> =  Oc # Oc\<^bsup>m\<^esup>"
       
  1695 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
       
  1696 done
       
  1697 lemma [simp]: "\<lbrakk>inv_locate_b (as, am) (q, aaa, Bk # xs) ires; \<exists>n. xs = Bk\<^bsup>n\<^esup>\<rbrakk>
       
  1698             \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires"
       
  1699 apply(simp add: inv_locate_b.simps inv_locate_a.simps)
       
  1700 apply(rule_tac disjI2, rule_tac disjI1)
       
  1701 apply(simp only: in_middle.simps at_begin_fst_bwtn.simps)
       
  1702 apply(erule_tac exE)+
       
  1703 apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = tn in exI, simp)
       
  1704 apply(subgoal_tac "mr = 0 \<and> lm2 = []")
       
  1705 defer
       
  1706 apply(rule_tac n = n and mr = mr and lm = "lm2" 
       
  1707                and rn = rn and n = n in zero_and_nil)
       
  1708 apply(auto simp: exponent_def)
       
  1709 apply(case_tac "lm1 = []", auto simp: tape_of_nat_def)
       
  1710 done
       
  1711 
       
  1712 lemma length_equal: "xs = ys \<Longrightarrow> length xs = length ys"
       
  1713 by auto
       
  1714 lemma [simp]: "a\<^bsup>0\<^esup> = []" 
       
  1715 by(simp add: exp_zero)
       
  1716 (*inv: from locate_b to locate_a*)
       
  1717 lemma [simp]: "length (a\<^bsup>b\<^esup>) = b"
       
  1718 apply(simp add: exponent_def)
       
  1719 done
       
  1720 
       
  1721 lemma [simp]: "\<lbrakk>inv_locate_b (as, am) (q, aaa, Bk # xs) ires; 
       
  1722                 \<not> (\<exists>n. xs = Bk\<^bsup>n\<^esup>)\<rbrakk> 
       
  1723        \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires"
       
  1724 apply(simp add: inv_locate_b.simps inv_locate_a.simps)
       
  1725 apply(rule_tac disjI1)
       
  1726 apply(simp only: in_middle.simps at_begin_norm.simps)
       
  1727 apply(erule_tac exE)+
       
  1728 apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = lm2 in exI, simp)
       
  1729 apply(subgoal_tac "tn = 0", simp add: exponent_def , auto split: if_splits)
       
  1730 apply(case_tac [!] mr, simp_all add: tape_of_nat_def, auto)
       
  1731 apply(case_tac lm2, simp, erule_tac x = rn in allE, simp)
       
  1732 apply(case_tac am, simp, simp)
       
  1733 apply(case_tac lm2, simp, erule_tac x = rn in allE, simp)
       
  1734 apply(drule_tac length_equal, simp)
       
  1735 done
       
  1736 
       
  1737 lemma locate_b_2_a[intro]: 
       
  1738        "inv_locate_b (as, am) (q, aaa, Bk # xs) ires
       
  1739     \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires"
       
  1740 apply(case_tac "\<exists> n. xs = Bk\<^bsup>n\<^esup>", simp, simp)
       
  1741 done
       
  1742 
       
  1743 lemma locate_b_2_locate_a[simp]: 
       
  1744     "\<lbrakk>\<not> a < start_of ly as; 
       
  1745       a < start_of ly as + 2 * n; 
       
  1746       (a - start_of ly as) mod 2 = Suc 0; 
       
  1747      inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, Bk # xs) ires\<rbrakk>
       
  1748    \<Longrightarrow> (Suc a < start_of ly as + 2 * n \<longrightarrow> inv_locate_a (as, am)
       
  1749        (Suc ((a - start_of ly as) div 2), Bk # aaa, xs) ires) \<and>
       
  1750        (\<not> Suc a < start_of ly as + 2 * n \<longrightarrow> 
       
  1751                 inv_locate_a (as, am) (n, Bk # aaa, xs) ires)"
       
  1752 apply(auto)
       
  1753 apply(subgoal_tac "n > 0")
       
  1754 apply(subgoal_tac "(a - start_of ly as) div 2 = n - 1")
       
  1755 apply(insert locate_b_2_a [of as am "n - 1" aaa xs], simp)
       
  1756 apply(arith)
       
  1757 apply(case_tac n, simp, simp)
       
  1758 done
       
  1759 
       
  1760 lemma [simp]:  "inv_locate_b (as, am) (q, l, []) ires 
       
  1761            \<Longrightarrow>  inv_locate_b (as, am) (q, l, [Bk]) ires"
       
  1762 apply(simp only: inv_locate_b.simps in_middle.simps)
       
  1763 apply(erule exE)+
       
  1764 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
       
  1765       rule_tac x = tn in exI, rule_tac x = m in exI, 
       
  1766       rule_tac x = ml in exI, rule_tac x = mr in exI)
       
  1767 apply(auto)
       
  1768 done
       
  1769 
       
  1770 lemma locate_b_2_locate_a_B[simp]: 
       
  1771  "\<lbrakk>\<not> a < start_of ly as; 
       
  1772    a < start_of ly as + 2 * n; 
       
  1773    (a - start_of ly as) mod 2 = Suc 0; 
       
  1774    inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, []) ires\<rbrakk>
       
  1775    \<Longrightarrow> (Suc a < start_of ly as + 2 * n \<longrightarrow> 
       
  1776      inv_locate_a (as, am) 
       
  1777             (Suc ((a - start_of ly as) div 2), Bk # aaa, []) ires) 
       
  1778     \<and> (\<not> Suc a < start_of ly as + 2 * n \<longrightarrow> 
       
  1779                   inv_locate_a (as, am) (n, Bk # aaa, []) ires)"
       
  1780 apply(insert locate_b_2_locate_a [of a ly as n am aaa "[]"], simp)
       
  1781 done
       
  1782 
       
  1783 (*inv: from locate_b to after_write*)
       
  1784 lemma inv_locate_b_2_after_write[simp]: 
       
  1785       "inv_locate_b (as, am) (n, aaa, Bk # xs) ires
       
  1786       \<Longrightarrow> inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n)))
       
  1787           (Suc (Suc (2 * n)), aaa, Oc # xs) ires"
       
  1788 apply(auto simp: in_middle.simps inv_after_write.simps 
       
  1789                  abc_lm_v.simps abc_lm_s.simps  inv_locate_b.simps)
       
  1790 apply(subgoal_tac [!] "mr = 0", auto simp: exponent_def split: if_splits)
       
  1791 apply(subgoal_tac "lm2 = []", simp)
       
  1792 apply(rule_tac x = rn in exI, rule_tac x = "Suc m" in exI,
       
  1793       rule_tac x = "lm1" in exI, simp, rule_tac x = "[]" in exI, simp)
       
  1794 apply(case_tac "Suc (length lm1) - length am", simp, simp only: rep_ind, simp)
       
  1795 apply(subgoal_tac "length lm1 - length am = nat", simp, arith)
       
  1796 apply(drule_tac length_equal, simp)
       
  1797 done
       
  1798 
       
  1799 lemma [simp]: "inv_locate_b (as, am) (n, aaa, []) ires \<Longrightarrow> 
       
  1800      inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) 
       
  1801                      (Suc (Suc (2 * n)), aaa, [Oc]) ires"
       
  1802 apply(insert inv_locate_b_2_after_write [of as am n aaa "[]"])
       
  1803 by(simp)
       
  1804 
       
  1805 (*inv: from after_write to after_move*)
       
  1806 lemma [simp]: "inv_after_write (as, lm) (Suc (Suc (2 * n)), l, Oc # r) ires
       
  1807                 \<Longrightarrow> inv_after_move (as, lm) (2 * n + 3, Oc # l, r) ires"
       
  1808 apply(auto simp:inv_after_move.simps inv_after_write.simps split: if_splits)
       
  1809 done
       
  1810 
       
  1811 lemma [simp]: "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n)
       
  1812                 )) (Suc (Suc (2 * n)), aaa, Bk # xs) ires = False"
       
  1813 apply(simp add: inv_after_write.simps )
       
  1814 done
       
  1815 
       
  1816 lemma [simp]: 
       
  1817  "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) 
       
  1818                         (Suc (Suc (2 * n)), aaa, []) ires = False"
       
  1819 apply(simp add: inv_after_write.simps )
       
  1820 done
       
  1821 
       
  1822 (*inv: from after_move to after_clear*)
       
  1823 lemma [simp]: "inv_after_move (as, lm) (s, l, Oc # r) ires
       
  1824                 \<Longrightarrow> inv_after_clear (as, lm) (s', l, Bk # r) ires"
       
  1825 apply(auto simp: inv_after_move.simps inv_after_clear.simps split: if_splits)
       
  1826 done
       
  1827 
       
  1828 (*inv: from after_move to on_leftmoving*)
       
  1829 lemma inv_after_move_2_inv_on_left_moving[simp]:  
       
  1830    "inv_after_move (as, lm) (s, l, Bk # r) ires
       
  1831    \<Longrightarrow> (l = [] \<longrightarrow> 
       
  1832          inv_on_left_moving (as, lm) (s', [], Bk # Bk # r) ires) \<and>
       
  1833       (l \<noteq> [] \<longrightarrow> 
       
  1834          inv_on_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)"
       
  1835 apply(simp only: inv_after_move.simps inv_on_left_moving.simps)
       
  1836 apply(subgoal_tac "l \<noteq> []", rule conjI, simp, rule impI, 
       
  1837                 rule disjI1, simp only: inv_on_left_moving_norm.simps)
       
  1838 apply(erule exE)+
       
  1839 apply(subgoal_tac "lm2 = []")
       
  1840 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,  
       
  1841     rule_tac x = m in exI, rule_tac x = m in exI, 
       
  1842     rule_tac x = 1 in exI,  
       
  1843     rule_tac x = "rn - 1" in exI, simp, case_tac rn)
       
  1844 apply(auto simp: exponent_def  intro: BkCons_nil split: if_splits)
       
  1845 done
       
  1846 
       
  1847 lemma [elim]: "[] = <lm::nat list> \<Longrightarrow> lm = []"
       
  1848 using tape_of_nl_nil_eq[of lm]
       
  1849 by simp
       
  1850 
       
  1851 lemma inv_after_move_2_inv_on_left_moving_B[simp]: 
       
  1852     "inv_after_move (as, lm) (s, l, []) ires
       
  1853       \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s', [], [Bk]) ires) \<and>
       
  1854           (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s', tl l, [hd l]) ires)"
       
  1855 apply(simp only: inv_after_move.simps inv_on_left_moving.simps)
       
  1856 apply(subgoal_tac "l \<noteq> []", rule conjI, simp, rule impI, rule disjI1,
       
  1857         simp only: inv_on_left_moving_norm.simps)
       
  1858 apply(erule exE)+
       
  1859 apply(subgoal_tac "lm2 = []")
       
  1860 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,  
       
  1861       rule_tac x = m in exI, rule_tac x = m in exI, 
       
  1862       rule_tac x = 1 in exI, rule_tac x = "rn - 1" in exI, simp, case_tac rn)
       
  1863 apply(auto simp: exponent_def  tape_of_nl_nil_eq  intro: BkCons_nil  split: if_splits)
       
  1864 done
       
  1865 
       
  1866 (*inv: from after_clear to on_right_moving*)
       
  1867 lemma [simp]: "Oc # r = replicate rn Bk = False"
       
  1868 apply(case_tac rn, simp, simp)
       
  1869 done
       
  1870 
       
  1871 lemma inv_after_clear_2_inv_on_right_moving[simp]: 
       
  1872      "inv_after_clear (as, lm) (2 * n + 4, l, Bk # r) ires
       
  1873       \<Longrightarrow> inv_on_right_moving (as, lm) (2 * n + 5, Bk # l, r) ires"
       
  1874 apply(auto simp: inv_after_clear.simps inv_on_right_moving.simps )
       
  1875 apply(subgoal_tac "lm2 \<noteq> []")
       
  1876 apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI, 
       
  1877       rule_tac x = "hd lm2" in exI, simp)
       
  1878 apply(rule_tac x = 0 in exI, rule_tac x = "hd lm2" in exI)
       
  1879 apply(simp add: exponent_def, rule conjI)
       
  1880 apply(case_tac [!] "lm2::nat list", auto simp: exponent_def)
       
  1881 apply(case_tac rn, auto split: if_splits simp: tape_of_nat_def)
       
  1882 apply(case_tac list, 
       
  1883      simp add:  tape_of_nl_abv tape_of_nat_list.simps exponent_def)
       
  1884 apply(erule_tac x = "rn - 1" in allE, 
       
  1885       case_tac rn, auto simp: exponent_def)
       
  1886 apply(case_tac list, 
       
  1887      simp add:  tape_of_nl_abv tape_of_nat_list.simps exponent_def)
       
  1888 apply(erule_tac x = "rn - 1" in allE, 
       
  1889       case_tac rn, auto simp: exponent_def)
       
  1890 done
       
  1891 
       
  1892 
       
  1893 lemma [simp]: "inv_after_clear (as, lm) (2 * n + 4, l, []) ires\<Longrightarrow> 
       
  1894                inv_after_clear (as, lm) (2 * n + 4, l, [Bk]) ires" 
       
  1895 by(auto simp: inv_after_clear.simps)
       
  1896 
       
  1897 lemma [simp]: "inv_after_clear (as, lm) (2 * n + 4, l, []) ires
       
  1898              \<Longrightarrow> inv_on_right_moving (as, lm) (2 * n + 5, Bk # l, []) ires"
       
  1899 by(insert 
       
  1900     inv_after_clear_2_inv_on_right_moving[of as lm n l "[]"], simp)
       
  1901 
       
  1902 (*inv: from on_right_moving to on_right_movign*)
       
  1903 lemma [simp]: "inv_on_right_moving (as, lm) (2 * n + 5, l, Oc # r) ires
       
  1904       \<Longrightarrow> inv_on_right_moving (as, lm) (2 * n + 5, Oc # l, r) ires"
       
  1905 apply(auto simp: inv_on_right_moving.simps)
       
  1906 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
       
  1907            rule_tac x = "ml + mr" in exI, simp)
       
  1908 apply(rule_tac x = "Suc ml" in exI, 
       
  1909            rule_tac x = "mr - 1" in exI, simp)
       
  1910 apply(case_tac mr, auto simp: exponent_def )
       
  1911 apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, 
       
  1912       rule_tac x = "ml + mr" in exI, simp)
       
  1913 apply(rule_tac x = "Suc ml" in exI, 
       
  1914       rule_tac x = "mr - 1" in exI, simp)
       
  1915 apply(case_tac mr, auto split: if_splits simp: exponent_def)
       
  1916 done
       
  1917 
       
  1918 lemma inv_on_right_moving_2_inv_on_right_moving[simp]: 
       
  1919      "inv_on_right_moving (as, lm) (2 * n + 5, l, Bk # r) ires
       
  1920      \<Longrightarrow> inv_after_write (as, lm) (Suc (Suc (2 * n)), l, Oc # r) ires"
       
  1921 apply(auto simp: inv_on_right_moving.simps inv_after_write.simps )
       
  1922 apply(case_tac mr, auto simp: exponent_def split: if_splits)
       
  1923 apply(case_tac [!] mr, simp_all)
       
  1924 done
       
  1925       
       
  1926 lemma [simp]: "inv_on_right_moving (as, lm) (2 * n + 5, l, []) ires\<Longrightarrow> 
       
  1927              inv_on_right_moving (as, lm) (2 * n + 5, l, [Bk]) ires"
       
  1928 apply(auto simp: inv_on_right_moving.simps exponent_def)
       
  1929 apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, simp)
       
  1930 apply (rule_tac x = m in exI, auto split: if_splits simp: exponent_def)
       
  1931 done
       
  1932 
       
  1933 (*inv: from on_right_moving to after_write*)
       
  1934 lemma [simp]: "inv_on_right_moving (as, lm) (2 * n + 5, l, []) ires
       
  1935        \<Longrightarrow> inv_after_write (as, lm) (Suc (Suc (2 * n)), l, [Oc]) ires"
       
  1936 apply(rule_tac inv_on_right_moving_2_inv_on_right_moving, simp)
       
  1937 done
       
  1938 
       
  1939 (*inv: from on_left_moving to on_left_moving*)
       
  1940 lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) 
       
  1941                (s, l, Oc # r) ires = False"
       
  1942 apply(auto simp: inv_on_left_moving_in_middle_B.simps )
       
  1943 done
       
  1944 
       
  1945 lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, l, Bk # r) ires 
       
  1946              = False"
       
  1947 apply(auto simp: inv_on_left_moving_norm.simps)
       
  1948 apply(case_tac [!] mr, auto simp: )
       
  1949 done
       
  1950 
       
  1951 lemma [intro]: "\<exists>rna. Oc # Oc\<^bsup>m\<^esup> @ Bk # <lm> @ Bk\<^bsup>rn\<^esup> = <m # lm> @ Bk\<^bsup>rna\<^esup>"
       
  1952 apply(case_tac lm, simp add: tape_of_nl_abv tape_of_nat_list.simps)
       
  1953 apply(rule_tac x = "Suc rn" in exI, simp)
       
  1954 apply(case_tac list, simp_all add: tape_of_nl_abv tape_of_nat_list.simps, auto)
       
  1955 done
       
  1956 
       
  1957 
       
  1958 lemma [simp]: 
       
  1959   "\<lbrakk>inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires;
       
  1960     hd l = Bk; l \<noteq> []\<rbrakk> \<Longrightarrow> 
       
  1961      inv_on_left_moving_in_middle_B (as, lm) (s, tl l, Bk # Oc # r) ires"
       
  1962 apply(case_tac l, simp, simp)
       
  1963 apply(simp only: inv_on_left_moving_norm.simps 
       
  1964                  inv_on_left_moving_in_middle_B.simps)
       
  1965 apply(erule_tac exE)+
       
  1966 apply(rule_tac x = lm1 in exI, rule_tac x = "m # lm2" in exI, auto)
       
  1967 apply(case_tac [!] ml, auto)
       
  1968 apply(rule_tac [!] x = 0 in exI, simp_all add: tape_of_nl_abv tape_of_nat_list.simps)
       
  1969 done
       
  1970 
       
  1971 lemma [simp]: "\<lbrakk>inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires; 
       
  1972                 hd l = Oc; l \<noteq> []\<rbrakk>
       
  1973             \<Longrightarrow> inv_on_left_moving_norm (as, lm) 
       
  1974                                         (s, tl l, Oc # Oc # r) ires"
       
  1975 apply(simp only: inv_on_left_moving_norm.simps)
       
  1976 apply(erule exE)+
       
  1977 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
       
  1978       rule_tac x = m in exI, rule_tac x = "ml - 1" in exI,
       
  1979       rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI, simp)
       
  1980 apply(case_tac ml, auto simp: exponent_def split: if_splits)
       
  1981 done
       
  1982 
       
  1983 lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, [], Oc # r) ires
       
  1984      \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # Oc # r) ires"
       
  1985 apply(auto simp: inv_on_left_moving_norm.simps 
       
  1986                  inv_on_left_moving_in_middle_B.simps split: if_splits)
       
  1987 done
       
  1988 
       
  1989 lemma [simp]:"inv_on_left_moving (as, lm) (s, l, Oc # r) ires
       
  1990     \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s, [], Bk # Oc # r) ires)
       
  1991  \<and>  (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s, tl l, hd l # Oc # r) ires)"
       
  1992 apply(simp add: inv_on_left_moving.simps)
       
  1993 apply(case_tac "l \<noteq> []", rule conjI, simp, simp)
       
  1994 apply(case_tac "hd l", simp, simp, simp)
       
  1995 done
       
  1996 
       
  1997 (*inv: from on_left_moving to check_left_moving*)
       
  1998 lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) 
       
  1999                                       (s, Bk # list, Bk # r) ires
       
  2000           \<Longrightarrow> inv_check_left_moving_on_leftmost (as, lm) 
       
  2001                                       (s', list, Bk # Bk # r) ires"
       
  2002 apply(auto simp: inv_on_left_moving_in_middle_B.simps 
       
  2003                  inv_check_left_moving_on_leftmost.simps split: if_splits)
       
  2004 apply(case_tac [!] "rev lm1", simp_all)
       
  2005 apply(case_tac [!] lista, simp_all add: tape_of_nl_abv tape_of_nat_list.simps)
       
  2006 done
       
  2007 
       
  2008 lemma [simp]:
       
  2009     "inv_check_left_moving_in_middle (as, lm) (s, l, Bk # r) ires= False"
       
  2010 by(auto simp: inv_check_left_moving_in_middle.simps )
       
  2011 
       
  2012 lemma [simp]: 
       
  2013  "inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # r) ires\<Longrightarrow> 
       
  2014   inv_check_left_moving_on_leftmost (as, lm) (s', [], Bk # Bk # r) ires"
       
  2015 apply(auto simp: inv_on_left_moving_in_middle_B.simps 
       
  2016                  inv_check_left_moving_on_leftmost.simps split: if_splits)
       
  2017 done
       
  2018 
       
  2019 
       
  2020 lemma [simp]: "inv_check_left_moving_on_leftmost (as, lm) 
       
  2021                                        (s, list, Oc # r) ires= False"
       
  2022 by(auto simp: inv_check_left_moving_on_leftmost.simps split: if_splits)
       
  2023 
       
  2024 lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) 
       
  2025                                          (s, Oc # list, Bk # r) ires
       
  2026  \<Longrightarrow> inv_check_left_moving_in_middle (as, lm) (s', list, Oc # Bk # r) ires"
       
  2027 apply(auto simp: inv_on_left_moving_in_middle_B.simps 
       
  2028                  inv_check_left_moving_in_middle.simps  split: if_splits)
       
  2029 done
       
  2030 
       
  2031 lemma inv_on_left_moving_2_check_left_moving[simp]:
       
  2032  "inv_on_left_moving (as, lm) (s, l, Bk # r) ires
       
  2033  \<Longrightarrow> (l = [] \<longrightarrow> inv_check_left_moving (as, lm) (s', [], Bk # Bk # r) ires)
       
  2034  \<and> (l \<noteq> [] \<longrightarrow> 
       
  2035       inv_check_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)"
       
  2036 apply(simp add: inv_on_left_moving.simps inv_check_left_moving.simps)
       
  2037 apply(case_tac l, simp, simp)
       
  2038 apply(case_tac a, simp, simp)
       
  2039 done
       
  2040 
       
  2041 lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, l, []) ires = False"
       
  2042 apply(auto simp: inv_on_left_moving_norm.simps)
       
  2043 by(case_tac [!] mr, auto)
       
  2044 
       
  2045 lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires\<Longrightarrow> 
       
  2046      inv_on_left_moving (as, lm) (6 + 2 * n, l, [Bk]) ires"
       
  2047 apply(simp add: inv_on_left_moving.simps)
       
  2048 apply(auto simp: inv_on_left_moving_in_middle_B.simps)
       
  2049 done
       
  2050 
       
  2051 lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires = False"
       
  2052 apply(simp add: inv_on_left_moving.simps)
       
  2053 apply(simp add: inv_on_left_moving_in_middle_B.simps)
       
  2054 done
       
  2055 
       
  2056 lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires
       
  2057  \<Longrightarrow> (l = [] \<longrightarrow> inv_check_left_moving (as, lm) (s', [], [Bk]) ires) \<and>
       
  2058     (l \<noteq> [] \<longrightarrow> inv_check_left_moving (as, lm) (s', tl l, [hd l]) ires)"
       
  2059 by simp
       
  2060 
       
  2061 lemma Oc_Bk_Cons_ex[simp]: 
       
  2062  "Oc # Bk # list = <lm::nat list> @ Bk\<^bsup>ln\<^esup> \<Longrightarrow> 
       
  2063                              \<exists>ln. list = <tl (lm)> @ Bk\<^bsup>ln\<^esup>"
       
  2064 apply(case_tac "lm", simp)
       
  2065 apply(case_tac ln, simp_all add: exponent_def)
       
  2066 apply(case_tac lista, 
       
  2067       auto simp: tape_of_nl_abv tape_of_nat_list.simps exponent_def)
       
  2068 apply(case_tac [!] a, auto simp: )
       
  2069 apply(case_tac ln, simp, rule_tac x = nat in exI, simp)
       
  2070 done
       
  2071 
       
  2072 lemma [simp]:
       
  2073   "Oc # Bk # list = <rev lm1::nat list> @ Bk\<^bsup>ln\<^esup> \<Longrightarrow> 
       
  2074       \<exists>rna. Oc # Bk # <lm2> @ Bk\<^bsup>rn\<^esup> = <hd (rev lm1) # lm2> @ Bk\<^bsup>rna\<^esup>"
       
  2075 apply(frule Oc_Bk_Cons, simp)
       
  2076 apply(case_tac lm2, 
       
  2077      auto simp: tape_of_nl_abv tape_of_nat_list.simps  exponent_def )
       
  2078 apply(rule_tac x = "Suc rn" in exI, simp)
       
  2079 done
       
  2080 
       
  2081 (*inv: from check_left_moving to on_left_moving*)
       
  2082 lemma [intro]: "\<exists>rna. a # a\<^bsup>rn\<^esup> = a\<^bsup>rna\<^esup>"
       
  2083 apply(rule_tac x = "Suc rn" in exI, simp)
       
  2084 done
       
  2085 
       
  2086 lemma 
       
  2087 inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[simp]:
       
  2088 "inv_check_left_moving_in_middle (as, lm) (s, Bk # list, Oc # r) ires
       
  2089   \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s', list, Bk # Oc # r) ires"
       
  2090 apply(simp only: inv_check_left_moving_in_middle.simps 
       
  2091                  inv_on_left_moving_in_middle_B.simps)
       
  2092 apply(erule_tac exE)+
       
  2093 apply(rule_tac x = "rev (tl (rev lm1))" in exI, 
       
  2094       rule_tac x = "[hd (rev lm1)] @ lm2" in exI, auto)
       
  2095 apply(case_tac [!] "rev lm1",simp_all add: tape_of_nl_abv tape_of_nat_list.simps)
       
  2096 apply(case_tac [!] a, simp_all)
       
  2097 apply(case_tac [1] lm2, simp_all add: tape_of_nat_list.simps, auto)
       
  2098 apply(case_tac [3] lm2, simp_all add: tape_of_nat_list.simps, auto)
       
  2099 apply(case_tac [!] lista, simp_all add: tape_of_nat_list.simps)
       
  2100 done
       
  2101 
       
  2102 lemma [simp]: 
       
  2103  "inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires\<Longrightarrow>
       
  2104      inv_check_left_moving_in_middle (as, lm) (s', [Bk], Oc # r) ires"
       
  2105 apply(auto simp: inv_check_left_moving_in_middle.simps )
       
  2106 done
       
  2107 
       
  2108 lemma [simp]: 
       
  2109  "inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires
       
  2110    \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s', [], Bk # Oc # r) ires"
       
  2111 apply(insert 
       
  2112 inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[of 
       
  2113                   as lm n "[]" r], simp)
       
  2114 done 
       
  2115 
       
  2116 lemma [simp]: "a\<^bsup>0\<^esup> = []"
       
  2117 apply(simp add: exponent_def)
       
  2118 done
       
  2119 
       
  2120 lemma [simp]: "inv_check_left_moving_in_middle (as, lm) 
       
  2121                        (s, Oc # list, Oc # r) ires
       
  2122    \<Longrightarrow> inv_on_left_moving_norm (as, lm) (s', list, Oc # Oc # r) ires"
       
  2123 apply(auto simp: inv_check_left_moving_in_middle.simps 
       
  2124                  inv_on_left_moving_norm.simps)
       
  2125 apply(rule_tac x = "rev (tl (rev lm1))" in exI, 
       
  2126       rule_tac x = lm2 in exI, rule_tac x = "hd (rev lm1)" in exI)
       
  2127 apply(rule_tac conjI)
       
  2128 apply(case_tac "rev lm1", simp, simp)
       
  2129 apply(rule_tac x = "hd (rev lm1) - 1" in exI, auto)
       
  2130 apply(rule_tac [!] x = "Suc (Suc 0)" in exI, simp)
       
  2131 apply(case_tac [!] "rev lm1", simp_all)
       
  2132 apply(case_tac [!] a, simp_all add: tape_of_nl_abv tape_of_nat_list.simps, auto)
       
  2133 done 
       
  2134 
       
  2135 lemma [simp]: "inv_check_left_moving (as, lm) (s, l, Oc # r) ires
       
  2136 \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s', [], Bk # Oc # r) ires) \<and>
       
  2137    (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s', tl l, hd l # Oc # r) ires)"
       
  2138 apply(case_tac l, 
       
  2139       auto simp: inv_check_left_moving.simps inv_on_left_moving.simps)
       
  2140 apply(case_tac a, simp, simp)
       
  2141 done
       
  2142 
       
  2143 (*inv: check_left_moving to after_left_moving*)
       
  2144 lemma [simp]: "inv_check_left_moving (as, lm) (s, l, Bk # r) ires
       
  2145                 \<Longrightarrow> inv_after_left_moving (as, lm) (s', Bk # l, r) ires"
       
  2146 apply(auto simp: inv_check_left_moving.simps 
       
  2147  inv_check_left_moving_on_leftmost.simps inv_after_left_moving.simps)
       
  2148 done
       
  2149 
       
  2150 
       
  2151 lemma [simp]:"inv_check_left_moving (as, lm) (s, l, []) ires
       
  2152       \<Longrightarrow> inv_after_left_moving (as, lm) (s', Bk # l, []) ires"
       
  2153 by(simp add: inv_check_left_moving.simps  
       
  2154 inv_check_left_moving_in_middle.simps 
       
  2155 inv_check_left_moving_on_leftmost.simps)
       
  2156 
       
  2157 (*inv: after_left_moving to inv_stop*)
       
  2158 lemma [simp]: "inv_after_left_moving (as, lm) (s, l, Bk # r) ires
       
  2159        \<Longrightarrow> inv_stop (as, lm) (s', Bk # l, r) ires"
       
  2160 apply(auto simp: inv_after_left_moving.simps inv_stop.simps)
       
  2161 done
       
  2162 
       
  2163 lemma [simp]: "inv_after_left_moving (as, lm) (s, l, []) ires
       
  2164              \<Longrightarrow> inv_stop (as, lm) (s', Bk # l, []) ires"
       
  2165 by(auto simp: inv_after_left_moving.simps)
       
  2166 
       
  2167 (*inv: stop to stop*)
       
  2168 lemma [simp]: "inv_stop (as, lm) (x, l, r) ires \<Longrightarrow> 
       
  2169                inv_stop (as, lm) (y, l, r) ires"
       
  2170 apply(simp add: inv_stop.simps)
       
  2171 done
       
  2172 
       
  2173 lemma [simp]: "inv_after_clear (as, lm) (s, aaa, Oc # xs) ires= False"
       
  2174 apply(auto simp: inv_after_clear.simps )
       
  2175 done
       
  2176 
       
  2177 lemma [simp]: 
       
  2178   "inv_after_left_moving (as, lm) (s, aaa, Oc # xs) ires = False"
       
  2179 by(auto simp: inv_after_left_moving.simps  )
       
  2180 
       
  2181 lemma start_of_not0: "as \<noteq> 0 \<Longrightarrow> start_of ly as > 0"
       
  2182 apply(rule startof_not0)
       
  2183 done
       
  2184 
       
  2185 text {*
       
  2186   The single step currectness of the TM complied from Abacus instruction @{text "Inc n"}.
       
  2187   It shows every single step execution of this TM keeps the invariant.
       
  2188 *}
       
  2189 
       
  2190 lemma inc_inv_step: 
       
  2191   assumes 
       
  2192   -- {* Invariant holds on the start *}
       
  2193       h11: "inc_inv ly n (as, am) tc ires" 
       
  2194   -- {* The layout of Abacus program @{text "aprog"} is @{text "ly"} *}
       
  2195   and h12: "ly = layout_of aprog" 
       
  2196   -- {* The instruction at position @{text "as"} is @{text "Inc n"} *}
       
  2197   and h21: "abc_fetch as aprog = Some (Inc n)" 
       
  2198   -- {* TM not yet reach the final state, where @{text "start_of ly as + 2*n + 9"} is the state
       
  2199         where the current TM stops and the next TM starts. *}
       
  2200   and h22: "(\<lambda> (s, l, r). s \<noteq> start_of ly as + 2*n + 9) tc"
       
  2201   shows 
       
  2202   -- {* 
       
  2203   Single step execution of the TM keeps the invaraint, where 
       
  2204   the TM compiled from @{text "Inc n"} is @{text "(ci ly (start_of ly as) (Inc n))"}
       
  2205   @{text "start_of ly as - Suc 0)"} is the offset used to execute this {\em shifted}
       
  2206   TM.
       
  2207   *}
       
  2208   "inc_inv ly n (as, am) (t_step tc (ci ly (start_of ly as) (Inc n), start_of ly as - Suc 0)) ires"
       
  2209 proof -
       
  2210   from h21 h22  have h3 : "start_of (layout_of aprog) as > 0"
       
  2211     apply(case_tac as, simp add: start_of.simps abc_fetch.simps)
       
  2212     apply(insert start_of_not0[of as "layout_of aprog"], simp)
       
  2213     done    
       
  2214   from h11 h12 and h21 h22 and this show ?thesis 
       
  2215     apply(case_tac tc, simp)
       
  2216     apply(case_tac "a = 0", 
       
  2217       auto split:if_splits simp add:t_step.simps,
       
  2218       tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *})
       
  2219     apply (simp_all add:fetch_simps new_tape.simps)
       
  2220     done
       
  2221 qed
       
  2222 
       
  2223 
       
  2224 lemma t_steps_ind: "t_steps tc (p, off) (Suc n)
       
  2225                  = t_step (t_steps tc (p, off) n) (p, off)"
       
  2226 apply(induct n arbitrary: tc)
       
  2227 apply(simp add: t_steps.simps)
       
  2228 apply(simp add: t_steps.simps)
       
  2229 done
       
  2230 
       
  2231 definition lex_pair :: "((nat \<times> nat) \<times> (nat \<times> nat)) set"
       
  2232   where 
       
  2233   "lex_pair \<equiv> less_than <*lex*> less_than"
       
  2234 
       
  2235 definition lex_triple :: 
       
  2236    "((nat \<times> (nat \<times> nat)) \<times> (nat \<times> (nat \<times> nat))) set"
       
  2237   where "lex_triple \<equiv> less_than <*lex*> lex_pair"
       
  2238 
       
  2239 definition lex_square :: 
       
  2240     "((nat \<times> nat \<times> nat \<times> nat) \<times> (nat \<times> nat \<times> nat \<times> nat)) set"
       
  2241   where "lex_square \<equiv> less_than <*lex*> lex_triple"
       
  2242 
       
  2243 fun abc_inc_stage1 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
       
  2244   where
       
  2245   "abc_inc_stage1 (s, l, r) ss n = 
       
  2246             (if s = 0 then 0
       
  2247              else if s \<le> ss+2*n+1 then 5
       
  2248              else if s\<le> ss+2*n+5 then 4
       
  2249              else if s \<le> ss+2*n+7 then 3
       
  2250              else if s = ss+2*n+8 then 2
       
  2251              else 1)"
       
  2252 
       
  2253 fun abc_inc_stage2 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
       
  2254   where
       
  2255   "abc_inc_stage2 (s, l, r) ss n =
       
  2256                 (if s \<le> ss + 2*n + 1 then 0
       
  2257                  else if s = ss + 2*n + 2 then length r
       
  2258                  else if s = ss + 2*n + 3 then length r
       
  2259                  else if s = ss + 2*n + 4 then length r
       
  2260                  else if s = ss + 2*n + 5 then 
       
  2261                                   if r \<noteq> [] then length r
       
  2262                                   else 1
       
  2263                  else if s = ss+2*n+6 then length l
       
  2264                  else if s = ss+2*n+7 then length l
       
  2265                  else 0)"
       
  2266 
       
  2267 fun abc_inc_stage3 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> block list \<Rightarrow>  nat"
       
  2268   where
       
  2269   "abc_inc_stage3 (s, l, r) ss n ires = (
       
  2270               if s = ss + 2*n + 3 then 4
       
  2271               else if s = ss + 2*n + 4 then 3
       
  2272               else if s = ss + 2*n + 5 then 
       
  2273                    if r \<noteq> [] \<and> hd r = Oc then 2
       
  2274                    else 1
       
  2275               else if s = ss + 2*n + 2 then 0
       
  2276               else if s = ss + 2*n + 6 then 
       
  2277                       if l = Bk # ires \<and> r \<noteq> [] \<and>  hd r = Oc then 2
       
  2278                       else 1
       
  2279               else if s = ss + 2*n + 7 then 
       
  2280                       if r \<noteq> [] \<and> hd r = Oc then 3
       
  2281                       else 0
       
  2282               else ss+2*n+9 - s)"
       
  2283 
       
  2284 fun abc_inc_stage4 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> block list \<Rightarrow> nat"
       
  2285   where
       
  2286   "abc_inc_stage4 (s, l, r) ss n ires = 
       
  2287             (if s \<le> ss+2*n+1 \<and> (s - ss) mod 2 = 0 then 
       
  2288                 if (r\<noteq>[] \<and> hd r = Oc) then 0
       
  2289                 else 1
       
  2290              else if (s \<le> ss+2*n+1 \<and> (s - ss) mod 2 = Suc 0) 
       
  2291                                                  then length r
       
  2292              else if s = ss + 2*n + 6 then 
       
  2293                   if l = Bk # ires \<and> hd r = Bk then 0
       
  2294                   else Suc (length l)
       
  2295              else 0)"
       
  2296  
       
  2297 fun abc_inc_measure :: "(t_conf \<times> nat \<times> nat \<times> block list) \<Rightarrow> 
       
  2298                         (nat \<times> nat \<times> nat \<times> nat)"
       
  2299   where
       
  2300   "abc_inc_measure (c, ss, n, ires) = 
       
  2301      (abc_inc_stage1 c ss n, abc_inc_stage2 c ss n, 
       
  2302       abc_inc_stage3 c ss n ires, abc_inc_stage4 c ss n ires)"
       
  2303 
       
  2304 definition abc_inc_LE :: "(((nat \<times> block list \<times> block list) \<times> nat \<times> 
       
  2305        nat \<times> block list) \<times> ((nat \<times> block list \<times> block list) \<times> nat \<times> nat \<times> block list)) set"
       
  2306   where "abc_inc_LE \<equiv> (inv_image lex_square abc_inc_measure)"
       
  2307 
       
  2308 lemma wf_lex_triple: "wf lex_triple"
       
  2309 by (auto intro:wf_lex_prod simp:lex_triple_def lex_pair_def)
       
  2310 
       
  2311 lemma wf_lex_square: "wf lex_square"
       
  2312 by (auto intro:wf_lex_triple simp:lex_triple_def lex_square_def lex_pair_def)
       
  2313 
       
  2314 lemma wf_abc_inc_le[intro]: "wf abc_inc_LE"
       
  2315 by(auto intro:wf_inv_image wf_lex_square simp:abc_inc_LE_def)
       
  2316 
       
  2317 (********************************************************************)
       
  2318 declare inc_inv.simps[simp del]
       
  2319 
       
  2320 lemma halt_lemma2': 
       
  2321   "\<lbrakk>wf LE;  \<forall> n. ((\<not> P (f n) \<and> Q (f n)) \<longrightarrow> 
       
  2322     (Q (f (Suc n)) \<and> (f (Suc n), (f n)) \<in> LE)); Q (f 0)\<rbrakk> 
       
  2323       \<Longrightarrow> \<exists> n. P (f n)"
       
  2324 apply(intro exCI, simp)
       
  2325 apply(subgoal_tac "\<forall> n. Q (f n)", simp)
       
  2326 apply(drule_tac f = f in wf_inv_image)
       
  2327 apply(simp add: inv_image_def)
       
  2328 apply(erule wf_induct, simp)
       
  2329 apply(erule_tac x = x in allE)
       
  2330 apply(erule_tac x = n in allE, erule_tac x = n in allE)
       
  2331 apply(erule_tac x = "Suc x" in allE, simp)
       
  2332 apply(rule_tac allI)
       
  2333 apply(induct_tac n, simp)
       
  2334 apply(erule_tac x = na in allE, simp)
       
  2335 done
       
  2336 
       
  2337 lemma halt_lemma2'': 
       
  2338   "\<lbrakk>P (f n); \<not> P (f (0::nat))\<rbrakk> \<Longrightarrow> 
       
  2339          \<exists> n. (P (f n) \<and> (\<forall> i < n. \<not> P (f i)))"
       
  2340 apply(induct n rule: nat_less_induct, auto)
       
  2341 done
       
  2342 
       
  2343 lemma halt_lemma2''':
       
  2344  "\<lbrakk>\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> LE;
       
  2345                  Q (f 0);  \<forall>i<na. \<not> P (f i)\<rbrakk> \<Longrightarrow> Q (f na)"
       
  2346 apply(induct na, simp, simp)
       
  2347 done
       
  2348 
       
  2349 lemma halt_lemma2: 
       
  2350   "\<lbrakk>wf LE;  
       
  2351     \<forall> n. ((\<not> P (f n) \<and> Q (f n)) \<longrightarrow> (Q (f (Suc n)) \<and> (f (Suc n), (f n)) \<in> LE)); 
       
  2352     Q (f 0); \<not> P (f 0)\<rbrakk> 
       
  2353   \<Longrightarrow> \<exists> n. P (f n) \<and> Q (f n)"
       
  2354 apply(insert halt_lemma2' [of LE P f Q], simp, erule_tac exE)
       
  2355 apply(subgoal_tac "\<exists> n. (P (f n) \<and> (\<forall> i < n. \<not> P (f i)))")
       
  2356 apply(erule_tac exE)+
       
  2357 apply(rule_tac x = na in exI, auto)
       
  2358 apply(rule halt_lemma2''', simp, simp, simp)
       
  2359 apply(erule_tac halt_lemma2'', simp)
       
  2360 done
       
  2361 
       
  2362 lemma [simp]: 
       
  2363   "\<lbrakk>ly = layout_of aprog; abc_fetch as aprog = Some (Inc n)\<rbrakk>
       
  2364     \<Longrightarrow> start_of ly (Suc as) = start_of ly as + 2*n +9"
       
  2365 apply(case_tac as, auto simp: abc_fetch.simps start_of.simps 
       
  2366           layout_of.simps length_of.simps split: if_splits)
       
  2367 done
       
  2368 
       
  2369 lemma inc_inv_init: 
       
  2370  "\<lbrakk>abc_fetch as aprog = Some (Inc n); 
       
  2371    crsp_l ly (as, am) (start_of ly as, l, r) ires; ly = layout_of aprog\<rbrakk>
       
  2372   \<Longrightarrow> inc_inv ly n (as, am) (start_of ly as, l, r) ires"
       
  2373 apply(auto simp: crsp_l.simps inc_inv.simps 
       
  2374       inv_locate_a.simps at_begin_fst_bwtn.simps 
       
  2375       at_begin_fst_awtn.simps at_begin_norm.simps )
       
  2376 apply(auto intro: startof_not0)
       
  2377 done
       
  2378      
       
  2379 lemma inc_inv_stop_pre[simp]: 
       
  2380    "\<lbrakk>ly = layout_of aprog; inc_inv ly n (as, am) (s, l, r) ires; 
       
  2381      s = start_of ly as; abc_fetch as aprog = Some (Inc n)\<rbrakk>
       
  2382     \<Longrightarrow>  (\<forall>na. \<not> (\<lambda>((s, l, r), ss, n', ires'). s = start_of ly (Suc as)) 
       
  2383          (t_steps (s, l, r) (ci ly (start_of ly as) 
       
  2384           (Inc n), start_of ly as - Suc 0) na, s, n, ires) \<and>
       
  2385        (\<lambda>((s, l, r), ss, n', ires'). inc_inv ly n (as, am) (s, l, r) ires')
       
  2386          (t_steps (s, l, r) (ci ly (start_of ly as) 
       
  2387              (Inc n), start_of ly as - Suc 0) na, s, n, ires) \<longrightarrow>
       
  2388        (\<lambda>((s, l, r), ss, n', ires'). inc_inv ly n (as, am) (s, l, r) ires') 
       
  2389       (t_steps (s, l, r) (ci ly (start_of ly as) 
       
  2390               (Inc n), start_of ly as - Suc 0) (Suc na), s, n, ires) \<and>
       
  2391      ((t_steps (s, l, r) (ci ly (start_of ly as) (Inc n), 
       
  2392         start_of ly as - Suc 0) (Suc na), s, n, ires), 
       
  2393       t_steps (s, l, r) (ci ly (start_of ly as) 
       
  2394         (Inc n), start_of ly as - Suc 0) na, s, n, ires) \<in> abc_inc_LE)"
       
  2395 apply(rule allI, rule impI, simp add: t_steps_ind,
       
  2396        rule conjI, erule_tac conjE)
       
  2397 apply(rule_tac inc_inv_step, simp, simp, simp)
       
  2398 apply(case_tac "t_steps (start_of (layout_of aprog) as, l, r) (ci (layout_of aprog) 
       
  2399   (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na", simp)
       
  2400 proof -
       
  2401   fix na
       
  2402   assume h1: "abc_fetch as aprog = Some (Inc n)"
       
  2403     "\<not> (\<lambda>(s, l, r) (ss, n', ires'). s = start_of (layout_of aprog) as + 2 * n + 9)
       
  2404     (t_steps (start_of (layout_of aprog) as, l, r) (ci (layout_of aprog) 
       
  2405     (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na) 
       
  2406     (start_of (layout_of aprog) as, n, ires) \<and>
       
  2407     inc_inv (layout_of aprog) n (as, am) (t_steps (start_of (layout_of aprog) as, l, r)
       
  2408     (ci (layout_of aprog) (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na) ires"
       
  2409   from h1 have h2: "start_of (layout_of aprog) as > 0"
       
  2410     apply(rule_tac startof_not0)
       
  2411     done
       
  2412   from h1 and h2 show "((t_step (t_steps (start_of (layout_of aprog) as, l, r) (ci (layout_of aprog)
       
  2413     (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na)
       
  2414     (ci (layout_of aprog) (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0),
       
  2415     start_of (layout_of aprog) as, n, ires),
       
  2416     t_steps (start_of (layout_of aprog) as, l, r) 
       
  2417     (ci (layout_of aprog) (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na, 
       
  2418     start_of (layout_of aprog) as, n, ires)
       
  2419             \<in> abc_inc_LE"
       
  2420     apply(case_tac "(t_steps (start_of (layout_of aprog) as, l, r) 
       
  2421                (ci (layout_of aprog)
       
  2422          (start_of (layout_of aprog) as) (Inc n), 
       
  2423            start_of (layout_of aprog) as - Suc 0) na)", simp)
       
  2424     apply(case_tac "a = 0", 
       
  2425      auto split:if_splits simp add:t_step.simps inc_inv.simps, 
       
  2426        tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *})
       
  2427     apply(simp_all add:fetch_simps new_tape.simps)
       
  2428     apply(auto simp add: abc_inc_LE_def  
       
  2429     lex_square_def lex_triple_def lex_pair_def
       
  2430       inv_after_write.simps inv_after_move.simps inv_after_clear.simps
       
  2431       inv_on_left_moving.simps inv_on_left_moving_norm.simps split: if_splits)
       
  2432     done
       
  2433 qed
       
  2434 
       
  2435 lemma inc_inv_stop_pre1: 
       
  2436   "\<lbrakk>
       
  2437   ly = layout_of aprog; 
       
  2438   abc_fetch as aprog = Some (Inc n);
       
  2439   s = start_of ly as; 
       
  2440   inc_inv ly n (as, am) (s, l, r) ires
       
  2441   \<rbrakk> \<Longrightarrow> 
       
  2442   (\<exists> stp > 0. (\<lambda> (s', l', r').
       
  2443            s' = start_of ly (Suc as) \<and> 
       
  2444            inc_inv ly n (as, am) (s', l', r') ires) 
       
  2445                (t_steps (s, l, r) (ci ly (start_of ly as) (Inc n), 
       
  2446                         start_of ly as - Suc 0) stp))"
       
  2447 apply(insert halt_lemma2[of abc_inc_LE 
       
  2448     "\<lambda> ((s, l, r), ss, n', ires'). s = start_of ly (Suc as)" 
       
  2449     "(\<lambda> stp. (t_steps (s, l, r) 
       
  2450      (ci ly (start_of ly as) (Inc n), 
       
  2451      start_of ly as - Suc 0) stp, s, n, ires))" 
       
  2452     "\<lambda> ((s, l, r), ss, n'). inc_inv ly n (as, am) (s, l, r) ires"])
       
  2453 apply(insert  wf_abc_inc_le)
       
  2454 apply(insert inc_inv_stop_pre[of ly aprog n as am s l r ires], simp)
       
  2455 apply(simp only: t_steps.simps, auto)
       
  2456 apply(rule_tac x = na in exI)
       
  2457 apply(case_tac "(t_steps (start_of (layout_of aprog) as, l, r) 
       
  2458    (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  2459    (Inc n), start_of (layout_of aprog) as - Suc 0) na)", simp)
       
  2460 apply(case_tac na, simp add: t_steps.simps, simp)
       
  2461 done
       
  2462 
       
  2463 lemma inc_inv_stop: 
       
  2464   assumes program_and_layout: 
       
  2465   -- {* There is an Abacus program @{text "aprog"} and its layout is @{text "ly"}: *}
       
  2466   "ly = layout_of aprog"
       
  2467   and an_instruction:
       
  2468   -- {* There is an instruction @{text "Inc n"} at postion @{text "as"} of @{text "aprog"} *}
       
  2469   "abc_fetch as aprog = Some (Inc n)"
       
  2470   and the_start_state:
       
  2471   -- {* According to @{text "ly"} and @{text "as"}, 
       
  2472         the start state of the TM compiled from this
       
  2473         @{text "Inc n"} instruction should be @{text "s"}:
       
  2474      *}
       
  2475   "s = start_of ly as"
       
  2476   and inv:
       
  2477   -- {* Invariant holds on configuration @{text "(s, l, r)"} *}
       
  2478   "inc_inv ly n (as, am) (s, l, r) ires"
       
  2479   shows  -- {* After @{text "stp"} steps of execution, the compiled 
       
  2480             TM reaches the start state of next compiled TM and the invariant 
       
  2481             still holds.
       
  2482             *}
       
  2483       "(\<exists> stp > 0. (\<lambda> (s', l', r').
       
  2484            s' = start_of ly (Suc as) \<and> 
       
  2485            inc_inv ly n (as, am) (s', l', r') ires) 
       
  2486                (t_steps (s, l, r) (ci ly (start_of ly as) (Inc n), 
       
  2487                         start_of ly as - Suc 0) stp))"
       
  2488 proof -
       
  2489   from inc_inv_stop_pre1 [OF  program_and_layout an_instruction the_start_state inv] 
       
  2490   show ?thesis .
       
  2491 qed
       
  2492 
       
  2493 lemma inc_inv_stop_cond: 
       
  2494   "\<lbrakk>ly = layout_of aprog; 
       
  2495     s' = start_of ly (as + 1); 
       
  2496     inc_inv ly n (as, lm) (s', (l', r')) ires; 
       
  2497     abc_fetch as aprog = Some (Inc n)\<rbrakk> \<Longrightarrow>
       
  2498     crsp_l ly (Suc as, abc_lm_s lm n (Suc (abc_lm_v lm n))) 
       
  2499                                                 (s', l', r') ires"
       
  2500 apply(subgoal_tac "s' = start_of ly as + 2*n + 9", simp)
       
  2501 apply(auto simp: inc_inv.simps inv_stop.simps crsp_l.simps )
       
  2502 done
       
  2503 
       
  2504 lemma inc_crsp_ex_pre:
       
  2505   "\<lbrakk>ly = layout_of aprog; 
       
  2506     crsp_l ly (as, am) tc ires;  
       
  2507     abc_fetch as aprog = Some (Inc n)\<rbrakk>
       
  2508  \<Longrightarrow> \<exists>stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Inc n))) 
       
  2509                 (t_steps tc (ci ly (start_of ly as) (Inc n), 
       
  2510                                 start_of ly as - Suc 0) stp) ires"
       
  2511 proof(case_tac tc, simp add: abc_step_l.simps)
       
  2512   fix a b c
       
  2513   assume h1: "ly = layout_of aprog" 
       
  2514          "crsp_l (layout_of aprog) (as, am) (a, b, c) ires"
       
  2515          "abc_fetch as aprog = Some (Inc n)"
       
  2516   hence h2: "a = start_of ly as"
       
  2517     by(auto simp: crsp_l.simps)
       
  2518   from h1 and h2 have h3: 
       
  2519        "inc_inv ly n (as, am) (start_of ly as, b, c) ires"
       
  2520     by(rule_tac inc_inv_init, simp, simp, simp)
       
  2521   from h1 and h2 and h3 have h4:
       
  2522        "(\<exists> stp > 0. (\<lambda> (s', l', r'). s' = 
       
  2523            start_of ly (Suc as) \<and> inc_inv ly n (as, am) (s', l', r') ires)
       
  2524          (t_steps (a, b, c) (ci ly (start_of ly as) 
       
  2525                  (Inc n), start_of ly as - Suc 0) stp))"
       
  2526     apply(rule_tac inc_inv_stop, auto)
       
  2527     done
       
  2528   from h1 and h2 and h3 and h4 show 
       
  2529      "\<exists>stp > 0. crsp_l (layout_of aprog) 
       
  2530         (Suc as, abc_lm_s am n (Suc (abc_lm_v am n)))
       
  2531        (t_steps (a, b, c) (ci (layout_of aprog) 
       
  2532           (start_of (layout_of aprog) as) (Inc n), 
       
  2533               start_of (layout_of aprog) as - Suc 0) stp) ires"
       
  2534     apply(erule_tac exE)
       
  2535     apply(rule_tac x = stp in exI)
       
  2536     apply(case_tac "(t_steps (a, b, c) (ci (layout_of aprog) 
       
  2537          (start_of (layout_of aprog) as) (Inc n), 
       
  2538              start_of (layout_of aprog) as - Suc 0) stp)", simp)
       
  2539     apply(rule_tac inc_inv_stop_cond, auto)
       
  2540     done
       
  2541 qed
       
  2542 
       
  2543 text {*
       
  2544   The total correctness of the compilaton of @{text "Inc n"} instruction.
       
  2545 *}
       
  2546 
       
  2547 lemma inc_crsp_ex:
       
  2548   assumes layout: 
       
  2549   -- {* For any Abacus program @{text "aprog"}, assuming its layout is @{text "ly"} *}
       
  2550   "ly = layout_of aprog"
       
  2551   and corresponds: 
       
  2552   -- {* Abacus configuration @{text "(as, am)"} is in correspondence with 
       
  2553         TM configuration @{text "tc"} *}
       
  2554   "crsp_l ly (as, am) tc ires"
       
  2555   and inc:
       
  2556   -- {* There is an instruction @{text "Inc n"} at postion @{text "as"} of @{text "aprog"} *}
       
  2557   "abc_fetch as aprog = Some (Inc n)"
       
  2558   shows
       
  2559   -- {* 
       
  2560   After @{text "stp"} steps of execution, the TM compiled from this @{text "Inc n"}
       
  2561   stops with a configuration which corresponds to the Abacus configuration obtained
       
  2562   from the execution of this @{text "Inc n"} instruction.
       
  2563   *} 
       
  2564   "\<exists>stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Inc n))) 
       
  2565                        (t_steps tc (ci ly (start_of ly as) (Inc n), 
       
  2566                                 start_of ly as - Suc 0) stp) ires"
       
  2567 proof -
       
  2568   from inc_crsp_ex_pre [OF layout corresponds inc] show ?thesis .
       
  2569 qed
       
  2570 
       
  2571 (*******End: inc crsp********)
       
  2572 
       
  2573 (*******Begin: dec crsp******)
       
  2574 
       
  2575 subsubsection {* The compilation of @{text "Dec n e"} *}
       
  2576 
       
  2577 
       
  2578 text {*
       
  2579   The lemmas in this section lead to the correctness of the compilation 
       
  2580   of @{text "Dec n e"} instruction using the same techniques as 
       
  2581   @{text "Inc n"}.
       
  2582 *}
       
  2583 
       
  2584 type_synonym dec_inv_t = "(nat * nat list) \<Rightarrow> t_conf \<Rightarrow> block list \<Rightarrow>  bool"
       
  2585 
       
  2586 fun dec_first_on_right_moving :: "nat \<Rightarrow> dec_inv_t"
       
  2587   where
       
  2588   "dec_first_on_right_moving n (as, lm) (s, l, r) ires = 
       
  2589                (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
       
  2590          ml + mr = Suc m \<and> length lm1 = n \<and> ml > 0 \<and> m > 0 \<and>
       
  2591              (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
       
  2592                           else  l = (Oc\<^bsup>ml\<^esup>) @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
       
  2593     ((r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ <lm2> @ (Bk\<^bsup>rn\<^esup>)) \<or> (r = (Oc\<^bsup>mr\<^esup>) \<and> lm2 = [])))"
       
  2594 
       
  2595 fun dec_on_right_moving :: "dec_inv_t"
       
  2596   where
       
  2597   "dec_on_right_moving (as, lm) (s, l, r) ires =  
       
  2598    (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
       
  2599                              ml + mr = Suc (Suc m) \<and>
       
  2600    (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
       
  2601                 else  l = (Oc\<^bsup>ml\<^esup>) @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
       
  2602    ((r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ <lm2> @ (Bk\<^bsup>rn\<^esup>)) \<or> (r = (Oc\<^bsup>mr\<^esup>) \<and> lm2 = [])))"
       
  2603 
       
  2604 fun dec_after_clear :: "dec_inv_t"
       
  2605   where
       
  2606   "dec_after_clear (as, lm) (s, l, r) ires = 
       
  2607               (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
       
  2608                 ml + mr = Suc m \<and> ml = Suc m \<and> r \<noteq> [] \<and> r \<noteq> [] \<and>
       
  2609                (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
       
  2610                             else l = (Oc\<^bsup>ml \<^esup>) @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
       
  2611                (tl r = Bk # <lm2> @ (Bk\<^bsup>rn\<^esup>) \<or> tl r = [] \<and> lm2 = []))"
       
  2612 
       
  2613 fun dec_after_write :: "dec_inv_t"
       
  2614   where
       
  2615   "dec_after_write (as, lm) (s, l, r) ires = 
       
  2616          (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
       
  2617        ml + mr = Suc m \<and> ml = Suc m \<and> lm2 \<noteq> [] \<and>
       
  2618        (if lm1 = [] then l = Bk # Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
       
  2619                     else l = Bk # (Oc\<^bsup>ml \<^esup>) @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
       
  2620        tl r = <lm2> @ (Bk\<^bsup>rn\<^esup>))"
       
  2621 
       
  2622 fun dec_right_move :: "dec_inv_t"
       
  2623   where
       
  2624   "dec_right_move (as, lm) (s, l, r) ires = 
       
  2625         (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 
       
  2626             \<and> ml = Suc m \<and> mr = (0::nat) \<and> 
       
  2627               (if lm1 = [] then l = Bk # Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
       
  2628                           else l = Bk # Oc\<^bsup>ml\<^esup>@ [Bk] @ <rev lm1> @ Bk # Bk # ires) 
       
  2629            \<and> (r = Bk # <lm2> @ Bk\<^bsup>rn\<^esup>\<or> r = [] \<and> lm2 = []))"
       
  2630 
       
  2631 fun dec_check_right_move :: "dec_inv_t"
       
  2632   where
       
  2633   "dec_check_right_move (as, lm) (s, l, r) ires = 
       
  2634         (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
       
  2635            ml = Suc m \<and> mr = (0::nat) \<and> 
       
  2636            (if lm1 = [] then l = Bk # Bk # Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
       
  2637                        else l = Bk # Bk # Oc\<^bsup>ml \<^esup>@ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
       
  2638            r = <lm2> @ Bk\<^bsup>rn\<^esup>)"
       
  2639 
       
  2640 fun dec_left_move :: "dec_inv_t"
       
  2641   where
       
  2642   "dec_left_move (as, lm) (s, l, r) ires = 
       
  2643     (\<exists> lm1 m rn. (lm::nat list) = lm1 @ [m::nat] \<and>   
       
  2644     rn > 0 \<and> 
       
  2645    (if lm1 = [] then l = Bk # Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # ires
       
  2646     else l = Bk # Oc\<^bsup>Suc m\<^esup> @ Bk # <rev lm1> @ Bk # Bk # ires) \<and> r = Bk\<^bsup>rn\<^esup>)"
       
  2647 
       
  2648 declare
       
  2649   dec_on_right_moving.simps[simp del] dec_after_clear.simps[simp del] 
       
  2650   dec_after_write.simps[simp del] dec_left_move.simps[simp del] 
       
  2651   dec_check_right_move.simps[simp del] dec_right_move.simps[simp del] 
       
  2652   dec_first_on_right_moving.simps[simp del]
       
  2653 
       
  2654 fun inv_locate_n_b :: "inc_inv_t"
       
  2655   where
       
  2656   "inv_locate_n_b (as, lm) (s, l, r) ires= 
       
  2657     (\<exists> lm1 lm2 tn m ml mr rn. lm @ 0\<^bsup>tn\<^esup> = lm1 @ [m] @ lm2 \<and> 
       
  2658      length lm1 = s \<and> m + 1 = ml + mr \<and> 
       
  2659      ml = 1 \<and> tn = s + 1 - length lm \<and>
       
  2660      (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires
       
  2661       else l = Oc\<^bsup>ml\<^esup>@Bk#<rev lm1>@Bk#Bk#ires) \<and> 
       
  2662      (r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ <lm2>@ (Bk\<^bsup>rn\<^esup>) \<or> (lm2 = [] \<and> r = (Oc\<^bsup>mr\<^esup>)))
       
  2663   )"
       
  2664 
       
  2665 fun dec_inv_1 :: "layout \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> dec_inv_t"
       
  2666   where
       
  2667   "dec_inv_1 ly n e (as, am) (s, l, r) ires = 
       
  2668            (let ss = start_of ly as in
       
  2669             let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in
       
  2670             let am'' = abc_lm_s am n (abc_lm_v am n) in
       
  2671               if s = start_of ly e then  inv_stop (as, am'') (s, l, r) ires
       
  2672               else if s = ss then False
       
  2673               else if ss \<le> s \<and> s < ss + 2*n then
       
  2674                    if (s - ss) mod 2 = 0 then 
       
  2675                         inv_locate_a (as, am) ((s - ss) div 2, l, r) ires
       
  2676                     \<or> inv_locate_a (as, am'') ((s - ss) div 2, l, r) ires
       
  2677                    else 
       
  2678                      inv_locate_b (as, am) ((s - ss) div 2, l, r) ires
       
  2679                   \<or> inv_locate_b (as, am'') ((s - ss) div 2, l, r) ires
       
  2680               else if s = ss + 2 * n then 
       
  2681                   inv_locate_a (as, am) (n, l, r) ires
       
  2682                 \<or> inv_locate_a (as, am'') (n, l, r) ires
       
  2683               else if s = ss + 2 * n + 1 then 
       
  2684                   inv_locate_b (as, am) (n, l, r) ires
       
  2685               else if s = ss + 2 * n + 13 then 
       
  2686                   inv_on_left_moving (as, am'') (s, l, r) ires
       
  2687               else if s = ss + 2 * n + 14 then 
       
  2688                   inv_check_left_moving (as, am'') (s, l, r) ires
       
  2689               else if s = ss + 2 * n + 15 then 
       
  2690                   inv_after_left_moving (as, am'') (s, l, r) ires
       
  2691               else False)"
       
  2692 
       
  2693 fun dec_inv_2 :: "layout \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> dec_inv_t"
       
  2694   where
       
  2695   "dec_inv_2 ly n e (as, am) (s, l, r) ires =
       
  2696            (let ss = start_of ly as in
       
  2697             let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in
       
  2698             let am'' = abc_lm_s am n (abc_lm_v am n) in
       
  2699               if s = 0 then False
       
  2700               else if s = ss then False
       
  2701               else if ss \<le> s \<and> s < ss + 2*n then
       
  2702                    if (s - ss) mod 2 = 0 then 
       
  2703                       inv_locate_a (as, am) ((s - ss) div 2, l, r) ires
       
  2704                    else inv_locate_b (as, am) ((s - ss) div 2, l, r) ires
       
  2705               else if s = ss + 2 * n then 
       
  2706                       inv_locate_a (as, am) (n, l, r) ires
       
  2707               else if s = ss + 2 * n + 1 then 
       
  2708                       inv_locate_n_b (as, am) (n, l, r) ires
       
  2709               else if s = ss + 2 * n + 2 then 
       
  2710                       dec_first_on_right_moving n (as, am'') (s, l, r) ires
       
  2711               else if s = ss + 2 * n + 3 then 
       
  2712                       dec_after_clear (as, am') (s, l, r) ires
       
  2713               else if s = ss + 2 * n + 4 then 
       
  2714                       dec_right_move (as, am') (s, l, r) ires
       
  2715               else if s = ss + 2 * n + 5 then 
       
  2716                       dec_check_right_move (as, am') (s, l, r) ires
       
  2717               else if s = ss + 2 * n + 6 then 
       
  2718                       dec_left_move (as, am') (s, l, r) ires
       
  2719               else if s = ss + 2 * n + 7 then 
       
  2720                       dec_after_write (as, am') (s, l, r) ires
       
  2721               else if s = ss + 2 * n + 8 then 
       
  2722                       dec_on_right_moving (as, am') (s, l, r) ires
       
  2723               else if s = ss + 2 * n + 9 then 
       
  2724                       dec_after_clear (as, am') (s, l, r) ires
       
  2725               else if s = ss + 2 * n + 10 then 
       
  2726                       inv_on_left_moving (as, am') (s, l, r) ires
       
  2727               else if s = ss + 2 * n + 11 then 
       
  2728                       inv_check_left_moving (as, am') (s, l, r) ires
       
  2729               else if s = ss + 2 * n + 12 then 
       
  2730                       inv_after_left_moving (as, am') (s, l, r) ires
       
  2731               else if s = ss + 2 * n + 16 then 
       
  2732                       inv_stop (as, am') (s, l, r) ires
       
  2733               else False)"
       
  2734 
       
  2735 (*begin: dec_fetch lemmas*)
       
  2736 
       
  2737 lemma dec_fetch_locate_a_o: 
       
  2738       "\<lbrakk>start_of ly as \<le> a;
       
  2739         a < start_of ly as + 2 * n; start_of ly as > 0;
       
  2740         a - start_of ly as = 2 * q\<rbrakk>
       
  2741        \<Longrightarrow> fetch (ci (layout_of aprog) 
       
  2742          (start_of ly as) (Dec n e)) (Suc (2 * q))  Oc = (R, a + 1)"
       
  2743 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  2744                   nth_of.simps tshift.simps nth_append Suc_pre)
       
  2745 apply(subgoal_tac "(findnth n ! Suc (4 * q)) = 
       
  2746                           findnth (Suc q) ! (4 * q + 1)")
       
  2747 apply(simp add: findnth.simps nth_append)
       
  2748 apply(subgoal_tac " findnth n !(4 * q + 1) = 
       
  2749                           findnth (Suc q) ! (4 * q + 1)", simp)
       
  2750 apply(rule_tac findnth_nth, auto)
       
  2751 done
       
  2752 
       
  2753 lemma  dec_fetch_locate_a_b:
       
  2754        "\<lbrakk>start_of ly as \<le> a; 
       
  2755          a < start_of ly as + 2 * n; 
       
  2756          start_of ly as > 0;
       
  2757          a - start_of ly as = 2 * q\<rbrakk>
       
  2758        \<Longrightarrow> fetch (ci (layout_of aprog) (start_of ly as) (Dec n e)) 
       
  2759               (Suc (2 * q))  Bk = (W1, a)"
       
  2760 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  2761                   nth_of.simps tshift.simps nth_append)
       
  2762 apply(subgoal_tac "(findnth n ! (4 * q)) = 
       
  2763                        findnth (Suc q) ! (4 * q )")
       
  2764 apply(simp add: findnth.simps nth_append)
       
  2765 apply(subgoal_tac " findnth n !(4 * q + 0) = 
       
  2766                        findnth (Suc q) ! (4 * q + 0)", simp)
       
  2767 apply(rule_tac findnth_nth, auto)
       
  2768 done
       
  2769 
       
  2770 lemma dec_fetch_locate_b_o:
       
  2771       "\<lbrakk>start_of ly as \<le> a; 
       
  2772         a < start_of ly as + 2 * n; 
       
  2773         (a - start_of ly as) mod 2 = Suc 0; 
       
  2774         start_of ly as> 0\<rbrakk>
       
  2775        \<Longrightarrow> fetch (ci (layout_of aprog) (start_of ly as) (Dec n e)) 
       
  2776                        (Suc (a - start_of ly as)) Oc = (R, a)"
       
  2777 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  2778                   nth_of.simps tshift.simps nth_append)
       
  2779 apply(subgoal_tac "\<exists> q. (a - start_of ly as) = 2 * q + 1", auto)
       
  2780 apply(subgoal_tac "(findnth n ! Suc (Suc (Suc (4 * q)))) = 
       
  2781                                 findnth (Suc q) ! (4 * q + 3)")
       
  2782 apply(simp add: findnth.simps nth_append)
       
  2783 apply(subgoal_tac " findnth n ! (4 * q + 3) = 
       
  2784                  findnth (Suc q) ! (4 * q + 3)", simp add: add3_Suc)
       
  2785 apply(rule_tac findnth_nth, auto)
       
  2786 done
       
  2787 
       
  2788 lemma dec_fetch_locate_b_b: 
       
  2789       "\<lbrakk>\<not> a < start_of ly as; 
       
  2790         a < start_of ly as + 2 * n; 
       
  2791        (a - start_of ly as) mod 2 = Suc 0; 
       
  2792         start_of ly as > 0\<rbrakk>
       
  2793        \<Longrightarrow> fetch (ci (layout_of aprog) (start_of ly as) (Dec n e)) 
       
  2794               (Suc (a - start_of ly as))  Bk = (R, a + 1)"
       
  2795 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  2796                   nth_of.simps tshift.simps nth_append)
       
  2797 apply(subgoal_tac "\<exists> q. (a - start_of ly as) = 2 * q + 1", auto)
       
  2798 apply(subgoal_tac "(findnth n ! Suc ((Suc (4 * q)))) = 
       
  2799                           findnth (Suc q) ! (4 * q + 2)")
       
  2800 apply(simp add: findnth.simps nth_append)
       
  2801 apply(subgoal_tac " findnth n ! (4 * q + 2) = 
       
  2802                           findnth (Suc q) ! (4 * q + 2)", simp)
       
  2803 apply(rule_tac findnth_nth, auto)
       
  2804 done
       
  2805 
       
  2806 lemma dec_fetch_locate_n_a_o: 
       
  2807        "start_of ly as > 0 \<Longrightarrow> fetch (ci (layout_of aprog) 
       
  2808                        (start_of ly as) (Dec n e)) (Suc (2 * n))  Oc
       
  2809        = (R, start_of ly as + 2*n + 1)"
       
  2810 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  2811                   nth_of.simps tshift.simps nth_append tdec_b_def)
       
  2812 done
       
  2813 
       
  2814 lemma dec_fetch_locate_n_a_b: 
       
  2815        "start_of ly as > 0 \<Longrightarrow> fetch (ci (layout_of aprog) 
       
  2816                        (start_of ly as) (Dec n e)) (Suc (2 * n))  Bk
       
  2817        = (W1, start_of ly as + 2*n)"
       
  2818 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  2819                   nth_of.simps tshift.simps nth_append tdec_b_def)
       
  2820 done
       
  2821 
       
  2822 lemma dec_fetch_locate_n_b_o: 
       
  2823        "start_of ly as > 0 \<Longrightarrow> 
       
  2824             fetch (ci (layout_of aprog) 
       
  2825                 (start_of ly as) (Dec n e)) (Suc (Suc (2 * n)))  Oc
       
  2826       = (R, start_of ly as + 2*n + 2)"
       
  2827 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  2828                   nth_of.simps tshift.simps nth_append tdec_b_def)
       
  2829 done
       
  2830 
       
  2831 
       
  2832 lemma dec_fetch_locate_n_b_b: 
       
  2833        "start_of ly as > 0 \<Longrightarrow> 
       
  2834        fetch (ci (layout_of aprog) 
       
  2835                   (start_of ly as) (Dec n e)) (Suc (Suc (2 * n))) Bk
       
  2836       = (L, start_of ly as + 2*n + 13)"
       
  2837 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  2838                   nth_of.simps tshift.simps nth_append tdec_b_def)
       
  2839 done
       
  2840 
       
  2841 lemma dec_fetch_first_on_right_move_o: 
       
  2842       "start_of ly as > 0 \<Longrightarrow> 
       
  2843        fetch (ci (layout_of aprog) 
       
  2844              (start_of ly as) (Dec n e)) (Suc (Suc (Suc (2 * n))))  Oc
       
  2845      = (R, start_of ly as + 2*n + 2)"
       
  2846 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  2847                   nth_of.simps tshift.simps nth_append tdec_b_def)
       
  2848 done
       
  2849 
       
  2850 lemma dec_fetch_first_on_right_move_b: 
       
  2851       "start_of ly as > 0 \<Longrightarrow> 
       
  2852       fetch (ci (layout_of aprog) (start_of ly as) (Dec n e)) 
       
  2853                              (Suc (Suc (Suc (2 * n))))  Bk
       
  2854      = (L, start_of ly as + 2*n + 3)"
       
  2855 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  2856                   nth_of.simps tshift.simps nth_append tdec_b_def)
       
  2857 done
       
  2858 
       
  2859 lemma [simp]: "fetch x (a + 2 * n) b = fetch x (2 * n + a) b"
       
  2860 thm arg_cong
       
  2861 apply(rule_tac x = "a + 2*n" and y = "2*n + a" in arg_cong, simp)
       
  2862 done
       
  2863 
       
  2864 lemma dec_fetch_first_after_clear_o: 
       
  2865      "start_of ly as > 0 \<Longrightarrow> fetch (ci (layout_of aprog) 
       
  2866                       (start_of ly as) (Dec n e)) (2 * n + 4) Oc
       
  2867     = (W0, start_of ly as + 2*n + 3)"
       
  2868 apply(auto simp: ci.simps findnth.simps tshift.simps 
       
  2869                           tdec_b_def add3_Suc)
       
  2870 apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps)
       
  2871 apply(auto simp: nth_of.simps nth_append)
       
  2872 done
       
  2873 
       
  2874 lemma dec_fetch_first_after_clear_b: 
       
  2875      "start_of ly as > 0 \<Longrightarrow> 
       
  2876      fetch (ci (layout_of aprog) 
       
  2877                    (start_of ly as) (Dec n e)) (2 * n + 4) Bk
       
  2878     = (R, start_of ly as + 2*n + 4)"
       
  2879 apply(auto simp: ci.simps findnth.simps 
       
  2880                tshift.simps tdec_b_def add3_Suc)
       
  2881 apply(subgoal_tac "2*n + 4= Suc (2*n + 3)", simp only: fetch.simps)
       
  2882 apply(auto simp: nth_of.simps nth_append)
       
  2883 done
       
  2884 
       
  2885 lemma dec_fetch_right_move_b: 
       
  2886      "start_of ly as > 0 \<Longrightarrow> fetch (ci (layout_of aprog) 
       
  2887                           (start_of ly as) (Dec n e)) (2 * n + 5) Bk
       
  2888     = (R, start_of ly as + 2*n + 5)"
       
  2889 apply(auto simp: ci.simps findnth.simps 
       
  2890                  tshift.simps tdec_b_def add3_Suc)
       
  2891 apply(subgoal_tac "2*n + 5= Suc (2*n + 4)", simp only: fetch.simps)
       
  2892 apply(auto simp: nth_of.simps nth_append)
       
  2893 done
       
  2894 
       
  2895 lemma dec_fetch_check_right_move_b: 
       
  2896      "start_of ly as > 0 \<Longrightarrow> 
       
  2897       fetch (ci (layout_of aprog)
       
  2898                 (start_of ly as) (Dec n e)) (2 * n + 6) Bk
       
  2899     = (L, start_of ly as + 2*n + 6)"
       
  2900 apply(auto simp: ci.simps findnth.simps 
       
  2901                tshift.simps tdec_b_def add3_Suc)
       
  2902 apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps)
       
  2903 apply(auto simp: nth_of.simps nth_append)
       
  2904 done
       
  2905 
       
  2906 lemma dec_fetch_check_right_move_o: 
       
  2907      "start_of ly as > 0 \<Longrightarrow> 
       
  2908     fetch (ci (layout_of aprog) (start_of ly as) 
       
  2909                       (Dec n e)) (2 * n + 6) Oc
       
  2910     = (L, start_of ly as + 2*n + 7)"
       
  2911 apply(auto simp: ci.simps findnth.simps 
       
  2912                  tshift.simps tdec_b_def add3_Suc)
       
  2913 apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps)
       
  2914 apply(auto simp: nth_of.simps nth_append)
       
  2915 done
       
  2916 
       
  2917 lemma dec_fetch_left_move_b: 
       
  2918      "start_of ly as > 0 \<Longrightarrow> 
       
  2919      fetch (ci (layout_of aprog) 
       
  2920              (start_of ly as) (Dec n e)) (2 * n + 7) Bk
       
  2921     = (L, start_of ly as + 2*n + 10)"
       
  2922 apply(auto simp: ci.simps findnth.simps 
       
  2923                  tshift.simps tdec_b_def add3_Suc)
       
  2924 apply(subgoal_tac "2*n + 7 = Suc (2*n + 6)", simp only: fetch.simps)
       
  2925 apply(auto simp: nth_of.simps nth_append)
       
  2926 done
       
  2927 
       
  2928 lemma dec_fetch_after_write_b: 
       
  2929      "start_of ly as > 0 \<Longrightarrow> 
       
  2930     fetch (ci (layout_of aprog) 
       
  2931                    (start_of ly as) (Dec n e)) (2 * n + 8) Bk
       
  2932     = (W1, start_of ly as + 2*n + 7)"
       
  2933 apply(auto simp: ci.simps findnth.simps 
       
  2934                  tshift.simps tdec_b_def add3_Suc)
       
  2935 apply(subgoal_tac "2*n + 8 = Suc (2*n + 7)", simp only: fetch.simps)
       
  2936 apply(auto simp: nth_of.simps nth_append)
       
  2937 done
       
  2938 
       
  2939 lemma dec_fetch_after_write_o: 
       
  2940      "start_of ly as > 0 \<Longrightarrow> 
       
  2941      fetch (ci (layout_of aprog) 
       
  2942                    (start_of ly as) (Dec n e)) (2 * n + 8) Oc
       
  2943     = (R, start_of ly as + 2*n + 8)"
       
  2944 apply(auto simp: ci.simps findnth.simps 
       
  2945                  tshift.simps tdec_b_def add3_Suc)
       
  2946 apply(subgoal_tac "2*n + 8 = Suc (2*n + 7)", simp only: fetch.simps)
       
  2947 apply(auto simp: nth_of.simps nth_append)
       
  2948 done
       
  2949 
       
  2950 lemma dec_fetch_on_right_move_b: 
       
  2951      "start_of ly as > 0 \<Longrightarrow> 
       
  2952      fetch (ci (layout_of aprog) 
       
  2953                    (start_of ly as) (Dec n e)) (2 * n + 9) Bk
       
  2954     = (L, start_of ly as + 2*n + 9)"
       
  2955 apply(auto simp: ci.simps findnth.simps 
       
  2956                  tshift.simps tdec_b_def add3_Suc)
       
  2957 apply(subgoal_tac "2*n + 9 = Suc (2*n + 8)", simp only: fetch.simps)
       
  2958 apply(auto simp: nth_of.simps nth_append)
       
  2959 done
       
  2960 
       
  2961 lemma dec_fetch_on_right_move_o: 
       
  2962      "start_of ly as > 0 \<Longrightarrow> 
       
  2963      fetch (ci (layout_of aprog) 
       
  2964              (start_of ly as) (Dec n e)) (2 * n + 9) Oc
       
  2965     = (R, start_of ly as + 2*n + 8)"
       
  2966 apply(auto simp: ci.simps findnth.simps 
       
  2967                  tshift.simps tdec_b_def add3_Suc)
       
  2968 apply(subgoal_tac "2*n + 9 = Suc (2*n + 8)", simp only: fetch.simps)
       
  2969 apply(auto simp: nth_of.simps nth_append)
       
  2970 done
       
  2971 
       
  2972 lemma dec_fetch_after_clear_b: 
       
  2973      "start_of ly as > 0 \<Longrightarrow> 
       
  2974      fetch (ci (layout_of aprog) 
       
  2975             (start_of ly as) (Dec n e)) (2 * n + 10) Bk
       
  2976     = (R, start_of ly as + 2*n + 4)" 
       
  2977 apply(auto simp: ci.simps findnth.simps 
       
  2978                  tshift.simps tdec_b_def add3_Suc)
       
  2979 apply(subgoal_tac "2*n + 10 = Suc (2*n + 9)", simp only: fetch.simps)
       
  2980 apply(auto simp: nth_of.simps nth_append)
       
  2981 done
       
  2982 
       
  2983 lemma dec_fetch_after_clear_o: 
       
  2984      "start_of ly as > 0 \<Longrightarrow> 
       
  2985     fetch (ci (layout_of aprog) 
       
  2986              (start_of ly as) (Dec n e)) (2 * n + 10) Oc
       
  2987     = (W0, start_of ly as + 2*n + 9)"
       
  2988 apply(auto simp: ci.simps findnth.simps 
       
  2989                  tshift.simps tdec_b_def add3_Suc)
       
  2990 apply(subgoal_tac "2*n + 10= Suc (2*n + 9)", simp only: fetch.simps)
       
  2991 apply(auto simp: nth_of.simps nth_append)
       
  2992 done
       
  2993 
       
  2994 lemma dec_fetch_on_left_move1_o:
       
  2995       "start_of ly as > 0 \<Longrightarrow> 
       
  2996     fetch (ci (layout_of aprog) 
       
  2997            (start_of ly as) (Dec n e)) (2 * n + 11) Oc
       
  2998     = (L, start_of ly as + 2*n + 10)"
       
  2999 apply(auto simp: ci.simps findnth.simps 
       
  3000                  tshift.simps tdec_b_def add3_Suc)
       
  3001 apply(subgoal_tac "2*n + 11= Suc (2*n + 10)", simp only: fetch.simps)
       
  3002 apply(auto simp: nth_of.simps nth_append)
       
  3003 done
       
  3004 
       
  3005 lemma dec_fetch_on_left_move1_b:
       
  3006      "start_of ly as > 0 \<Longrightarrow> 
       
  3007     fetch (ci (layout_of aprog) 
       
  3008              (start_of ly as) (Dec n e)) (2 * n + 11) Bk
       
  3009     = (L, start_of ly as + 2*n + 11)"
       
  3010 apply(auto simp: ci.simps findnth.simps 
       
  3011                  tshift.simps tdec_b_def add3_Suc)
       
  3012 apply(subgoal_tac "2*n + 11 = Suc (2*n + 10)", 
       
  3013       simp only: fetch.simps)
       
  3014 apply(auto simp: nth_of.simps nth_append)
       
  3015 done
       
  3016 
       
  3017 lemma dec_fetch_check_left_move1_o: 
       
  3018     "start_of ly as > 0 \<Longrightarrow> 
       
  3019   fetch (ci (layout_of aprog) 
       
  3020              (start_of ly as) (Dec n e)) (2 * n + 12) Oc
       
  3021     = (L, start_of ly as + 2*n + 10)"
       
  3022 apply(auto simp: ci.simps findnth.simps 
       
  3023                  tshift.simps tdec_b_def add3_Suc)
       
  3024 apply(subgoal_tac "2*n + 12= Suc (2*n + 11)", simp only: fetch.simps)
       
  3025 apply(auto simp: nth_of.simps nth_append)
       
  3026 done
       
  3027 
       
  3028 lemma dec_fetch_check_left_move1_b: 
       
  3029     "start_of ly as > 0 \<Longrightarrow> 
       
  3030    fetch (ci (layout_of aprog) 
       
  3031                   (start_of ly as) (Dec n e)) (2 * n + 12) Bk
       
  3032     = (R, start_of ly as + 2*n + 12)"
       
  3033 apply(auto simp: ci.simps findnth.simps 
       
  3034                  tshift.simps tdec_b_def add3_Suc)
       
  3035 apply(subgoal_tac "2*n + 12 = Suc (2*n + 11)", 
       
  3036       simp only: fetch.simps)
       
  3037 apply(auto simp: nth_of.simps nth_append)
       
  3038 done
       
  3039 
       
  3040 lemma dec_fetch_after_left_move1_b: 
       
  3041   "start_of ly as > 0 \<Longrightarrow> 
       
  3042   fetch (ci (layout_of aprog) 
       
  3043                 (start_of ly as) (Dec n e)) (2 * n + 13) Bk
       
  3044     = (R, start_of ly as + 2*n + 16)"
       
  3045 apply(auto simp: ci.simps findnth.simps 
       
  3046                  tshift.simps tdec_b_def add3_Suc)
       
  3047 apply(subgoal_tac "2*n + 13 = Suc (2*n + 12)", 
       
  3048       simp only: fetch.simps)
       
  3049 apply(auto simp: nth_of.simps nth_append)
       
  3050 done
       
  3051 
       
  3052 lemma dec_fetch_on_left_move2_o:
       
  3053   "start_of ly as > 0 \<Longrightarrow> 
       
  3054   fetch (ci (layout_of aprog) 
       
  3055            (start_of ly as) (Dec n e)) (2 * n + 14) Oc
       
  3056    = (L, start_of ly as + 2*n + 13)"
       
  3057 apply(auto simp: ci.simps findnth.simps 
       
  3058                  tshift.simps tdec_b_def add3_Suc)
       
  3059 apply(subgoal_tac "2*n + 14 = Suc (2*n + 13)", 
       
  3060       simp only: fetch.simps)
       
  3061 apply(auto simp: nth_of.simps nth_append)
       
  3062 done
       
  3063 
       
  3064 lemma dec_fetch_on_left_move2_b:
       
  3065   "start_of ly as > 0 \<Longrightarrow> 
       
  3066   fetch (ci (layout_of aprog) 
       
  3067               (start_of ly as) (Dec n e)) (2 * n + 14) Bk
       
  3068  = (L, start_of ly as + 2*n + 14)"
       
  3069 apply(auto simp: ci.simps findnth.simps 
       
  3070                  tshift.simps tdec_b_def add3_Suc)
       
  3071 apply(subgoal_tac "2*n + 14 = Suc (2*n + 13)", 
       
  3072       simp only: fetch.simps)
       
  3073 apply(auto simp: nth_of.simps nth_append)
       
  3074 done
       
  3075 
       
  3076 lemma dec_fetch_check_left_move2_o:
       
  3077   "start_of ly as > 0 \<Longrightarrow> 
       
  3078   fetch (ci (layout_of aprog) 
       
  3079                 (start_of ly as) (Dec n e)) (2 * n + 15)  Oc
       
  3080  = (L, start_of ly as + 2*n + 13)"
       
  3081 apply(auto simp: ci.simps findnth.simps 
       
  3082                  tshift.simps tdec_b_def add3_Suc)
       
  3083 apply(subgoal_tac "2*n + 15 = Suc (2*n + 14)", 
       
  3084       simp only: fetch.simps)
       
  3085 apply(auto simp: nth_of.simps nth_append)
       
  3086 done
       
  3087 
       
  3088 lemma dec_fetch_check_left_move2_b:
       
  3089   "start_of ly as > 0 \<Longrightarrow> 
       
  3090   fetch (ci (layout_of aprog) 
       
  3091                 (start_of ly as) (Dec n e)) (2 * n + 15)  Bk
       
  3092  = (R, start_of ly as + 2*n + 15)"
       
  3093 apply(auto simp: ci.simps findnth.simps 
       
  3094                  tshift.simps tdec_b_def add3_Suc)
       
  3095 apply(subgoal_tac "2*n + 15= Suc (2*n + 14)", simp only: fetch.simps)
       
  3096 apply(auto simp: nth_of.simps nth_append)
       
  3097 done
       
  3098 
       
  3099 lemma dec_fetch_after_left_move2_b: 
       
  3100   "\<lbrakk>ly = layout_of aprog; 
       
  3101     abc_fetch as aprog = Some (Dec n e); 
       
  3102     start_of ly as > 0\<rbrakk> \<Longrightarrow> 
       
  3103      fetch (ci (layout_of aprog) (start_of ly as) 
       
  3104               (Dec n e)) (2 * n + 16)  Bk
       
  3105  = (R, start_of ly e)" 
       
  3106 apply(auto simp: ci.simps findnth.simps 
       
  3107                  tshift.simps tdec_b_def add3_Suc)
       
  3108 apply(subgoal_tac "2*n + 16 = Suc (2*n + 15)",
       
  3109       simp only: fetch.simps)
       
  3110 apply(auto simp: nth_of.simps nth_append)
       
  3111 done
       
  3112 
       
  3113 lemma dec_fetch_next_state: 
       
  3114     "start_of ly as > 0 \<Longrightarrow> 
       
  3115     fetch (ci (layout_of aprog) 
       
  3116            (start_of ly as) (Dec n e)) (2* n + 17)  b
       
  3117     = (Nop, 0)"
       
  3118 apply(case_tac b)
       
  3119 apply(auto simp: ci.simps findnth.simps 
       
  3120                  tshift.simps tdec_b_def add3_Suc)
       
  3121 apply(subgoal_tac [!] "2*n + 17 = Suc (2*n + 16)", 
       
  3122       simp_all only: fetch.simps)
       
  3123 apply(auto simp: nth_of.simps nth_append)
       
  3124 done
       
  3125 
       
  3126 (*End: dec_fetch lemmas*)
       
  3127 lemmas dec_fetch_simps = 
       
  3128  dec_fetch_locate_a_o dec_fetch_locate_a_b dec_fetch_locate_b_o 
       
  3129  dec_fetch_locate_b_b dec_fetch_locate_n_a_o 
       
  3130  dec_fetch_locate_n_a_b dec_fetch_locate_n_b_o 
       
  3131  dec_fetch_locate_n_b_b dec_fetch_first_on_right_move_o 
       
  3132  dec_fetch_first_on_right_move_b dec_fetch_first_after_clear_b
       
  3133  dec_fetch_first_after_clear_o dec_fetch_right_move_b 
       
  3134  dec_fetch_on_right_move_b dec_fetch_on_right_move_o 
       
  3135  dec_fetch_after_clear_b dec_fetch_after_clear_o
       
  3136  dec_fetch_check_right_move_b dec_fetch_check_right_move_o 
       
  3137  dec_fetch_left_move_b dec_fetch_on_left_move1_b 
       
  3138  dec_fetch_on_left_move1_o dec_fetch_check_left_move1_b 
       
  3139  dec_fetch_check_left_move1_o dec_fetch_after_left_move1_b 
       
  3140  dec_fetch_on_left_move2_b dec_fetch_on_left_move2_o
       
  3141  dec_fetch_check_left_move2_o dec_fetch_check_left_move2_b 
       
  3142  dec_fetch_after_left_move2_b dec_fetch_after_write_b 
       
  3143  dec_fetch_after_write_o dec_fetch_next_state
       
  3144 
       
  3145 lemma [simp]:
       
  3146   "\<lbrakk>start_of ly as \<le> a; 
       
  3147     a < start_of ly as + 2 * n; 
       
  3148     (a - start_of ly as) mod 2 = Suc 0; 
       
  3149     inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, Bk # xs) ires\<rbrakk>
       
  3150      \<Longrightarrow> \<not> Suc a < start_of ly as + 2 * n \<longrightarrow> 
       
  3151                 inv_locate_a (as, am) (n, Bk # aaa, xs) ires"
       
  3152 apply(insert locate_b_2_locate_a[of a ly as n am aaa xs], simp)
       
  3153 done
       
  3154  
       
  3155 lemma [simp]: 
       
  3156   "\<lbrakk>start_of ly as \<le> a; 
       
  3157     a < start_of ly as + 2 * n; 
       
  3158     (a - start_of ly as) mod 2 = Suc 0; 
       
  3159     inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, []) ires\<rbrakk>
       
  3160    \<Longrightarrow> \<not> Suc a < start_of ly as + 2 * n \<longrightarrow> 
       
  3161                   inv_locate_a (as, am) (n, Bk # aaa, []) ires"
       
  3162 apply(insert locate_b_2_locate_a_B[of a ly as n am aaa], simp)
       
  3163 done
       
  3164 
       
  3165 (*
       
  3166 lemma [simp]: "a\<^bsup>0\<^esup>=[]"
       
  3167 apply(simp add: exponent_def)
       
  3168 done
       
  3169 *)
       
  3170 
       
  3171 lemma exp_ind: "a\<^bsup>Suc b\<^esup> =  a\<^bsup>b\<^esup> @ [a]"
       
  3172 apply(simp only: exponent_def rep_ind)
       
  3173 done
       
  3174 
       
  3175 lemma [simp]:
       
  3176   "inv_locate_b (as, am) (n, l, Oc # r) ires
       
  3177   \<Longrightarrow> dec_first_on_right_moving n (as,  abc_lm_s am n (abc_lm_v am n))
       
  3178                       (Suc (Suc (start_of ly as + 2 * n)), Oc # l, r) ires"
       
  3179 apply(simp only: inv_locate_b.simps 
       
  3180      dec_first_on_right_moving.simps in_middle.simps 
       
  3181      abc_lm_s.simps abc_lm_v.simps)
       
  3182 apply(erule_tac exE)+
       
  3183 apply(erule conjE)+
       
  3184 apply(case_tac "n < length am", simp)
       
  3185 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
       
  3186       rule_tac x = m in exI, simp)
       
  3187 apply(rule_tac x = "Suc ml" in exI, rule_tac conjI, rule_tac [1-2] impI)
       
  3188 prefer 3
       
  3189 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
       
  3190       rule_tac x = m in exI, simp)
       
  3191 apply(subgoal_tac "Suc n - length am = Suc (n - length am)",
       
  3192       simp only:exponent_def rep_ind, simp)
       
  3193 apply(rule_tac x = "Suc ml" in exI, simp_all)
       
  3194 apply(rule_tac [!] x = "mr - 1" in exI, simp_all)
       
  3195 apply(case_tac [!] mr, auto)
       
  3196 done
       
  3197 
       
  3198 lemma [simp]: 
       
  3199   "\<lbrakk>inv_locate_b (as, am) (n, l, r) ires; l \<noteq> []\<rbrakk> \<Longrightarrow> 
       
  3200   \<not> inv_on_left_moving_in_middle_B (as, abc_lm_s am n (abc_lm_v am n)) 
       
  3201     (s, tl l, hd l # r) ires"
       
  3202 apply(auto simp: inv_locate_b.simps 
       
  3203                  inv_on_left_moving_in_middle_B.simps in_middle.simps)
       
  3204 apply(case_tac [!] ml, auto split: if_splits)
       
  3205 done
       
  3206 
       
  3207 lemma [simp]: "inv_locate_b (as, am) (n, l, r) ires \<Longrightarrow> l \<noteq> []"
       
  3208 apply(auto simp: inv_locate_b.simps in_middle.simps split: if_splits)
       
  3209 done
       
  3210 
       
  3211 lemma [simp]: "\<lbrakk>inv_locate_b (as, am) (n, l, Bk # r) ires; n < length am\<rbrakk>
       
  3212      \<Longrightarrow> inv_on_left_moving_norm (as, am) (s, tl l, hd l # Bk # r) ires"
       
  3213 apply(simp only: inv_locate_b.simps inv_on_left_moving_norm.simps 
       
  3214                  in_middle.simps)
       
  3215 apply(erule_tac exE)+
       
  3216 apply(erule_tac conjE)+
       
  3217 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
       
  3218       rule_tac x = m in exI, simp)
       
  3219 apply(rule_tac x = "ml - 1" in exI, auto)
       
  3220 apply(rule_tac [!] x = "Suc mr" in exI)
       
  3221 apply(case_tac [!] mr, auto)
       
  3222 done
       
  3223 
       
  3224 lemma [simp]: "\<lbrakk>inv_locate_b (as, am) (n, l, Bk # r) ires; \<not> n < length am\<rbrakk>
       
  3225     \<Longrightarrow> inv_on_left_moving_norm (as, am @ 
       
  3226         replicate (n - length am) 0 @ [0]) (s, tl l, hd l # Bk # r) ires"
       
  3227 apply(simp only: inv_locate_b.simps inv_on_left_moving_norm.simps 
       
  3228                  in_middle.simps)
       
  3229 apply(erule_tac exE)+
       
  3230 apply(erule_tac conjE)+
       
  3231 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
       
  3232       rule_tac x = m in exI, simp)
       
  3233 apply(subgoal_tac "Suc n - length am = Suc (n - length am)", simp only: rep_ind exponent_def, simp_all)
       
  3234 apply(rule_tac x = "Suc mr" in exI, auto)
       
  3235 done
       
  3236 
       
  3237 lemma inv_locate_b_2_on_left_moving[simp]: 
       
  3238   "\<lbrakk>inv_locate_b (as, am) (n, l, Bk # r) ires\<rbrakk>
       
  3239    \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as,  
       
  3240             abc_lm_s am n (abc_lm_v am n)) (s, [], Bk # Bk # r) ires) \<and>
       
  3241        (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as,  
       
  3242             abc_lm_s am n (abc_lm_v am n)) (s, tl l, hd l # Bk # r) ires)"
       
  3243 apply(subgoal_tac "l\<noteq>[]")
       
  3244 apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B 
       
  3245       (as,  abc_lm_s am n (abc_lm_v am n)) (s, tl l, hd l # Bk # r) ires")
       
  3246 apply(simp add:inv_on_left_moving.simps 
       
  3247           abc_lm_s.simps abc_lm_v.simps split: if_splits, auto)
       
  3248 done
       
  3249 
       
  3250 lemma [simp]: 
       
  3251   "inv_locate_b (as, am) (n, l, []) ires \<Longrightarrow> 
       
  3252                    inv_locate_b (as, am) (n, l, [Bk]) ires" 
       
  3253 apply(auto simp: inv_locate_b.simps in_middle.simps)
       
  3254 apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI,
       
  3255       rule_tac x = "Suc (length lm1) - length am" in exI, 
       
  3256       rule_tac x = m in exI, simp)
       
  3257 apply(rule_tac x = ml in exI, rule_tac x = mr in exI)
       
  3258 apply(auto)
       
  3259 done
       
  3260 
       
  3261 lemma nil_2_nil: "<lm::nat list> = [] \<Longrightarrow> lm = []"
       
  3262 apply(auto simp: tape_of_nl_abv)
       
  3263 apply(case_tac lm, simp)
       
  3264 apply(case_tac list, auto simp: tape_of_nat_list.simps)
       
  3265 done
       
  3266 
       
  3267 lemma  inv_locate_b_2_on_left_moving_b[simp]: 
       
  3268    "inv_locate_b (as, am) (n, l, []) ires
       
  3269      \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, 
       
  3270                   abc_lm_s am n (abc_lm_v am n)) (s, [], [Bk]) ires) \<and>
       
  3271          (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, abc_lm_s am n 
       
  3272                   (abc_lm_v am n)) (s, tl l, [hd l]) ires)"
       
  3273 apply(insert inv_locate_b_2_on_left_moving[of as am n l "[]" ires s])
       
  3274 apply(simp only: inv_on_left_moving.simps, simp)
       
  3275 apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B 
       
  3276          (as, abc_lm_s am n (abc_lm_v am n)) (s, tl l, [hd l]) ires", simp)
       
  3277 apply(simp only: inv_on_left_moving_norm.simps)
       
  3278 apply(erule_tac exE)+
       
  3279 apply(erule_tac conjE)+
       
  3280 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
       
  3281       rule_tac x = m in exI, rule_tac x = ml in exI, 
       
  3282       rule_tac x = mr in exI, simp)
       
  3283 apply(case_tac mr, simp, simp, case_tac nat, auto intro: nil_2_nil)
       
  3284 done
       
  3285 
       
  3286 lemma [simp]: 
       
  3287  "\<lbrakk>dec_first_on_right_moving n (as, am) (s, aaa, Oc # xs) ires\<rbrakk>
       
  3288    \<Longrightarrow> dec_first_on_right_moving n (as, am) (s', Oc # aaa, xs) ires"
       
  3289 apply(simp only: dec_first_on_right_moving.simps)
       
  3290 apply(erule exE)+
       
  3291 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
       
  3292       rule_tac x = m in exI, simp)
       
  3293 apply(rule_tac x = "Suc ml" in exI, 
       
  3294       rule_tac x = "mr - 1" in exI, auto)
       
  3295 apply(case_tac [!] mr, auto)
       
  3296 done
       
  3297 
       
  3298 lemma [simp]: 
       
  3299   "dec_first_on_right_moving n (as, am) (s, l, Bk # xs) ires \<Longrightarrow> l \<noteq> []"
       
  3300 apply(auto simp: dec_first_on_right_moving.simps split: if_splits)
       
  3301 done
       
  3302 
       
  3303 lemma [elim]: 
       
  3304   "\<lbrakk>\<not> length lm1 < length am; 
       
  3305     am @ replicate (length lm1 - length am) 0 @ [0::nat] = 
       
  3306                                                 lm1 @ m # lm2;
       
  3307     0 < m\<rbrakk>
       
  3308    \<Longrightarrow> RR"
       
  3309 apply(subgoal_tac "lm2 = []", simp)
       
  3310 apply(drule_tac length_equal, simp)
       
  3311 done
       
  3312 
       
  3313 lemma [simp]: 
       
  3314  "\<lbrakk>dec_first_on_right_moving n (as, 
       
  3315                    abc_lm_s am n (abc_lm_v am n)) (s, l, Bk # xs) ires\<rbrakk>
       
  3316 \<Longrightarrow> dec_after_clear (as, abc_lm_s am n 
       
  3317                  (abc_lm_v am n - Suc 0)) (s', tl l, hd l # Bk # xs) ires"
       
  3318 apply(simp only: dec_first_on_right_moving.simps 
       
  3319                  dec_after_clear.simps abc_lm_s.simps abc_lm_v.simps)
       
  3320 apply(erule_tac exE)+
       
  3321 apply(case_tac "n < length am")
       
  3322 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
       
  3323       rule_tac x = "m - 1" in exI, auto simp: )
       
  3324 apply(case_tac [!] mr, auto)
       
  3325 done
       
  3326 
       
  3327 lemma [simp]: 
       
  3328  "\<lbrakk>dec_first_on_right_moving n (as, 
       
  3329                    abc_lm_s am n (abc_lm_v am n)) (s, l, []) ires\<rbrakk>
       
  3330 \<Longrightarrow> (l = [] \<longrightarrow> dec_after_clear (as, 
       
  3331              abc_lm_s am n (abc_lm_v am n - Suc 0)) (s', [], [Bk]) ires) \<and>
       
  3332     (l \<noteq> [] \<longrightarrow> dec_after_clear (as, abc_lm_s am n 
       
  3333                       (abc_lm_v am n - Suc 0)) (s', tl l, [hd l]) ires)"
       
  3334 apply(subgoal_tac "l \<noteq> []", 
       
  3335       simp only: dec_first_on_right_moving.simps 
       
  3336                  dec_after_clear.simps abc_lm_s.simps abc_lm_v.simps)
       
  3337 apply(erule_tac exE)+
       
  3338 apply(case_tac "n < length am", simp)
       
  3339 apply(rule_tac x = lm1 in exI, rule_tac x = "m - 1" in exI, auto)
       
  3340 apply(case_tac [1-2] mr, auto)
       
  3341 apply(case_tac [1-2] m, auto simp: dec_first_on_right_moving.simps split: if_splits)
       
  3342 done
       
  3343 
       
  3344 lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, Oc # r) ires\<rbrakk>
       
  3345                 \<Longrightarrow> dec_after_clear (as, am) (s', l, Bk # r) ires"
       
  3346 apply(auto simp: dec_after_clear.simps)
       
  3347 done
       
  3348 
       
  3349 lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, Bk # r) ires\<rbrakk>
       
  3350                 \<Longrightarrow> dec_right_move (as, am) (s', Bk # l, r) ires"
       
  3351 apply(auto simp: dec_after_clear.simps dec_right_move.simps split: if_splits)
       
  3352 done
       
  3353 
       
  3354 lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, []) ires\<rbrakk>
       
  3355              \<Longrightarrow> dec_right_move (as, am) (s', Bk # l, []) ires"
       
  3356 apply(auto simp: dec_after_clear.simps dec_right_move.simps )
       
  3357 done
       
  3358 
       
  3359 lemma [simp]: "\<exists>rn. a::block\<^bsup>rn\<^esup> = []"
       
  3360 apply(rule_tac x = 0 in exI, simp)
       
  3361 done
       
  3362 
       
  3363 lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, []) ires\<rbrakk>
       
  3364              \<Longrightarrow> dec_right_move (as, am) (s', Bk # l, [Bk]) ires"
       
  3365 apply(auto simp: dec_after_clear.simps dec_right_move.simps split: if_splits)
       
  3366 done
       
  3367 
       
  3368 lemma [simp]:"dec_right_move (as, am) (s, l, Oc # r) ires = False"
       
  3369 apply(auto simp: dec_right_move.simps)
       
  3370 done
       
  3371               
       
  3372 lemma dec_right_move_2_check_right_move[simp]:
       
  3373      "\<lbrakk>dec_right_move (as, am) (s, l, Bk # r) ires\<rbrakk>
       
  3374       \<Longrightarrow> dec_check_right_move (as, am) (s', Bk # l, r) ires"
       
  3375 apply(auto simp: dec_right_move.simps dec_check_right_move.simps split: if_splits)
       
  3376 done
       
  3377 
       
  3378 lemma [simp]: 
       
  3379  "dec_right_move (as, am) (s, l, []) ires= 
       
  3380   dec_right_move (as, am) (s, l, [Bk]) ires"
       
  3381 apply(simp add: dec_right_move.simps)
       
  3382 apply(rule_tac iffI)
       
  3383 apply(erule_tac [!] exE)+
       
  3384 apply(erule_tac [2] exE)
       
  3385 apply(rule_tac [!] x = lm1 in exI, rule_tac x = "[]" in exI, 
       
  3386       rule_tac [!] x = m in exI, auto)
       
  3387 apply(auto intro: nil_2_nil)
       
  3388 done
       
  3389 
       
  3390 lemma [simp]: "\<lbrakk>dec_right_move (as, am) (s, l, []) ires\<rbrakk>
       
  3391              \<Longrightarrow> dec_check_right_move (as, am) (s, Bk # l, []) ires"
       
  3392 apply(insert dec_right_move_2_check_right_move[of as am s l "[]" s'], 
       
  3393       simp)
       
  3394 done
       
  3395 
       
  3396 lemma [simp]: "dec_check_right_move (as, am) (s, l, r) ires\<Longrightarrow> l \<noteq> []"
       
  3397 apply(auto simp: dec_check_right_move.simps split: if_splits)
       
  3398 done
       
  3399  
       
  3400 lemma [simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, Oc # r) ires\<rbrakk>
       
  3401              \<Longrightarrow> dec_after_write (as, am) (s', tl l, hd l # Oc # r) ires"
       
  3402 apply(auto simp: dec_check_right_move.simps dec_after_write.simps)
       
  3403 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
       
  3404       rule_tac x = m in exI, auto)
       
  3405 done
       
  3406 
       
  3407 lemma [simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, Bk # r) ires\<rbrakk>
       
  3408                 \<Longrightarrow> dec_left_move (as, am) (s', tl l, hd l # Bk # r) ires"
       
  3409 apply(auto simp: dec_check_right_move.simps 
       
  3410                  dec_left_move.simps inv_after_move.simps)
       
  3411 apply(rule_tac x = lm1 in exI, rule_tac x = m in exI, auto)
       
  3412 apply(auto intro: BkCons_nil nil_2_nil dest: BkCons_nil)
       
  3413 apply(rule_tac x = "Suc rn" in exI)
       
  3414 apply(auto intro: BkCons_nil nil_2_nil dest: BkCons_nil)
       
  3415 done
       
  3416 
       
  3417 lemma [simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, []) ires\<rbrakk>
       
  3418              \<Longrightarrow> dec_left_move (as, am) (s', tl l, [hd l]) ires"
       
  3419 apply(auto simp: dec_check_right_move.simps 
       
  3420                  dec_left_move.simps inv_after_move.simps)
       
  3421 apply(rule_tac x = lm1 in exI, rule_tac x = m in exI, auto)
       
  3422 apply(auto intro: BkCons_nil nil_2_nil dest: BkCons_nil)
       
  3423 done
       
  3424 
       
  3425 lemma [simp]: "dec_left_move (as, am) (s, aaa, Oc # xs) ires = False"
       
  3426 apply(auto simp: dec_left_move.simps inv_after_move.simps)
       
  3427 apply(case_tac [!] rn, auto)
       
  3428 done
       
  3429 
       
  3430 lemma [simp]: "dec_left_move (as, am) (s, l, r) ires
       
  3431              \<Longrightarrow> l \<noteq> []"
       
  3432 apply(auto simp: dec_left_move.simps split: if_splits)
       
  3433 done
       
  3434 
       
  3435 lemma tape_of_nl_abv_cons_ex[simp]: 
       
  3436    "\<exists>lna. Oc # Oc\<^bsup>m\<^esup> @ Bk # <rev lm1> @ Bk\<^bsup>ln\<^esup> = <m # rev lm1> @ Bk\<^bsup>lna\<^esup>"
       
  3437 apply(case_tac "lm1=[]", auto simp: tape_of_nl_abv 
       
  3438                                     tape_of_nat_list.simps)
       
  3439 apply(rule_tac x = "ln" in exI, simp)
       
  3440 apply(simp add:  tape_of_nat_list_cons exponent_def)
       
  3441 done
       
  3442 
       
  3443 (*
       
  3444 lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm1 @ [m])
       
  3445                  (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # <rev lm1> @ Bk\<^bsup>ln\<^esup>, Bk # Bk\<^bsup>rn\<^esup>) ires"
       
  3446 apply(simp only: inv_on_left_moving_in_middle_B.simps)
       
  3447 apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "[]" in exI, auto)
       
  3448 done    
       
  3449 *)
       
  3450 lemma [simp]: "inv_on_left_moving_in_middle_B (as, [m])
       
  3451   (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # Bk # ires, Bk # Bk\<^bsup>rn\<^esup>) ires"
       
  3452 apply(simp add: inv_on_left_moving_in_middle_B.simps)
       
  3453 apply(rule_tac x = "[m]" in exI, simp, auto simp: tape_of_nat_def)
       
  3454 done
       
  3455 
       
  3456 lemma [simp]: "inv_on_left_moving_in_middle_B (as, [m])
       
  3457   (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # Bk # ires, [Bk]) ires"
       
  3458 apply(simp add: inv_on_left_moving_in_middle_B.simps)
       
  3459 apply(rule_tac x = "[m]" in exI, simp, auto simp: tape_of_nat_def)
       
  3460 done
       
  3461 
       
  3462 lemma [simp]: "lm1 \<noteq> [] \<Longrightarrow> 
       
  3463   inv_on_left_moving_in_middle_B (as, lm1 @ [m]) (s', 
       
  3464   Oc # Oc\<^bsup>m\<^esup> @ Bk # <rev lm1> @ Bk # Bk # ires, Bk # Bk\<^bsup>rn\<^esup>) ires"
       
  3465 apply(simp only: inv_on_left_moving_in_middle_B.simps)
       
  3466 apply(rule_tac x = "lm1 @ [m ]" in exI, rule_tac x = "[]" in exI, simp, auto)
       
  3467 done
       
  3468 
       
  3469 lemma [simp]: "lm1 \<noteq> [] \<Longrightarrow> 
       
  3470   inv_on_left_moving_in_middle_B (as, lm1 @ [m]) (s', 
       
  3471   Oc # Oc\<^bsup>m\<^esup> @ Bk # <rev lm1> @ Bk # Bk # ires, [Bk]) ires"
       
  3472 apply(simp only: inv_on_left_moving_in_middle_B.simps)
       
  3473 apply(rule_tac x = "lm1 @ [m ]" in exI, rule_tac x = "[]" in exI, simp, auto)
       
  3474 done
       
  3475 
       
  3476 lemma [simp]: "dec_left_move (as, am) (s, l, Bk # r) ires
       
  3477        \<Longrightarrow> inv_on_left_moving (as, am) (s', tl l, hd l # Bk # r) ires"
       
  3478 apply(auto simp: dec_left_move.simps inv_on_left_moving.simps split: if_splits)
       
  3479 done
       
  3480 
       
  3481 (*
       
  3482 lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm1 @ [m]) 
       
  3483                         (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # <rev lm1> @ Bk\<^bsup>ln\<^esup>, [Bk])  ires"
       
  3484 apply(auto simp: inv_on_left_moving_in_middle_B.simps)
       
  3485 apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "[]" in exI, auto)
       
  3486 done
       
  3487 *)
       
  3488 
       
  3489 lemma [simp]: "dec_left_move (as, am) (s, l, []) ires
       
  3490              \<Longrightarrow> inv_on_left_moving (as, am) (s', tl l, [hd l]) ires"
       
  3491 apply(auto simp: dec_left_move.simps inv_on_left_moving.simps split: if_splits)
       
  3492 done
       
  3493 
       
  3494 lemma [simp]: "dec_after_write (as, am) (s, l, Oc # r) ires
       
  3495        \<Longrightarrow> dec_on_right_moving (as, am) (s', Oc # l, r) ires"
       
  3496 apply(auto simp: dec_after_write.simps dec_on_right_moving.simps)
       
  3497 apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI, 
       
  3498       rule_tac x = "hd lm2" in exI, simp)
       
  3499 apply(rule_tac x = "Suc 0" in exI,rule_tac x =  "Suc (hd lm2)" in exI)
       
  3500 apply(case_tac lm2, simp, simp)
       
  3501 apply(case_tac "list = []", 
       
  3502       auto simp: tape_of_nl_abv tape_of_nat_list.simps split: if_splits )
       
  3503 apply(case_tac rn, auto)
       
  3504 apply(case_tac "rev lm1", simp, simp add: tape_of_nat_list.simps)
       
  3505 apply(case_tac rn, auto)
       
  3506 apply(case_tac list, simp_all add: tape_of_nat_list.simps, auto)
       
  3507 apply(case_tac "rev lm1", simp, simp add: tape_of_nat_list.simps)
       
  3508 apply(case_tac list, simp_all add: tape_of_nat_list.simps, auto)
       
  3509 done
       
  3510 
       
  3511 lemma [simp]: "dec_after_write (as, am) (s, l, Bk # r) ires
       
  3512        \<Longrightarrow> dec_after_write (as, am) (s', l, Oc # r) ires"
       
  3513 apply(auto simp: dec_after_write.simps)
       
  3514 done
       
  3515 
       
  3516 lemma [simp]: "dec_after_write (as, am) (s, aaa, []) ires
       
  3517              \<Longrightarrow> dec_after_write (as, am) (s', aaa, [Oc]) ires"
       
  3518 apply(auto simp: dec_after_write.simps)
       
  3519 done
       
  3520 
       
  3521 lemma [simp]: "dec_on_right_moving (as, am) (s, l, Oc # r) ires
       
  3522        \<Longrightarrow> dec_on_right_moving (as, am) (s', Oc # l, r) ires"
       
  3523 apply(simp only: dec_on_right_moving.simps)
       
  3524 apply(erule_tac exE)+
       
  3525 apply(erule conjE)+
       
  3526 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,
       
  3527       rule_tac x = "m" in exI, rule_tac x = "Suc ml" in exI, 
       
  3528       rule_tac x = "mr - 1" in exI, simp)
       
  3529 apply(case_tac mr, auto)
       
  3530 done
       
  3531 
       
  3532 lemma [simp]: "dec_on_right_moving (as, am) (s, l, r) ires\<Longrightarrow>  l \<noteq> []"
       
  3533 apply(auto simp: dec_on_right_moving.simps split: if_splits)
       
  3534 done
       
  3535 
       
  3536 lemma [simp]: "dec_on_right_moving (as, am) (s, l, Bk # r) ires
       
  3537       \<Longrightarrow>  dec_after_clear (as, am) (s', tl l, hd l # Bk # r) ires"
       
  3538 apply(auto simp: dec_on_right_moving.simps dec_after_clear.simps)
       
  3539 apply(case_tac [!] mr, auto split: if_splits)
       
  3540 done
       
  3541 
       
  3542 lemma [simp]: "dec_on_right_moving (as, am) (s, l, []) ires
       
  3543              \<Longrightarrow> dec_after_clear (as, am) (s', tl l, [hd l]) ires"
       
  3544 apply(auto simp: dec_on_right_moving.simps dec_after_clear.simps)
       
  3545 apply(case_tac mr, simp_all split: if_splits)
       
  3546 apply(rule_tac x = lm1 in exI, simp)
       
  3547 done
       
  3548 
       
  3549 lemma start_of_le: "a < b \<Longrightarrow> start_of ly a \<le> start_of ly b"
       
  3550 proof(induct b arbitrary: a, simp, case_tac "a = b", simp)
       
  3551   fix b a
       
  3552   show "start_of ly b \<le> start_of ly (Suc b)"
       
  3553     apply(case_tac "b::nat", 
       
  3554           simp add: start_of.simps, simp add: start_of.simps)
       
  3555     done
       
  3556 next
       
  3557   fix b a
       
  3558   assume h1: "\<And>a. a < b \<Longrightarrow> start_of ly a \<le> start_of ly b" 
       
  3559              "a < Suc b" "a \<noteq> b"
       
  3560   hence "a < b"
       
  3561     by(simp)
       
  3562   from h1 and this have h2: "start_of ly a \<le> start_of ly b"
       
  3563     by(drule_tac h1, simp)
       
  3564   from h2 show "start_of ly a \<le> start_of ly (Suc b)"
       
  3565   proof -
       
  3566     have "start_of ly b \<le> start_of ly (Suc b)"
       
  3567       apply(case_tac "b::nat", 
       
  3568             simp add: start_of.simps, simp add: start_of.simps)
       
  3569       done
       
  3570     from h2 and this show "start_of ly a \<le> start_of ly (Suc b)"
       
  3571       by simp
       
  3572   qed
       
  3573 qed
       
  3574 
       
  3575 lemma start_of_dec_length[simp]: 
       
  3576   "\<lbrakk>abc_fetch a aprog = Some (Dec n e)\<rbrakk> \<Longrightarrow> 
       
  3577     start_of (layout_of aprog) (Suc a)
       
  3578           = start_of (layout_of aprog) a + 2*n + 16"
       
  3579 apply(case_tac a, auto simp: abc_fetch.simps start_of.simps 
       
  3580                              layout_of.simps length_of.simps 
       
  3581                        split: if_splits)
       
  3582 done
       
  3583 
       
  3584 lemma start_of_ge: 
       
  3585  "\<lbrakk>abc_fetch a aprog = Some (Dec n e); a < e\<rbrakk> \<Longrightarrow>
       
  3586   start_of (layout_of aprog) e > 
       
  3587               start_of (layout_of aprog) a + 2*n + 15"
       
  3588 apply(case_tac "e = Suc a", 
       
  3589       simp add: start_of.simps abc_fetch.simps layout_of.simps 
       
  3590                 length_of.simps split: if_splits)
       
  3591 apply(subgoal_tac "Suc a < e", drule_tac a = "Suc a" 
       
  3592              and ly = "layout_of aprog" in start_of_le)
       
  3593 apply(subgoal_tac "start_of (layout_of aprog) (Suc a)
       
  3594          = start_of (layout_of aprog) a + 2*n + 16", simp)
       
  3595 apply(rule_tac start_of_dec_length, simp)
       
  3596 apply(arith)
       
  3597 done
       
  3598 
       
  3599 lemma starte_not_equal[simp]: 
       
  3600  "\<lbrakk>abc_fetch as aprog = Some (Dec n e); ly = layout_of aprog\<rbrakk>
       
  3601    \<Longrightarrow> (start_of ly e \<noteq> Suc (Suc (start_of ly as + 2 * n)) \<and>  
       
  3602         start_of ly e \<noteq> start_of ly as + 2 * n + 3 \<and> 
       
  3603         start_of ly e \<noteq> start_of ly as + 2 * n + 4 \<and>
       
  3604         start_of ly e \<noteq> start_of ly as + 2 * n + 5 \<and> 
       
  3605         start_of ly e \<noteq> start_of ly as + 2 * n + 6 \<and> 
       
  3606         start_of ly e \<noteq> start_of ly as + 2 * n + 7 \<and>
       
  3607         start_of ly e \<noteq> start_of ly as + 2 * n + 8 \<and> 
       
  3608         start_of ly e \<noteq> start_of ly as + 2 * n + 9 \<and> 
       
  3609         start_of ly e \<noteq> start_of ly as + 2 * n + 10 \<and>
       
  3610         start_of ly e \<noteq> start_of ly as + 2 * n + 11 \<and> 
       
  3611         start_of ly e \<noteq> start_of ly as + 2 * n + 12 \<and> 
       
  3612         start_of ly e \<noteq> start_of ly as + 2 * n + 13 \<and>
       
  3613         start_of ly e \<noteq> start_of ly as + 2 * n + 14 \<and> 
       
  3614         start_of ly e \<noteq> start_of ly as + 2 * n + 15)" 
       
  3615 apply(case_tac "e = as", simp) 
       
  3616 apply(case_tac "e < as")
       
  3617 apply(drule_tac a = e and b = as and ly = ly in start_of_le, simp)
       
  3618 apply(drule_tac a = as and e = e in start_of_ge, simp, simp)
       
  3619 done
       
  3620 
       
  3621 lemma [simp]: "\<lbrakk>abc_fetch as aprog = Some (Dec n e); ly = layout_of aprog\<rbrakk>
       
  3622       \<Longrightarrow> (Suc (Suc (start_of ly as + 2 * n)) \<noteq> start_of ly e \<and> 
       
  3623           start_of ly as + 2 * n + 3 \<noteq> start_of ly e \<and> 
       
  3624           start_of ly as + 2 * n + 4 \<noteq> start_of ly e \<and>
       
  3625           start_of ly as + 2 * n + 5 \<noteq>start_of ly e \<and> 
       
  3626           start_of ly as + 2 * n + 6 \<noteq> start_of ly e \<and>
       
  3627           start_of ly as + 2 * n + 7 \<noteq> start_of ly e \<and> 
       
  3628           start_of ly as + 2 * n + 8 \<noteq> start_of ly e \<and> 
       
  3629           start_of ly as + 2 * n + 9 \<noteq> start_of ly e \<and> 
       
  3630           start_of ly as + 2 * n + 10 \<noteq> start_of ly e \<and>
       
  3631           start_of ly as + 2 * n + 11 \<noteq> start_of ly e \<and> 
       
  3632           start_of ly as + 2 * n + 12 \<noteq> start_of ly e \<and> 
       
  3633           start_of ly as + 2 * n + 13 \<noteq> start_of ly e \<and> 
       
  3634           start_of ly as + 2 * n + 14 \<noteq> start_of ly e \<and> 
       
  3635           start_of ly as + 2 * n + 15 \<noteq> start_of ly e)"
       
  3636 apply(insert starte_not_equal[of as aprog n e ly], 
       
  3637                             simp del: starte_not_equal)
       
  3638 apply(erule_tac conjE)+
       
  3639 apply(rule_tac conjI, simp del: starte_not_equal)+
       
  3640 apply(rule not_sym, simp)
       
  3641 done
       
  3642 
       
  3643 lemma [simp]: "start_of (layout_of aprog) as > 0 \<Longrightarrow> 
       
  3644   fetch (ci (layout_of aprog) (start_of (layout_of aprog) as)
       
  3645                        (Dec n as)) (Suc 0) Oc =
       
  3646  (R, Suc (start_of (layout_of aprog) as))"
       
  3647 
       
  3648 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  3649                  nth_of.simps tshift.simps nth_append 
       
  3650                  Suc_pre tdec_b_def)
       
  3651 apply(insert findnth_nth[of 0 n "Suc 0"], simp)
       
  3652 apply(simp add: findnth.simps)
       
  3653 done
       
  3654 
       
  3655 lemma start_of_inj[simp]: 
       
  3656   "\<lbrakk>abc_fetch as aprog = Some (Dec n e); e \<noteq> as; ly = layout_of aprog\<rbrakk>
       
  3657    \<Longrightarrow> start_of ly as \<noteq> start_of ly e"
       
  3658 apply(case_tac "e < as")
       
  3659 apply(case_tac "as", simp, simp)
       
  3660 apply(case_tac "e = nat", simp add: start_of.simps 
       
  3661                                     layout_of.simps length_of.simps)
       
  3662 apply(subgoal_tac "e < length aprog", simp add: length_of.simps 
       
  3663                                          split: abc_inst.splits)
       
  3664 apply(simp add: abc_fetch.simps split: if_splits)
       
  3665 apply(subgoal_tac "e < nat", drule_tac a = e and b = nat 
       
  3666                                    and ly =ly in start_of_le, simp)
       
  3667 apply(subgoal_tac "start_of ly nat < start_of ly (Suc nat)", 
       
  3668           simp, simp add: start_of.simps layout_of.simps)
       
  3669 apply(subgoal_tac "nat < length aprog", simp)
       
  3670 apply(case_tac "aprog ! nat", auto simp: length_of.simps)
       
  3671 apply(simp add: abc_fetch.simps split: if_splits)
       
  3672 apply(subgoal_tac "e > as", drule_tac start_of_ge, auto)
       
  3673 done
       
  3674 
       
  3675 lemma [simp]: "\<lbrakk>abc_fetch as aprog = Some (Dec n e); e < as\<rbrakk>
       
  3676     \<Longrightarrow> Suc (start_of (layout_of aprog) e) - 
       
  3677                                start_of (layout_of aprog) as = 0"
       
  3678 apply(frule_tac ly = "layout_of aprog" in start_of_le, simp)
       
  3679 apply(subgoal_tac "start_of (layout_of aprog) as \<noteq> 
       
  3680                             start_of (layout_of aprog) e", arith)
       
  3681 apply(rule start_of_inj, auto)
       
  3682 done
       
  3683 
       
  3684 lemma [simp]:
       
  3685    "\<lbrakk>abc_fetch as aprog = Some (Dec n e); 
       
  3686      0 < start_of (layout_of aprog) as\<rbrakk>
       
  3687  \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  3688      (Dec n e)) (Suc (start_of (layout_of aprog) e) - 
       
  3689                  start_of (layout_of aprog) as) Oc)
       
  3690     = (if e = as then (R, start_of (layout_of aprog) as + 1)
       
  3691                  else (Nop, 0))"
       
  3692 apply(auto split: if_splits)
       
  3693 apply(case_tac "e < as", simp add: fetch.simps)
       
  3694 apply(subgoal_tac " e > as")
       
  3695 apply(drule start_of_ge, simp,
       
  3696       auto simp: fetch.simps ci_length nth_of.simps)
       
  3697 apply(subgoal_tac 
       
  3698  "length (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  3699                         (Dec n e)) div 2= length_of (Dec n e)")
       
  3700 defer
       
  3701 apply(simp add: ci_length)
       
  3702 apply(subgoal_tac 
       
  3703  "length (ci (layout_of aprog) (start_of (layout_of aprog) as)
       
  3704                   (Dec n e)) mod 2 = 0", auto simp: length_of.simps)
       
  3705 done
       
  3706 
       
  3707 lemma [simp]:
       
  3708     "start_of (layout_of aprog) as > 0 \<Longrightarrow> 
       
  3709  fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  3710                                           (Dec n as)) (Suc 0)  Bk 
       
  3711       = (W1, start_of (layout_of aprog) as)"
       
  3712 apply(auto simp: ci.simps findnth.simps fetch.simps nth_of.simps 
       
  3713                  tshift.simps nth_append Suc_pre tdec_b_def)
       
  3714 apply(insert findnth_nth[of 0 n "0"], simp)
       
  3715 apply(simp add: findnth.simps)
       
  3716 done
       
  3717 
       
  3718 lemma [simp]: 
       
  3719  "\<lbrakk>abc_fetch as aprog = Some (Dec n e); 
       
  3720    0 < start_of (layout_of aprog) as\<rbrakk>
       
  3721 \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as)
       
  3722          (Dec n e)) (Suc (start_of (layout_of aprog) e) - 
       
  3723               start_of (layout_of aprog) as)  Bk)
       
  3724    = (if e = as then (W1, start_of (layout_of aprog) as)
       
  3725                   else (Nop, 0))"
       
  3726 apply(auto split: if_splits)
       
  3727 apply(case_tac "e < as", simp add: fetch.simps)
       
  3728 apply(subgoal_tac " e > as")
       
  3729 apply(drule start_of_ge, simp, auto simp: fetch.simps 
       
  3730                                           ci_length nth_of.simps)
       
  3731 apply(subgoal_tac 
       
  3732  "length (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  3733                             (Dec n e)) div 2= length_of (Dec n e)")
       
  3734 defer
       
  3735 apply(simp add: ci_length)
       
  3736 apply(subgoal_tac 
       
  3737  "length (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  3738                    (Dec n e)) mod 2 = 0", auto simp: length_of.simps)
       
  3739 apply(simp add: ci.simps tshift.simps tdec_b_def)
       
  3740 done
       
  3741 
       
  3742 lemma [simp]: 
       
  3743  "inv_stop (as, abc_lm_s am n (abc_lm_v am n)) (s, l, r) ires \<Longrightarrow> l \<noteq> []"
       
  3744 apply(auto simp: inv_stop.simps)
       
  3745 done
       
  3746 
       
  3747 lemma [simp]: 
       
  3748  "\<lbrakk>abc_fetch as aprog = Some (Dec n e); e \<noteq> as; ly = layout_of aprog\<rbrakk>
       
  3749   \<Longrightarrow> (\<not> (start_of ly as \<le> start_of ly e \<and> 
       
  3750       start_of ly e < start_of ly as + 2 * n))
       
  3751     \<and> start_of ly e \<noteq> start_of ly as + 2*n \<and> 
       
  3752       start_of ly e \<noteq> Suc (start_of ly as + 2*n) "
       
  3753 apply(case_tac "e < as")
       
  3754 apply(drule_tac ly = ly in start_of_le, simp)
       
  3755 apply(case_tac n, simp, drule start_of_inj, simp, simp, simp, simp)
       
  3756 apply(drule_tac start_of_ge, simp, simp)
       
  3757 done
       
  3758 
       
  3759 lemma [simp]: 
       
  3760    "\<lbrakk>abc_fetch as aprog = Some (Dec n e); start_of ly as \<le> s; 
       
  3761      s < start_of ly as + 2 * n; ly = layout_of aprog\<rbrakk>
       
  3762      \<Longrightarrow> Suc s \<noteq> start_of ly e "
       
  3763 apply(case_tac "e = as", simp)
       
  3764 apply(case_tac "e < as")
       
  3765 apply(drule_tac a = e and b = as and ly = ly in start_of_le, simp)
       
  3766 apply(drule_tac start_of_ge, auto)
       
  3767 done
       
  3768 
       
  3769 lemma [simp]: "\<lbrakk>abc_fetch as aprog = Some (Dec n e); 
       
  3770                 ly = layout_of aprog\<rbrakk>
       
  3771          \<Longrightarrow> Suc (start_of ly as + 2 * n) \<noteq> start_of ly e"
       
  3772 apply(case_tac "e = as", simp)
       
  3773 apply(case_tac "e < as")
       
  3774 apply(drule_tac a = e and b = as and ly = ly in start_of_le, simp)
       
  3775 apply(drule_tac start_of_ge, auto)
       
  3776 done
       
  3777 
       
  3778 lemma dec_false_1[simp]:
       
  3779  "\<lbrakk>abc_lm_v am n = 0; inv_locate_b (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
       
  3780   \<Longrightarrow> False"
       
  3781 apply(auto simp: inv_locate_b.simps in_middle.simps exponent_def)
       
  3782 apply(case_tac "length lm1 \<ge> length am", auto)
       
  3783 apply(subgoal_tac "lm2 = []", simp, subgoal_tac "m = 0", simp)
       
  3784 apply(case_tac mr, auto simp: )
       
  3785 apply(subgoal_tac "Suc (length lm1) - length am = 
       
  3786                    Suc (length lm1 - length am)", 
       
  3787       simp add: rep_ind del: replicate.simps, simp)
       
  3788 apply(drule_tac xs = "am @ replicate (Suc (length lm1) - length am) 0"
       
  3789                 and ys = "lm1 @ m # lm2" in length_equal, simp)
       
  3790 apply(case_tac mr, auto simp: abc_lm_v.simps)
       
  3791 apply(case_tac "mr = 0", simp_all add:  exponent_def split: if_splits)
       
  3792 apply(subgoal_tac "Suc (length lm1) - length am = 
       
  3793                        Suc (length lm1 - length am)", 
       
  3794       simp add: rep_ind del: replicate.simps, simp)
       
  3795 done
       
  3796 
       
  3797 lemma [simp]: 
       
  3798  "\<lbrakk>inv_locate_b (as, am) (n, aaa, Bk # xs) ires; 
       
  3799    abc_lm_v am n = 0\<rbrakk>
       
  3800    \<Longrightarrow> inv_on_left_moving (as, abc_lm_s am n 0) 
       
  3801                          (s, tl aaa, hd aaa # Bk # xs) ires" 
       
  3802 apply(insert inv_locate_b_2_on_left_moving[of as am n aaa xs ires s], simp)
       
  3803 done
       
  3804 
       
  3805 lemma [simp]:
       
  3806  "\<lbrakk>abc_lm_v am n = 0; inv_locate_b (as, am) (n, aaa, []) ires\<rbrakk>
       
  3807    \<Longrightarrow> inv_on_left_moving (as, abc_lm_s am n 0) (s, tl aaa, [hd aaa]) ires"
       
  3808 apply(insert inv_locate_b_2_on_left_moving_b[of as am n aaa ires s], simp)
       
  3809 done
       
  3810 
       
  3811 lemma [simp]: "\<lbrakk>am ! n = (0::nat); n < length am\<rbrakk> \<Longrightarrow> am[n := 0] = am"
       
  3812 apply(simp add: list_update_same_conv)
       
  3813 done
       
  3814 
       
  3815 lemma [simp]: "\<lbrakk>abc_lm_v am n = 0; 
       
  3816                 inv_locate_b (as, abc_lm_s am n 0) (n, Oc # aaa, xs) ires\<rbrakk>
       
  3817      \<Longrightarrow> inv_locate_b (as, am) (n, Oc # aaa, xs) ires"
       
  3818 apply(simp only: inv_locate_b.simps in_middle.simps abc_lm_s.simps 
       
  3819                  abc_lm_v.simps)
       
  3820 apply(erule_tac exE)+
       
  3821 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, simp)
       
  3822 apply(case_tac "n < length am", simp_all)
       
  3823 apply(erule_tac conjE)+
       
  3824 apply(rule_tac x = tn in exI, rule_tac x = m in exI, simp)
       
  3825 apply(rule_tac x = ml in exI, rule_tac x = mr in exI, simp)
       
  3826 defer
       
  3827 apply(rule_tac x = "Suc n - length am" in exI, rule_tac x = m in exI)
       
  3828 apply(subgoal_tac "Suc n - length am = Suc (n - length am)")
       
  3829 apply(simp add: exponent_def rep_ind del: replicate.simps, auto)
       
  3830 done
       
  3831 
       
  3832 lemma  [intro]: "\<lbrakk>abc_lm_v (a # list) 0 = 0\<rbrakk> \<Longrightarrow> a = 0"
       
  3833 apply(simp add: abc_lm_v.simps split: if_splits)
       
  3834 done
       
  3835 
       
  3836 lemma [simp]: 
       
  3837  "inv_stop (as, abc_lm_s am n 0) 
       
  3838           (start_of (layout_of aprog) e, aaa, Oc # xs) ires
       
  3839   \<Longrightarrow> inv_locate_a (as, abc_lm_s am n 0) (0, aaa, Oc # xs) ires"
       
  3840 apply(simp add: inv_locate_a.simps)
       
  3841 apply(rule disjI1)
       
  3842 apply(auto simp: inv_stop.simps at_begin_norm.simps)
       
  3843 done
       
  3844 
       
  3845 lemma [simp]: 
       
  3846  "\<lbrakk>abc_lm_v am 0 = 0; 
       
  3847   inv_stop (as, abc_lm_s am 0 0) 
       
  3848       (start_of (layout_of aprog) e, aaa, Oc # xs) ires\<rbrakk> \<Longrightarrow> 
       
  3849   inv_locate_b (as, am) (0, Oc # aaa, xs) ires"
       
  3850 apply(auto simp: inv_stop.simps inv_locate_b.simps 
       
  3851                  in_middle.simps abc_lm_s.simps)
       
  3852 apply(case_tac "am = []", simp)
       
  3853 apply(rule_tac x = "[]" in exI, rule_tac x = "Suc 0" in exI, 
       
  3854       rule_tac x = 0 in exI, simp)
       
  3855 apply(rule_tac x = "Suc 0" in exI, rule_tac x = 0 in exI, 
       
  3856   simp add: tape_of_nl_abv tape_of_nat_list.simps, auto)
       
  3857 apply(case_tac rn, auto)
       
  3858 apply(rule_tac x = "tl am" in exI, rule_tac x = 0 in exI, 
       
  3859       rule_tac x = "hd am" in exI, simp)
       
  3860 apply(rule_tac x = "Suc 0" in exI, rule_tac x = "hd am" in exI, simp)
       
  3861 apply(case_tac am, simp, simp)
       
  3862 apply(subgoal_tac "a = 0", case_tac list, 
       
  3863       auto simp: tape_of_nat_list.simps tape_of_nl_abv)
       
  3864 apply(case_tac rn, auto)
       
  3865 done
       
  3866 
       
  3867 lemma [simp]: 
       
  3868  "\<lbrakk>inv_stop (as, abc_lm_s am n 0) 
       
  3869           (start_of (layout_of aprog) e, aaa, Oc # xs) ires\<rbrakk>
       
  3870   \<Longrightarrow> inv_locate_b (as, am) (0, Oc # aaa, xs) ires \<or> 
       
  3871       inv_locate_b (as, abc_lm_s am n 0) (0, Oc # aaa, xs) ires"
       
  3872 apply(simp)
       
  3873 done
       
  3874 
       
  3875 lemma [simp]: 
       
  3876 "\<lbrakk>abc_lm_v am n = 0; 
       
  3877   inv_stop (as, abc_lm_s am n 0) 
       
  3878           (start_of (layout_of aprog) e, aaa, Oc # xs) ires\<rbrakk>
       
  3879  \<Longrightarrow> \<not> Suc 0 < 2 * n \<longrightarrow> e = as \<longrightarrow> 
       
  3880             inv_locate_b (as, am) (n, Oc # aaa, xs) ires"
       
  3881 apply(case_tac n, simp, simp)
       
  3882 done
       
  3883 
       
  3884 lemma dec_false2: 
       
  3885  "inv_stop (as, abc_lm_s am n 0) 
       
  3886   (start_of (layout_of aprog) e, aaa, Bk # xs) ires = False"
       
  3887 apply(auto simp: inv_stop.simps abc_lm_s.simps)
       
  3888 apply(case_tac "am", simp, case_tac n, simp add: tape_of_nl_abv)
       
  3889 apply(case_tac list, simp add: tape_of_nat_list.simps )
       
  3890 apply(simp add: tape_of_nat_list.simps , simp)
       
  3891 apply(case_tac "list[nat := 0]", 
       
  3892       simp add: tape_of_nat_list.simps  tape_of_nl_abv)
       
  3893 apply(simp add: tape_of_nat_list.simps )
       
  3894 apply(case_tac "am @ replicate (n - length am) 0 @ [0]", simp)
       
  3895 apply(case_tac list, auto simp: tape_of_nl_abv 
       
  3896                                 tape_of_nat_list.simps )
       
  3897 done	
       
  3898 
       
  3899 lemma dec_false3:
       
  3900    "inv_stop (as, abc_lm_s am n 0) 
       
  3901               (start_of (layout_of aprog) e, aaa, []) ires = False"
       
  3902 apply(auto simp: inv_stop.simps abc_lm_s.simps)
       
  3903 apply(case_tac "am", case_tac n, auto)
       
  3904 apply(case_tac n, auto simp: tape_of_nl_abv)
       
  3905 apply(case_tac "list::nat list",
       
  3906             simp add: tape_of_nat_list.simps tape_of_nat_list.simps)
       
  3907 apply(simp add: tape_of_nat_list.simps)
       
  3908 apply(case_tac "list[nat := 0]", 
       
  3909             simp add: tape_of_nat_list.simps tape_of_nat_list.simps)
       
  3910 apply(simp add: tape_of_nat_list.simps)
       
  3911 apply(case_tac "(am @ replicate (n - length am) 0 @ [0])", simp)
       
  3912 apply(case_tac list, auto simp: tape_of_nat_list.simps)
       
  3913 done
       
  3914 
       
  3915 lemma [simp]:
       
  3916   "fetch (ci (layout_of aprog) 
       
  3917        (start_of (layout_of aprog) as) (Dec n e)) 0 b = (Nop, 0)"
       
  3918 by(simp add: fetch.simps)
       
  3919 
       
  3920 declare dec_inv_1.simps[simp del]
       
  3921 
       
  3922 declare inv_locate_n_b.simps [simp del]
       
  3923 
       
  3924 lemma [simp]:
       
  3925 "\<lbrakk>0 < abc_lm_v am n; 0 < n; 
       
  3926   at_begin_norm (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
       
  3927   \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires"
       
  3928 apply(simp only: at_begin_norm.simps inv_locate_n_b.simps)
       
  3929 apply(erule_tac exE)+
       
  3930 apply(rule_tac x = lm1 in exI, simp)
       
  3931 apply(case_tac "length lm2", simp)
       
  3932 apply(case_tac rn, simp, simp)
       
  3933 apply(rule_tac x = "tl lm2" in exI, rule_tac x = "hd lm2" in exI, simp)
       
  3934 apply(rule conjI)
       
  3935 apply(case_tac "lm2", simp, simp)
       
  3936 apply(case_tac "lm2", auto simp: tape_of_nl_abv tape_of_nat_list.simps)
       
  3937 apply(case_tac [!] "list", auto simp: tape_of_nl_abv tape_of_nat_list.simps)
       
  3938 apply(case_tac rn, auto)
       
  3939 done 
       
  3940 lemma [simp]: "(\<exists>rn. Oc # xs = Bk\<^bsup>rn\<^esup>) = False"
       
  3941 apply(auto)
       
  3942 apply(case_tac rn, auto simp: )
       
  3943 done
       
  3944 
       
  3945 lemma [simp]:
       
  3946   "\<lbrakk>0 < abc_lm_v am n; 0 < n; 
       
  3947     at_begin_fst_bwtn (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
       
  3948  \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires"
       
  3949 apply(simp add: at_begin_fst_bwtn.simps inv_locate_n_b.simps )
       
  3950 done
       
  3951  
       
  3952 lemma Suc_minus:"length am + tn = n
       
  3953        \<Longrightarrow> Suc tn = Suc n - length am "
       
  3954 apply(arith)
       
  3955 done
       
  3956 
       
  3957 lemma [simp]: 
       
  3958  "\<lbrakk>0 < abc_lm_v am n; 0 < n; 
       
  3959    at_begin_fst_awtn (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
       
  3960  \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires"
       
  3961 apply(simp only: at_begin_fst_awtn.simps inv_locate_n_b.simps )
       
  3962 apply(erule exE)+
       
  3963 apply(erule conjE)+
       
  3964 apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, 
       
  3965       rule_tac x = "Suc tn" in exI, rule_tac x = 0 in exI)
       
  3966 apply(simp add: exponent_def rep_ind del: replicate.simps)
       
  3967 apply(rule conjI)+
       
  3968 apply(auto)
       
  3969 apply(case_tac [!] rn, auto)
       
  3970 done
       
  3971 
       
  3972 lemma [simp]: 
       
  3973  "\<lbrakk>0 < abc_lm_v am n; 0 < n; inv_locate_a (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
       
  3974  \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc#aaa, xs) ires"
       
  3975 apply(auto simp: inv_locate_a.simps)
       
  3976 done
       
  3977 
       
  3978 lemma [simp]:
       
  3979  "\<lbrakk>inv_locate_n_b (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
       
  3980  \<Longrightarrow> dec_first_on_right_moving n (as, abc_lm_s am n (abc_lm_v am n))  
       
  3981                                       (s, Oc # aaa, xs) ires"
       
  3982 apply(auto simp: inv_locate_n_b.simps dec_first_on_right_moving.simps 
       
  3983                  abc_lm_s.simps abc_lm_v.simps)
       
  3984 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
       
  3985       rule_tac x = m in exI, simp)
       
  3986 apply(rule_tac x = "Suc (Suc 0)" in exI, 
       
  3987       rule_tac x = "m - 1" in exI, simp)
       
  3988 apply(case_tac m, auto simp:  exponent_def)
       
  3989 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
       
  3990       rule_tac x = m in exI, 
       
  3991       simp add: Suc_diff_le rep_ind del: replicate.simps)
       
  3992 apply(rule_tac x = "Suc (Suc 0)" in exI, 
       
  3993       rule_tac x = "m - 1" in exI, simp)
       
  3994 apply(case_tac m, auto simp:  exponent_def)
       
  3995 apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, 
       
  3996       rule_tac x = m in exI, simp)
       
  3997 apply(rule_tac x = "Suc (Suc 0)" in exI, 
       
  3998       rule_tac x = "m - 1" in exI, simp)
       
  3999 apply(case_tac m, auto)
       
  4000 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
       
  4001       rule_tac x = m in exI, 
       
  4002       simp add: Suc_diff_le rep_ind del: replicate.simps, simp)
       
  4003 done
       
  4004 
       
  4005 lemma dec_false_2: 
       
  4006  "\<lbrakk>0 < abc_lm_v am n; inv_locate_n_b (as, am) (n, aaa, Bk # xs) ires\<rbrakk> 
       
  4007  \<Longrightarrow> False"
       
  4008 apply(auto simp: inv_locate_n_b.simps abc_lm_v.simps split: if_splits)
       
  4009 apply(case_tac [!] m, auto)
       
  4010 done
       
  4011  
       
  4012 lemma dec_false_2_b:
       
  4013  "\<lbrakk>0 < abc_lm_v am n; inv_locate_n_b (as, am) 
       
  4014                                 (n, aaa, []) ires\<rbrakk> \<Longrightarrow> False"
       
  4015 apply(auto simp: inv_locate_n_b.simps abc_lm_v.simps split: if_splits)
       
  4016 apply(case_tac [!] m, auto simp: )
       
  4017 done
       
  4018 
       
  4019 
       
  4020 (*begin: dec halt1 lemmas*)
       
  4021 thm abc_inc_stage1.simps
       
  4022 fun abc_dec_1_stage1:: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
       
  4023   where
       
  4024   "abc_dec_1_stage1 (s, l, r) ss n = 
       
  4025        (if s > ss \<and> s \<le> ss + 2*n + 1 then 4
       
  4026         else if s = ss + 2 * n + 13 \<or> s = ss + 2*n + 14 then 3
       
  4027         else if s = ss + 2*n + 15 then 2
       
  4028         else 0)"
       
  4029 
       
  4030 fun abc_dec_1_stage2:: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
       
  4031   where
       
  4032   "abc_dec_1_stage2 (s, l, r) ss n = 
       
  4033        (if s \<le> ss + 2 * n + 1 then (ss + 2 * n + 16 - s)
       
  4034         else if s = ss + 2*n + 13 then length l
       
  4035         else if s = ss + 2*n + 14 then length l
       
  4036         else 0)"
       
  4037 
       
  4038 fun abc_dec_1_stage3 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> block list \<Rightarrow> nat"
       
  4039   where
       
  4040   "abc_dec_1_stage3 (s, l, r) ss n ires = 
       
  4041         (if s \<le> ss + 2*n + 1 then 
       
  4042              if (s - ss) mod 2 = 0 then 
       
  4043                          if r \<noteq> [] \<and> hd r = Oc then 0 else 1  
       
  4044                          else length r
       
  4045          else if s = ss + 2 * n + 13 then 
       
  4046              if l = Bk # ires \<and> r \<noteq> [] \<and> hd r = Oc then 2 
       
  4047              else 1
       
  4048          else if s = ss + 2 * n + 14 then 
       
  4049              if r \<noteq> [] \<and> hd r = Oc then 3 else 0 
       
  4050          else 0)"
       
  4051 
       
  4052 fun abc_dec_1_measure :: "(t_conf \<times> nat \<times> nat \<times> block list) \<Rightarrow> (nat \<times> nat \<times> nat)"
       
  4053   where
       
  4054   "abc_dec_1_measure (c, ss, n, ires) = (abc_dec_1_stage1 c ss n, 
       
  4055                    abc_dec_1_stage2 c ss n, abc_dec_1_stage3 c ss n ires)"
       
  4056 
       
  4057 definition abc_dec_1_LE ::
       
  4058   "(((nat \<times> block list \<times> block list) \<times> nat \<times>
       
  4059   nat \<times> block list) \<times> ((nat \<times> block list \<times> block list) \<times> nat \<times> nat \<times> block list)) set"
       
  4060   where "abc_dec_1_LE \<equiv> (inv_image lex_triple abc_dec_1_measure)"
       
  4061 
       
  4062 lemma wf_dec_le: "wf abc_dec_1_LE"
       
  4063 by(auto intro:wf_inv_image wf_lex_triple simp:abc_dec_1_LE_def)
       
  4064 
       
  4065 declare dec_inv_1.simps[simp del] dec_inv_2.simps[simp del]
       
  4066 
       
  4067 lemma [elim]: 
       
  4068  "\<lbrakk>abc_fetch as aprog = Some (Dec n e); 
       
  4069    start_of (layout_of aprog) as < start_of (layout_of aprog) e;
       
  4070    start_of (layout_of aprog) e \<le> 
       
  4071          Suc (start_of (layout_of aprog) as + 2 * n)\<rbrakk> \<Longrightarrow> False"
       
  4072 apply(case_tac "e = as", simp)
       
  4073 apply(case_tac "e < as")
       
  4074 apply(drule_tac a = e and b = as and ly = "layout_of aprog" in 
       
  4075                                                  start_of_le, simp)
       
  4076 apply(drule_tac start_of_ge, auto)
       
  4077 done
       
  4078 
       
  4079 lemma [elim]: "\<lbrakk>abc_fetch as aprog = Some (Dec n e); 
       
  4080                                 start_of (layout_of aprog) e 
       
  4081     = start_of (layout_of aprog) as + 2 * n + 13\<rbrakk>
       
  4082          \<Longrightarrow> False"
       
  4083 apply(insert starte_not_equal[of as aprog n e "layout_of aprog"], 
       
  4084       simp)
       
  4085 done
       
  4086 
       
  4087 lemma [elim]: "\<lbrakk>abc_fetch as aprog = Some (Dec n e);
       
  4088                  start_of (layout_of aprog) e = 
       
  4089                start_of (layout_of aprog) as + 2 * n + 14\<rbrakk>
       
  4090         \<Longrightarrow> False"
       
  4091 apply(insert starte_not_equal[of as aprog n e "layout_of aprog"],
       
  4092       simp)
       
  4093 done
       
  4094 
       
  4095 lemma [elim]: 
       
  4096  "\<lbrakk>abc_fetch as aprog = Some (Dec n e);
       
  4097    start_of (layout_of aprog) as < start_of (layout_of aprog) e;
       
  4098    start_of (layout_of aprog) e \<le> 
       
  4099               Suc (start_of (layout_of aprog) as + 2 * n)\<rbrakk>
       
  4100    \<Longrightarrow> False"
       
  4101 apply(case_tac "e = as", simp)
       
  4102 apply(case_tac "e < as")
       
  4103 apply(drule_tac a = e and b = as and ly = "layout_of aprog" in 
       
  4104                                                     start_of_le, simp)
       
  4105 apply(drule_tac start_of_ge, auto)
       
  4106 done
       
  4107 
       
  4108 lemma [elim]: 
       
  4109  "\<lbrakk>abc_fetch as aprog = Some (Dec n e); 
       
  4110    start_of (layout_of aprog) e = 
       
  4111                start_of (layout_of aprog) as + 2 * n + 13\<rbrakk>
       
  4112     \<Longrightarrow> False"
       
  4113 apply(insert starte_not_equal[of as aprog n e "layout_of aprog"], 
       
  4114       simp)
       
  4115 done
       
  4116 
       
  4117 lemma [simp]: 
       
  4118  "abc_fetch as aprog = Some (Dec n e) \<Longrightarrow> 
       
  4119    Suc (start_of (layout_of aprog) as) \<noteq> start_of (layout_of aprog) e"
       
  4120 apply(case_tac "e = as", simp) 
       
  4121 apply(case_tac "e < as")
       
  4122 apply(drule_tac a = e and b = as and ly = "(layout_of aprog)" in 
       
  4123                                                  start_of_le, simp)
       
  4124 apply(drule_tac a = as and e = e in start_of_ge, simp, simp)
       
  4125 done
       
  4126 
       
  4127 lemma [simp]: "inv_on_left_moving (as, am) (s, [], r) ires 
       
  4128   = False"
       
  4129 apply(simp add: inv_on_left_moving.simps inv_on_left_moving_norm.simps
       
  4130                 inv_on_left_moving_in_middle_B.simps)
       
  4131 done
       
  4132 
       
  4133 lemma [simp]: 
       
  4134   "inv_check_left_moving (as, abc_lm_s am n 0)
       
  4135   (start_of (layout_of aprog) as + 2 * n + 14, [], Oc # xs) ires
       
  4136  = False"
       
  4137 apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps)
       
  4138 done
       
  4139 
       
  4140 lemma dec_inv_stop1_pre: 
       
  4141     "\<lbrakk>abc_fetch as aprog = Some (Dec n e); abc_lm_v am n = 0;
       
  4142       start_of (layout_of aprog) as > 0\<rbrakk>
       
  4143  \<Longrightarrow> \<forall>na. \<not> (\<lambda>(s, l, r) (ss, n', ires'). s = start_of (layout_of aprog) e)
       
  4144             (t_steps (Suc (start_of (layout_of aprog) as), l, r) 
       
  4145               (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  4146                  (Dec n e), start_of (layout_of aprog) as - Suc 0) na)
       
  4147                       (start_of (layout_of aprog) as, n, ires) \<and>
       
  4148            dec_inv_1 (layout_of aprog) n e (as, am)
       
  4149             (t_steps (Suc (start_of (layout_of aprog) as), l, r) 
       
  4150               (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  4151                 (Dec n e), start_of (layout_of aprog) as - Suc 0) na) ires
       
  4152        \<longrightarrow> dec_inv_1 (layout_of aprog) n e (as, am)
       
  4153             (t_steps (Suc (start_of (layout_of aprog) as), l, r) 
       
  4154               (ci (layout_of aprog) (start_of (layout_of aprog) as)
       
  4155                  (Dec n e), start_of (layout_of aprog) as - Suc 0) 
       
  4156                     (Suc na)) ires \<and>
       
  4157             ((t_steps (Suc (start_of (layout_of aprog) as), l, r) 
       
  4158             (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  4159            (Dec n e), start_of (layout_of aprog) as - Suc 0) (Suc na),
       
  4160              start_of (layout_of aprog) as, n, ires),
       
  4161          t_steps (Suc (start_of (layout_of aprog) as), l, r) 
       
  4162             (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  4163                (Dec n e), start_of (layout_of aprog) as - Suc 0) na,
       
  4164             start_of (layout_of aprog) as, n, ires)
       
  4165            \<in> abc_dec_1_LE"
       
  4166 apply(rule allI, rule impI, simp add: t_steps_ind)
       
  4167 apply(case_tac "(t_steps (Suc (start_of (layout_of aprog) as), l, r) 
       
  4168 (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), 
       
  4169 start_of (layout_of aprog) as - Suc 0) na)", simp)
       
  4170 apply(auto split:if_splits simp add:t_step.simps dec_inv_1.simps, 
       
  4171           tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *})
       
  4172 apply(simp_all add:dec_fetch_simps new_tape.simps dec_inv_1.simps)
       
  4173 apply(auto simp add: abc_dec_1_LE_def lex_square_def 
       
  4174                      lex_triple_def lex_pair_def  
       
  4175                 split: if_splits)
       
  4176 apply(rule dec_false_1, simp, simp)
       
  4177 done
       
  4178 
       
  4179 lemma dec_inv_stop1: 
       
  4180   "\<lbrakk>ly = layout_of aprog; 
       
  4181     dec_inv_1 ly n e (as, am) (start_of ly as + 1, l, r) ires; 
       
  4182     abc_fetch as aprog = Some (Dec n e); abc_lm_v am n = 0\<rbrakk> \<Longrightarrow> 
       
  4183   (\<exists> stp. (\<lambda> (s', l', r'). s' = start_of ly e \<and> 
       
  4184            dec_inv_1 ly n e (as, am) (s', l' , r') ires) 
       
  4185   (t_steps (start_of ly as + 1, l, r)
       
  4186      (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp))"
       
  4187 apply(insert halt_lemma2[of abc_dec_1_LE 
       
  4188     "\<lambda> ((s, l, r), ss, n', ires'). s = start_of ly e" 
       
  4189      "(\<lambda> stp. (t_steps (start_of ly as + 1, l, r) 
       
  4190           (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) 
       
  4191                stp, start_of ly as, n, ires))"
       
  4192      "\<lambda> ((s, l, r), ss, n, ires'). dec_inv_1 ly n e (as, am) (s, l, r) ires'"],
       
  4193      simp)
       
  4194 apply(insert wf_dec_le, simp)
       
  4195 apply(insert dec_inv_stop1_pre[of as aprog n e am l r], simp)
       
  4196 apply(subgoal_tac "start_of (layout_of aprog) as > 0", 
       
  4197                                       simp add: t_steps.simps)
       
  4198 apply(erule_tac exE, rule_tac x = na in exI)
       
  4199 apply(case_tac
       
  4200      "(t_steps (Suc (start_of (layout_of aprog) as), l, r) 
       
  4201          (ci (layout_of aprog) (start_of (layout_of aprog) as)
       
  4202            (Dec n e), start_of (layout_of aprog) as - Suc 0) na)",
       
  4203       case_tac b, auto)
       
  4204 apply(rule startof_not0)
       
  4205 done
       
  4206 
       
  4207 (*begin: dec halt2 lemmas*)
       
  4208 
       
  4209 lemma [simp]:
       
  4210   "\<lbrakk>abc_fetch as aprog = Some (Dec n e); 
       
  4211     ly = layout_of aprog\<rbrakk> \<Longrightarrow> 
       
  4212               start_of ly (Suc as) = start_of ly as + 2*n + 16"
       
  4213 by simp
       
  4214 
       
  4215 fun abc_dec_2_stage1 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
       
  4216   where
       
  4217   "abc_dec_2_stage1 (s, l, r) ss n = 
       
  4218               (if s \<le> ss + 2*n + 1 then 7
       
  4219                else if s = ss + 2*n + 2 then 6 
       
  4220                else if s = ss + 2*n + 3 then 5
       
  4221                else if s \<ge> ss + 2*n + 4 \<and> s \<le> ss + 2*n + 9 then 4
       
  4222                else if s = ss + 2*n + 6 then 3
       
  4223                else if s = ss + 2*n + 10 \<or> s = ss + 2*n + 11 then 2
       
  4224                else if s = ss + 2*n + 12 then 1
       
  4225                else 0)"
       
  4226 
       
  4227 thm new_tape.simps
       
  4228 
       
  4229 fun abc_dec_2_stage2 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
       
  4230   where
       
  4231   "abc_dec_2_stage2 (s, l, r) ss n = 
       
  4232        (if s \<le> ss + 2 * n + 1 then (ss + 2 * n + 16 - s)
       
  4233         else if s = ss + 2*n + 10 then length l
       
  4234         else if s = ss + 2*n + 11 then length l
       
  4235         else if s = ss + 2*n + 4 then length r - 1
       
  4236         else if s = ss + 2*n + 5 then length r 
       
  4237         else if s = ss + 2*n + 7 then length r - 1
       
  4238         else if s = ss + 2*n + 8 then  
       
  4239               length r + length (takeWhile (\<lambda> a. a = Oc) l) - 1
       
  4240         else if s = ss + 2*n + 9 then 
       
  4241               length r + length (takeWhile (\<lambda> a. a = Oc) l) - 1
       
  4242         else 0)"
       
  4243 
       
  4244 fun abc_dec_2_stage3 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> block list \<Rightarrow> nat"
       
  4245   where
       
  4246   "abc_dec_2_stage3 (s, l, r) ss n ires =
       
  4247         (if s \<le> ss + 2*n + 1 then 
       
  4248             if (s - ss) mod 2 = 0 then if r \<noteq> [] \<and> 
       
  4249                                           hd r = Oc then 0 else 1  
       
  4250             else length r
       
  4251          else if s = ss + 2 * n + 10 then 
       
  4252              if l = Bk # ires \<and> r \<noteq> [] \<and> hd r = Oc then 2
       
  4253              else 1
       
  4254          else if s = ss + 2 * n + 11 then 
       
  4255              if r \<noteq> [] \<and> hd r = Oc then 3 
       
  4256              else 0 
       
  4257          else (ss + 2 * n + 16 - s))"
       
  4258 
       
  4259 fun abc_dec_2_stage4 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
       
  4260   where
       
  4261   "abc_dec_2_stage4 (s, l, r) ss n = 
       
  4262           (if s = ss + 2*n + 2 then length r
       
  4263            else if s = ss + 2*n + 8 then length r
       
  4264            else if s = ss + 2*n + 3 then 
       
  4265                if r \<noteq> [] \<and> hd r = Oc then 1
       
  4266                else 0
       
  4267            else if s = ss + 2*n + 7 then 
       
  4268                if r \<noteq> [] \<and> hd r = Oc then 0 
       
  4269                else 1
       
  4270            else if s = ss + 2*n + 9 then 
       
  4271                if r \<noteq> [] \<and> hd r = Oc then 1
       
  4272                else 0 
       
  4273            else 0)"
       
  4274 
       
  4275 fun abc_dec_2_measure :: "(t_conf \<times> nat \<times> nat \<times> block list) \<Rightarrow> 
       
  4276                                     (nat \<times> nat \<times> nat \<times> nat)"
       
  4277   where
       
  4278   "abc_dec_2_measure (c, ss, n, ires) = 
       
  4279        (abc_dec_2_stage1 c ss n, abc_dec_2_stage2 c ss n,
       
  4280         abc_dec_2_stage3 c ss n ires, abc_dec_2_stage4 c ss n)"
       
  4281 
       
  4282 definition abc_dec_2_LE :: 
       
  4283        "(((nat \<times> block list \<times> block list) \<times> nat \<times> nat \<times> block list) \<times> 
       
  4284         ((nat \<times> block list \<times> block list) \<times> nat \<times> nat \<times> block list)) set"
       
  4285   where "abc_dec_2_LE \<equiv> (inv_image lex_square abc_dec_2_measure)"
       
  4286 
       
  4287 lemma wf_dec_2_le: "wf abc_dec_2_LE"
       
  4288 by(auto intro:wf_inv_image wf_lex_triple wf_lex_square 
       
  4289    simp:abc_dec_2_LE_def)
       
  4290 
       
  4291 lemma [simp]: "dec_after_write (as, am) (s, aa, r) ires
       
  4292            \<Longrightarrow> takeWhile (\<lambda>a. a = Oc) aa = []"
       
  4293 apply(simp only : dec_after_write.simps)
       
  4294 apply(erule exE)+
       
  4295 apply(erule_tac conjE)+
       
  4296 apply(case_tac aa, simp)
       
  4297 apply(case_tac a, simp only: takeWhile.simps , simp, simp split: if_splits)
       
  4298 done
       
  4299 
       
  4300 lemma [simp]: 
       
  4301      "\<lbrakk>dec_on_right_moving (as, lm) (s, aa, []) ires; 
       
  4302        length (takeWhile (\<lambda>a. a = Oc) (tl aa)) 
       
  4303            \<noteq> length (takeWhile (\<lambda>a. a = Oc) aa) - Suc 0\<rbrakk>
       
  4304     \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (tl aa)) < 
       
  4305                        length (takeWhile (\<lambda>a. a = Oc) aa) - Suc 0"
       
  4306 apply(simp only: dec_on_right_moving.simps)
       
  4307 apply(erule_tac exE)+
       
  4308 apply(erule_tac conjE)+
       
  4309 apply(case_tac mr, auto split: if_splits)
       
  4310 done
       
  4311 
       
  4312 lemma [simp]: 
       
  4313   "dec_after_clear (as, abc_lm_s am n (abc_lm_v am n - Suc 0)) 
       
  4314              (start_of (layout_of aprog) as + 2 * n + 9, aa, Bk # xs) ires
       
  4315  \<Longrightarrow> length xs - Suc 0 < length xs + 
       
  4316                              length (takeWhile (\<lambda>a. a = Oc) aa)"
       
  4317 apply(simp only: dec_after_clear.simps)
       
  4318 apply(erule_tac exE)+
       
  4319 apply(erule conjE)+
       
  4320 apply(simp split: if_splits )
       
  4321 done
       
  4322 
       
  4323 lemma [simp]: 
       
  4324  "\<lbrakk>dec_after_clear (as, abc_lm_s am n (abc_lm_v am n - Suc 0))
       
  4325        (start_of (layout_of aprog) as + 2 * n + 9, aa, []) ires\<rbrakk>
       
  4326     \<Longrightarrow> Suc 0 < length (takeWhile (\<lambda>a. a = Oc) aa)"
       
  4327 apply(simp add: dec_after_clear.simps split: if_splits)
       
  4328 done
       
  4329 
       
  4330 lemma [simp]: 
       
  4331  "\<lbrakk>dec_on_right_moving (as, am) (s, aa, Bk # xs) ires; 
       
  4332    Suc (length (takeWhile (\<lambda>a. a = Oc) (tl aa)))
       
  4333    \<noteq> length (takeWhile (\<lambda>a. a = Oc) aa)\<rbrakk>
       
  4334   \<Longrightarrow> Suc (length (takeWhile (\<lambda>a. a = Oc) (tl aa))) 
       
  4335     < length (takeWhile (\<lambda>a. a = Oc) aa)"
       
  4336 apply(simp only: dec_on_right_moving.simps)
       
  4337 apply(erule exE)+
       
  4338 apply(erule conjE)+
       
  4339 apply(case_tac ml, auto split: if_splits )
       
  4340 done
       
  4341 
       
  4342 (*
       
  4343 lemma abc_dec_2_wf: 
       
  4344      "\<lbrakk>ly = layout_of aprog; dec_inv_2 ly n e (as, am) (start_of ly as + 1, l, r);  abc_fetch as aprog = Dec n e; abc_lm_v am n > 0\<rbrakk>
       
  4345        \<Longrightarrow> \<forall>na. \<not> (\<lambda>(s, l, r) (ss, n'). s = start_of (layout_of aprog) as + 2*n + 16)
       
  4346         (t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na)
       
  4347            (start_of (layout_of aprog) as, n) \<longrightarrow>
       
  4348         ((t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) (Suc na),
       
  4349             start_of (layout_of aprog) as, n),
       
  4350           t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na,
       
  4351            start_of (layout_of aprog) as, n)
       
  4352         \<in> abc_dec_2_LE"
       
  4353 proof(rule allI, rule impI, simp add: t_steps_ind)
       
  4354   fix na
       
  4355   assume h1 :"ly = layout_of aprog" "dec_inv_2 (layout_of aprog) n e (as, am) (Suc (start_of (layout_of aprog) as), l, r)" 
       
  4356           "abc_fetch as aprog = Dec n e" "abc_lm_v am n > 0"
       
  4357          "\<not> (\<lambda>(s, l, r) (ss, n'). s = start_of (layout_of aprog) as + 2*n + 16)
       
  4358              (t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na)
       
  4359              (start_of (layout_of aprog) as, n)"
       
  4360   thus "((t_step (t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na)
       
  4361                (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0),
       
  4362               start_of (layout_of aprog) as, n),
       
  4363              t_steps (Suc (start_of (layout_of aprog) as), l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na,
       
  4364              start_of (layout_of aprog) as, n)
       
  4365             \<in> abc_dec_2_LE"
       
  4366   proof(insert dec_inv_2_steps[of "layout_of aprog" n e as am "(start_of (layout_of aprog) as + 1, l, r)" aprog na], 
       
  4367         case_tac "(t_steps (start_of (layout_of aprog) as + 1, l, r) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0) na)", case_tac b, simp)
       
  4368     fix a b aa ba
       
  4369     assume "dec_inv_2 (layout_of aprog) n e (as, am) (a, aa, ba)" " a \<noteq> start_of (layout_of aprog) as + 2*n + 16" "abc_lm_v am n > 0" "abc_fetch as aprog = Dec n e "
       
  4370     thus "((t_step (a, aa, ba) (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), start_of (layout_of aprog) as - Suc 0), start_of (layout_of aprog) as, n), (a, aa, ba),
       
  4371                     start_of (layout_of aprog) as, n)
       
  4372                    \<in> abc_dec_2_LE"
       
  4373       apply(case_tac "a = 0", auto split:if_splits simp add:t_step.simps dec_inv_2.simps, 
       
  4374                 tactic {* ALLGOALS (resolve_tac (thms "fetch_intro")) *})
       
  4375       apply(simp_all add:dec_fetch_simps new_tape.simps)
       
  4376       apply(auto simp add: abc_dec_2_LE_def  lex_square_def lex_triple_def lex_pair_def  
       
  4377                            split: if_splits)
       
  4378       
       
  4379       done
       
  4380   qed
       
  4381 qed
       
  4382 *)
       
  4383 
       
  4384 lemma [simp]: "inv_check_left_moving (as, abc_lm_s am n (abc_lm_v am n - Suc 0)) 
       
  4385   (start_of (layout_of aprog) as + 2 * n + 11, [], Oc # xs) ires = False"
       
  4386 apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps)
       
  4387 done
       
  4388 
       
  4389 lemma dec_inv_stop2_pre: 
       
  4390   "\<lbrakk>abc_fetch as aprog = Some (Dec n e); abc_lm_v am n > 0\<rbrakk> \<Longrightarrow> 
       
  4391     \<forall>na. \<not> (\<lambda>(s, l, r) (ss, n', ires'). 
       
  4392                      s = start_of (layout_of aprog) as + 2 * n + 16)
       
  4393    (t_steps (Suc (start_of (layout_of aprog) as), l, r) 
       
  4394       (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e),
       
  4395             start_of (layout_of aprog) as - Suc 0) na)
       
  4396     (start_of (layout_of aprog) as, n, ires) \<and>
       
  4397  dec_inv_2 (layout_of aprog) n e (as, am)
       
  4398      (t_steps (Suc (start_of (layout_of aprog) as), l, r) 
       
  4399       (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e),
       
  4400           start_of (layout_of aprog) as - Suc 0) na) ires
       
  4401  \<longrightarrow>
       
  4402  dec_inv_2 (layout_of aprog) n e (as, am)
       
  4403      (t_steps (Suc (start_of (layout_of aprog) as), l, r) 
       
  4404       (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e),
       
  4405               start_of (layout_of aprog) as - Suc 0) (Suc na)) ires \<and>
       
  4406  ((t_steps (Suc (start_of (layout_of aprog) as), l, r) 
       
  4407       (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e),
       
  4408             start_of (layout_of aprog) as - Suc 0) (Suc na),  
       
  4409               start_of (layout_of aprog) as, n, ires),
       
  4410   t_steps (Suc (start_of (layout_of aprog) as), l, r) 
       
  4411       (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e),
       
  4412              start_of (layout_of aprog) as - Suc 0) na,
       
  4413                           start_of (layout_of aprog) as, n, ires)
       
  4414    \<in> abc_dec_2_LE"
       
  4415 apply(subgoal_tac "start_of (layout_of aprog) as > 0")
       
  4416 apply(rule allI, rule impI, simp add: t_steps_ind)
       
  4417 apply(case_tac "(t_steps (Suc (start_of (layout_of aprog) as), l, r)
       
  4418      (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), 
       
  4419              start_of (layout_of aprog) as - Suc 0) na)", simp)
       
  4420 apply(auto split:if_splits simp add:t_step.simps dec_inv_2.simps, 
       
  4421            tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *})
       
  4422 apply(simp_all add:dec_fetch_simps new_tape.simps dec_inv_2.simps)
       
  4423 apply(auto simp add: abc_dec_2_LE_def lex_square_def lex_triple_def 
       
  4424                      lex_pair_def split: if_splits)
       
  4425 apply(auto intro: dec_false_2_b dec_false_2)
       
  4426 apply(rule startof_not0)
       
  4427 done
       
  4428 
       
  4429 lemma dec_stop2: 
       
  4430  "\<lbrakk>ly = layout_of aprog; 
       
  4431    dec_inv_2 ly n e (as, am) (start_of ly as + 1, l, r) ires; 
       
  4432    abc_fetch as aprog = Some (Dec n e); 
       
  4433    abc_lm_v am n > 0\<rbrakk> \<Longrightarrow> 
       
  4434   (\<exists> stp. (\<lambda> (s', l', r'). s' = start_of ly (Suc as) \<and> 
       
  4435    dec_inv_2 ly n e (as, am) (s', l', r') ires)
       
  4436        (t_steps (start_of ly as+1, l, r) (ci ly (start_of ly as)
       
  4437                            (Dec n e), start_of ly as - Suc 0) stp))"
       
  4438 apply(insert halt_lemma2[of abc_dec_2_LE 
       
  4439       "\<lambda> ((s, l, r), ss, n', ires'). s = start_of ly (Suc as)"
       
  4440       "(\<lambda> stp. (t_steps (start_of ly as + 1, l, r) 
       
  4441        (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp,
       
  4442                  start_of ly as, n, ires))"
       
  4443       "(\<lambda> ((s, l, r), ss, n, ires'). dec_inv_2 ly n e (as, am) (s, l, r) ires')"])
       
  4444 apply(insert wf_dec_2_le, simp)
       
  4445 apply(insert dec_inv_stop2_pre[of as aprog n e am l r], 
       
  4446       simp add: t_steps.simps)
       
  4447 apply(erule_tac exE)
       
  4448 apply(rule_tac x = na in exI)
       
  4449 apply(case_tac "(t_steps (Suc (start_of (layout_of aprog) as), l, r) 
       
  4450 (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), 
       
  4451             start_of (layout_of aprog) as - Suc 0) na)",
       
  4452       case_tac b, auto)
       
  4453 done
       
  4454 
       
  4455 lemma dec_inv_stop_cond1: 
       
  4456   "\<lbrakk>ly = layout_of aprog; 
       
  4457     dec_inv_1 ly n e (as, lm) (s, (l, r)) ires; s = start_of ly e;
       
  4458     abc_fetch as aprog = Some (Dec n e); abc_lm_v lm n = 0\<rbrakk> 
       
  4459    \<Longrightarrow> crsp_l ly (e, abc_lm_s lm n 0) (s, l, r) ires"
       
  4460 apply(simp add: dec_inv_1.simps split: if_splits)
       
  4461 apply(auto simp: crsp_l.simps inv_stop.simps )
       
  4462 done
       
  4463 
       
  4464 lemma dec_inv_stop_cond2: 
       
  4465    "\<lbrakk>ly = layout_of aprog; s = start_of ly (Suc as); 
       
  4466      dec_inv_2 ly n e (as, lm) (s, (l, r)) ires;
       
  4467      abc_fetch as aprog = Some (Dec n e); 
       
  4468      abc_lm_v lm n > 0\<rbrakk>
       
  4469    \<Longrightarrow> crsp_l ly (Suc as,
       
  4470                   abc_lm_s lm n (abc_lm_v lm n - Suc 0)) (s, l, r) ires"
       
  4471 apply(simp add: dec_inv_2.simps split: if_splits)
       
  4472 apply(auto simp: crsp_l.simps inv_stop.simps )
       
  4473 done
       
  4474 
       
  4475 lemma [simp]: "(case Bk\<^bsup>rn\<^esup> of [] \<Rightarrow> Bk |
       
  4476                  Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc) = Bk"
       
  4477 apply(case_tac rn, auto)
       
  4478 done
       
  4479 
       
  4480 lemma [simp]: "t_steps tc (p,off) (m + n) = 
       
  4481                    t_steps (t_steps tc (p, off) m) (p, off) n"
       
  4482 apply(induct m arbitrary: n)
       
  4483 apply(simp add: t_steps.simps)
       
  4484 proof -
       
  4485   fix m n
       
  4486   assume h1: "\<And>n. t_steps tc (p, off) (m + n) =
       
  4487                      t_steps (t_steps tc (p, off) m) (p, off) n"
       
  4488   hence h2: "t_steps tc (p, off) (Suc m + n) = 
       
  4489                      t_steps tc (p, off) (m + Suc n)"
       
  4490     by simp
       
  4491   from h1 and this show 
       
  4492     "t_steps tc (p, off) (Suc m + n) = 
       
  4493          t_steps (t_steps tc (p, off) (Suc m)) (p, off) n"
       
  4494   proof(simp only: h2, simp add: t_steps.simps)
       
  4495     have h3: "(t_step (t_steps tc (p, off) m) (p, off)) = 
       
  4496                       (t_steps (t_step tc (p, off)) (p, off) m)"
       
  4497       apply(simp add: t_steps.simps[THEN sym] t_steps_ind[THEN sym])
       
  4498       done
       
  4499     from h3 show 
       
  4500       "t_steps (t_step (t_steps tc (p, off) m) (p, off)) (p, off) n =          t_steps (t_steps (t_step tc (p, off)) (p, off) m) (p, off) n"
       
  4501       by simp
       
  4502   qed
       
  4503 qed
       
  4504 
       
  4505 lemma [simp]: " abc_fetch as aprog = Some (Dec n e) \<Longrightarrow> 
       
  4506           Suc (start_of (layout_of aprog) as) \<noteq> 
       
  4507                            start_of (layout_of aprog) e"
       
  4508 apply(case_tac "e = as", simp)
       
  4509 apply(case_tac "e < as")
       
  4510 apply(drule_tac a = e and b = as and ly = "layout_of aprog" 
       
  4511                                            in start_of_le, simp)
       
  4512 apply(drule_tac start_of_ge, auto)
       
  4513 done
       
  4514 
       
  4515 lemma [simp]: "inv_locate_b (as, []) (0, Oc # Bk # Bk # ires, Bk\<^bsup>rn - Suc 0\<^esup>) ires"
       
  4516 apply(auto simp: inv_locate_b.simps in_middle.simps)
       
  4517 apply(rule_tac x = "[]" in exI, rule_tac x = "Suc 0" in exI, 
       
  4518       rule_tac x = 0 in exI, simp)
       
  4519 apply(rule_tac x = "Suc 0" in exI, rule_tac x = 0 in exI, auto)
       
  4520 apply(case_tac rn, simp, case_tac nat, auto)
       
  4521 done
       
  4522 
       
  4523 lemma [simp]: 
       
  4524        "inv_locate_n_b (as, []) (0, Oc # Bk # Bk # ires, Bk\<^bsup>rn - Suc 0\<^esup>) ires"
       
  4525 apply(auto simp: inv_locate_n_b.simps in_middle.simps)
       
  4526 apply(case_tac rn, simp, case_tac nat, auto)
       
  4527 done 
       
  4528 
       
  4529 lemma [simp]:
       
  4530 "abc_fetch as aprog = Some (Dec n e) \<Longrightarrow>
       
  4531    dec_inv_1 (layout_of aprog) n e (as, []) 
       
  4532     (Suc (start_of (layout_of aprog) as), Oc # Bk # Bk # ires, Bk\<^bsup>rn - Suc 0\<^esup>) ires
       
  4533 \<and>
       
  4534    dec_inv_2 (layout_of aprog) n e (as, []) 
       
  4535     (Suc (start_of (layout_of aprog) as), Oc # Bk # Bk # ires, Bk\<^bsup>rn - Suc 0\<^esup>) ires"
       
  4536 apply(simp add: dec_inv_1.simps dec_inv_2.simps)
       
  4537 apply(case_tac n, auto)
       
  4538 done
       
  4539 
       
  4540 lemma [simp]: 
       
  4541  "\<lbrakk>am \<noteq> []; <am> = Oc # r'; 
       
  4542    abc_fetch as aprog = Some (Dec n e)\<rbrakk> 
       
  4543  \<Longrightarrow> inv_locate_b (as, am) (0, Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>) ires"
       
  4544 apply(auto simp: inv_locate_b.simps in_middle.simps)
       
  4545 apply(rule_tac x = "tl am" in exI, rule_tac x = 0 in exI,
       
  4546       rule_tac x = "hd am" in exI, simp)
       
  4547 apply(rule_tac x = "Suc 0" in exI)
       
  4548 apply(rule_tac x = "hd am" in exI, simp)
       
  4549 apply(case_tac am, simp, case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps)
       
  4550 apply(case_tac rn, auto)
       
  4551 done
       
  4552 
       
  4553 lemma [simp]: 
       
  4554   "\<lbrakk><am> = Oc # r'; abc_fetch as aprog = Some (Dec n e)\<rbrakk> \<Longrightarrow> 
       
  4555   inv_locate_n_b (as, am) (0, Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>) ires"
       
  4556 apply(auto simp: inv_locate_n_b.simps)
       
  4557 apply(rule_tac x = "tl am" in exI, rule_tac x = "hd am" in exI, auto)
       
  4558 apply(case_tac [!] am, auto simp: tape_of_nl_abv tape_of_nat_list.simps )
       
  4559 apply(case_tac [!]list, auto simp: tape_of_nl_abv tape_of_nat_list.simps)
       
  4560 apply(case_tac rn, simp, simp)
       
  4561 apply(erule_tac x = nat in allE, simp)
       
  4562 done
       
  4563 
       
  4564 lemma [simp]:
       
  4565    "\<lbrakk>am \<noteq> [];  
       
  4566      <am> = Oc # r'; 
       
  4567      abc_fetch as aprog = Some (Dec n e)\<rbrakk> \<Longrightarrow>
       
  4568     dec_inv_1 (layout_of aprog) n e (as, am) 
       
  4569       (Suc (start_of (layout_of aprog) as), 
       
  4570            Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>) ires \<and>
       
  4571     dec_inv_2 (layout_of aprog) n e (as, am) 
       
  4572       (Suc (start_of (layout_of aprog) as), 
       
  4573            Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>) ires"
       
  4574 apply(simp add: dec_inv_1.simps dec_inv_2.simps)
       
  4575 apply(case_tac n, auto)
       
  4576 done
       
  4577 
       
  4578 lemma [simp]: "am \<noteq> [] \<Longrightarrow>  \<exists>r'. <am::nat list> = Oc # r'"
       
  4579 apply(case_tac am, simp, case_tac list)
       
  4580 apply(auto simp: tape_of_nl_abv tape_of_nat_list.simps )
       
  4581 done
       
  4582 
       
  4583 lemma [simp]: "start_of (layout_of aprog) as > 0 \<Longrightarrow> 
       
  4584       (fetch (ci (layout_of aprog) 
       
  4585            (start_of (layout_of aprog) as) (Dec n e)) (Suc 0)  Bk)
       
  4586     = (W1, start_of (layout_of aprog) as)"
       
  4587 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  4588              nth_of.simps tshift.simps nth_append Suc_pre tdec_b_def)
       
  4589 thm findnth_nth
       
  4590 apply(insert findnth_nth[of 0 n 0], simp)
       
  4591 apply(simp add: findnth.simps)
       
  4592 done
       
  4593 
       
  4594 lemma [simp]:
       
  4595     "start_of (layout_of aprog) as > 0
       
  4596    \<Longrightarrow> (t_step (start_of (layout_of aprog) as, Bk # Bk # ires, Bk\<^bsup>rn\<^esup>)
       
  4597     (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), 
       
  4598                              start_of (layout_of aprog) as - Suc 0))
       
  4599    = (start_of (layout_of aprog) as, Bk # Bk # ires, Oc # Bk\<^bsup>rn- Suc 0\<^esup>)"
       
  4600 apply(simp add: t_step.simps)
       
  4601 apply(case_tac "start_of (layout_of aprog) as",
       
  4602       auto simp: new_tape.simps)
       
  4603 apply(case_tac rn, auto)
       
  4604 done
       
  4605 
       
  4606 lemma [simp]: "start_of (layout_of aprog) as > 0 \<Longrightarrow> 
       
  4607  (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  4608          (Dec n e)) (Suc 0)  Oc)
       
  4609   = (R, Suc (start_of (layout_of aprog) as))"
       
  4610 
       
  4611 apply(auto simp: ci.simps findnth.simps fetch.simps
       
  4612                  nth_of.simps tshift.simps nth_append 
       
  4613                  Suc_pre tdec_b_def)
       
  4614 apply(insert findnth_nth[of 0 n "Suc 0"], simp)
       
  4615 apply(simp add: findnth.simps)
       
  4616 done
       
  4617 
       
  4618 lemma [simp]: "start_of (layout_of aprog) as > 0 \<Longrightarrow> 
       
  4619  (t_step (start_of (layout_of aprog) as, Bk # Bk # ires, Oc # Bk\<^bsup>rn - Suc 0\<^esup>)
       
  4620      (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e), 
       
  4621         start_of (layout_of aprog) as - Suc 0)) =
       
  4622   (Suc (start_of (layout_of aprog) as), Oc # Bk # Bk # ires, Bk\<^bsup>rn-Suc 0\<^esup>)"
       
  4623 apply(simp add: t_step.simps)
       
  4624 apply(case_tac "start_of (layout_of aprog) as", 
       
  4625       auto simp: new_tape.simps)
       
  4626 done
       
  4627 
       
  4628 lemma [simp]: "start_of (layout_of aprog) as > 0 \<Longrightarrow> 
       
  4629  t_step (start_of (layout_of aprog) as, Bk # Bk # ires, Oc # r' @ Bk\<^bsup>rn\<^esup>) 
       
  4630       (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e),
       
  4631                  start_of (layout_of aprog) as - Suc 0) =
       
  4632       (Suc (start_of (layout_of aprog) as), Oc # Bk # Bk # ires, r' @ Bk\<^bsup>rn\<^esup>)"
       
  4633 apply(simp add: t_step.simps)
       
  4634 apply(case_tac "start_of (layout_of aprog) as", 
       
  4635       auto simp: new_tape.simps)
       
  4636 done
       
  4637 
       
  4638 lemma crsp_next_state:
       
  4639   "\<lbrakk>crsp_l (layout_of aprog) (as, am) tc ires; 
       
  4640     abc_fetch as aprog = Some (Dec n e)\<rbrakk>
       
  4641   \<Longrightarrow> \<exists> stp' > 0. (\<lambda> (s, l, r). 
       
  4642            (s = Suc (start_of (layout_of aprog) as) 
       
  4643  \<and> (dec_inv_1 (layout_of aprog) n e (as, am) (s, l, r) ires) 
       
  4644  \<and> (dec_inv_2 (layout_of aprog) n e (as, am) (s, l, r)) ires)) 
       
  4645      (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as)
       
  4646              (Dec n e), start_of (layout_of aprog) as - Suc 0) stp')"
       
  4647 apply(subgoal_tac "start_of (layout_of aprog) as > 0")
       
  4648 apply(case_tac tc, case_tac b, auto simp: crsp_l.simps)
       
  4649 apply(case_tac "am = []", simp)
       
  4650 apply(rule_tac x = "Suc (Suc 0)" in exI, simp add: t_steps.simps)
       
  4651 proof-
       
  4652   fix  rn
       
  4653   assume h1: "am \<noteq> []" "abc_fetch as aprog = Some (Dec n e)" 
       
  4654              "start_of (layout_of aprog) as > 0"
       
  4655   hence h2: "\<exists> r'. <am> = Oc # r'"
       
  4656     by simp
       
  4657   from h1 and h2 show 
       
  4658    "\<exists>stp'>0. case t_steps (start_of (layout_of aprog) as, Bk # Bk # ires, <am> @ Bk\<^bsup>rn\<^esup>)
       
  4659     (ci (layout_of aprog) (start_of (layout_of aprog) as) (Dec n e),
       
  4660     start_of (layout_of aprog) as - Suc 0) stp' of
       
  4661     (s, ab) \<Rightarrow> s = Suc (start_of (layout_of aprog) as) \<and>
       
  4662     dec_inv_1 (layout_of aprog) n e (as, am) (s, ab) ires \<and> 
       
  4663     dec_inv_2 (layout_of aprog) n e (as, am) (s, ab) ires"
       
  4664   proof(erule_tac exE, simp, rule_tac x = "Suc 0" in exI, 
       
  4665         simp add: t_steps.simps)
       
  4666   qed
       
  4667 next
       
  4668   assume "abc_fetch as aprog = Some (Dec n e)"
       
  4669   thus "0 < start_of (layout_of aprog) as"
       
  4670    apply(insert startof_not0[of "layout_of aprog" as], simp)
       
  4671    done
       
  4672 qed
       
  4673 
       
  4674 lemma dec_crsp_ex1: 
       
  4675   "\<lbrakk>crsp_l (layout_of aprog) (as, am) tc ires;
       
  4676   abc_fetch as aprog = Some (Dec n e); 
       
  4677   abc_lm_v am n = 0\<rbrakk>
       
  4678   \<Longrightarrow> \<exists>stp > 0. crsp_l (layout_of aprog) (e, abc_lm_s am n 0) 
       
  4679     (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as)
       
  4680  (Dec n e), start_of (layout_of aprog) as - Suc 0) stp) ires"
       
  4681 proof -
       
  4682   assume h1: "crsp_l (layout_of aprog) (as, am) tc ires" 
       
  4683        "abc_fetch as aprog = Some (Dec n e)" "abc_lm_v am n = 0"
       
  4684   hence h2: "\<exists> stp' > 0. (\<lambda> (s, l, r). 
       
  4685     (s = Suc (start_of (layout_of aprog) as) \<and> 
       
  4686  (dec_inv_1 (layout_of aprog) n e (as, am) (s, l, r)) ires)) 
       
  4687    (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  4688       (Dec n e), start_of (layout_of aprog) as - Suc 0) stp')"
       
  4689     apply(insert crsp_next_state[of aprog as am tc ires n e], auto)
       
  4690     done
       
  4691   from h1 and h2 show 
       
  4692  "\<exists>stp > 0. crsp_l (layout_of aprog) (e, abc_lm_s am n 0) 
       
  4693            (t_steps tc (ci (layout_of aprog) (start_of 
       
  4694                 (layout_of aprog) as) (Dec n e), 
       
  4695                     start_of (layout_of aprog) as - Suc 0) stp) ires" 
       
  4696   proof(erule_tac exE, case_tac "(t_steps tc (ci (layout_of aprog)
       
  4697        (start_of (layout_of aprog) as) (Dec n e), start_of 
       
  4698           (layout_of aprog) as - Suc 0) stp')",  simp)
       
  4699     fix stp' a b c
       
  4700     assume h3: "stp' > 0 \<and> a = Suc (start_of (layout_of aprog) as) \<and> 
       
  4701                dec_inv_1 (layout_of aprog) n e (as, am) (a, b, c) ires" 
       
  4702              "abc_fetch as aprog = Some (Dec n e)" "abc_lm_v am n = 0"
       
  4703      "t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as)
       
  4704           (Dec n e), start_of (layout_of aprog) as - Suc 0) stp' 
       
  4705         = (Suc (start_of (layout_of aprog) as), b, c)" 
       
  4706     thus "\<exists>stp > 0. crsp_l (layout_of aprog) (e, abc_lm_s am n 0) 
       
  4707      (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as)
       
  4708            (Dec n e), start_of (layout_of aprog) as - Suc 0) stp) ires"
       
  4709     proof(erule_tac conjE, simp)
       
  4710       assume "dec_inv_1 (layout_of aprog) n e (as, am) 
       
  4711                     (Suc (start_of (layout_of aprog) as), b, c) ires"     
       
  4712              "abc_fetch as aprog = Some (Dec n e)" 
       
  4713              "abc_lm_v am n = 0"
       
  4714              " t_steps tc (ci (layout_of aprog) 
       
  4715               (start_of (layout_of aprog) as) (Dec n e), 
       
  4716                start_of (layout_of aprog) as - Suc 0) stp' 
       
  4717              = (Suc (start_of (layout_of aprog) as), b, c)"
       
  4718       hence h4: "\<exists>stp. (\<lambda>(s', l', r'). s' = 
       
  4719                      start_of (layout_of aprog) e \<and> 
       
  4720                 dec_inv_1 (layout_of aprog) n e (as, am) (s', l', r') ires)
       
  4721                  (t_steps (start_of (layout_of aprog) as + 1, b, c) 
       
  4722                   (ci (layout_of aprog) 
       
  4723                       (start_of (layout_of aprog) as) (Dec n e), 
       
  4724                          start_of (layout_of aprog) as - Suc 0) stp)"
       
  4725 	apply(rule_tac dec_inv_stop1, auto)
       
  4726 	done
       
  4727       from  h3 and h4 show ?thesis
       
  4728 	apply(erule_tac exE)
       
  4729 	apply(rule_tac x = "stp' + stp" in exI, simp)
       
  4730 	apply(case_tac "(t_steps (Suc (start_of (layout_of aprog) as),
       
  4731                      b, c) (ci (layout_of aprog) 
       
  4732                      (start_of (layout_of aprog) as) (Dec n e), 
       
  4733                       start_of (layout_of aprog) as - Suc 0) stp)", 
       
  4734               simp)
       
  4735 	apply(rule_tac dec_inv_stop_cond1, auto)
       
  4736 	done
       
  4737     qed
       
  4738   qed
       
  4739 qed
       
  4740 	  
       
  4741 lemma dec_crsp_ex2:
       
  4742   "\<lbrakk>crsp_l (layout_of aprog) (as, am) tc ires; 
       
  4743     abc_fetch as aprog = Some (Dec n e);
       
  4744     0 < abc_lm_v am n\<rbrakk>
       
  4745  \<Longrightarrow> \<exists>stp > 0. crsp_l (layout_of aprog) 
       
  4746                (Suc as, abc_lm_s am n (abc_lm_v am n - Suc 0))
       
  4747    (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  4748               (Dec n e), start_of (layout_of aprog) as - Suc 0) stp) ires"
       
  4749 proof -
       
  4750   assume h1: 
       
  4751  "crsp_l (layout_of aprog) (as, am) tc ires" 
       
  4752  "abc_fetch as aprog = Some (Dec n e)"
       
  4753   "abc_lm_v am n > 0"
       
  4754   hence h2: 
       
  4755  "\<exists> stp' > 0. (\<lambda> (s, l, r). (s = Suc (start_of (layout_of aprog) as)
       
  4756  \<and> (dec_inv_2 (layout_of aprog) n e (as, am) (s, l, r)) ires)) 
       
  4757 (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  4758               (Dec n e), start_of (layout_of aprog) as - Suc 0) stp')"
       
  4759     apply(insert crsp_next_state[of aprog as am tc ires n e], auto)
       
  4760     done
       
  4761   from h1 and h2 show 
       
  4762  "\<exists>stp >0. crsp_l (layout_of aprog) 
       
  4763    (Suc as, abc_lm_s am n (abc_lm_v am n - Suc 0))
       
  4764    (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  4765                (Dec n e), start_of (layout_of aprog) as - Suc 0) stp) ires"
       
  4766   proof(erule_tac exE, 
       
  4767         case_tac 
       
  4768  "(t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  4769       (Dec n e), start_of (layout_of aprog) as - Suc 0) stp')",  simp)
       
  4770     fix stp' a b c
       
  4771     assume h3: "0 < stp' \<and> a = Suc (start_of (layout_of aprog) as) \<and>
       
  4772                dec_inv_2 (layout_of aprog) n e (as, am) (a, b, c) ires" 
       
  4773                "abc_fetch as aprog = Some (Dec n e)" 
       
  4774                "abc_lm_v am n > 0"
       
  4775                "t_steps tc (ci (layout_of aprog) 
       
  4776                    (start_of (layout_of aprog) as) (Dec n e), 
       
  4777                      start_of (layout_of aprog) as - Suc 0) stp' 
       
  4778                   = (Suc (start_of (layout_of aprog) as), b, c)"
       
  4779     thus "?thesis"
       
  4780     proof(erule_tac conjE, simp)
       
  4781       assume 
       
  4782     "dec_inv_2 (layout_of aprog) n e (as, am) 
       
  4783       (Suc (start_of (layout_of aprog) as), b, c) ires" 
       
  4784     "abc_fetch as aprog = Some (Dec n e)" "abc_lm_v am n > 0"
       
  4785     "t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as)
       
  4786          (Dec n e), start_of (layout_of aprog) as - Suc 0) stp'
       
  4787              = (Suc (start_of (layout_of aprog) as), b, c)"
       
  4788       hence h4: 
       
  4789    "\<exists>stp. (\<lambda>(s', l', r'). s' = start_of (layout_of aprog) (Suc as) \<and>
       
  4790            dec_inv_2 (layout_of aprog) n e (as, am) (s', l', r') ires)
       
  4791              (t_steps (start_of (layout_of aprog) as + 1, b, c) 
       
  4792               (ci (layout_of aprog) (start_of (layout_of aprog) as) 
       
  4793                (Dec n e), start_of (layout_of aprog) as - Suc 0) stp)"
       
  4794 	apply(rule_tac dec_stop2, auto)
       
  4795 	done
       
  4796       from  h3 and h4 show ?thesis
       
  4797 	apply(erule_tac exE)
       
  4798 	apply(rule_tac x = "stp' + stp" in exI, simp)
       
  4799 	apply(case_tac 
       
  4800          "(t_steps (Suc (start_of (layout_of aprog) as), b, c) 
       
  4801            (ci (layout_of aprog) (start_of (layout_of aprog) as)
       
  4802              (Dec n e), start_of (layout_of aprog) as - Suc 0) stp)"
       
  4803               ,simp)
       
  4804 	apply(rule_tac dec_inv_stop_cond2, auto)
       
  4805 	done
       
  4806     qed
       
  4807   qed
       
  4808 qed
       
  4809 
       
  4810 lemma dec_crsp_ex_pre:
       
  4811   "\<lbrakk>ly = layout_of aprog; crsp_l ly (as, am) tc ires; 
       
  4812      abc_fetch as aprog = Some (Dec n e)\<rbrakk>
       
  4813  \<Longrightarrow> \<exists>stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Dec n e))) 
       
  4814       (t_steps tc (ci (layout_of aprog) (start_of ly as) (Dec n e),
       
  4815                                        start_of ly as - Suc 0) stp) ires"
       
  4816 apply(auto simp: abc_step_l.simps intro: dec_crsp_ex2 dec_crsp_ex1)
       
  4817 done
       
  4818 
       
  4819 lemma dec_crsp_ex:
       
  4820   assumes layout: -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"} *}
       
  4821   "ly = layout_of aprog"
       
  4822   and dec: -- {* There is an @{text "Dec n e"} instruction at postion @{text "as"} of @{text "aprog"} *}
       
  4823       "abc_fetch as aprog = Some (Dec n e)"
       
  4824   and correspond: 
       
  4825   -- {* Abacus configuration @{text "(as, am)"} is in correspondence with TM 
       
  4826          configuration @{text "tc"}
       
  4827       *}
       
  4828   "crsp_l ly (as, am) tc ires"
       
  4829 shows 
       
  4830    "\<exists>stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Dec n e))) 
       
  4831       (t_steps tc (ci (layout_of aprog) (start_of ly as) (Dec n e),
       
  4832                                        start_of ly as - Suc 0) stp) ires"
       
  4833 proof -
       
  4834   from dec_crsp_ex_pre layout dec correspond  show ?thesis by blast
       
  4835 qed
       
  4836 
       
  4837 
       
  4838 (*******End: dec crsp********)
       
  4839 
       
  4840 
       
  4841 subsubsection {* Compilation of @{text "Goto n"}*}
       
  4842 
       
  4843 
       
  4844 (*******Begin: goto crsp********)
       
  4845 lemma goto_fetch: 
       
  4846      "fetch (ci (layout_of aprog) 
       
  4847          (start_of (layout_of aprog) as) (Goto n)) (Suc 0)  b
       
  4848      = (Nop, start_of (layout_of aprog) n)"
       
  4849 apply(auto simp: ci.simps fetch.simps nth_of.simps 
       
  4850            split: block.splits)
       
  4851 done
       
  4852 
       
  4853 text {*
       
  4854   Correctness of complied @{text "Goto n"}
       
  4855   *}
       
  4856 
       
  4857 lemma goto_crsp_ex_pre: 
       
  4858   "\<lbrakk>ly = layout_of aprog; 
       
  4859     crsp_l ly (as, am) tc ires;
       
  4860     abc_fetch as aprog = Some (Goto n)\<rbrakk>
       
  4861  \<Longrightarrow> \<exists>stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Goto n))) 
       
  4862       (t_steps tc (ci (layout_of aprog) (start_of ly as) (Goto n), 
       
  4863                                         start_of ly as - Suc 0) stp) ires"
       
  4864 apply(rule_tac x = 1 in exI)
       
  4865 apply(simp add: abc_step_l.simps t_steps.simps t_step.simps)
       
  4866 apply(case_tac tc, simp)
       
  4867 apply(subgoal_tac "a = start_of (layout_of aprog) as", auto)
       
  4868 apply(subgoal_tac "start_of (layout_of aprog) as > 0", simp)
       
  4869 apply(auto simp: goto_fetch new_tape.simps crsp_l.simps)
       
  4870 apply(rule startof_not0)
       
  4871 done
       
  4872 
       
  4873 lemma goto_crsp_ex:
       
  4874   assumes layout: "ly = layout_of aprog"
       
  4875   and goto: "abc_fetch as aprog = Some (Goto n)"
       
  4876   and correspondence: "crsp_l ly (as, am) tc ires"
       
  4877   shows "\<exists>stp>0. crsp_l ly (abc_step_l (as, am) (Some (Goto n))) 
       
  4878               (t_steps tc (ci (layout_of aprog) (start_of ly as) (Goto n),
       
  4879                                            start_of ly as - Suc 0) stp) ires"
       
  4880 proof -
       
  4881   from goto_crsp_ex_pre and layout goto correspondence show "?thesis" by blast
       
  4882 qed
       
  4883 (*******End : goto crsp*********)
       
  4884   
       
  4885 subsubsection {*
       
  4886   The correctness of the compiler
       
  4887   *}
       
  4888 
       
  4889 declare abc_step_l.simps[simp del]
       
  4890 
       
  4891 lemma tm_crsp_ex: 
       
  4892          "\<lbrakk>ly = layout_of aprog;
       
  4893            crsp_l ly (as, am) tc ires; 
       
  4894            as < length aprog;
       
  4895            abc_fetch as aprog = Some ins\<rbrakk>
       
  4896       \<Longrightarrow> \<exists> n > 0. crsp_l ly (abc_step_l (as,am) (Some ins))
       
  4897                (t_steps tc (ci (layout_of aprog) (start_of ly as) 
       
  4898                   (ins), (start_of ly as) - (Suc 0)) n) ires"
       
  4899 apply(case_tac "ins", simp)
       
  4900 apply(auto intro: inc_crsp_ex_pre dec_crsp_ex goto_crsp_ex)
       
  4901 done
       
  4902 
       
  4903 lemma start_of_pre: 
       
  4904   "n < length aprog \<Longrightarrow> start_of (layout_of aprog) n
       
  4905                      = start_of (layout_of (butlast aprog)) n"
       
  4906 apply(induct n, simp add: start_of.simps, simp)
       
  4907 apply(simp add: layout_of.simps start_of.simps)
       
  4908 apply(subgoal_tac "n < length aprog - Suc 0", simp)
       
  4909 apply(subgoal_tac "(aprog ! n) = (butlast aprog ! n)", simp)
       
  4910 proof -
       
  4911   fix n
       
  4912   assume h1: "Suc n < length aprog"
       
  4913   thus "aprog ! n = butlast aprog ! n"
       
  4914     apply(case_tac "length aprog", simp, simp)
       
  4915     apply(insert nth_append[of "butlast aprog" "[last aprog]" n])
       
  4916     apply(subgoal_tac "(butlast aprog @ [last aprog]) = aprog")
       
  4917     apply(simp split: if_splits)
       
  4918     apply(rule append_butlast_last_id, case_tac aprog, simp, simp)
       
  4919     done
       
  4920 next
       
  4921   fix n
       
  4922   assume "Suc n < length aprog"
       
  4923   thus "n < length aprog - Suc 0"
       
  4924     apply(case_tac aprog, simp, simp)
       
  4925     done
       
  4926 qed
       
  4927     
       
  4928 lemma zip_eq: "xs = ys \<Longrightarrow> zip xs zs = zip ys zs"
       
  4929 by simp
       
  4930 
       
  4931 lemma tpairs_of_append_iff: "length aprog = Suc n \<Longrightarrow> 
       
  4932          tpairs_of aprog = tpairs_of (butlast aprog) @ 
       
  4933                      [(start_of (layout_of aprog) n, aprog ! n)]"
       
  4934 apply(simp add: tpairs_of.simps)
       
  4935 apply(insert zip_append[of "map (start_of (layout_of aprog)) [0..<n]"
       
  4936      "butlast aprog" "[start_of (layout_of aprog) n]" "[last aprog]"])
       
  4937 apply(simp del: zip_append)
       
  4938 apply(subgoal_tac "(butlast aprog @ [last aprog]) = aprog", auto)
       
  4939 apply(rule_tac zip_eq, auto)
       
  4940 apply(rule_tac start_of_pre, simp)
       
  4941 apply(insert last_conv_nth[of aprog], case_tac aprog, simp, simp)
       
  4942 apply(rule append_butlast_last_id, case_tac aprog, simp, simp)
       
  4943 done
       
  4944 
       
  4945 lemma [simp]: "list_all (\<lambda>(n, tm). abacus.t_ncorrect (ci layout n tm))
       
  4946          (zip (map (start_of layout) [0..<length aprog]) aprog)"
       
  4947 proof(induct "length aprog" arbitrary: aprog, simp)
       
  4948   fix x aprog
       
  4949   assume ind: "\<And>aprog. x = length aprog \<Longrightarrow> 
       
  4950         list_all (\<lambda>(n, tm). abacus.t_ncorrect (ci layout n tm))
       
  4951            (zip (map (start_of layout) [0..<length aprog]) aprog)"
       
  4952   and h: "Suc x = length (aprog::abc_inst list)"
       
  4953   have g1: "list_all (\<lambda>(n, tm). abacus.t_ncorrect (ci layout n tm)) 
       
  4954     (zip (map (start_of layout) [0..<length (butlast aprog)]) 
       
  4955                                                  (butlast aprog))"
       
  4956     using h
       
  4957     apply(rule_tac ind, auto)
       
  4958     done
       
  4959   have g2: "(map (start_of layout) [0..<length aprog]) = 
       
  4960                      map (start_of layout) ([0..<length aprog - 1] 
       
  4961          @ [length aprog - 1])"
       
  4962     using h
       
  4963     apply(case_tac aprog, simp, simp)
       
  4964     done
       
  4965   have "\<exists> xs a. aprog = xs @ [a]"
       
  4966     using h
       
  4967     apply(rule_tac x = "butlast aprog" in exI, 
       
  4968           rule_tac x = "last aprog" in exI)
       
  4969     apply(case_tac "aprog = []", simp, simp)
       
  4970     done
       
  4971   from this obtain xs where "\<exists> a. aprog = xs @ [a]" ..
       
  4972   from this obtain a where g3: "aprog = xs @ [a]" ..
       
  4973   from g1 and g2 and g3 show "list_all (\<lambda>(n, tm). 
       
  4974                               abacus.t_ncorrect (ci layout n tm)) 
       
  4975               (zip (map (start_of layout) [0..<length aprog]) aprog)"
       
  4976     apply(simp)
       
  4977     apply(auto simp: t_ncorrect.simps ci.simps  tshift.simps 
       
  4978           tinc_b_def tdec_b_def split: abc_inst.splits)
       
  4979     apply arith+
       
  4980     done
       
  4981 qed
       
  4982 
       
  4983 lemma [intro]: "abc2t_correct aprog"
       
  4984 apply(simp add: abc2t_correct.simps tpairs_of.simps 
       
  4985           split: abc_inst.splits)
       
  4986 done
       
  4987 
       
  4988 lemma as_out: "\<lbrakk>ly = layout_of aprog; tprog = tm_of aprog; 
       
  4989                 crsp_l ly (as, am) tc ires; length aprog \<le> as\<rbrakk> 
       
  4990             \<Longrightarrow> abc_step_l (as, am) (abc_fetch as aprog) = (as, am)"
       
  4991 apply(simp add: abc_fetch.simps abc_step_l.simps)
       
  4992 done
       
  4993 
       
  4994 lemma tm_merge_ex: 
       
  4995   "\<lbrakk>crsp_l (layout_of aprog) (as, am) tc ires; 
       
  4996     as < length aprog; 
       
  4997     abc_fetch as aprog = Some a; 
       
  4998     abc2t_correct aprog;
       
  4999     crsp_l (layout_of aprog) (abc_step_l (as, am) (Some a))
       
  5000      (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as)
       
  5001          a, start_of (layout_of aprog) as - Suc 0) n) ires; 
       
  5002     n > 0\<rbrakk>
       
  5003    \<Longrightarrow> \<exists>stp > 0. crsp_l (layout_of aprog) (abc_step_l (as, am) 
       
  5004                        (Some a)) (t_steps tc (tm_of aprog, 0) stp) ires"
       
  5005 apply(case_tac "(t_steps tc (ci (layout_of aprog) 
       
  5006            (start_of (layout_of aprog) as) a, 
       
  5007             start_of (layout_of aprog) as - Suc 0) n)",  simp)
       
  5008 apply(case_tac "(abc_step_l (as, am) (Some a))", simp)
       
  5009 proof -
       
  5010   fix aa b c aaa ba 
       
  5011   assume h: 
       
  5012   "crsp_l (layout_of aprog) (as, am) tc ires" 
       
  5013   "as < length aprog" 
       
  5014   "abc_fetch as aprog = Some a" 
       
  5015   "crsp_l (layout_of aprog) (aaa, ba) (aa, b, c) ires" 
       
  5016   "abc2t_correct aprog" 
       
  5017   "n>0"
       
  5018   "t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) a,
       
  5019       start_of (layout_of aprog) as - Suc 0) n = (aa, b, c)" 
       
  5020    "abc_step_l (as, am) (Some a) = (aaa, ba)"
       
  5021   hence "aa = start_of (layout_of aprog) aaa"
       
  5022     apply(simp add: crsp_l.simps)
       
  5023     done
       
  5024   from this and h show 
       
  5025   "\<exists>stp > 0. crsp_l (layout_of aprog) (aaa, ba) 
       
  5026                           (t_steps tc (tm_of aprog, 0) stp) ires"
       
  5027     apply(insert tms_out_ex[of "layout_of aprog" aprog 
       
  5028                 "tm_of aprog" as am tc ires a n aa b c aaa ba], auto)
       
  5029     done
       
  5030 qed
       
  5031  
       
  5032 lemma crsp_inside: 
       
  5033   "\<lbrakk>ly = layout_of aprog; 
       
  5034     tprog = tm_of aprog;
       
  5035     crsp_l ly (as, am) tc ires;
       
  5036     as < length aprog\<rbrakk> \<Longrightarrow> 
       
  5037     (\<exists> stp > 0. crsp_l ly (abc_step_l (as,am) (abc_fetch as aprog)) 
       
  5038                                          (t_steps tc (tprog, 0) stp) ires)"
       
  5039 apply(case_tac "abc_fetch as aprog", simp add: abc_fetch.simps)
       
  5040 proof -
       
  5041   fix a
       
  5042   assume "ly = layout_of aprog" 
       
  5043      "tprog = tm_of aprog" 
       
  5044      "crsp_l ly (as, am) tc ires" 
       
  5045      "as < length aprog" 
       
  5046      "abc_fetch as aprog = Some a"
       
  5047   thus "\<exists>stp > 0. crsp_l ly (abc_step_l (as, am) 
       
  5048                  (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires"
       
  5049     proof(insert tm_crsp_ex[of ly aprog as am tc ires a], 
       
  5050           auto intro: tm_merge_ex)
       
  5051   qed
       
  5052 qed
       
  5053 
       
  5054 lemma crsp_outside: 
       
  5055   "\<lbrakk>ly = layout_of aprog; tprog = tm_of aprog;
       
  5056     crsp_l ly (as, am) tc ires; as \<ge> length aprog\<rbrakk>
       
  5057     \<Longrightarrow> (\<exists> stp. crsp_l ly (abc_step_l (as,am) (abc_fetch as aprog)) 
       
  5058                                          (t_steps tc (tprog, 0) stp) ires)"
       
  5059 apply(subgoal_tac "abc_step_l (as, am) (abc_fetch as aprog)
       
  5060                 = (as, am)", simp)
       
  5061 apply(rule_tac x = 0 in exI, simp add: t_steps.simps)
       
  5062 apply(rule as_out, simp+)
       
  5063 done
       
  5064 
       
  5065 text {*
       
  5066   Single-step correntess of the compiler.
       
  5067 *}
       
  5068 lemma astep_crsp_pre: 
       
  5069       "\<lbrakk>ly = layout_of aprog; 
       
  5070         tprog = tm_of aprog;
       
  5071         crsp_l ly (as, am) tc ires\<rbrakk>
       
  5072        \<Longrightarrow> (\<exists> stp. crsp_l ly (abc_step_l (as,am) 
       
  5073                   (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires)"
       
  5074 apply(case_tac "as < length aprog")
       
  5075 apply(drule_tac crsp_inside, auto)
       
  5076 apply(rule_tac crsp_outside, simp+)
       
  5077 done
       
  5078 
       
  5079 text {*
       
  5080   Single-step correntess of the compiler.
       
  5081 *}
       
  5082 lemma astep_crsp_pre1: 
       
  5083       "\<lbrakk>ly = layout_of aprog;
       
  5084         tprog = tm_of aprog;
       
  5085         crsp_l ly (as, am) tc ires\<rbrakk>
       
  5086        \<Longrightarrow> (\<exists> stp. crsp_l ly (abc_step_l (as,am) 
       
  5087                   (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires)"
       
  5088 apply(case_tac "as < length aprog")
       
  5089 apply(drule_tac crsp_inside, auto)
       
  5090 apply(rule_tac crsp_outside, simp+)
       
  5091 done
       
  5092 
       
  5093 lemma astep_crsp:
       
  5094   assumes layout: 
       
  5095   -- {* There is a Abacus program @{text "aprog"} with layout @{text "ly"} *}
       
  5096   "ly = layout_of aprog"
       
  5097   and compiled: 
       
  5098   -- {* @{text "tprog"} is the TM compiled from @{text "aprog"} *}
       
  5099   "tprog = tm_of aprog"
       
  5100   and corresponds: 
       
  5101   -- {* Abacus configuration @{text "(as, am)"} is in correspondence with TM configuration
       
  5102    @{text "tc"} *}
       
  5103   "crsp_l ly (as, am) tc ires"
       
  5104   -- {* One step execution of @{text "aprog"} can be simulated by multi-step execution 
       
  5105   of @{text "tprog"} *}
       
  5106   shows "(\<exists> stp. crsp_l ly (abc_step_l (as,am) 
       
  5107                   (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires)"
       
  5108 proof -
       
  5109   from astep_crsp_pre1 [OF layout compiled corresponds] show ?thesis .
       
  5110 qed
       
  5111 
       
  5112 lemma steps_crsp_pre: 
       
  5113     "\<lbrakk>ly = layout_of aprog; tprog = tm_of aprog; 
       
  5114       crsp_l ly ac tc ires; ac' = abc_steps_l ac aprog n\<rbrakk> \<Longrightarrow> 
       
  5115         (\<exists> n'. crsp_l ly ac' (t_steps tc (tprog, 0) n') ires)"
       
  5116 apply(induct n arbitrary: ac' ac tc, simp add: abc_steps_l.simps)
       
  5117 apply(rule_tac x = 0 in exI)
       
  5118 apply(case_tac ac, simp add: abc_steps_l.simps t_steps.simps)
       
  5119 apply(case_tac ac, simp add: abc_steps_l.simps)
       
  5120 apply(subgoal_tac 
       
  5121    "(\<exists> stp. crsp_l ly (abc_step_l (a, b)
       
  5122             (abc_fetch a aprog)) (t_steps tc (tprog, 0) stp) ires)")
       
  5123 apply(erule exE)
       
  5124 apply(subgoal_tac 
       
  5125    "\<exists>n'. crsp_l (layout_of aprog) 
       
  5126     (abc_steps_l (abc_step_l (a, b) (abc_fetch a aprog)) aprog n)
       
  5127          (t_steps ((t_steps tc (tprog, 0) stp)) (tm_of aprog, 0) n') ires")
       
  5128 apply(erule exE)
       
  5129 apply(subgoal_tac 
       
  5130     "t_steps (t_steps tc (tprog, 0) stp) (tm_of aprog, 0) n' =
       
  5131      t_steps tc (tprog, 0) (stp + n')")
       
  5132 apply(rule_tac x = "stp + n'" in exI, simp)
       
  5133 apply(auto intro: astep_crsp simp: t_step_add)
       
  5134 done
       
  5135 
       
  5136 text {*
       
  5137   Multi-step correctess of the compiler.
       
  5138 *}
       
  5139 
       
  5140 lemma steps_crsp: 
       
  5141   assumes layout: 
       
  5142   -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"} *}
       
  5143     "ly = layout_of aprog"
       
  5144   and compiled: 
       
  5145   -- {* @{text "tprog"} is the TM compiled from @{text "aprog"} *}
       
  5146   "tprog = tm_of aprog"
       
  5147   and correspond: 
       
  5148   -- {* Abacus configuration @{text "ac"} is in correspondence with TM configuration @{text "tc"} *}
       
  5149       "crsp_l ly ac tc ires"
       
  5150   and execution: 
       
  5151   -- {* @{text "ac'"} is the configuration obtained from @{text "n"}-step execution 
       
  5152       of @{text "aprog"} starting from configuration @{text "ac"} *}
       
  5153   "ac' = abc_steps_l ac aprog n" 
       
  5154   -- {* There exists steps @{text "n'"} steps, after these steps of execution, 
       
  5155   the Turing configuration such obtained is in correspondence with @{text "ac'"} *}
       
  5156   shows "(\<exists> n'. crsp_l ly ac' (t_steps tc (tprog, 0) n') ires)"
       
  5157 proof -
       
  5158   from steps_crsp_pre [OF layout compiled correspond execution] show ?thesis .
       
  5159 qed
       
  5160 
       
  5161 
       
  5162 subsubsection {* The Mop-up machine *}
       
  5163 
       
  5164 fun mop_bef :: "nat \<Rightarrow> tprog"
       
  5165   where
       
  5166   "mop_bef 0 = []" |
       
  5167   "mop_bef (Suc n) = mop_bef n @ 
       
  5168        [(R, 2*n + 3), (W0, 2*n + 2), (R, 2*n + 1), (W1, 2*n + 2)]"
       
  5169 
       
  5170 definition mp_up :: "tprog"
       
  5171   where
       
  5172   "mp_up \<equiv> [(R, 2), (R, 1), (L, 5), (W0, 3), (R, 4), (W0, 3),
       
  5173             (R, 2), (W0, 3), (L, 5), (L, 6), (R, 0), (L, 6)]"
       
  5174 
       
  5175 fun tMp :: "nat \<Rightarrow> nat \<Rightarrow> tprog"
       
  5176   where 
       
  5177   "tMp n off = tshift (mop_bef n @ tshift mp_up (2*n)) off"
       
  5178 
       
  5179 declare  mp_up_def[simp del]  tMp.simps[simp del] mop_bef.simps[simp del]
       
  5180 (**********Begin: equiv among aba and turing***********)
       
  5181 
       
  5182 lemma tm_append_step: 
       
  5183  "\<lbrakk>t_ncorrect tp1; t_step tc (tp1, 0) = (s, l, r); s \<noteq> 0\<rbrakk> 
       
  5184  \<Longrightarrow> t_step tc (tp1 @ tp2, 0) = (s, l, r)"
       
  5185 apply(simp add: t_step.simps)
       
  5186 apply(case_tac tc, simp)
       
  5187 apply(case_tac 
       
  5188        "(fetch tp1 a (case c of [] \<Rightarrow> Bk |
       
  5189                Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))", simp)
       
  5190 apply(case_tac a, simp add: fetch.simps)
       
  5191 apply(simp add: fetch.simps)
       
  5192 apply(case_tac c, simp)
       
  5193 apply(case_tac [!] "ab::block")
       
  5194 apply(auto simp: nth_of.simps nth_append t_ncorrect.simps 
       
  5195            split: if_splits)
       
  5196 done
       
  5197 
       
  5198 lemma state0_ind: "t_steps (0, l, r) (tp, 0) stp = (0, l, r)"
       
  5199 apply(induct stp, simp add: t_steps.simps)
       
  5200 apply(simp add: t_steps.simps t_step.simps fetch.simps new_tape.simps)
       
  5201 done
       
  5202 
       
  5203 lemma tm_append_steps:  
       
  5204  "\<lbrakk>t_ncorrect tp1; t_steps tc (tp1, 0) stp = (s, l ,r); s \<noteq> 0\<rbrakk>
       
  5205   \<Longrightarrow> t_steps tc (tp1 @ tp2, 0) stp = (s, l, r)"
       
  5206 apply(induct stp arbitrary: tc s l r)
       
  5207 apply(case_tac tc,  simp)
       
  5208 apply(simp add: t_steps.simps)
       
  5209 proof -
       
  5210   fix stp tc s l r
       
  5211   assume h1: "\<And>tc s l r. \<lbrakk>t_ncorrect tp1; t_steps tc (tp1, 0) stp = 
       
  5212    (s, l, r); s \<noteq> 0\<rbrakk> \<Longrightarrow> t_steps tc (tp1 @ tp2, 0) stp = (s, l, r)"
       
  5213     and h2: "t_steps tc (tp1, 0) (Suc stp) = (s, l, r)" "s \<noteq> 0" 
       
  5214             "t_ncorrect tp1"
       
  5215   thus "t_steps tc (tp1 @ tp2, 0) (Suc stp) = (s, l, r)"
       
  5216     apply(simp add: t_steps.simps)
       
  5217     apply(case_tac "(t_step tc (tp1, 0))", simp)
       
  5218     proof-
       
  5219       fix a b c 
       
  5220       assume g1: "\<And>tc s l r. \<lbrakk>t_steps tc (tp1, 0) stp = (s, l, r); 
       
  5221                 0 < s\<rbrakk> \<Longrightarrow> t_steps tc (tp1 @ tp2, 0) stp = (s, l, r)"
       
  5222 	and g2: "t_step tc (tp1, 0) = (a, b, c)" 
       
  5223                 "t_steps (a, b, c) (tp1, 0) stp = (s, l, r)" 
       
  5224                 "0 < s" 
       
  5225                 "t_ncorrect tp1"
       
  5226       hence g3: "a > 0"
       
  5227 	apply(case_tac "a::nat", auto simp: t_steps.simps)
       
  5228 	apply(simp add: state0_ind)
       
  5229 	done
       
  5230       from g1 and g2 and this have g4: 
       
  5231                     "(t_step tc (tp1 @ tp2, 0)) = (a, b, c)"
       
  5232 	apply(rule_tac tm_append_step, simp, simp, simp)
       
  5233 	done
       
  5234       from g1 and g2 and g3 and g4 show 
       
  5235           "t_steps (t_step tc (tp1 @ tp2, 0)) (tp1 @ tp2, 0) stp
       
  5236                                                          = (s, l, r)"
       
  5237 	apply(simp)
       
  5238 	done
       
  5239     qed
       
  5240 qed
       
  5241 
       
  5242 lemma shift_fetch: 
       
  5243  "\<lbrakk>n < length tp; 
       
  5244   (tp:: (taction \<times> nat) list) ! n = (aa, ba);
       
  5245    ba \<noteq> 0\<rbrakk> 
       
  5246    \<Longrightarrow> (tshift tp (length tp div 2)) ! n = 
       
  5247                      (aa , ba + length tp div 2)"
       
  5248 apply(simp add: tshift.simps)
       
  5249 done
       
  5250 
       
  5251 lemma tshift_length_equal: "length (tshift tp q) = length tp"
       
  5252 apply(auto simp: tshift.simps)
       
  5253 done
       
  5254 
       
  5255 thm nth_of.simps
       
  5256 
       
  5257 lemma [simp]: "t_ncorrect tp \<Longrightarrow> 2 * (length tp div 2) = length tp"
       
  5258 apply(auto simp: t_ncorrect.simps)
       
  5259 done
       
  5260 
       
  5261 lemma  tm_append_step_equal': 
       
  5262    "\<lbrakk>t_ncorrect tp; t_ncorrect tp'; off = length tp div 2\<rbrakk> \<Longrightarrow> 
       
  5263     (\<lambda> (s, l, r). ((\<lambda> (s', l', r'). 
       
  5264       (s'\<noteq> 0 \<longrightarrow> (s = s' + off \<and> l = l' \<and> r = r'))) 
       
  5265          (t_step (a, b, c) (tp', 0))))
       
  5266                (t_step (a + off, b, c) (tp @ tshift tp' off, 0))"
       
  5267 apply(simp add: t_step.simps)
       
  5268 apply(case_tac a, simp add: fetch.simps)
       
  5269 apply(case_tac 
       
  5270 "(fetch tp' a (case c of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))",
       
  5271  simp)
       
  5272 apply(case_tac 
       
  5273 "(fetch (tp @ tshift tp' (length tp div 2))
       
  5274         (Suc (nat + length tp div 2)) 
       
  5275            (case c of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))", 
       
  5276  simp)
       
  5277 apply(case_tac "(new_tape aa (b, c))",
       
  5278       case_tac "(new_tape aaa (b, c))", simp, 
       
  5279       rule impI, simp add: fetch.simps split: block.splits option.splits)
       
  5280 apply (auto simp: nth_of.simps t_ncorrect.simps 
       
  5281                       nth_append tshift_length_equal tshift.simps split: if_splits)
       
  5282 done
       
  5283 
       
  5284 
       
  5285 lemma  tm_append_step_equal: 
       
  5286  "\<lbrakk>t_ncorrect tp; t_ncorrect tp'; off = length tp div 2; 
       
  5287    t_step (a, b, c) (tp', 0) = (aa, ab, bb);  aa \<noteq> 0\<rbrakk>
       
  5288  \<Longrightarrow> t_step (a + length tp div 2, b, c) 
       
  5289         (tp @ tshift tp' (length tp div 2), 0)
       
  5290                           = (aa + length tp div 2, ab, bb)"
       
  5291 apply(insert tm_append_step_equal'[of tp tp' off a b c], simp)
       
  5292 apply(case_tac "(t_step (a + length tp div 2, b, c) 
       
  5293                    (tp @ tshift tp' (length tp div 2), 0))", simp)
       
  5294 done
       
  5295 
       
  5296 lemma tm_append_steps_equal: 
       
  5297  "\<lbrakk>t_ncorrect tp; t_ncorrect tp'; off = length tp div 2\<rbrakk> \<Longrightarrow> 
       
  5298    (\<lambda> (s, l, r). ((\<lambda> (s', l', r'). ((s'\<noteq> 0 \<longrightarrow> s = s' + off \<and> l = l'
       
  5299                      \<and> r = r'))) (t_steps (a, b, c) (tp', 0) stp)))
       
  5300    (t_steps (a + off, b, c) (tp @ tshift tp' off, 0) stp)"
       
  5301 apply(induct stp arbitrary : a b c, simp add: t_steps.simps)
       
  5302 apply(simp add: t_steps.simps)
       
  5303 apply(case_tac "(t_step (a, b, c) (tp', 0))", simp)
       
  5304 apply(case_tac "aa = 0", simp add: state0_ind)
       
  5305 apply(subgoal_tac "(t_step (a + length tp div 2, b, c) 
       
  5306                       (tp @ tshift tp' (length tp div 2), 0)) 
       
  5307   = (aa + length tp div 2, ba, ca)", simp)
       
  5308 apply(rule tm_append_step_equal, auto)
       
  5309 done
       
  5310 
       
  5311 (*********Begin: mop_up***************)
       
  5312 type_synonym mopup_type = "t_conf \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> block list \<Rightarrow> bool"
       
  5313 
       
  5314 fun mopup_stop :: "mopup_type"
       
  5315   where
       
  5316   "mopup_stop (s, l, r) lm n ires= 
       
  5317         (\<exists> ln rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \<and> r = <abc_lm_v lm n> @ Bk\<^bsup>rn\<^esup>)"
       
  5318 
       
  5319 fun mopup_bef_erase_a :: "mopup_type"
       
  5320   where
       
  5321   "mopup_bef_erase_a (s, l, r) lm n ires= 
       
  5322          (\<exists> ln m rn. l = Bk \<^bsup>ln\<^esup> @ Bk # Bk # ires \<and> 
       
  5323                   r = Oc\<^bsup>m \<^esup>@ Bk # <(drop ((s + 1) div 2) lm)> @ Bk\<^bsup>rn\<^esup>)"
       
  5324 
       
  5325 fun mopup_bef_erase_b :: "mopup_type"
       
  5326   where
       
  5327   "mopup_bef_erase_b (s, l, r) lm n ires = 
       
  5328       (\<exists> ln m rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \<and> r = Bk # Oc\<^bsup>m\<^esup> @ Bk # 
       
  5329                                       <(drop (s div 2) lm)> @ Bk\<^bsup>rn\<^esup>)"
       
  5330 
       
  5331 
       
  5332 fun mopup_jump_over1 :: "mopup_type"
       
  5333   where
       
  5334   "mopup_jump_over1 (s, l, r) lm n ires = 
       
  5335       (\<exists> ln m1 m2 rn. m1 + m2 = Suc (abc_lm_v lm n) \<and> 
       
  5336         l = Oc\<^bsup>m1\<^esup> @ Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \<and> 
       
  5337      (r = Oc\<^bsup>m2\<^esup> @ Bk # <(drop (Suc n) lm)> @ Bk\<^bsup>rn \<^esup>\<or> 
       
  5338      (r = Oc\<^bsup>m2\<^esup> \<and> (drop (Suc n) lm) = [])))"
       
  5339 
       
  5340 fun mopup_aft_erase_a :: "mopup_type"
       
  5341   where
       
  5342   "mopup_aft_erase_a (s, l, r) lm n ires = 
       
  5343       (\<exists> lnl lnr rn (ml::nat list) m. 
       
  5344           m = Suc (abc_lm_v lm n) \<and> l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \<and> 
       
  5345                                    (r = <ml> @ Bk\<^bsup>rn\<^esup>))"
       
  5346 
       
  5347 fun mopup_aft_erase_b :: "mopup_type"
       
  5348   where
       
  5349   "mopup_aft_erase_b (s, l, r) lm n ires= 
       
  5350    (\<exists> lnl lnr rn (ml::nat list) m. 
       
  5351       m = Suc (abc_lm_v lm n) \<and> 
       
  5352       l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \<and> 
       
  5353      (r = Bk # <ml> @ Bk\<^bsup>rn \<^esup>\<or>
       
  5354       r = Bk # Bk # <ml> @ Bk\<^bsup>rn\<^esup>))"
       
  5355 
       
  5356 fun mopup_aft_erase_c :: "mopup_type"
       
  5357   where
       
  5358   "mopup_aft_erase_c (s, l, r) lm n ires = 
       
  5359  (\<exists> lnl lnr rn (ml::nat list) m. 
       
  5360      m = Suc (abc_lm_v lm n) \<and> 
       
  5361      l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \<and> 
       
  5362     (r = <ml> @ Bk\<^bsup>rn \<^esup>\<or> r = Bk # <ml> @ Bk\<^bsup>rn\<^esup>))"
       
  5363 
       
  5364 fun mopup_left_moving :: "mopup_type"
       
  5365   where
       
  5366   "mopup_left_moving (s, l, r) lm n ires = 
       
  5367   (\<exists> lnl lnr rn m.
       
  5368      m = Suc (abc_lm_v lm n) \<and> 
       
  5369    ((l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \<and> r = Bk\<^bsup>rn\<^esup>) \<or>
       
  5370     (l = Oc\<^bsup>m - 1\<^esup> @ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \<and> r = Oc # Bk\<^bsup>rn\<^esup>)))"
       
  5371 
       
  5372 fun mopup_jump_over2 :: "mopup_type"
       
  5373   where
       
  5374   "mopup_jump_over2 (s, l, r) lm n ires = 
       
  5375      (\<exists> ln rn m1 m2.
       
  5376           m1 + m2 = Suc (abc_lm_v lm n) 
       
  5377         \<and> r \<noteq> [] 
       
  5378         \<and> (hd r = Oc \<longrightarrow> (l = Oc\<^bsup>m1\<^esup> @ Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \<and> r = Oc\<^bsup>m2\<^esup> @ Bk\<^bsup>rn\<^esup>)) 
       
  5379         \<and> (hd r = Bk \<longrightarrow> (l = Bk\<^bsup>ln\<^esup> @ Bk # ires \<and> r = Bk # Oc\<^bsup>m1 + m2\<^esup> @ Bk\<^bsup>rn\<^esup>)))"
       
  5380 
       
  5381 
       
  5382 fun mopup_inv :: "mopup_type"
       
  5383   where
       
  5384   "mopup_inv (s, l, r) lm n ires = 
       
  5385       (if s = 0 then mopup_stop (s, l, r) lm n ires
       
  5386        else if s \<le> 2*n then
       
  5387                if s mod 2 = 1 then mopup_bef_erase_a (s, l, r) lm n ires
       
  5388                    else mopup_bef_erase_b (s, l, r) lm n ires
       
  5389             else if s = 2*n + 1 then 
       
  5390                 mopup_jump_over1 (s, l, r) lm n ires
       
  5391             else if s = 2*n + 2 then mopup_aft_erase_a (s, l, r) lm n ires
       
  5392             else if s = 2*n + 3 then mopup_aft_erase_b (s, l, r) lm n ires
       
  5393             else if s = 2*n + 4 then mopup_aft_erase_c (s, l, r) lm n ires
       
  5394             else if s = 2*n + 5 then mopup_left_moving (s, l, r) lm n ires
       
  5395             else if s = 2*n + 6 then mopup_jump_over2 (s, l, r) lm n ires
       
  5396             else False)"
       
  5397 
       
  5398 declare 
       
  5399   mopup_jump_over2.simps[simp del] mopup_left_moving.simps[simp del]
       
  5400   mopup_aft_erase_c.simps[simp del] mopup_aft_erase_b.simps[simp del] 
       
  5401   mopup_aft_erase_a.simps[simp del] mopup_jump_over1.simps[simp del]
       
  5402   mopup_bef_erase_a.simps[simp del] mopup_bef_erase_b.simps[simp del]
       
  5403   mopup_stop.simps[simp del]
       
  5404 
       
  5405 lemma mopup_fetch_0[simp]: 
       
  5406      "(fetch (mop_bef n @ tshift mp_up (2 * n)) 0 b) = (Nop, 0)"
       
  5407 by(simp add: fetch.simps)
       
  5408 
       
  5409 lemma mop_bef_length[simp]: "length (mop_bef n) = 4 * n"
       
  5410 apply(induct n, simp add: mop_bef.simps, simp add: mop_bef.simps)
       
  5411 done
       
  5412 
       
  5413 thm findnth_nth
       
  5414 lemma mop_bef_nth: 
       
  5415   "\<lbrakk>q < n; x < 4\<rbrakk> \<Longrightarrow> mop_bef n ! (4 * q + x) = 
       
  5416                              mop_bef (Suc q) ! ((4 * q) + x)"
       
  5417 apply(induct n, simp)
       
  5418 apply(case_tac "q < n", simp add: mop_bef.simps, auto)
       
  5419 apply(simp add: nth_append)
       
  5420 apply(subgoal_tac "q = n", simp)
       
  5421 apply(arith)
       
  5422 done
       
  5423 
       
  5424 lemma fetch_bef_erase_a_o[simp]: 
       
  5425  "\<lbrakk>0 < s; s \<le> 2 * n; s mod 2 = Suc 0\<rbrakk>
       
  5426   \<Longrightarrow> (fetch (mop_bef n @ tshift mp_up (2 * n)) s Oc) = (W0, s + 1)"
       
  5427 apply(subgoal_tac "\<exists> q. s = 2*q + 1", auto)
       
  5428 apply(subgoal_tac "length (mop_bef n) = 4*n")
       
  5429 apply(auto simp: fetch.simps nth_of.simps nth_append)
       
  5430 apply(subgoal_tac "mop_bef n ! (4 * q + 1) = 
       
  5431                       mop_bef (Suc q) ! ((4 * q) + 1)", 
       
  5432       simp add: mop_bef.simps nth_append)
       
  5433 apply(rule mop_bef_nth, auto)
       
  5434 done
       
  5435 
       
  5436 lemma fetch_bef_erase_a_b[simp]:
       
  5437   "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = Suc 0\<rbrakk>
       
  5438    \<Longrightarrow>  (fetch (mop_bef n @ tshift mp_up (2 * n)) s Bk) = (R, s + 2)"
       
  5439 apply(subgoal_tac "\<exists> q. s = 2*q + 1", auto)
       
  5440 apply(subgoal_tac "length (mop_bef n) = 4*n")
       
  5441 apply(auto simp: fetch.simps nth_of.simps nth_append)
       
  5442 apply(subgoal_tac "mop_bef n ! (4 * q + 0) = 
       
  5443                        mop_bef (Suc q) ! ((4 * q + 0))", 
       
  5444       simp add: mop_bef.simps nth_append)
       
  5445 apply(rule mop_bef_nth, auto)
       
  5446 done
       
  5447 
       
  5448 lemma fetch_bef_erase_b_b: 
       
  5449   "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = 0\<rbrakk> \<Longrightarrow> 
       
  5450      (fetch (mop_bef n @ tshift mp_up (2 * n)) s Bk) = (R, s - 1)"
       
  5451 apply(subgoal_tac "\<exists> q. s = 2 * q", auto)
       
  5452 apply(case_tac qa, simp, simp)
       
  5453 apply(auto simp: fetch.simps nth_of.simps nth_append)
       
  5454 apply(subgoal_tac "mop_bef n ! (4 * nat + 2) = 
       
  5455                      mop_bef (Suc nat) ! ((4 * nat) + 2)", 
       
  5456       simp add: mop_bef.simps nth_append)
       
  5457 apply(rule mop_bef_nth, auto)
       
  5458 done
       
  5459 
       
  5460 lemma fetch_jump_over1_o: 
       
  5461  "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (2 * n)) Oc
       
  5462   = (R, Suc (2 * n))"
       
  5463 apply(subgoal_tac "length (mop_bef n) = 4 * n")
       
  5464 apply(auto simp: fetch.simps nth_of.simps mp_up_def nth_append 
       
  5465                  tshift.simps)
       
  5466 done
       
  5467 
       
  5468 lemma fetch_jump_over1_b: 
       
  5469  "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (2 * n)) Bk 
       
  5470  = (R, Suc (Suc (2 * n)))"
       
  5471 apply(subgoal_tac "length (mop_bef n) = 4 * n")
       
  5472 apply(auto simp: fetch.simps nth_of.simps mp_up_def 
       
  5473                  nth_append tshift.simps)
       
  5474 done
       
  5475 
       
  5476 lemma fetch_aft_erase_a_o: 
       
  5477  "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (Suc (2 * n))) Oc 
       
  5478  = (W0, Suc (2 * n + 2))"
       
  5479 apply(subgoal_tac "length (mop_bef n) = 4 * n")
       
  5480 apply(auto simp: fetch.simps nth_of.simps mp_up_def 
       
  5481                  nth_append tshift.simps)
       
  5482 done
       
  5483 
       
  5484 lemma fetch_aft_erase_a_b: 
       
  5485  "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (Suc (2 * n))) Bk
       
  5486   = (L, Suc (2 * n + 4))"
       
  5487 apply(subgoal_tac "length (mop_bef n) = 4 * n")
       
  5488 apply(auto simp: fetch.simps nth_of.simps mp_up_def 
       
  5489                  nth_append tshift.simps)
       
  5490 done
       
  5491 
       
  5492 lemma fetch_aft_erase_b_b: 
       
  5493  "fetch (mop_bef n @ tshift mp_up (2 * n)) (2*n + 3) Bk
       
  5494   = (R, Suc (2 * n + 3))"
       
  5495 apply(subgoal_tac "length (mop_bef n) = 4 * n")
       
  5496 apply(subgoal_tac "2*n + 3 = Suc (2*n + 2)", simp only: fetch.simps)
       
  5497 apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps)
       
  5498 done
       
  5499 
       
  5500 lemma fetch_aft_erase_c_o: 
       
  5501  "fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 4) Oc 
       
  5502  = (W0, Suc (2 * n + 2))"
       
  5503 apply(subgoal_tac "length (mop_bef n) = 4 * n")
       
  5504 apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps)
       
  5505 apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps)
       
  5506 done
       
  5507 
       
  5508 lemma fetch_aft_erase_c_b: 
       
  5509  "fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 4) Bk 
       
  5510  = (R, Suc (2 * n + 1))"
       
  5511 apply(subgoal_tac "length (mop_bef n) = 4 * n")
       
  5512 apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps)
       
  5513 apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps)
       
  5514 done
       
  5515 
       
  5516 lemma fetch_left_moving_o: 
       
  5517  "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 5) Oc) 
       
  5518  = (L, 2*n + 6)"
       
  5519 apply(subgoal_tac "length (mop_bef n) = 4 * n")
       
  5520 apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps)
       
  5521 apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps)
       
  5522 done
       
  5523 
       
  5524 lemma fetch_left_moving_b: 
       
  5525  "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 5) Bk)
       
  5526   = (L, 2*n + 5)"
       
  5527 apply(subgoal_tac "length (mop_bef n) = 4 * n")
       
  5528 apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps)
       
  5529 apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps)
       
  5530 done
       
  5531 
       
  5532 lemma fetch_jump_over2_b:
       
  5533   "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 6) Bk) 
       
  5534  = (R, 0)"
       
  5535 apply(subgoal_tac "length (mop_bef n) = 4 * n")
       
  5536 apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps)
       
  5537 apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps)
       
  5538 done
       
  5539 
       
  5540 lemma fetch_jump_over2_o: 
       
  5541 "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 6) Oc) 
       
  5542  = (L, 2*n + 6)"
       
  5543 apply(subgoal_tac "length (mop_bef n) = 4 * n")
       
  5544 apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps)
       
  5545 apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps)
       
  5546 done
       
  5547 
       
  5548 lemmas mopupfetchs = 
       
  5549 fetch_bef_erase_a_o fetch_bef_erase_a_b fetch_bef_erase_b_b 
       
  5550 fetch_jump_over1_o fetch_jump_over1_b fetch_aft_erase_a_o 
       
  5551 fetch_aft_erase_a_b fetch_aft_erase_b_b fetch_aft_erase_c_o 
       
  5552 fetch_aft_erase_c_b fetch_left_moving_o fetch_left_moving_b 
       
  5553 fetch_jump_over2_b fetch_jump_over2_o
       
  5554 
       
  5555 lemma [simp]: 
       
  5556 "\<lbrakk>n < length lm; 0 < s; s mod 2 = Suc 0; 
       
  5557   mopup_bef_erase_a (s, l, Oc # xs) lm n ires; 
       
  5558   Suc s \<le> 2 * n\<rbrakk> \<Longrightarrow> 
       
  5559   mopup_bef_erase_b (Suc s, l, Bk # xs) lm n ires"
       
  5560 apply(auto simp: mopup_bef_erase_a.simps mopup_bef_erase_b.simps )
       
  5561 apply(rule_tac x = "m - 1" in exI, rule_tac x = rn in exI)
       
  5562 apply(case_tac m, simp, simp)
       
  5563 done
       
  5564 
       
  5565 lemma mopup_false1:
       
  5566   "\<lbrakk>0 < s; s \<le> 2 * n; s mod 2 = Suc 0;  \<not> Suc s \<le> 2 * n\<rbrakk> 
       
  5567   \<Longrightarrow> RR"
       
  5568 apply(arith)
       
  5569 done
       
  5570 
       
  5571 lemma [simp]: 
       
  5572  "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = Suc 0; 
       
  5573    mopup_bef_erase_a (s, l, Oc # xs) lm n ires; r = Oc # xs\<rbrakk>
       
  5574  \<Longrightarrow> (Suc s \<le> 2 * n \<longrightarrow> mopup_bef_erase_b (Suc s, l, Bk # xs) lm n ires)  \<and>
       
  5575      (\<not> Suc s \<le> 2 * n \<longrightarrow> mopup_jump_over1 (Suc s, l, Bk # xs) lm n ires) "
       
  5576 apply(auto elim: mopup_false1)
       
  5577 done
       
  5578 
       
  5579 lemma drop_abc_lm_v_simp[simp]: 
       
  5580    "n < length lm \<Longrightarrow> drop n lm = abc_lm_v lm n # drop (Suc n) lm"
       
  5581 apply(auto simp: abc_lm_v.simps)
       
  5582 apply(drule drop_Suc_conv_tl, simp)
       
  5583 done
       
  5584 lemma [simp]: "(\<exists>rna. Bk\<^bsup>rn\<^esup> = Bk # Bk\<^bsup>rna\<^esup>) \<or> Bk\<^bsup>rn\<^esup> = []"
       
  5585 apply(case_tac rn, simp, auto)
       
  5586 done
       
  5587 
       
  5588 lemma [simp]: "\<exists>lna. Bk # Bk\<^bsup>ln\<^esup> = Bk\<^bsup>lna\<^esup>"
       
  5589 apply(rule_tac x = "Suc ln" in exI, auto)
       
  5590 done
       
  5591 
       
  5592 lemma mopup_bef_erase_a_2_jump_over[simp]: 
       
  5593  "\<lbrakk>n < length lm; 0 < s; s mod 2 = Suc 0; 
       
  5594    mopup_bef_erase_a (s, l, Bk # xs) lm n ires; Suc s = 2 * n\<rbrakk> 
       
  5595 \<Longrightarrow> mopup_jump_over1 (Suc (2 * n), Bk # l, xs) lm n ires"
       
  5596 apply(auto simp: mopup_bef_erase_a.simps mopup_jump_over1.simps)
       
  5597 apply(case_tac m, simp)
       
  5598 apply(rule_tac x = "Suc ln" in exI, rule_tac x = 0 in exI, 
       
  5599       simp add: tape_of_nl_abv)
       
  5600 apply(case_tac "drop (Suc n) lm", auto simp: tape_of_nat_list.simps )
       
  5601 done
       
  5602 
       
  5603 lemma Suc_Suc_div:  "\<lbrakk>0 < s; s mod 2 = Suc 0; Suc (Suc s) \<le> 2 * n\<rbrakk>
       
  5604            \<Longrightarrow> (Suc (Suc (s div 2))) \<le> n"
       
  5605 apply(arith)
       
  5606 done
       
  5607 
       
  5608 lemma mopup_bef_erase_a_2_a[simp]: 
       
  5609  "\<lbrakk>n < length lm; 0 < s; s mod 2 = Suc 0; 
       
  5610    mopup_bef_erase_a (s, l, Bk # xs) lm n ires; 
       
  5611    Suc (Suc s) \<le> 2 * n\<rbrakk> \<Longrightarrow> 
       
  5612    mopup_bef_erase_a (Suc (Suc s), Bk # l, xs) lm n ires"
       
  5613 apply(auto simp: mopup_bef_erase_a.simps )
       
  5614 apply(subgoal_tac "drop (Suc (Suc (s div 2))) lm \<noteq> []")
       
  5615 apply(case_tac m, simp)
       
  5616 apply(rule_tac x = "Suc (abc_lm_v lm (Suc (s div 2)))" in exI, 
       
  5617       rule_tac x = rn in exI, simp, simp)
       
  5618 apply(subgoal_tac "(Suc (Suc (s div 2))) \<le> n", simp, 
       
  5619       rule_tac Suc_Suc_div, auto)
       
  5620 done
       
  5621 
       
  5622 lemma mopup_false2: 
       
  5623  "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; 
       
  5624    s mod 2 = Suc 0; Suc s \<noteq> 2 * n;
       
  5625    \<not> Suc (Suc s) \<le> 2 * n\<rbrakk> \<Longrightarrow> RR"
       
  5626 apply(arith)
       
  5627 done
       
  5628 
       
  5629 lemma [simp]:
       
  5630   "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; 
       
  5631    s mod 2 = Suc 0; 
       
  5632    mopup_bef_erase_a (s, l, Bk # xs) lm n ires; 
       
  5633    r = Bk # xs\<rbrakk>
       
  5634  \<Longrightarrow> (Suc s = 2 * n \<longrightarrow> 
       
  5635              mopup_jump_over1 (Suc (2 * n), Bk # l, xs) lm n ires) \<and>
       
  5636      (Suc s \<noteq> 2 * n \<longrightarrow> 
       
  5637        (Suc (Suc s) \<le> 2 * n \<longrightarrow> 
       
  5638           mopup_bef_erase_a (Suc (Suc s), Bk # l, xs) lm n ires) \<and> 
       
  5639        (\<not> Suc (Suc s) \<le> 2 * n \<longrightarrow> 
       
  5640           mopup_aft_erase_a (Suc (Suc s), Bk # l, xs) lm n ires))"
       
  5641 apply(auto elim: mopup_false2)
       
  5642 done
       
  5643 
       
  5644 lemma [simp]: "mopup_bef_erase_a (s, l, []) lm n ires \<Longrightarrow> 
       
  5645                         mopup_bef_erase_a (s, l, [Bk]) lm n ires"
       
  5646 apply(auto simp: mopup_bef_erase_a.simps)
       
  5647 done
       
  5648 
       
  5649 lemma [simp]:
       
  5650    "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = Suc 0;
       
  5651      mopup_bef_erase_a (s, l, []) lm n ires; r = []\<rbrakk>
       
  5652     \<Longrightarrow> (Suc s = 2 * n \<longrightarrow> 
       
  5653               mopup_jump_over1 (Suc (2 * n), Bk # l, []) lm n ires) \<and>
       
  5654         (Suc s \<noteq> 2 * n \<longrightarrow> 
       
  5655              (Suc (Suc s) \<le> 2 * n \<longrightarrow> 
       
  5656                  mopup_bef_erase_a (Suc (Suc s), Bk # l, []) lm n ires) \<and>
       
  5657              (\<not> Suc (Suc s) \<le> 2 * n \<longrightarrow> 
       
  5658                  mopup_aft_erase_a (Suc (Suc s), Bk # l, []) lm n ires))"
       
  5659 apply(auto)
       
  5660 done
       
  5661 
       
  5662 lemma "mopup_bef_erase_b (s, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []"
       
  5663 apply(auto simp: mopup_bef_erase_b.simps)
       
  5664 done
       
  5665 
       
  5666 lemma [simp]: "mopup_bef_erase_b (s, l, Oc # xs) lm n ires = False"
       
  5667 apply(auto simp: mopup_bef_erase_b.simps )
       
  5668 done
       
  5669  
       
  5670 lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow> 
       
  5671                                       (s - Suc 0) mod 2 = Suc 0"
       
  5672 apply(arith)
       
  5673 done
       
  5674 
       
  5675 lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow>
       
  5676                                        s - Suc 0 \<le> 2 * n"
       
  5677 apply(simp)
       
  5678 done
       
  5679 
       
  5680 lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow> \<not> s \<le> Suc 0"
       
  5681 apply(arith)
       
  5682 done
       
  5683 
       
  5684 lemma [simp]: "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; 
       
  5685                s mod 2 \<noteq> Suc 0; 
       
  5686                mopup_bef_erase_b (s, l, Bk # xs) lm n ires; r = Bk # xs\<rbrakk> 
       
  5687            \<Longrightarrow> mopup_bef_erase_a (s - Suc 0, Bk # l, xs) lm n ires"
       
  5688 apply(auto simp: mopup_bef_erase_b.simps mopup_bef_erase_a.simps)
       
  5689 done
       
  5690 
       
  5691 lemma [simp]: "\<lbrakk>mopup_bef_erase_b (s, l, []) lm n ires\<rbrakk> \<Longrightarrow> 
       
  5692                    mopup_bef_erase_a (s - Suc 0, Bk # l, []) lm n ires"
       
  5693 apply(auto simp: mopup_bef_erase_b.simps mopup_bef_erase_a.simps)
       
  5694 done
       
  5695 
       
  5696 lemma [simp]: 
       
  5697    "\<lbrakk>n < length lm;
       
  5698     mopup_jump_over1 (Suc (2 * n), l, Oc # xs) lm n ires;
       
  5699     r = Oc # xs\<rbrakk>
       
  5700   \<Longrightarrow> mopup_jump_over1 (Suc (2 * n), Oc # l, xs) lm n ires"
       
  5701 apply(auto simp: mopup_jump_over1.simps)
       
  5702 apply(rule_tac x = ln in exI, rule_tac x = "Suc m1" in exI,
       
  5703        rule_tac x = "m2 - 1" in exI)
       
  5704 apply(case_tac "m2", simp, simp, rule_tac x = rn in exI, simp)
       
  5705 apply(rule_tac x = ln in exI, rule_tac x = "Suc m1" in exI, 
       
  5706       rule_tac x = "m2 - 1" in exI)
       
  5707 apply(case_tac m2, simp, simp)
       
  5708 done
       
  5709 
       
  5710 lemma mopup_jump_over1_2_aft_erase_a[simp]:  
       
  5711  "\<lbrakk>n < length lm; mopup_jump_over1 (Suc (2 * n), l, Bk # xs) lm n ires\<rbrakk>
       
  5712   \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, xs) lm n ires"
       
  5713 apply(simp only: mopup_jump_over1.simps mopup_aft_erase_a.simps)
       
  5714 apply(erule_tac exE)+
       
  5715 apply(rule_tac x = ln in exI, rule_tac x = "Suc 0" in exI)
       
  5716 apply(case_tac m2, simp)
       
  5717 apply(rule_tac x = rn in exI, rule_tac x = "drop (Suc n) lm" in exI, 
       
  5718       simp)
       
  5719 apply(simp)
       
  5720 done
       
  5721 
       
  5722 lemma [simp]: 
       
  5723  "\<lbrakk>n < length lm; mopup_jump_over1 (Suc (2 * n), l, []) lm n ires\<rbrakk> \<Longrightarrow> 
       
  5724     mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, []) lm n ires"
       
  5725 apply(rule mopup_jump_over1_2_aft_erase_a, simp)
       
  5726 apply(auto simp: mopup_jump_over1.simps)
       
  5727 apply(rule_tac x = ln in exI, rule_tac x = m1 in exI, 
       
  5728       rule_tac x = m2 in exI, simp add: )
       
  5729 apply(rule_tac x = 0 in exI, auto)
       
  5730 done
       
  5731 
       
  5732 lemma [simp]: 
       
  5733  "\<lbrakk>n < length lm; 
       
  5734    mopup_aft_erase_a (Suc (Suc (2 * n)), l, Oc # xs) lm n ires\<rbrakk> 
       
  5735  \<Longrightarrow> mopup_aft_erase_b (Suc (Suc (Suc (2 * n))), l, Bk # xs) lm n ires"
       
  5736 apply(auto simp: mopup_aft_erase_a.simps mopup_aft_erase_b.simps )
       
  5737 apply(case_tac ml, simp, case_tac rn, simp, simp)
       
  5738 apply(case_tac list, auto simp: tape_of_nl_abv 
       
  5739                                 tape_of_nat_list.simps )
       
  5740 apply(case_tac a, simp, rule_tac x = rn in exI, 
       
  5741       rule_tac x = "[]" in exI,
       
  5742        simp add: tape_of_nat_list.simps, simp)
       
  5743 apply(rule_tac x = rn in exI, rule_tac x = "[nat]" in exI, 
       
  5744       simp add: tape_of_nat_list.simps )
       
  5745 apply(case_tac a, simp, rule_tac x = rn in exI, 
       
  5746        rule_tac x = "aa # lista" in exI, simp, simp)
       
  5747 apply(rule_tac x = rn in exI, rule_tac x = "nat # aa # lista" in exI, 
       
  5748        simp add: tape_of_nat_list.simps )
       
  5749 done
       
  5750 
       
  5751 lemma [simp]:
       
  5752   "mopup_aft_erase_a (Suc (Suc (2 * n)), l, Bk # xs) lm n ires \<Longrightarrow> l \<noteq> []"
       
  5753 apply(auto simp: mopup_aft_erase_a.simps)
       
  5754 done
       
  5755 
       
  5756 lemma [simp]:
       
  5757   "\<lbrakk>n < length lm;
       
  5758     mopup_aft_erase_a (Suc (Suc (2 * n)), l, Bk # xs) lm n ires\<rbrakk>
       
  5759   \<Longrightarrow> mopup_left_moving (5 + 2 * n, tl l, hd l # Bk # xs) lm n ires"
       
  5760 apply(simp only: mopup_aft_erase_a.simps mopup_left_moving.simps)
       
  5761 apply(erule exE)+
       
  5762 apply(case_tac lnr, simp)
       
  5763 apply(rule_tac x = lnl in exI, simp, rule_tac x = rn in exI, simp)
       
  5764 apply(subgoal_tac "ml = []", simp)
       
  5765 apply(rule_tac xs = xs and rn = rn in BkCons_nil, simp, auto)
       
  5766 apply(subgoal_tac "ml = []", auto)
       
  5767 apply(rule_tac xs = xs and rn = rn in BkCons_nil, simp)
       
  5768 done
       
  5769 
       
  5770 lemma [simp]:
       
  5771   "mopup_aft_erase_a (Suc (Suc (2 * n)), l, []) lm n ires \<Longrightarrow> l \<noteq> []"
       
  5772 apply(simp only: mopup_aft_erase_a.simps)
       
  5773 apply(erule exE)+
       
  5774 apply(auto)
       
  5775 done
       
  5776 
       
  5777 lemma [simp]:
       
  5778   "\<lbrakk>n < length lm; mopup_aft_erase_a (Suc (Suc (2 * n)), l, []) lm n ires\<rbrakk>
       
  5779   \<Longrightarrow> mopup_left_moving (5 + 2 * n, tl l, [hd l]) lm n ires"
       
  5780 apply(simp only: mopup_aft_erase_a.simps mopup_left_moving.simps)
       
  5781 apply(erule exE)+
       
  5782 apply(subgoal_tac "ml = [] \<and> rn = 0", erule conjE, erule conjE, simp)
       
  5783 apply(case_tac lnr, simp, rule_tac x = lnl in exI, simp, 
       
  5784       rule_tac x = 0 in exI, simp)
       
  5785 apply(rule_tac x = lnl in exI, rule_tac x = nat in exI, 
       
  5786       rule_tac x = "Suc 0" in exI, simp)
       
  5787 apply(case_tac ml, simp, case_tac rn, simp, simp)
       
  5788 apply(case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps)
       
  5789 done
       
  5790 
       
  5791 lemma [simp]: "mopup_aft_erase_b (2 * n + 3, l, Oc # xs) lm n ires = False"
       
  5792 apply(auto simp: mopup_aft_erase_b.simps )
       
  5793 done
       
  5794 
       
  5795 lemma [simp]: 
       
  5796  "\<lbrakk>n < length lm; 
       
  5797    mopup_aft_erase_c (2 * n + 4, l, Oc # xs) lm n ires\<rbrakk>
       
  5798   \<Longrightarrow> mopup_aft_erase_b (Suc (Suc (Suc (2 * n))), l, Bk # xs) lm n ires"
       
  5799 apply(auto simp: mopup_aft_erase_c.simps mopup_aft_erase_b.simps )
       
  5800 apply(case_tac ml, simp, case_tac rn, simp, simp, simp)
       
  5801 apply(case_tac list, auto simp: tape_of_nl_abv 
       
  5802                         tape_of_nat_list.simps tape_of_nat_abv )
       
  5803 apply(case_tac a, rule_tac x = rn in exI, 
       
  5804       rule_tac x = "[]" in exI, simp add: tape_of_nat_list.simps)
       
  5805 apply(rule_tac x = rn in exI, rule_tac x = "[nat]" in exI, 
       
  5806       simp add: tape_of_nat_list.simps )
       
  5807 apply(case_tac a, simp, rule_tac x = rn in exI, 
       
  5808       rule_tac x = "aa # lista" in exI, simp)
       
  5809 apply(rule_tac x = rn in exI, rule_tac x = "nat # aa # lista" in exI, 
       
  5810       simp add: tape_of_nat_list.simps )
       
  5811 done
       
  5812 
       
  5813 lemma mopup_aft_erase_c_aft_erase_a[simp]: 
       
  5814  "\<lbrakk>n < length lm; mopup_aft_erase_c (2 * n + 4, l, Bk # xs) lm n ires\<rbrakk> 
       
  5815  \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, xs) lm n ires"
       
  5816 apply(simp only: mopup_aft_erase_c.simps mopup_aft_erase_a.simps )
       
  5817 apply(erule_tac exE)+
       
  5818 apply(erule conjE, erule conjE, erule disjE)
       
  5819 apply(subgoal_tac "ml = []", simp, case_tac rn, 
       
  5820       simp, simp, rule conjI)
       
  5821 apply(rule_tac x = lnl in exI, rule_tac x = "Suc lnr" in exI, simp)
       
  5822 apply(rule_tac x = nat in exI, rule_tac x = "[]" in exI, simp)
       
  5823 apply(rule_tac xs = xs and rn = rn in BkCons_nil, simp, simp, 
       
  5824       rule conjI)
       
  5825 apply(rule_tac x = lnl in exI, rule_tac x = "Suc lnr" in exI, simp)
       
  5826 apply(rule_tac x = rn in exI, rule_tac x = "ml" in exI, simp)
       
  5827 done
       
  5828 
       
  5829 lemma [simp]: 
       
  5830  "\<lbrakk>n < length lm; mopup_aft_erase_c (2 * n + 4, l, []) lm n ires\<rbrakk> 
       
  5831  \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, []) lm n ires"
       
  5832 apply(rule mopup_aft_erase_c_aft_erase_a, simp)
       
  5833 apply(simp only: mopup_aft_erase_c.simps)
       
  5834 apply(erule exE)+
       
  5835 apply(rule_tac x = lnl in exI, rule_tac x = lnr in exI, simp add: )
       
  5836 apply(rule_tac x = 0 in exI, rule_tac x = "[]" in exI, simp)
       
  5837 done
       
  5838 
       
  5839 lemma mopup_aft_erase_b_2_aft_erase_c[simp]:
       
  5840   "\<lbrakk>n < length lm; mopup_aft_erase_b (2 * n + 3, l, Bk # xs) lm n ires\<rbrakk>  
       
  5841  \<Longrightarrow> mopup_aft_erase_c (4 + 2 * n, Bk # l, xs) lm n ires"
       
  5842 apply(auto simp: mopup_aft_erase_b.simps mopup_aft_erase_c.simps)
       
  5843 apply(rule_tac x = "lnl" in exI, rule_tac x = "Suc lnr" in exI, simp)
       
  5844 apply(rule_tac x = "lnl" in exI, rule_tac x = "Suc lnr" in exI, simp)
       
  5845 done
       
  5846 
       
  5847 lemma [simp]: 
       
  5848  "\<lbrakk>n < length lm; mopup_aft_erase_b (2 * n + 3, l, []) lm n ires\<rbrakk> 
       
  5849  \<Longrightarrow> mopup_aft_erase_c (4 + 2 * n, Bk # l, []) lm n ires"
       
  5850 apply(rule_tac mopup_aft_erase_b_2_aft_erase_c, simp)
       
  5851 apply(simp add: mopup_aft_erase_b.simps)
       
  5852 done
       
  5853 
       
  5854 lemma [simp]: 
       
  5855     "mopup_left_moving (2 * n + 5, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []"
       
  5856 apply(auto simp: mopup_left_moving.simps)
       
  5857 done
       
  5858 
       
  5859 lemma [simp]:  
       
  5860  "\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, Oc # xs) lm n ires\<rbrakk>
       
  5861   \<Longrightarrow> mopup_jump_over2 (2 * n + 6, tl l, hd l # Oc # xs) lm n ires"
       
  5862 apply(simp only: mopup_left_moving.simps mopup_jump_over2.simps)
       
  5863 apply(erule_tac exE)+
       
  5864 apply(erule conjE, erule disjE, erule conjE)
       
  5865 apply(case_tac rn, simp, simp add: )
       
  5866 apply(case_tac "hd l", simp add:  )
       
  5867 apply(case_tac "abc_lm_v lm n", simp)
       
  5868 apply(rule_tac x = "lnl" in exI, rule_tac x = rn in exI, 
       
  5869       rule_tac x = "Suc 0" in exI, rule_tac x = 0 in exI)
       
  5870 apply(case_tac lnl, simp, simp, simp add: exp_ind[THEN sym], simp)
       
  5871 apply(case_tac "abc_lm_v lm n", simp)
       
  5872 apply(case_tac lnl, simp, simp)
       
  5873 apply(rule_tac x = lnl in exI, rule_tac x = rn in exI)
       
  5874 apply(rule_tac x = nat in exI, rule_tac x = "Suc (Suc 0)" in exI, simp)
       
  5875 done
       
  5876 
       
  5877 lemma [simp]: "mopup_left_moving (2 * n + 5, l, xs) lm n ires \<Longrightarrow> l \<noteq> []"
       
  5878 apply(auto simp: mopup_left_moving.simps)
       
  5879 done
       
  5880 
       
  5881 lemma [simp]:
       
  5882   "\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, Bk # xs) lm n ires\<rbrakk> 
       
  5883  \<Longrightarrow> mopup_left_moving (2 * n + 5, tl l, hd l # Bk # xs) lm n ires"
       
  5884 apply(simp only: mopup_left_moving.simps)
       
  5885 apply(erule exE)+
       
  5886 apply(case_tac lnr, simp)
       
  5887 apply(rule_tac x = lnl in exI, rule_tac x = 0 in exI, 
       
  5888       rule_tac x = rn in exI, simp, simp)
       
  5889 apply(rule_tac x = lnl in exI, rule_tac x = nat in exI, simp)
       
  5890 done
       
  5891 
       
  5892 lemma [simp]: 
       
  5893 "\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, []) lm n ires\<rbrakk>
       
  5894     \<Longrightarrow> mopup_left_moving (2 * n + 5, tl l, [hd l]) lm n ires"
       
  5895 apply(simp only: mopup_left_moving.simps)
       
  5896 apply(erule exE)+
       
  5897 apply(case_tac lnr, simp)
       
  5898 apply(rule_tac x = lnl in exI, rule_tac x = 0 in exI, 
       
  5899       rule_tac x = 0 in exI, simp, auto)
       
  5900 done
       
  5901 
       
  5902 lemma [simp]: 
       
  5903  "mopup_jump_over2 (2 * n + 6, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []"
       
  5904 apply(auto simp: mopup_jump_over2.simps )
       
  5905 done
       
  5906 
       
  5907 lemma [intro]: "\<exists>lna. Bk # Bk\<^bsup>ln\<^esup> = Bk\<^bsup>lna\<^esup> @ [Bk]"
       
  5908 apply(simp only: exp_ind[THEN sym], auto)
       
  5909 done
       
  5910 
       
  5911 lemma [simp]: 
       
  5912 "\<lbrakk>n < length lm; mopup_jump_over2 (2 * n + 6, l, Oc # xs) lm n ires\<rbrakk>
       
  5913  \<Longrightarrow>  mopup_jump_over2 (2 * n + 6, tl l, hd l # Oc # xs) lm n ires"
       
  5914 apply(simp only: mopup_jump_over2.simps)
       
  5915 apply(erule_tac exE)+
       
  5916 apply(simp add:  , erule conjE, erule_tac conjE)
       
  5917 apply(case_tac m1, simp)
       
  5918 apply(rule_tac x = ln in exI, rule_tac x = rn in exI, 
       
  5919       rule_tac x = 0 in exI, simp)
       
  5920 apply(case_tac ln, simp, simp, simp only: exp_ind[THEN sym], simp)
       
  5921 apply(rule_tac x = ln in exI, rule_tac x = rn in exI, 
       
  5922       rule_tac x = nat in exI, rule_tac x = "Suc m2" in exI, simp)
       
  5923 done
       
  5924 
       
  5925 lemma [simp]: "\<exists>rna. Oc # Oc\<^bsup>a\<^esup> @ Bk\<^bsup>rn\<^esup> = <a> @ Bk\<^bsup>rna\<^esup>"
       
  5926 apply(case_tac a, auto simp: tape_of_nat_abv )
       
  5927 done
       
  5928 
       
  5929 lemma [simp]: 
       
  5930  "\<lbrakk>n < length lm; mopup_jump_over2 (2 * n + 6, l, Bk # xs) lm n ires\<rbrakk> 
       
  5931   \<Longrightarrow> mopup_stop (0, Bk # l, xs) lm n ires"
       
  5932 apply(auto simp: mopup_jump_over2.simps mopup_stop.simps)
       
  5933 done
       
  5934 
       
  5935 lemma [simp]: "mopup_jump_over2 (2 * n + 6, l, []) lm n ires = False"
       
  5936 apply(simp only: mopup_jump_over2.simps, simp)
       
  5937 done
       
  5938 
       
  5939 lemma mopup_inv_step:
       
  5940   "\<lbrakk>n < length lm; mopup_inv (s, l, r) lm n ires\<rbrakk>
       
  5941   \<Longrightarrow> mopup_inv (t_step (s, l, r) 
       
  5942        ((mop_bef n @ tshift mp_up (2 * n)), 0)) lm n ires"
       
  5943 apply(auto split:if_splits simp add:t_step.simps,
       
  5944       tactic {* ALLGOALS (resolve_tac [@{thm "fetch_intro"}]) *})
       
  5945 apply(simp_all add: mopupfetchs new_tape.simps)
       
  5946 done
       
  5947 
       
  5948 declare mopup_inv.simps[simp del]
       
  5949 
       
  5950 lemma mopup_inv_steps: 
       
  5951 "\<lbrakk>n < length lm; mopup_inv (s, l, r) lm n ires\<rbrakk> \<Longrightarrow> 
       
  5952      mopup_inv (t_steps (s, l, r) 
       
  5953                    ((mop_bef n @ tshift mp_up (2 * n)), 0) stp) lm n ires"
       
  5954 apply(induct stp, simp add: t_steps.simps)
       
  5955 apply(simp add: t_steps_ind)
       
  5956 apply(case_tac "(t_steps (s, l, r) 
       
  5957                 (mop_bef n @ tshift mp_up (2 * n), 0) stp)", simp)
       
  5958 apply(rule_tac mopup_inv_step, simp, simp)
       
  5959 done
       
  5960 
       
  5961 lemma [simp]: 
       
  5962  "\<lbrakk>n < length lm; Suc 0 \<le> n\<rbrakk> \<Longrightarrow> 
       
  5963             mopup_bef_erase_a (Suc 0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>) lm n ires"
       
  5964 apply(auto simp: mopup_bef_erase_a.simps  abc_lm_v.simps)
       
  5965 apply(case_tac lm, simp, case_tac list, simp, simp)
       
  5966 apply(rule_tac x = "Suc a" in exI, rule_tac x = rn in exI, simp)
       
  5967 done
       
  5968   
       
  5969 lemma [simp]:
       
  5970   "lm \<noteq> [] \<Longrightarrow> mopup_jump_over1 (Suc 0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>) lm 0  ires"
       
  5971 apply(auto simp: mopup_jump_over1.simps)
       
  5972 apply(rule_tac x = ln in exI, rule_tac x = 0 in exI, simp add: )
       
  5973 apply(case_tac lm, simp, simp add: abc_lm_v.simps)
       
  5974 apply(case_tac rn, simp)
       
  5975 apply(case_tac list, rule_tac disjI2, 
       
  5976       simp add: tape_of_nl_abv tape_of_nat_list.simps)
       
  5977 apply(rule_tac disjI1,
       
  5978       simp add: tape_of_nl_abv tape_of_nat_list.simps )
       
  5979 apply(rule_tac disjI1, case_tac list, 
       
  5980       simp add: tape_of_nl_abv tape_of_nat_list.simps, 
       
  5981       rule_tac x = nat in exI, simp)
       
  5982 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps )
       
  5983 done
       
  5984 
       
  5985 lemma mopup_init: 
       
  5986  "\<lbrakk>n < length lm; crsp_l ly (as, lm) (ac, l, r) ires\<rbrakk> \<Longrightarrow> 
       
  5987                                mopup_inv (Suc 0, l, r) lm n ires"
       
  5988 apply(auto simp: crsp_l.simps mopup_inv.simps)
       
  5989 apply(case_tac n, simp, auto simp: mopup_bef_erase_a.simps )
       
  5990 apply(rule_tac x = "Suc (hd lm)" in exI, rule_tac x = rn in exI, simp)
       
  5991 apply(case_tac lm, simp, case_tac list, simp, case_tac lista, simp add: abc_lm_v.simps)
       
  5992 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps)
       
  5993 apply(simp add: mopup_jump_over1.simps)
       
  5994 apply(rule_tac x = 0 in exI, rule_tac x = 0 in exI, auto)
       
  5995 apply(case_tac [!] n, simp_all)
       
  5996 apply(case_tac [!] lm, simp, case_tac list, simp)
       
  5997 apply(case_tac rn, simp add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps)
       
  5998 apply(erule_tac x = nat in allE, simp add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps)
       
  5999 apply(simp add: abc_lm_v.simps, auto)
       
  6000 apply(case_tac list, simp_all add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps) 
       
  6001 apply(erule_tac x = rn in allE, simp_all)
       
  6002 done
       
  6003 
       
  6004 (***Begin: mopup stop***)
       
  6005 fun abc_mopup_stage1 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat"
       
  6006   where
       
  6007   "abc_mopup_stage1 (s, l, r) n = 
       
  6008            (if s > 0 \<and> s \<le> 2*n then 6
       
  6009             else if s = 2*n + 1 then 4
       
  6010             else if s \<ge> 2*n + 2 \<and> s \<le> 2*n + 4 then 3
       
  6011             else if s = 2*n + 5 then 2
       
  6012             else if s = 2*n + 6 then 1
       
  6013             else 0)"
       
  6014 
       
  6015 fun abc_mopup_stage2 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat"
       
  6016   where
       
  6017   "abc_mopup_stage2 (s, l, r) n = 
       
  6018            (if s > 0 \<and> s \<le> 2*n then length r
       
  6019             else if s = 2*n + 1 then length r
       
  6020             else if s = 2*n + 5 then length l
       
  6021             else if s = 2*n + 6 then length l
       
  6022             else if s \<ge> 2*n + 2 \<and> s \<le> 2*n + 4 then length r
       
  6023             else 0)"
       
  6024 
       
  6025 fun abc_mopup_stage3 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat"
       
  6026   where
       
  6027   "abc_mopup_stage3 (s, l, r) n = 
       
  6028           (if s > 0 \<and> s \<le> 2*n then 
       
  6029               if hd r = Bk then 0
       
  6030               else 1
       
  6031            else if s = 2*n + 2 then 1 
       
  6032            else if s = 2*n + 3 then 0
       
  6033            else if s = 2*n + 4 then 2
       
  6034            else 0)"
       
  6035 
       
  6036 fun abc_mopup_measure :: "(t_conf \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)"
       
  6037   where
       
  6038   "abc_mopup_measure (c, n) = 
       
  6039     (abc_mopup_stage1 c n, abc_mopup_stage2 c n, 
       
  6040                                        abc_mopup_stage3 c n)"
       
  6041 
       
  6042 definition abc_mopup_LE ::
       
  6043    "(((nat \<times> block list \<times> block list) \<times> nat) \<times> 
       
  6044     ((nat \<times> block list \<times> block list) \<times> nat)) set"
       
  6045   where
       
  6046   "abc_mopup_LE \<equiv> (inv_image lex_triple abc_mopup_measure)"
       
  6047 
       
  6048 lemma wf_abc_mopup_le[intro]: "wf abc_mopup_LE"
       
  6049 by(auto intro:wf_inv_image wf_lex_triple simp:abc_mopup_LE_def)
       
  6050 
       
  6051 lemma [simp]: "mopup_bef_erase_a (a, aa, []) lm n ires = False"
       
  6052 apply(auto simp: mopup_bef_erase_a.simps)
       
  6053 done
       
  6054 
       
  6055 lemma [simp]: "mopup_bef_erase_b (a, aa, []) lm n ires = False"
       
  6056 apply(auto simp: mopup_bef_erase_b.simps) 
       
  6057 done
       
  6058 
       
  6059 lemma [simp]: "mopup_aft_erase_b (2 * n + 3, aa, []) lm n ires = False"
       
  6060 apply(auto simp: mopup_aft_erase_b.simps)
       
  6061 done
       
  6062 
       
  6063 lemma mopup_halt_pre: 
       
  6064  "\<lbrakk>n < length lm; mopup_inv (Suc 0, l, r) lm n ires; wf abc_mopup_LE\<rbrakk>
       
  6065  \<Longrightarrow>  \<forall>na. \<not> (\<lambda>(s, l, r) n. s = 0) (t_steps (Suc 0, l, r)
       
  6066       (mop_bef n @ tshift mp_up (2 * n), 0) na) n \<longrightarrow>
       
  6067        ((t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) 
       
  6068         (Suc na), n),
       
  6069        t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0)
       
  6070          na, n) \<in> abc_mopup_LE"
       
  6071 apply(rule allI, rule impI, simp add: t_steps_ind)
       
  6072 apply(subgoal_tac "mopup_inv (t_steps (Suc 0, l, r) 
       
  6073                      (mop_bef n @ tshift mp_up (2 * n), 0) na) lm n ires")
       
  6074 apply(case_tac "(t_steps (Suc 0, l, r) 
       
  6075                (mop_bef n @ tshift mp_up (2 * n), 0) na)",  simp)
       
  6076 proof -
       
  6077   fix na a b c
       
  6078   assume  "n < length lm" "mopup_inv (a, b, c) lm n ires" "0 < a"
       
  6079   thus "((t_step (a, b, c) (mop_bef n @ tshift mp_up (2 * n), 0), n),
       
  6080          (a, b, c), n) \<in> abc_mopup_LE"
       
  6081     apply(auto split:if_splits simp add:t_step.simps mopup_inv.simps,
       
  6082       tactic {* ALLGOALS (resolve_tac [@{thm "fetch_intro"}]) *})
       
  6083     apply(simp_all add: mopupfetchs new_tape.simps abc_mopup_LE_def 
       
  6084                    lex_triple_def lex_pair_def )
       
  6085     done
       
  6086 next
       
  6087   fix na
       
  6088   assume "n < length lm" "mopup_inv (Suc 0, l, r) lm n ires"
       
  6089   thus "mopup_inv (t_steps (Suc 0, l, r) 
       
  6090        (mop_bef n @ tshift mp_up (2 * n), 0) na) lm n ires"
       
  6091     apply(rule mopup_inv_steps)
       
  6092     done
       
  6093 qed
       
  6094 
       
  6095 lemma mopup_halt: "\<lbrakk>n < length lm; crsp_l ly (as, lm) (s, l, r) ires\<rbrakk> \<Longrightarrow> 
       
  6096   \<exists> stp. (\<lambda> (s, l, r). s = 0) (t_steps (Suc 0, l, r) 
       
  6097         ((mop_bef n @ tshift mp_up (2 * n)), 0) stp)"
       
  6098 apply(subgoal_tac "mopup_inv (Suc 0, l, r) lm n ires")
       
  6099 apply(insert wf_abc_mopup_le)
       
  6100 apply(insert halt_lemma[of abc_mopup_LE 
       
  6101     "\<lambda> ((s, l, r), n). s = 0" 
       
  6102     "\<lambda> stp. (t_steps (Suc 0, l, r) ((mop_bef n @ tshift mp_up (2 * n))
       
  6103            , 0) stp, n)"], auto)
       
  6104 apply(insert mopup_halt_pre[of n lm l r], simp, erule exE)
       
  6105 apply(rule_tac x = na in exI, case_tac "(t_steps (Suc 0, l, r) 
       
  6106           (mop_bef n @ tshift mp_up (2 * n), 0) na)", simp)
       
  6107 apply(rule_tac mopup_init, auto)
       
  6108 done
       
  6109 (***End: mopup stop****)
       
  6110 (*
       
  6111 lemma mopup_stop_cond: "mopup_inv (0, l, r) lm n ires \<Longrightarrow> 
       
  6112                                      (\<exists>ln rn. ?l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ?ires \<and> ?r = <abc_lm_v ?lm ?n> @ Bk\<^bsup>rn\<^esup>) "
       
  6113          t_halt_conf (0, l, r) \<and> t_result r = Suc (abc_lm_v lm n)"
       
  6114 apply(simp add: mopup_inv.simps mopup_stop.simps t_halt_conf.simps
       
  6115                 t_result.simps, auto simp: tape_of_nat_abv)
       
  6116 apply(rule_tac x = rn in exI, 
       
  6117       rule_tac x = "Suc (abc_lm_v lm n)" in exI,
       
  6118        simp add: tape_of_nat_abv)
       
  6119 apply(simp add: tape_of_nat_abv  exponent_def)
       
  6120 apply(subgoal_tac "takeWhile (\<lambda>a. a = Oc) 
       
  6121              (replicate (abc_lm_v lm n) Oc @ replicate rn Bk)
       
  6122        = replicate (abc_lm_v lm n) Oc @ takeWhile (\<lambda>a. a = Oc)
       
  6123                                           (replicate rn Bk)", simp)
       
  6124 apply(case_tac rn, simp, simp)
       
  6125 apply(rule takeWhile_append2)
       
  6126 apply(case_tac x, auto)
       
  6127 done
       
  6128 *)
       
  6129 
       
  6130 
       
  6131 lemma mopup_halt_conf_pre: 
       
  6132  "\<lbrakk>n < length lm; crsp_l ly (as, lm) (s, l, r) ires\<rbrakk> 
       
  6133   \<Longrightarrow> \<exists> na. (\<lambda> (s', l', r').  s' = 0 \<and> mopup_stop (s', l', r') lm n ires)
       
  6134       (t_steps (Suc 0, l, r) 
       
  6135             ((mop_bef n @ tshift mp_up (2 * n)), 0) na)"
       
  6136 apply(subgoal_tac "\<exists> stp. (\<lambda> (s, l, r). s = 0) 
       
  6137  (t_steps (Suc 0, l, r) ((mop_bef n @ tshift mp_up (2 * n)), 0) stp)",
       
  6138        erule exE)
       
  6139 apply(rule_tac x = stp in exI, 
       
  6140       case_tac "(t_steps (Suc 0, l, r) 
       
  6141           (mop_bef n @ tshift mp_up (2 * n), 0) stp)", simp)
       
  6142 apply(subgoal_tac " mopup_inv (Suc 0, l, r) lm n ires")
       
  6143 apply(subgoal_tac "mopup_inv (t_steps (Suc 0, l, r) 
       
  6144             (mop_bef n @ tshift mp_up (2 * n), 0) stp) lm n ires", simp)
       
  6145 apply(simp only: mopup_inv.simps)
       
  6146 apply(rule_tac mopup_inv_steps, simp, simp)
       
  6147 apply(rule_tac mopup_init, simp, simp)
       
  6148 apply(rule_tac mopup_halt, simp, simp)
       
  6149 done
       
  6150 
       
  6151 thm mopup_stop.simps
       
  6152 
       
  6153 lemma  mopup_halt_conf:
       
  6154   assumes len: "n < length lm"
       
  6155   and correspond: "crsp_l ly (as, lm) (s, l, r) ires"
       
  6156   shows 
       
  6157   "\<exists> na. (\<lambda> (s', l', r'). ((\<exists>ln rn. s' = 0 \<and> l' = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \<and> r' = Oc\<^bsup>Suc (abc_lm_v lm n)\<^esup> @ Bk\<^bsup>rn\<^esup>)))
       
  6158              (t_steps (Suc 0, l, r) 
       
  6159                   ((mop_bef n @ tshift mp_up (2 * n)), 0) na)"
       
  6160 using len correspond mopup_halt_conf_pre[of n lm ly as s l r ires]
       
  6161 apply(simp add: mopup_stop.simps tape_of_nat_abv tape_of_nat_list.simps)
       
  6162 done
       
  6163 (*********End: mop_up****************************)
       
  6164 
       
  6165 
       
  6166 subsubsection {* Final results about Abacus machine *}
       
  6167 
       
  6168 thm mopup_halt
       
  6169 lemma mopup_halt_bef: "\<lbrakk>n < length lm; crsp_l ly (as, lm) (s, l, r) ires\<rbrakk> 
       
  6170     \<Longrightarrow> \<exists>stp. (\<lambda>(s, l, r). s \<noteq> 0 \<and> ((\<lambda> (s', l', r'). s' = 0)
       
  6171    (t_step (s, l, r) (mop_bef n @ tshift mp_up (2 * n), 0))))
       
  6172     (t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) stp)"
       
  6173 apply(insert mopup_halt[of n lm ly as s l r ires], simp, erule_tac exE)
       
  6174 proof -
       
  6175   fix stp
       
  6176   assume "n < length lm" 
       
  6177          "crsp_l ly (as, lm) (s, l, r) ires" 
       
  6178          "(\<lambda>(s, l, r). s = 0) 
       
  6179             (t_steps (Suc 0, l, r) 
       
  6180               (mop_bef n @ tshift mp_up (2 * n), 0) stp)"
       
  6181   thus "\<exists>stp. (\<lambda>(s, ab). 0 < s \<and> (\<lambda>(s', l', r'). s' = 0) 
       
  6182    (t_step (s, ab) (mop_bef n @ tshift mp_up (2 * n), 0))) 
       
  6183     (t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) stp)"
       
  6184   proof(induct stp, simp add: t_steps.simps, simp)
       
  6185     fix stpa
       
  6186     assume h1: 
       
  6187       "(\<lambda>(s, l, r). s = 0) (t_steps (Suc 0, l, r) 
       
  6188            (mop_bef n @ tshift mp_up (2 * n), 0) stpa) \<Longrightarrow>
       
  6189        \<exists>stp. (\<lambda>(s, ab). 0 < s \<and> (\<lambda>(s', l', r'). s' = 0) 
       
  6190          (t_step (s, ab) (mop_bef n @ tshift mp_up (2 * n), 0))) 
       
  6191             (t_steps (Suc 0, l, r) 
       
  6192               (mop_bef n @ tshift mp_up (2 * n), 0) stp)"
       
  6193       and h2: 
       
  6194         "(\<lambda>(s, l, r). s = 0) (t_steps (Suc 0, l, r) 
       
  6195                     (mop_bef n @ tshift mp_up (2 * n), 0) (Suc stpa))"
       
  6196          "n < length lm" 
       
  6197          "crsp_l ly (as, lm) (s, l, r) ires"
       
  6198     thus "\<exists>stp. (\<lambda>(s, ab). 0 < s \<and> (\<lambda>(s', l', r'). s' = 0) 
       
  6199              (t_step (s, ab) (mop_bef n @ tshift mp_up (2 * n), 0))) (
       
  6200                 t_steps (Suc 0, l, r) 
       
  6201                   (mop_bef n @ tshift mp_up (2 * n), 0) stp)"
       
  6202       apply(case_tac "(\<lambda>(s, l, r). s = 0) (t_steps (Suc 0, l, r) 
       
  6203                      (mop_bef n @ tshift mp_up (2 * n), 0) stpa)", 
       
  6204             simp)
       
  6205       apply(rule_tac x = "stpa" in exI)
       
  6206       apply(case_tac "(t_steps (Suc 0, l, r) 
       
  6207                          (mop_bef n @ tshift mp_up (2 * n), 0) stpa)",
       
  6208             simp add: t_steps_ind)
       
  6209       done
       
  6210   qed
       
  6211 qed
       
  6212 
       
  6213 lemma tshift_nth_state0: "\<lbrakk>n < length tp; tp ! n = (a, 0)\<rbrakk>
       
  6214     \<Longrightarrow> tshift tp off ! n = (a, 0)"
       
  6215 apply(induct n, case_tac tp, simp)
       
  6216 apply(auto simp: tshift.simps)
       
  6217 done
       
  6218 
       
  6219 lemma shift_length: "length (tshift tp n) = length tp"
       
  6220 apply(auto simp: tshift.simps)
       
  6221 done
       
  6222 
       
  6223 lemma even_Suc_le: "\<lbrakk>y mod 2 = 0; 2 * x < y\<rbrakk> \<Longrightarrow> Suc (2 * x) < y"
       
  6224 by arith
       
  6225 
       
  6226 lemma [simp]: "(4::nat) * n mod 2 = 0"
       
  6227 by arith
       
  6228 
       
  6229 lemma tm_append_fetch_equal: 
       
  6230   "\<lbrakk>t_ncorrect (tm_of aprog); s'> 0;
       
  6231     fetch (mop_bef n @ tshift mp_up (2 * n)) s' b = (a, 0)\<rbrakk>
       
  6232 \<Longrightarrow> fetch (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) 
       
  6233     (length (tm_of aprog) div 2)) (s' + length (tm_of aprog) div 2) b 
       
  6234    = (a, 0)"
       
  6235 apply(case_tac s', simp)
       
  6236 apply(auto simp: fetch.simps nth_of.simps t_ncorrect.simps shift_length nth_append
       
  6237                  tshift.simps split: list.splits block.splits split: if_splits)
       
  6238 done
       
  6239 
       
  6240 lemma [simp]:
       
  6241   "\<lbrakk>t_ncorrect (tm_of aprog);
       
  6242     t_step (s', l', r') (mop_bef n @ tshift mp_up (2 * n), 0) = 
       
  6243                                                (0, l'', r''); s' > 0\<rbrakk>
       
  6244   \<Longrightarrow> t_step (s' + length (tm_of aprog) div 2, l', r') 
       
  6245         (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n))
       
  6246            (length (tm_of aprog) div 2), 0) = (0, l'', r'')"
       
  6247 apply(simp add: t_step.simps)
       
  6248 apply(subgoal_tac 
       
  6249    "(fetch (mop_bef n @ tshift mp_up (2 * n)) s' 
       
  6250               (case r' of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))
       
  6251   = (fetch (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) 
       
  6252        (length (tm_of aprog) div 2)) (s' + length (tm_of aprog) div 2)
       
  6253     (case r' of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))", simp)
       
  6254 apply(case_tac "(fetch (mop_bef n @ tshift mp_up (2 * n)) s' 
       
  6255        (case r' of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))", simp)
       
  6256 apply(drule_tac tm_append_fetch_equal, auto)
       
  6257 done
       
  6258 
       
  6259 lemma [intro]: 
       
  6260   "start_of (layout_of aprog) (length aprog) - Suc 0 = 
       
  6261                                       length (tm_of aprog) div 2"
       
  6262 apply(subgoal_tac  "abc2t_correct aprog")
       
  6263 apply(insert pre_lheq[of "concat (take (length aprog) 
       
  6264        (tms_of aprog))" "length aprog" aprog], simp add: tm_of.simps)
       
  6265 by auto
       
  6266 
       
  6267 lemma tm_append_stop_step: 
       
  6268   "\<lbrakk>t_ncorrect (tm_of aprog); 
       
  6269     t_ncorrect (mop_bef n @ tshift mp_up (2 * n)); n < length lm; 
       
  6270    (t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) stp) =
       
  6271                          (s', l', r');
       
  6272     s' \<noteq> 0;
       
  6273     t_step (s', l', r') (mop_bef n @ tshift mp_up (2 * n), 0) 
       
  6274                                                      = (0, l'', r'')\<rbrakk>
       
  6275      \<Longrightarrow>
       
  6276 (t_steps ((start_of (layout_of aprog) (length aprog), l, r)) 
       
  6277   (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n))
       
  6278    (start_of (layout_of aprog) (length aprog) - Suc 0), 0) (Suc stp))
       
  6279   = (0, l'', r'')"
       
  6280 apply(insert tm_append_steps_equal[of "tm_of aprog" 
       
  6281       "(mop_bef n @ tshift mp_up (2 * n))"
       
  6282       "(start_of (layout_of aprog) (length aprog) - Suc 0)" 
       
  6283       "Suc 0" l r stp], simp)
       
  6284 apply(subgoal_tac "(start_of (layout_of aprog) (length aprog) - Suc 0)
       
  6285               = (length (tm_of aprog) div 2)", simp add: t_steps_ind)
       
  6286 apply(case_tac 
       
  6287  "(t_steps (start_of (layout_of aprog) (length aprog), l, r) 
       
  6288       (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n))
       
  6289            (length (tm_of aprog) div 2), 0) stp)", simp)
       
  6290 apply(subgoal_tac "start_of (layout_of aprog) (length aprog) > 0", 
       
  6291       case_tac "start_of (layout_of aprog) (length aprog)", 
       
  6292       simp, simp)
       
  6293 apply(rule startof_not0, auto)
       
  6294 done
       
  6295 
       
  6296 (*
       
  6297 lemma stop_conf: "mopup_inv (0, aca, bc) am n
       
  6298     \<Longrightarrow> t_halt_conf (0, aca, bc) \<and> t_result bc = Suc (abc_lm_v am n)"
       
  6299 apply(case_tac n, 
       
  6300       auto simp: mopup_inv.simps mopup_stop.simps t_halt_conf.simps 
       
  6301                  t_result.simps tape_of_nl_abv tape_of_nat_abv )
       
  6302 apply(rule_tac x = "rn" in exI, 
       
  6303       rule_tac x = "Suc (abc_lm_v am 0)" in exI, simp) 
       
  6304 apply(subgoal_tac "takeWhile (\<lambda>a. a = Oc) (Oc\<^bsup>abc_lm_v am 0\<^esup> @ Bk\<^bsup>rn\<^esup>)
       
  6305               = Oc\<^bsup>abc_lm_v am 0\<^esup> @ takeWhile (\<lambda>a. a = Oc) (Bk\<^bsup>rn\<^esup>)", simp)
       
  6306 apply(simp add: exponent_def, case_tac rn, simp, simp)
       
  6307 apply(rule_tac takeWhile_append2, simp add: exponent_def)
       
  6308 apply(rule_tac x = rn in exI,
       
  6309       rule_tac x = "Suc (abc_lm_v am (Suc nat))" in exI, simp)
       
  6310 apply(subgoal_tac 
       
  6311  "takeWhile (\<lambda>a. a = Oc) (Oc\<^bsup>abc_lm_v am (Suc nat)\<^esup> @ Bk\<^bsup>rn\<^esup>) = 
       
  6312        Oc\<^bsup>abc_lm_v am (Suc nat)\<^esup> @ takeWhile (\<lambda>a. a = Oc) (Bk\<^bsup>rn\<^esup>)", simp)
       
  6313 apply(simp add: exponent_def, case_tac rn, simp, simp)
       
  6314 apply(rule_tac takeWhile_append2, simp add: exponent_def)
       
  6315 done
       
  6316 *)
       
  6317 
       
  6318 
       
  6319 lemma start_of_out_range: 
       
  6320 "as \<ge> length aprog \<Longrightarrow> 
       
  6321    start_of (layout_of aprog) as = 
       
  6322              start_of (layout_of aprog) (length aprog)"
       
  6323 apply(induct as, simp)
       
  6324 apply(case_tac "length aprog = Suc as", simp)
       
  6325 apply(simp add: start_of.simps)
       
  6326 done
       
  6327 
       
  6328 lemma [intro]: "t_ncorrect (tm_of aprog)"
       
  6329 apply(simp add: tm_of.simps)
       
  6330 apply(insert tms_mod2[of "length aprog" aprog], 
       
  6331                                 simp add: t_ncorrect.simps)
       
  6332 done
       
  6333 
       
  6334 lemma abacus_turing_eq_halt_case_pre: 
       
  6335    "\<lbrakk>ly = layout_of aprog; 
       
  6336      tprog = tm_of aprog; 
       
  6337      crsp_l ly ac tc ires;
       
  6338      n < length am;
       
  6339      abc_steps_l ac aprog stp = (as, am); 
       
  6340      mop_ss = start_of ly (length aprog);
       
  6341      as \<ge> length aprog\<rbrakk>
       
  6342      \<Longrightarrow> \<exists> stp. (\<lambda> (s, l, r). s = 0 \<and> mopup_inv (0, l, r) am n ires)
       
  6343                 (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp)"
       
  6344 apply(insert steps_crsp[of ly aprog tprog ac tc ires "(as, am)" stp], auto)
       
  6345 apply(case_tac "(t_steps tc (tm_of aprog, 0) n')",  
       
  6346       simp add: tMp.simps)
       
  6347 apply(subgoal_tac "t_ncorrect (mop_bef n @ tshift mp_up (2 * n))")
       
  6348 proof -
       
  6349   fix n' a b c
       
  6350   assume h1: 
       
  6351     "crsp_l (layout_of aprog) ac tc ires" 
       
  6352     "abc_steps_l ac aprog stp = (as, am)" 
       
  6353     "length aprog \<le> as"
       
  6354     "crsp_l (layout_of aprog) (as, am) (a, b, c) ires"
       
  6355     "t_steps tc (tm_of aprog, 0) n' = (a, b, c)"
       
  6356     "n < length am"
       
  6357     "t_ncorrect (mop_bef n @ tshift mp_up (2 * n))"
       
  6358   hence h2:
       
  6359   "t_steps tc (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n))
       
  6360     (start_of (layout_of aprog) (length aprog) - Suc 0), 0) n' 
       
  6361                                     = (a, b, c)" 
       
  6362     apply(rule_tac tm_append_steps, simp)
       
  6363     apply(simp add: crsp_l.simps, auto)
       
  6364     apply(simp add: crsp_l.simps)
       
  6365     apply(rule startof_not0)
       
  6366     done
       
  6367   from h1 have h3: 
       
  6368   "\<exists>stp. (\<lambda>(s, l, r). s \<noteq> 0 \<and> ((\<lambda> (s', l', r'). s' = 0) 
       
  6369            (t_step (s, l, r) (mop_bef n @ tshift mp_up (2 * n), 0))))
       
  6370          (t_steps (Suc 0, b, c) 
       
  6371                (mop_bef n @ tshift mp_up (2 * n), 0) stp)"
       
  6372     apply(rule_tac mopup_halt_bef, auto)
       
  6373     done
       
  6374   from h1 and h2 and h3 show 
       
  6375     "\<exists>stp. case t_steps tc (tm_of aprog @ abacus.tshift (mop_bef n @ abacus.tshift mp_up (2 * n))
       
  6376     (start_of (layout_of aprog) (length aprog) - Suc 0), 0) stp of (s, ab)
       
  6377     \<Rightarrow> s = 0 \<and> mopup_inv (0, ab) am n ires"
       
  6378   proof(erule_tac exE, 
       
  6379     case_tac "(t_steps (Suc 0, b, c) 
       
  6380               (mop_bef n @ tshift mp_up (2 * n), 0) stpa)", simp,
       
  6381     case_tac "(t_step (aa, ba, ca) 
       
  6382               (mop_bef n @ tshift mp_up (2 * n), 0))", simp)
       
  6383     fix stpa aa ba ca aaa baa caa
       
  6384     assume g1: "0 < aa \<and> aaa = 0" 
       
  6385       "t_steps (Suc 0, b, c) 
       
  6386       (mop_bef n @ tshift mp_up (2 * n), 0) stpa = (aa, ba,ca)" 
       
  6387       "t_step (aa, ba, ca) (mop_bef n @ tshift mp_up (2 * n), 0)
       
  6388       = (0, baa, caa)"
       
  6389     from h1 and this have g2: 
       
  6390       "t_steps (start_of (layout_of aprog) (length aprog), b, c) 
       
  6391          (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) 
       
  6392            (start_of (layout_of aprog) (length aprog) - Suc 0), 0) 
       
  6393                 (Suc stpa) = (0, baa, caa)"
       
  6394       apply(rule_tac tm_append_stop_step, auto)
       
  6395       done
       
  6396     from h1 and h2 and g1 and this show "?thesis"
       
  6397       apply(rule_tac x = "n' + Suc stpa" in exI)
       
  6398       apply(simp add: t_steps_ind del: t_steps.simps)
       
  6399       apply(subgoal_tac "a = start_of (layout_of aprog) 
       
  6400                                           (length aprog)", simp)
       
  6401       apply(insert mopup_inv_steps[of n am "Suc 0" b c ires "Suc stpa"],
       
  6402             simp add: t_steps_ind)
       
  6403       apply(subgoal_tac "mopup_inv (Suc 0, b, c) am n ires", simp)
       
  6404       apply(rule_tac mopup_init, simp, simp)
       
  6405       apply(simp add: crsp_l.simps)
       
  6406       apply(erule_tac start_of_out_range)
       
  6407       done
       
  6408   qed
       
  6409 next
       
  6410   show " t_ncorrect (mop_bef n @ tshift mp_up (2 * n))"
       
  6411     apply(auto simp: t_ncorrect.simps tshift.simps mp_up_def)
       
  6412     done   
       
  6413 qed
       
  6414 
       
  6415 text {*
       
  6416   One of the main theorems about Abacus compilation.
       
  6417 *}
       
  6418 
       
  6419 lemma abacus_turing_eq_halt_case: 
       
  6420   assumes layout: 
       
  6421   -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *}
       
  6422   "ly = layout_of aprog"
       
  6423   and complied: 
       
  6424   -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *}
       
  6425   "tprog = tm_of aprog"
       
  6426   and correspond: 
       
  6427   -- {* TM configuration @{text "tc"} and Abacus configuration @{text "ac"}
       
  6428   are in correspondence: *}
       
  6429   "crsp_l ly ac tc ires"
       
  6430   and halt_state: 
       
  6431   -- {* @{text "as"} is a program label outside the range of @{text "aprog"}. So 
       
  6432   if Abacus is in such a state, it is in halt state: *}
       
  6433   "as \<ge> length aprog"
       
  6434   and abc_exec: 
       
  6435   -- {* Supposing after @{text "stp"} step of execution, Abacus program @{text "aprog"}
       
  6436   reaches such a halt state: *}
       
  6437   "abc_steps_l ac aprog stp = (as, am)"
       
  6438   and rs_len: 
       
  6439   -- {* @{text "n"} is a memory address in the range of Abacus memory @{text "am"}: *}
       
  6440   "n < length am"
       
  6441   and mopup_start:
       
  6442   -- {* The startling label for mopup mahines, according to the layout and Abacus program 
       
  6443    should be @{text "mop_ss"}: *}
       
  6444   "mop_ss = start_of ly (length aprog)"
       
  6445   shows 
       
  6446   -- {* 
       
  6447   After @{text "stp"} steps of execution of the TM composed of @{text "tprog"} and the mopup 
       
  6448   TM @{text "(tMp n (mop_ss - 1))"} will halt and gives rise to a configuration which 
       
  6449   only hold the content of memory cell @{text "n"}:
       
  6450   *}
       
  6451   "\<exists> stp. (\<lambda> (s, l, r). \<exists> ln rn. s = 0 \<and>  l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \<and> r = Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>rn\<^esup>)
       
  6452            (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp)"
       
  6453 proof -
       
  6454   from layout complied correspond halt_state abc_exec rs_len mopup_start
       
  6455        and abacus_turing_eq_halt_case_pre [of ly aprog tprog ac tc ires n am stp as mop_ss]
       
  6456   show "?thesis" 
       
  6457     apply(simp add: mopup_inv.simps mopup_stop.simps tape_of_nat_abv)
       
  6458     done
       
  6459 qed
       
  6460 
       
  6461 lemma abc_unhalt_case_zero: 
       
  6462 "\<lbrakk>crsp_l (layout_of aprog) ac tc ires;
       
  6463   n < length am; 
       
  6464   \<forall>stp. (\<lambda>(as, am). as < length aprog) (abc_steps_l ac aprog stp)\<rbrakk>
       
  6465  \<Longrightarrow> \<exists>astp bstp. 0 \<le> bstp \<and> 
       
  6466           crsp_l (layout_of aprog) (abc_steps_l ac aprog astp) 
       
  6467                 (t_steps tc (tm_of aprog, 0) bstp) ires"
       
  6468 apply(rule_tac x = "Suc 0" in exI)
       
  6469 apply(case_tac " abc_steps_l ac aprog (Suc 0)", simp)
       
  6470 proof -
       
  6471   fix a b
       
  6472   assume "crsp_l (layout_of aprog) ac tc ires" 
       
  6473          "abc_steps_l ac aprog (Suc 0) = (a, b)"
       
  6474   thus "\<exists>bstp. crsp_l (layout_of aprog) (a, b) 
       
  6475                (t_steps tc (tm_of aprog, 0) bstp) ires"
       
  6476     apply(insert steps_crsp[of "layout_of aprog" aprog 
       
  6477                   "tm_of aprog" ac tc ires "(a, b)" "Suc 0"], auto)
       
  6478     done
       
  6479 qed
       
  6480 
       
  6481 declare abc_steps_l.simps[simp del]
       
  6482 
       
  6483 lemma abc_steps_ind: 
       
  6484  "let (as, am) = abc_steps_l ac aprog stp in 
       
  6485    abc_steps_l ac aprog (Suc stp) =
       
  6486               abc_step_l (as, am) (abc_fetch as aprog) "
       
  6487 proof(simp)
       
  6488   show "(\<lambda>(as, am). abc_steps_l ac aprog (Suc stp) = 
       
  6489         abc_step_l (as, am) (abc_fetch as aprog)) 
       
  6490               (abc_steps_l ac aprog stp)"
       
  6491   proof(induct stp arbitrary: ac)
       
  6492     fix ac
       
  6493     show "(\<lambda>(as, am). abc_steps_l ac aprog (Suc 0) = 
       
  6494             abc_step_l (as, am) (abc_fetch as aprog))  
       
  6495                     (abc_steps_l ac aprog 0)"
       
  6496       apply(case_tac "ac:: nat \<times> nat list", 
       
  6497             simp add: abc_steps_l.simps)
       
  6498       apply(case_tac "(abc_step_l (a, b) (abc_fetch a aprog))",
       
  6499             simp add: abc_steps_l.simps)
       
  6500       done
       
  6501   next
       
  6502     fix stp ac
       
  6503     assume h1:
       
  6504       "(\<And>ac. (\<lambda>(as, am). abc_steps_l ac aprog (Suc stp) =
       
  6505                             abc_step_l (as, am) (abc_fetch as aprog)) 
       
  6506              (abc_steps_l ac aprog stp))"
       
  6507     thus 
       
  6508       "(\<lambda>(as, am). abc_steps_l ac aprog (Suc (Suc stp)) =
       
  6509               abc_step_l (as, am) (abc_fetch as aprog)) 
       
  6510                              (abc_steps_l ac aprog (Suc stp))"
       
  6511       apply(case_tac "ac::nat \<times> nat list", simp)
       
  6512       apply(subgoal_tac 
       
  6513            "abc_steps_l (a, b) aprog (Suc (Suc stp)) =
       
  6514             abc_steps_l (abc_step_l (a, b) (abc_fetch a aprog)) 
       
  6515                                               aprog (Suc stp)", simp)
       
  6516       apply(case_tac "(abc_step_l (a, b) (abc_fetch a aprog))", simp)
       
  6517     proof -
       
  6518       fix a b aa ba
       
  6519       assume h2: "abc_step_l (a, b) (abc_fetch a aprog) = (aa, ba)"
       
  6520       from h1 and h2  show 
       
  6521       "(\<lambda>(as, am). abc_steps_l (aa, ba) aprog (Suc stp) = 
       
  6522           abc_step_l (as, am) (abc_fetch as aprog)) 
       
  6523                     (abc_steps_l (a, b) aprog (Suc stp))"
       
  6524 	apply(insert h1[of "(aa, ba)"])
       
  6525 	apply(simp add: abc_steps_l.simps)
       
  6526 	apply(insert h2, simp)
       
  6527 	done
       
  6528     next
       
  6529       fix a b
       
  6530       show 
       
  6531         "abc_steps_l (a, b) aprog (Suc (Suc stp)) = 
       
  6532          abc_steps_l (abc_step_l (a, b) (abc_fetch a aprog)) 
       
  6533                                                    aprog (Suc stp)"
       
  6534 	apply(simp only: abc_steps_l.simps)
       
  6535 	done
       
  6536     qed
       
  6537   qed
       
  6538 qed
       
  6539 
       
  6540 lemma abc_unhalt_case_induct: 
       
  6541   "\<lbrakk>crsp_l (layout_of aprog) ac tc ires;
       
  6542     n < length am; 
       
  6543     \<forall>stp. (\<lambda>(as, am). as < length aprog) (abc_steps_l ac aprog stp); 
       
  6544     stp \<le> bstp;
       
  6545     crsp_l (layout_of aprog) (abc_steps_l ac aprog astp) 
       
  6546                            (t_steps tc (tm_of aprog, 0) bstp) ires\<rbrakk>
       
  6547  \<Longrightarrow> \<exists>astp bstp. Suc stp \<le> bstp \<and> crsp_l (layout_of aprog) 
       
  6548        (abc_steps_l ac aprog astp) (t_steps tc (tm_of aprog, 0) bstp) ires"
       
  6549 apply(rule_tac x = "Suc astp" in exI)
       
  6550 apply(case_tac "abc_steps_l ac aprog astp")
       
  6551 proof -
       
  6552   fix a b
       
  6553   assume 
       
  6554     "\<forall>stp. (\<lambda>(as, am). as < length aprog)  
       
  6555                  (abc_steps_l ac aprog stp)" 
       
  6556     "stp \<le> bstp"
       
  6557     "crsp_l (layout_of aprog) (abc_steps_l ac aprog astp) 
       
  6558       (t_steps tc (tm_of aprog, 0) bstp) ires" 
       
  6559     "abc_steps_l ac aprog astp = (a, b)"
       
  6560   thus 
       
  6561  "\<exists>bstp\<ge>Suc stp. crsp_l (layout_of aprog)
       
  6562        (abc_steps_l ac aprog (Suc astp)) 
       
  6563    (t_steps tc (tm_of aprog, 0) bstp) ires"
       
  6564     apply(insert crsp_inside[of "layout_of aprog" aprog 
       
  6565       "tm_of aprog" a b "(t_steps tc (tm_of aprog, 0) bstp)" "ires"], auto)
       
  6566     apply(erule_tac x = astp in allE, auto)
       
  6567     apply(rule_tac x = "bstp + stpa" in exI, simp)
       
  6568     apply(insert abc_steps_ind[of ac aprog "astp"], simp)
       
  6569     done
       
  6570 qed   
       
  6571 
       
  6572 lemma abc_unhalt_case: 
       
  6573   "\<lbrakk>crsp_l (layout_of aprog) ac tc ires;  
       
  6574     \<forall>stp. (\<lambda>(as, am). as < length aprog) (abc_steps_l ac aprog stp)\<rbrakk>
       
  6575  \<Longrightarrow> (\<exists> astp bstp. bstp \<ge> stp \<and> 
       
  6576          crsp_l (layout_of aprog) (abc_steps_l ac aprog astp) 
       
  6577                                 (t_steps tc (tm_of aprog, 0) bstp) ires)"
       
  6578 apply(induct stp)
       
  6579 apply(rule_tac abc_unhalt_case_zero, auto)
       
  6580 apply(rule_tac abc_unhalt_case_induct, auto)
       
  6581 done
       
  6582   
       
  6583 lemma abacus_turing_eq_unhalt_case_pre: 
       
  6584   "\<lbrakk>ly = layout_of aprog; 
       
  6585     tprog = tm_of aprog;
       
  6586     crsp_l ly ac tc ires;
       
  6587     \<forall> stp. ((\<lambda> (as, am). as < length aprog)
       
  6588                        (abc_steps_l ac aprog stp));
       
  6589     mop_ss = start_of ly (length aprog)\<rbrakk>
       
  6590   \<Longrightarrow> (\<not> (\<exists> stp. (\<lambda> (s, l, r). s = 0)
       
  6591               (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp)))"
       
  6592   apply(auto)
       
  6593 proof -
       
  6594   fix stp a b
       
  6595   assume h1: 
       
  6596     "crsp_l (layout_of aprog) ac tc ires" 
       
  6597     "\<forall>stp. (\<lambda>(as, am). as < length aprog) (abc_steps_l ac aprog stp)"
       
  6598     "t_steps tc (tm_of aprog @ tMp n (start_of (layout_of aprog) 
       
  6599     (length aprog) - Suc 0), 0) stp = (0, a, b)"
       
  6600   thus "False"
       
  6601   proof(insert abc_unhalt_case[of aprog ac tc ires stp], auto, 
       
  6602         case_tac "(abc_steps_l ac aprog astp)", 
       
  6603         case_tac "(t_steps tc (tm_of aprog, 0) bstp)", simp)
       
  6604     fix astp bstp aa ba aaa baa c
       
  6605     assume h2: 
       
  6606       "abc_steps_l ac aprog astp = (aa, ba)" "stp \<le> bstp"
       
  6607       "t_steps tc (tm_of aprog, 0) bstp = (aaa, baa, c)" 
       
  6608       "crsp_l (layout_of aprog) (aa, ba) (aaa, baa, c) ires"
       
  6609     hence h3: 
       
  6610       "t_steps tc (tm_of aprog @ tMp n 
       
  6611        (start_of (layout_of aprog) (length aprog) - Suc 0), 0) bstp 
       
  6612                     = (aaa, baa, c)"
       
  6613       apply(intro tm_append_steps, auto)
       
  6614       apply(simp add: crsp_l.simps, rule startof_not0)
       
  6615       done
       
  6616     from h2 have h4: "\<exists> diff. bstp = stp + diff"
       
  6617       apply(rule_tac x = "bstp - stp" in exI, simp)
       
  6618       done
       
  6619     from h4 and h3 and h2  and h1 show "?thesis"
       
  6620       apply(auto)
       
  6621       apply(simp add: state0_ind crsp_l.simps)
       
  6622       apply(subgoal_tac "start_of (layout_of aprog) aa > 0", simp)
       
  6623       apply(rule startof_not0)
       
  6624       done
       
  6625   qed
       
  6626 qed
       
  6627 
       
  6628 lemma abacus_turing_eq_unhalt_case:
       
  6629   assumes layout: 
       
  6630   -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *}
       
  6631   "ly = layout_of aprog"
       
  6632   and compiled: 
       
  6633   -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *}
       
  6634   "tprog = tm_of aprog"
       
  6635   and correspond: 
       
  6636   -- {* 
       
  6637   TM configuration @{text "tc"} and Abacus configuration @{text "ac"}
       
  6638   are in correspondence: 
       
  6639   *}
       
  6640   "crsp_l ly ac tc ires"
       
  6641   and abc_unhalt: 
       
  6642   -- {*
       
  6643   If, no matter how many steps the Abacus program @{text "aprog"} executes, it
       
  6644   may never reach a halt state. 
       
  6645   *}
       
  6646   "\<forall> stp. ((\<lambda> (as, am). as < length aprog)
       
  6647                        (abc_steps_l ac aprog stp))"
       
  6648   and mopup_start: "mop_ss = start_of ly (length aprog)"
       
  6649   shows
       
  6650   -- {*
       
  6651   The the TM composed of TM @{text "tprog"} and the moupup TM may never reach a halt state as well.
       
  6652   *}
       
  6653   "\<not> (\<exists> stp. (\<lambda> (s, l, r). s = 0)
       
  6654               (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp))"
       
  6655   using layout compiled correspond abc_unhalt mopup_start
       
  6656   apply(rule_tac abacus_turing_eq_unhalt_case_pre, auto)
       
  6657   done
       
  6658 
       
  6659 definition abc_list_crsp:: "nat list \<Rightarrow> nat list \<Rightarrow> bool"
       
  6660   where
       
  6661   "abc_list_crsp xs ys = (\<exists> n. xs = ys @ 0\<^bsup>n\<^esup> \<or> ys = xs @ 0\<^bsup>n\<^esup>)"
       
  6662 lemma [intro]: "abc_list_crsp (lm @ 0\<^bsup>m\<^esup>) lm"
       
  6663 apply(auto simp: abc_list_crsp_def)
       
  6664 done
       
  6665 
       
  6666 thm abc_lm_v.simps
       
  6667 lemma abc_list_crsp_lm_v: 
       
  6668   "abc_list_crsp lma lmb \<Longrightarrow> abc_lm_v lma n = abc_lm_v lmb n"
       
  6669 apply(auto simp: abc_list_crsp_def abc_lm_v.simps 
       
  6670                  nth_append exponent_def)
       
  6671 done
       
  6672 
       
  6673 lemma  rep_app_cons_iff: 
       
  6674   "k < n \<Longrightarrow> replicate n a[k:=b] = 
       
  6675           replicate k a @ b # replicate (n - k - 1) a"
       
  6676 apply(induct n arbitrary: k, simp)
       
  6677 apply(simp split:nat.splits)
       
  6678 done
       
  6679 
       
  6680 lemma abc_list_crsp_lm_s: 
       
  6681   "abc_list_crsp lma lmb \<Longrightarrow> 
       
  6682       abc_list_crsp (abc_lm_s lma m n) (abc_lm_s lmb m n)"
       
  6683 apply(auto simp: abc_list_crsp_def abc_lm_v.simps abc_lm_s.simps)
       
  6684 apply(simp_all add: list_update_append, auto simp: exponent_def)
       
  6685 proof -
       
  6686   fix na
       
  6687   assume h: "m < length lmb + na" " \<not> m < length lmb"
       
  6688   hence "m - length lmb < na" by simp
       
  6689   hence "replicate na 0[(m- length lmb):= n] = 
       
  6690            replicate (m - length lmb) 0 @ n # 
       
  6691               replicate (na - (m - length lmb) - 1) 0"
       
  6692     apply(erule_tac rep_app_cons_iff)
       
  6693     done
       
  6694   thus "\<exists>nb. replicate na 0[m - length lmb := n] =
       
  6695                  replicate (m - length lmb) 0 @ n # replicate nb 0 \<or>
       
  6696                  replicate (m - length lmb) 0 @ [n] =
       
  6697                  replicate na 0[m - length lmb := n] @ replicate nb 0"
       
  6698     apply(auto)
       
  6699     done
       
  6700 next
       
  6701   fix na
       
  6702   assume h: "\<not> m < length lmb + na"
       
  6703   show 
       
  6704     "\<exists>nb. replicate na 0 @ replicate (m - (length lmb + na)) 0 @ [n] =
       
  6705            replicate (m - length lmb) 0 @ n # replicate nb 0 \<or>
       
  6706           replicate (m - length lmb) 0 @ [n] =
       
  6707             replicate na 0 @
       
  6708             replicate (m - (length lmb + na)) 0 @ n # replicate nb 0"
       
  6709     apply(rule_tac x = 0 in exI, simp, auto)
       
  6710     using h
       
  6711     apply(simp add: replicate_add[THEN sym])
       
  6712     done
       
  6713 next
       
  6714   fix na
       
  6715   assume h: "\<not> m < length lma" "m < length lma + na"
       
  6716   hence "m - length lma < na" by simp
       
  6717   hence 
       
  6718     "replicate na 0[(m- length lma):= n] = replicate (m - length lma) 
       
  6719                   0 @ n # replicate (na - (m - length lma) - 1) 0"
       
  6720     apply(erule_tac rep_app_cons_iff)
       
  6721     done
       
  6722   thus "\<exists>nb. replicate (m - length lma) 0 @ [n] =
       
  6723                  replicate na 0[m - length lma := n] @ replicate nb 0 
       
  6724            \<or> replicate na 0[m - length lma := n] =
       
  6725                  replicate (m - length lma) 0 @ n # replicate nb 0"
       
  6726     apply(auto)
       
  6727     done
       
  6728 next
       
  6729   fix na
       
  6730   assume "\<not> m < length lma + na"
       
  6731   thus " \<exists>nb. replicate (m - length lma) 0 @ [n] =
       
  6732             replicate na 0 @
       
  6733             replicate (m - (length lma + na)) 0 @ n # replicate nb 0 
       
  6734         \<or>   replicate na 0 @ 
       
  6735                replicate (m - (length lma + na)) 0 @ [n] =
       
  6736             replicate (m - length lma) 0 @ n # replicate nb 0"
       
  6737     apply(rule_tac x = 0 in exI, simp, auto)
       
  6738     apply(simp add: replicate_add[THEN sym])
       
  6739     done
       
  6740 qed
       
  6741 
       
  6742 lemma abc_list_crsp_step: 
       
  6743   "\<lbrakk>abc_list_crsp lma lmb; abc_step_l (aa, lma) i = (a, lma'); 
       
  6744     abc_step_l (aa, lmb) i = (a', lmb')\<rbrakk>
       
  6745     \<Longrightarrow> a' = a \<and> abc_list_crsp lma' lmb'"
       
  6746 apply(case_tac i, auto simp: abc_step_l.simps 
       
  6747        abc_list_crsp_lm_s abc_list_crsp_lm_v Let_def 
       
  6748                        split: abc_inst.splits if_splits)
       
  6749 done
       
  6750 
       
  6751 thm abc_step_l.simps
       
  6752 
       
  6753 lemma abc_steps_red: 
       
  6754   "abc_steps_l ac aprog stp = (as, am) \<Longrightarrow>
       
  6755      abc_steps_l ac aprog (Suc stp) = 
       
  6756            abc_step_l (as, am) (abc_fetch as aprog)"
       
  6757 using abc_steps_ind[of ac aprog stp]
       
  6758 apply(simp)
       
  6759 done
       
  6760 
       
  6761 lemma abc_list_crsp_steps: 
       
  6762   "\<lbrakk>abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog stp = (a, lm'); aprog \<noteq> []\<rbrakk> 
       
  6763       \<Longrightarrow> \<exists> lma. abc_steps_l (0, lm) aprog stp = (a, lma) \<and> 
       
  6764                                           abc_list_crsp lm' lma"
       
  6765 apply(induct stp arbitrary: a lm', simp add: abc_steps_l.simps, auto)
       
  6766 apply(case_tac "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog stp", 
       
  6767       simp add: abc_steps_ind)
       
  6768 proof -
       
  6769   fix stp a lm' aa b
       
  6770   assume ind:
       
  6771     "\<And>a lm'. aa = a \<and> b = lm' \<Longrightarrow> 
       
  6772      \<exists>lma. abc_steps_l (0, lm) aprog stp = (a, lma) \<and>
       
  6773                                           abc_list_crsp lm' lma"
       
  6774     and h: "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog (Suc stp) = (a, lm')" 
       
  6775            "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog stp = (aa, b)" 
       
  6776            "aprog \<noteq> []"
       
  6777   hence g1: "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog (Suc stp)
       
  6778           = abc_step_l (aa, b) (abc_fetch aa aprog)"
       
  6779     apply(rule_tac abc_steps_red, simp)
       
  6780     done
       
  6781   have "\<exists>lma. abc_steps_l (0, lm) aprog stp = (aa, lma) \<and> 
       
  6782               abc_list_crsp b lma"
       
  6783     apply(rule_tac ind, simp)
       
  6784     done
       
  6785   from this obtain lma where g2: 
       
  6786     "abc_steps_l (0, lm) aprog stp = (aa, lma) \<and> 
       
  6787      abc_list_crsp b lma"   ..
       
  6788   hence g3: "abc_steps_l (0, lm) aprog (Suc stp)
       
  6789           = abc_step_l (aa, lma) (abc_fetch aa aprog)"
       
  6790     apply(rule_tac abc_steps_red, simp)
       
  6791     done
       
  6792   show "\<exists>lma. abc_steps_l (0, lm) aprog (Suc stp) = (a, lma) \<and>  
       
  6793               abc_list_crsp lm' lma"
       
  6794     using g1 g2 g3 h
       
  6795     apply(auto)
       
  6796     apply(case_tac "abc_step_l (aa, b) (abc_fetch aa aprog)",
       
  6797           case_tac "abc_step_l (aa, lma) (abc_fetch aa aprog)", simp)
       
  6798     apply(rule_tac abc_list_crsp_step, auto)
       
  6799     done
       
  6800 qed
       
  6801 
       
  6802 text {* Begin: equvilence between steps and t_steps*}
       
  6803 lemma [simp]: "(case ca of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc) =
       
  6804                 (case ca of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)"
       
  6805 by(case_tac ca, simp_all, case_tac a, simp, simp)
       
  6806 
       
  6807 text {* needed to interpret*}
       
  6808 
       
  6809 lemma steps_eq: "length t mod 2 = 0 \<Longrightarrow> 
       
  6810                     t_steps c (t, 0) stp = steps c t stp"
       
  6811 apply(induct stp)
       
  6812 apply(simp add: steps.simps t_steps.simps)
       
  6813 apply(simp add:tstep_red t_steps_ind)
       
  6814 apply(case_tac "steps c t stp", simp)
       
  6815 apply(auto simp: t_step.simps tstep.simps)
       
  6816 done
       
  6817 
       
  6818 text{* end: equvilence between steps and t_steps*}
       
  6819 
       
  6820 lemma crsp_l_start: "crsp_l ly (0, lm) (Suc 0, Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>) ires"
       
  6821 apply(simp add: crsp_l.simps, auto simp: start_of.simps)
       
  6822 done
       
  6823 
       
  6824 lemma t_ncorrect_app: "\<lbrakk>t_ncorrect t1; t_ncorrect t2\<rbrakk> \<Longrightarrow> 
       
  6825                                           t_ncorrect (t1 @ t2)"
       
  6826 apply(simp add: t_ncorrect.simps, auto)
       
  6827 done
       
  6828 
       
  6829 lemma [simp]: 
       
  6830   "(length (tm_of aprog) + 
       
  6831     length (tMp n (start_of ly (length aprog) - Suc 0))) mod 2 = 0"
       
  6832 apply(subgoal_tac 
       
  6833  "t_ncorrect (tm_of aprog @ tMp n 
       
  6834              (start_of ly (length aprog) - Suc 0))")
       
  6835 apply(simp add: t_ncorrect.simps)
       
  6836 apply(rule_tac t_ncorrect_app, 
       
  6837       auto simp: tMp.simps t_ncorrect.simps tshift.simps mp_up_def)
       
  6838 apply(subgoal_tac
       
  6839        "t_ncorrect (tm_of aprog)", simp add: t_ncorrect.simps)
       
  6840 apply(auto)
       
  6841 done
       
  6842 
       
  6843 lemma [simp]: "takeWhile (\<lambda>a. a = Oc) 
       
  6844               (replicate rs Oc @ replicate rn Bk) = replicate rs Oc"
       
  6845 apply(induct rs, auto)
       
  6846 apply(induct rn, auto)
       
  6847 done
       
  6848 
       
  6849 lemma abacus_turing_eq_halt': 
       
  6850   "\<lbrakk>ly = layout_of aprog; 
       
  6851     tprog = tm_of aprog; 
       
  6852     n < length am;
       
  6853     abc_steps_l (0, lm) aprog stp = (as, am); 
       
  6854     mop_ss = start_of ly (length aprog);
       
  6855     as \<ge> length aprog\<rbrakk>
       
  6856     \<Longrightarrow> \<exists> stp m l. steps (Suc 0, Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>) 
       
  6857                 (tprog @ (tMp n (mop_ss - 1))) stp
       
  6858                   = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>l\<^esup>)"
       
  6859 apply(drule_tac tc = "(Suc 0, Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>)" in 
       
  6860                abacus_turing_eq_halt_case, auto intro: crsp_l_start)
       
  6861 apply(subgoal_tac 
       
  6862          "length (tm_of aprog @ tMp n 
       
  6863                   (start_of ly (length aprog) - Suc 0)) mod 2 = 0")
       
  6864 apply(simp add: steps_eq)
       
  6865 apply(rule_tac x = stpa in exI, 
       
  6866        simp add:  exponent_def, auto)
       
  6867 done
       
  6868 
       
  6869 
       
  6870 thm tinres_steps
       
  6871 lemma list_length: "xs = ys \<Longrightarrow> length xs = length ys"
       
  6872 by simp
       
  6873 lemma [elim]: "tinres (Bk\<^bsup>m\<^esup>) b \<Longrightarrow> \<exists>m. b = Bk\<^bsup>m\<^esup>"
       
  6874 apply(auto simp: tinres_def)
       
  6875 apply(rule_tac x = "m-n" in exI, 
       
  6876              auto simp: exponent_def replicate_add[THEN sym]) 
       
  6877 apply(case_tac "m < n", auto)
       
  6878 apply(drule_tac list_length, auto)
       
  6879 apply(subgoal_tac "\<exists> d. m = d + n", auto simp: replicate_add)
       
  6880 apply(rule_tac x = "m - n" in exI, simp)
       
  6881 done
       
  6882 lemma [intro]: "tinres [Bk] (Bk\<^bsup>k\<^esup>) "
       
  6883 apply(auto simp: tinres_def exponent_def)
       
  6884 apply(case_tac k, auto)
       
  6885 apply(rule_tac x = "Suc 0" in exI, simp)
       
  6886 done
       
  6887 
       
  6888 lemma abacus_turing_eq_halt_pre: 
       
  6889  "\<lbrakk>ly = layout_of aprog; 
       
  6890    tprog = tm_of aprog; 
       
  6891    n < length am;     
       
  6892    abc_steps_l (0, lm) aprog stp = (as, am);  
       
  6893    mop_ss = start_of ly (length aprog);
       
  6894    as \<ge> length aprog\<rbrakk>
       
  6895   \<Longrightarrow> \<exists> stp m l. steps  (Suc 0, Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>)
       
  6896                (tprog @ (tMp n (mop_ss - 1))) stp
       
  6897                  = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>l\<^esup>)"
       
  6898 using abacus_turing_eq_halt'
       
  6899 apply(simp)
       
  6900 done
       
  6901 
       
  6902 
       
  6903 text {*
       
  6904   Main theorem for the case when the original Abacus program does halt.
       
  6905 *}
       
  6906 lemma abacus_turing_eq_halt: 
       
  6907   assumes layout:
       
  6908   "ly = layout_of aprog"
       
  6909   -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *}
       
  6910   and compiled: "tprog = tm_of aprog"
       
  6911   -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *}
       
  6912   and halt_state: 
       
  6913    -- {* @{text "as"} is a program label outside the range of @{text "aprog"}. So 
       
  6914   if Abacus is in such a state, it is in halt state: *}
       
  6915   "as \<ge> length aprog"
       
  6916   and abc_exec: 
       
  6917   -- {* Supposing after @{text "stp"} step of execution, Abacus program @{text "aprog"}
       
  6918   reaches such a halt state: *}
       
  6919   "abc_steps_l (0, lm) aprog stp = (as, am)"
       
  6920   and rs_locate: 
       
  6921    -- {* @{text "n"} is a memory address in the range of Abacus memory @{text "am"}: *}
       
  6922   "n < length am"  
       
  6923   and mopup_start: 
       
  6924    -- {* The startling label for mopup mahines, according to the layout and Abacus program 
       
  6925    should be @{text "mop_ss"}: *}
       
  6926   "mop_ss = start_of ly (length aprog)"
       
  6927   shows 
       
  6928   -- {* 
       
  6929   After @{text "stp"} steps of execution of the TM composed of @{text "tprog"} and the mopup 
       
  6930   TM @{text "(tMp n (mop_ss - 1))"} will halt and gives rise to a configuration which 
       
  6931   only hold the content of memory cell @{text "n"}:
       
  6932   *}
       
  6933   "\<exists> stp m l. steps (Suc 0, Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>) (tprog @ (tMp n (mop_ss - 1))) stp
       
  6934                       = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>l\<^esup>)"
       
  6935   using layout compiled halt_state abc_exec rs_locate mopup_start
       
  6936   by(rule_tac abacus_turing_eq_halt_pre, auto)
       
  6937 
       
  6938 lemma abacus_turing_eq_uhalt': 
       
  6939  "\<lbrakk>ly = layout_of aprog; 
       
  6940    tprog = tm_of aprog; 
       
  6941    \<forall> stp. ((\<lambda> (as, am). as < length aprog) 
       
  6942                    (abc_steps_l (0, lm) aprog stp));
       
  6943    mop_ss = start_of ly (length aprog)\<rbrakk>
       
  6944   \<Longrightarrow> (\<not> (\<exists> stp. isS0 (steps (Suc 0, [Bk, Bk], <lm>) 
       
  6945                       (tprog @ (tMp n (mop_ss - 1))) stp)))"
       
  6946 apply(drule_tac tc = "(Suc 0, [Bk, Bk], <lm>)" and n = n and ires = "[]" in 
       
  6947          abacus_turing_eq_unhalt_case, auto intro: crsp_l_start)
       
  6948 apply(simp add: crsp_l.simps start_of.simps)
       
  6949 apply(erule_tac x = stp in allE, erule_tac x = stp in allE)
       
  6950 apply(subgoal_tac 
       
  6951    "length (tm_of aprog @ tMp n 
       
  6952          (start_of ly (length aprog) - Suc 0)) mod 2 = 0")
       
  6953 apply(simp add: steps_eq, auto simp: isS0_def)
       
  6954 done
       
  6955 (*
       
  6956 lemma abacus_turing_eq_uhalt_pre: 
       
  6957   "\<lbrakk>ly = layout_of aprog; 
       
  6958     tprog = tm_of aprog;
       
  6959     \<forall> stp. ((\<lambda> (as, am). as < length aprog) 
       
  6960                       (abc_steps_l (0, lm) aprog stp));
       
  6961     mop_ss = start_of ly (length aprog)\<rbrakk>
       
  6962   \<Longrightarrow> (\<not> (\<exists> stp. isS0 (steps (Suc 0, [Bk, Bk], <lm>) 
       
  6963                     (tprog @ (tMp n (mop_ss - 1))) stp)))"
       
  6964 apply(drule_tac k = 0 and n = n  in abacus_turing_eq_uhalt', auto)
       
  6965 apply(erule_tac x = stp in allE, erule_tac x = stp in allE)
       
  6966 apply(subgoal_tac "tinres ([Bk]) (Bk\<^bsup>k\<^esup>)")
       
  6967 apply(case_tac "steps (Suc 0, Bk\<^bsup>k\<^esup>, <lm>)
       
  6968       (tm_of aprog @ tMp n (start_of ly (length aprog) - Suc 0)) stp")
       
  6969 apply(case_tac 
       
  6970   "steps (Suc 0, [Bk], <lm>)
       
  6971     (tm_of aprog @ tMp n (start_of ly (length aprog) - Suc 0)) stp")
       
  6972 apply(drule_tac tinres_steps, auto simp: isS0_def)
       
  6973 done
       
  6974 *)
       
  6975 text {*
       
  6976   Main theorem for the case when the original Abacus program does not halt.
       
  6977   *}
       
  6978 lemma abacus_turing_eq_uhalt:
       
  6979   assumes layout: 
       
  6980   -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *}
       
  6981   "ly = layout_of aprog"
       
  6982   and compiled:
       
  6983    -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *}
       
  6984   "tprog = tm_of aprog"
       
  6985   and abc_unhalt:
       
  6986   -- {*
       
  6987   If, no matter how many steps the Abacus program @{text "aprog"} executes, it
       
  6988   may never reach a halt state. 
       
  6989   *}
       
  6990   "\<forall> stp. ((\<lambda> (as, am). as < length aprog) 
       
  6991                       (abc_steps_l (0, lm) aprog stp))"
       
  6992   and mop_start: "mop_ss = start_of ly (length aprog)"
       
  6993   shows 
       
  6994    -- {*
       
  6995   The the TM composed of TM @{text "tprog"} and the moupup TM may never reach a halt state as well.
       
  6996   *}
       
  6997   "\<not> (\<exists> stp. isS0 (steps (Suc 0, [Bk, Bk], <lm>) 
       
  6998                     (tprog @ (tMp n (mop_ss - 1))) stp))"
       
  6999   using abacus_turing_eq_uhalt'
       
  7000         layout compiled abc_unhalt mop_start
       
  7001   by(auto)
       
  7002 
       
  7003 
       
  7004 end
       
  7005