|         |      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 (*  FIXME: doubly defined | 
|         |    181 text {* | 
|         |    182   @{text "tshift tm off"} shifts @{text "tm"} by offset @{text "off"}, leaving  | 
|         |    183   instructions concerning state @{text "0"} unchanged, because state @{text "0"}  | 
|         |    184   is the end state, which needs not be changed with shift operation. | 
|         |    185   *} | 
|         |    186  | 
|         |    187 fun tshift :: "tprog \<Rightarrow> nat \<Rightarrow> tprog" | 
|         |    188   where | 
|         |    189   "tshift tp off = (map (\<lambda> (action, state).  | 
|         |    190        (action, (if state = 0 then 0 | 
|         |    191                  else state + off))) tp)" | 
|         |    192 *) | 
|         |    193  | 
|         |    194 text {* | 
|         |    195   @{text "tinc ss n"} returns the TM which simulates the execution of  | 
|         |    196   Abacus instruction @{text "Inc n"}, assuming that TM is located at | 
|         |    197   location @{text "ss"} in the final TM complied from the whole | 
|         |    198   Abacus program. | 
|         |    199 *} | 
|         |    200  | 
|         |    201 fun tinc :: "nat \<Rightarrow> nat \<Rightarrow> tprog" | 
|         |    202   where | 
|         |    203   "tinc ss n = tshift (findnth n @ tshift tinc_b (2 * n)) (ss - 1)" | 
|         |    204  | 
|         |    205 text {* | 
|         |    206   @{text "tinc_b"} returns the TM which decrements the representation  | 
|         |    207   of the memory cell under rw-head by one and move the representation  | 
|         |    208   of cells afterwards to the left accordingly. | 
|         |    209   *} | 
|         |    210  | 
|         |    211 definition tdec_b :: "tprog" | 
|         |    212   where | 
|         |    213   "tdec_b \<equiv>  [(W1, 1), (R, 2), (L, 14), (R, 3), (L, 4), (R, 3), | 
|         |    214               (R, 5), (W0, 4), (R, 6), (W0, 5), (L, 7), (L, 8), | 
|         |    215               (L, 11), (W0, 7), (W1, 8), (R, 9), (L, 10), (R, 9), | 
|         |    216               (R, 5), (W0, 10), (L, 12), (L, 11), (R, 13), (L, 11), | 
|         |    217               (R, 17), (W0, 13), (L, 15), (L, 14), (R, 16), (L, 14), | 
|         |    218               (R, 0), (W0, 16)]" | 
|         |    219  | 
|         |    220 text {* | 
|         |    221   @{text "sete tp e"} attaches the termination edges (edges leading to state @{text "0"})  | 
|         |    222   of TM @{text "tp"} to the intruction labelled by @{text "e"}. | 
|         |    223   *} | 
|         |    224  | 
|         |    225 fun sete :: "tprog \<Rightarrow> nat \<Rightarrow> tprog" | 
|         |    226   where | 
|         |    227   "sete tp e = map (\<lambda> (action, state). (action, if state = 0 then e else state)) tp" | 
|         |    228  | 
|         |    229 text {* | 
|         |    230   @{text "tdec ss n label"} returns the TM which simulates the execution of  | 
|         |    231   Abacus instruction @{text "Dec n label"}, assuming that TM is located at | 
|         |    232   location @{text "ss"} in the final TM complied from the whole | 
|         |    233   Abacus program. | 
|         |    234 *} | 
|         |    235  | 
|         |    236 fun tdec :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> tprog" | 
|         |    237   where | 
|         |    238   "tdec ss n e = sete (tshift (findnth n @ tshift tdec_b (2 * n))  | 
|         |    239                  (ss - 1)) e" | 
|         |    240   | 
|         |    241 text {* | 
|         |    242   @{text "tgoto f(label)"} returns the TM simulating the execution of Abacus instruction | 
|         |    243   @{text "Goto label"}, where @{text "f(label)"} is the corresponding location of | 
|         |    244   @{text "label"} in the final TM compiled from the overall Abacus program. | 
|         |    245 *} | 
|         |    246  | 
|         |    247 fun tgoto :: "nat \<Rightarrow> tprog" | 
|         |    248   where | 
|         |    249   "tgoto n = [(Nop, n), (Nop, n)]" | 
|         |    250  | 
|         |    251 text {* | 
|         |    252   The layout of the final TM compiled from an Abacus program is represented | 
|         |    253   as a list of natural numbers, where the list element at index @{text "n"} represents the  | 
|         |    254   starting state of the TM simulating the execution of @{text "n"}-th instruction | 
|         |    255   in the Abacus program. | 
|         |    256 *} | 
|         |    257  | 
|         |    258 type_synonym layout = "nat list" | 
|         |    259  | 
|         |    260 text {* | 
|         |    261   @{text "length_of i"} is the length of the  | 
|         |    262   TM simulating the Abacus instruction @{text "i"}. | 
|         |    263 *} | 
|         |    264 fun length_of :: "abc_inst \<Rightarrow> nat" | 
|         |    265   where | 
|         |    266   "length_of i = (case i of  | 
|         |    267                     Inc n   \<Rightarrow> 2 * n + 9 | | 
|         |    268                     Dec n e \<Rightarrow> 2 * n + 16 | | 
|         |    269                     Goto n  \<Rightarrow> 1)" | 
|         |    270  | 
|         |    271 text {* | 
|         |    272   @{text "layout_of ap"} returns the layout of Abacus program @{text "ap"}. | 
|         |    273 *} | 
|         |    274 fun layout_of :: "abc_prog \<Rightarrow> layout" | 
|         |    275   where "layout_of ap = map length_of ap" | 
|         |    276  | 
|         |    277  | 
|         |    278 text {* | 
|         |    279   @{text "start_of layout n"} looks out the starting state of @{text "n"}-th | 
|         |    280   TM in the finall TM. | 
|         |    281 *} | 
|         |    282  | 
|         |    283 fun start_of :: "nat list \<Rightarrow> nat \<Rightarrow> nat" | 
|         |    284   where | 
|         |    285   "start_of ly 0 = Suc 0" | | 
|         |    286   "start_of ly (Suc as) =  | 
|         |    287         (if as < length ly then start_of ly as + (ly ! as) | 
|         |    288          else start_of ly as)" | 
|         |    289  | 
|         |    290 text {* | 
|         |    291   @{text "ci lo ss i"} complies Abacus instruction @{text "i"} | 
|         |    292   assuming the TM of @{text "i"} starts from state @{text "ss"}  | 
|         |    293   within the overal layout @{text "lo"}. | 
|         |    294 *} | 
|         |    295  | 
|         |    296 fun ci :: "layout \<Rightarrow> nat \<Rightarrow> abc_inst \<Rightarrow> tprog" | 
|         |    297   where | 
|         |    298   "ci ly ss i = (case i of  | 
|         |    299                     Inc n   \<Rightarrow> tinc ss n | | 
|         |    300                     Dec n e \<Rightarrow> tdec ss n (start_of ly e) | | 
|         |    301                     Goto n  \<Rightarrow> tgoto (start_of ly n))" | 
|         |    302  | 
|         |    303 text {* | 
|         |    304   @{text "tpairs_of ap"} transfroms Abacus program @{text "ap"} pairing | 
|         |    305   every instruction with its starting state. | 
|         |    306 *} | 
|         |    307 fun tpairs_of :: "abc_prog \<Rightarrow> (nat \<times> abc_inst) list" | 
|         |    308   where "tpairs_of ap = (zip (map (start_of (layout_of ap))  | 
|         |    309                          [0..<(length ap)]) ap)" | 
|         |    310  | 
|         |    311  | 
|         |    312 text {* | 
|         |    313   @{text "tms_of ap"} returns the list of TMs, where every one of them simulates | 
|         |    314   the corresponding Abacus intruction in @{text "ap"}. | 
|         |    315 *} | 
|         |    316  | 
|         |    317 fun tms_of :: "abc_prog \<Rightarrow> tprog list" | 
|         |    318   where "tms_of ap = map (\<lambda> (n, tm). ci (layout_of ap) n tm)  | 
|         |    319                          (tpairs_of ap)" | 
|         |    320  | 
|         |    321 text {* | 
|         |    322   @{text "tm_of ap"} returns the final TM machine compiled from Abacus program @{text "ap"}. | 
|         |    323 *} | 
|         |    324 fun tm_of :: "abc_prog \<Rightarrow> tprog" | 
|         |    325   where "tm_of ap = concat (tms_of ap)" | 
|         |    326  | 
|         |    327 text {* | 
|         |    328   The following two functions specify the well-formedness of complied TM. | 
|         |    329 *} | 
|         |    330 fun t_ncorrect :: "tprog \<Rightarrow> bool" | 
|         |    331   where | 
|         |    332   "t_ncorrect tp = (length tp mod 2 = 0)" | 
|         |    333  | 
|         |    334 fun abc2t_correct :: "abc_prog \<Rightarrow> bool" | 
|         |    335   where  | 
|         |    336   "abc2t_correct ap = list_all (\<lambda> (n, tm).  | 
|         |    337              t_ncorrect (ci (layout_of ap) n tm)) (tpairs_of ap)" | 
|         |    338  | 
|         |    339 lemma findnth_length: "length (findnth n) div 2 = 2 * n" | 
|         |    340 apply(induct n, simp, simp) | 
|         |    341 done | 
|         |    342  | 
|         |    343 lemma ci_length : "length (ci ns n ai) div 2 = length_of ai" | 
|         |    344 apply(auto simp: ci.simps tinc_b_def tdec_b_def findnth_length | 
|         |    345                  split: abc_inst.splits) | 
|         |    346 done | 
|         |    347  | 
|         |    348 subsection {* | 
|         |    349   Representation of Abacus memory by TM tape | 
|         |    350 *} | 
|         |    351  | 
|         |    352 consts tape_of :: "'a \<Rightarrow> block list" ("<_>" 100) | 
|         |    353  | 
|         |    354 text {* | 
|         |    355   @{text "tape_of_nat_list am"} returns the TM tape representing | 
|         |    356   Abacus memory @{text "am"}. | 
|         |    357   *} | 
|         |    358  | 
|         |    359 fun tape_of_nat_list :: "nat list \<Rightarrow> block list" | 
|         |    360   where  | 
|         |    361   "tape_of_nat_list [] = []" | | 
|         |    362   "tape_of_nat_list [n] = Oc\<^bsup>n+1\<^esup>" | | 
|         |    363   "tape_of_nat_list (n#ns) = (Oc\<^bsup>n+1\<^esup>) @ [Bk] @ (tape_of_nat_list ns)" | 
|         |    364  | 
|         |    365 defs (overloaded) | 
|         |    366   tape_of_nl_abv: "<am> \<equiv> tape_of_nat_list am" | 
|         |    367   tape_of_nat_abv : "<(n::nat)> \<equiv> Oc\<^bsup>n+1\<^esup>" | 
|         |    368  | 
|         |    369 text {* | 
|         |    370   @{text "crsp_l acf tcf"} meams the abacus configuration @{text "acf"} | 
|         |    371   is corretly represented by the TM configuration @{text "tcf"}. | 
|         |    372 *} | 
|         |    373  | 
|         |    374 fun crsp_l :: "layout \<Rightarrow> abc_conf_l \<Rightarrow> t_conf \<Rightarrow> block list \<Rightarrow> bool" | 
|         |    375   where  | 
|         |    376   "crsp_l ly (as, lm) (ts, (l, r)) inres =  | 
|         |    377            (ts = start_of ly as \<and> (\<exists> rn. r = <lm> @ Bk\<^bsup>rn\<^esup>) | 
|         |    378             \<and> l = Bk # Bk # inres)" | 
|         |    379  | 
|         |    380 declare crsp_l.simps[simp del] | 
|         |    381  | 
|         |    382 subsection {* | 
|         |    383   A more general definition of TM execution.  | 
|         |    384 *} | 
|         |    385  | 
|         |    386 (* | 
|         |    387 fun nnth_of :: "(taction \<times> nat) list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> (taction \<times> nat)" | 
|         |    388   where | 
|         |    389   "nnth_of p s b = (if 2*s < length p  | 
|         |    390                     then (p ! (2*s + b)) else (Nop, 0))" | 
|         |    391  | 
|         |    392 thm nth_of.simps | 
|         |    393  | 
|         |    394 fun nfetch :: "tprog \<Rightarrow> nat \<Rightarrow> block \<Rightarrow> taction \<times> nat" | 
|         |    395   where | 
|         |    396   "nfetch p 0 b = (Nop, 0)" | | 
|         |    397   "nfetch p (Suc s) b =  | 
|         |    398              (case b of  | 
|         |    399                 Bk \<Rightarrow> nnth_of p s 0 | | 
|         |    400                 Oc \<Rightarrow> nnth_of p s 1)" | 
|         |    401 *) | 
|         |    402  | 
|         |    403 text {* | 
|         |    404   @{text "t_step tcf (tp, ss)"} returns the result of one step exection of TM @{text "tp"} | 
|         |    405   assuming @{text "tp"} starts from instial state @{text "ss"}. | 
|         |    406 *} | 
|         |    407  | 
|         |    408 fun t_step :: "t_conf \<Rightarrow> (tprog \<times> nat) \<Rightarrow> t_conf" | 
|         |    409   where  | 
|         |    410   "t_step c (p, off) =  | 
|         |    411            (let (state, leftn, rightn) = c in | 
|         |    412             let (action, next_state) = fetch p (state-off) | 
|         |    413                              (case rightn of  | 
|         |    414                                 [] \<Rightarrow> Bk |  | 
|         |    415                                 Bk # xs \<Rightarrow> Bk | | 
|         |    416                                 Oc # xs \<Rightarrow> Oc | 
|         |    417                              )  | 
|         |    418              in  | 
|         |    419             (next_state, new_tape action (leftn, rightn)))" | 
|         |    420  | 
|         |    421  | 
|         |    422 text {* | 
|         |    423   @{text "t_steps tcf (tp, ss) n"} returns the result of @{text "n"}-step exection  | 
|         |    424   of TM @{text "tp"} assuming @{text "tp"} starts from instial state @{text "ss"}. | 
|         |    425 *} | 
|         |    426  | 
|         |    427 fun t_steps :: "t_conf \<Rightarrow> (tprog \<times> nat) \<Rightarrow> nat \<Rightarrow> t_conf" | 
|         |    428   where | 
|         |    429   "t_steps c (p, off) 0 = c" | | 
|         |    430   "t_steps c (p, off) (Suc n) = t_steps  | 
|         |    431                      (t_step c (p, off)) (p, off) n"  | 
|         |    432  | 
|         |    433 lemma stepn: "t_steps c (p, off) (Suc n) =  | 
|         |    434               t_step (t_steps c (p, off) n) (p, off)" | 
|         |    435 apply(induct n arbitrary: c, simp add: t_steps.simps) | 
|         |    436 apply(simp add: t_steps.simps) | 
|         |    437 done | 
|         |    438  | 
|         |    439 text {* | 
|         |    440   The type of invarints expressing correspondence between  | 
|         |    441   Abacus configuration and TM configuration. | 
|         |    442 *} | 
|         |    443  | 
|         |    444 type_synonym inc_inv_t = "abc_conf_l \<Rightarrow> t_conf \<Rightarrow> block list \<Rightarrow> bool" | 
|         |    445  | 
|         |    446 declare tms_of.simps[simp del] tm_of.simps[simp del] | 
|         |    447         layout_of.simps[simp del] abc_fetch.simps [simp del]   | 
|         |    448         t_step.simps[simp del] t_steps.simps[simp del]  | 
|         |    449         tpairs_of.simps[simp del] start_of.simps[simp del] | 
|         |    450         fetch.simps [simp del] t_ncorrect.simps[simp del] | 
|         |    451         new_tape.simps [simp del] ci.simps [simp del] length_of.simps[simp del]  | 
|         |    452         layout_of.simps[simp del] crsp_l.simps[simp del] | 
|         |    453         abc2t_correct.simps[simp del] | 
|         |    454  | 
|         |    455 lemma tct_div2: "t_ncorrect tp \<Longrightarrow> (length tp) mod 2 = 0" | 
|         |    456 apply(simp add: t_ncorrect.simps) | 
|         |    457 done | 
|         |    458  | 
|         |    459 lemma t_shift_fetch:  | 
|         |    460     "\<lbrakk>t_ncorrect tp1; t_ncorrect tp;  | 
|         |    461       length tp1 div 2 < a \<and> a \<le> length tp1 div 2 + length tp div 2\<rbrakk> | 
|         |    462     \<Longrightarrow> fetch tp (a - length tp1 div 2) b =  | 
|         |    463          fetch (tp1 @ tp @ tp2) a b" | 
|         |    464 apply(subgoal_tac "\<exists> x. a = length tp1 div 2 + x", erule exE, simp) | 
|         |    465 apply(case_tac x, simp) | 
|         |    466 apply(subgoal_tac "length tp1 div 2 + Suc nat =  | 
|         |    467              Suc (length tp1 div 2 + nat)") | 
|         |    468 apply(simp only: fetch.simps nth_of.simps, auto) | 
|         |    469 apply(case_tac b, simp) | 
|         |    470 apply(subgoal_tac "2 * (length tp1 div 2) = length tp1", simp) | 
|         |    471 apply(subgoal_tac "2 * nat < length tp", simp add: nth_append, simp) | 
|         |    472 apply(simp add: t_ncorrect.simps, auto) | 
|         |    473 apply(subgoal_tac "2 * (length tp1 div 2) = length tp1", simp) | 
|         |    474 apply(subgoal_tac "2 * nat < length tp", simp add: nth_append, auto) | 
|         |    475 apply(simp add: t_ncorrect.simps, auto) | 
|         |    476 apply(rule_tac x = "a - length tp1 div 2" in exI, simp) | 
|         |    477 done | 
|         |    478  | 
|         |    479 lemma t_shift_in_step: | 
|         |    480       "\<lbrakk>t_step (a, aa, ba) (tp, length tp1 div 2) = (s, l, r); | 
|         |    481         t_ncorrect tp1; t_ncorrect tp; | 
|         |    482         length tp1 div 2 < a \<and> a \<le> length tp1 div 2 + length tp div 2\<rbrakk> | 
|         |    483        \<Longrightarrow> t_step (a, aa, ba) (tp1 @ tp @ tp2, 0) = (s, l, r)" | 
|         |    484 apply(simp add: t_step.simps) | 
|         |    485 apply(subgoal_tac "fetch tp (a - length tp1 div 2) (case ba of [] \<Rightarrow>  | 
|         |    486                    Bk | x # xs \<Rightarrow> x) | 
|         |    487              = fetch (tp1 @ tp @ tp2) a (case ba of [] \<Rightarrow> Bk | x # xs | 
|         |    488                    \<Rightarrow> x)") | 
|         |    489 apply(case_tac "fetch tp (a - length tp1 div 2) (case ba of [] \<Rightarrow> Bk | 
|         |    490                 | x # xs \<Rightarrow> x)") | 
|         |    491 apply(auto intro: t_shift_fetch) | 
|         |    492 apply(case_tac ba, simp, simp) | 
|         |    493 apply(case_tac aaa, simp, simp) | 
|         |    494 done | 
|         |    495  | 
|         |    496 declare add_Suc_right[simp del] | 
|         |    497 lemma t_step_add: "t_steps c (p, off) (m + n) =  | 
|         |    498           t_steps (t_steps c (p, off) m) (p, off) n" | 
|         |    499 apply(induct m arbitrary: n,  simp add: t_steps.simps, simp) | 
|         |    500 apply(subgoal_tac "t_steps c (p, off) (Suc (m + n)) =  | 
|         |    501                          t_steps c (p, off) (m + Suc n)", simp) | 
|         |    502 apply(subgoal_tac "t_steps (t_steps c (p, off) m) (p, off) (Suc n) = | 
|         |    503                 t_steps (t_step (t_steps c (p, off) m) (p, off))  | 
|         |    504                          (p, off) n") | 
|         |    505 apply(simp, simp add: stepn) | 
|         |    506 apply(simp only: t_steps.simps) | 
|         |    507 apply(simp only: add_Suc_right) | 
|         |    508 done | 
|         |    509 declare add_Suc_right[simp] | 
|         |    510  | 
|         |    511 lemma s_out_fetch: "\<lbrakk>t_ncorrect tp;  | 
|         |    512         \<not> (length tp1 div 2 < a \<and> a \<le> length tp1 div 2 +  | 
|         |    513          length tp div 2)\<rbrakk> | 
|         |    514       \<Longrightarrow> fetch tp (a - length tp1 div 2) b = (Nop, 0)" | 
|         |    515 apply(auto) | 
|         |    516 apply(simp add: fetch.simps) | 
|         |    517 apply(subgoal_tac "\<exists> x. a - length tp1 div 2 = length tp div 2 + x") | 
|         |    518 apply(erule exE, simp) | 
|         |    519 apply(case_tac x, simp) | 
|         |    520 apply(auto simp add: fetch.simps) | 
|         |    521 apply(subgoal_tac "2 * (length tp div 2) =  length tp") | 
|         |    522 apply(auto simp: t_ncorrect.simps split: block.splits) | 
|         |    523 apply(rule_tac x = "a - length tp1 div 2 - length tp div 2" in exI | 
|         |    524      , simp) | 
|         |    525 done  | 
|         |    526  | 
|         |    527 lemma conf_keep_step:  | 
|         |    528       "\<lbrakk>t_ncorrect tp;  | 
|         |    529         \<not> (length tp1 div 2 < a \<and> a \<le> length tp1 div 2 +  | 
|         |    530        length tp div 2)\<rbrakk> | 
|         |    531       \<Longrightarrow> t_step (a, aa, ba) (tp, length tp1 div 2) = (0, aa, ba)" | 
|         |    532 apply(simp add: t_step.simps) | 
|         |    533 apply(subgoal_tac "fetch tp (a - length tp1 div 2) (case ba of [] \<Rightarrow>  | 
|         |    534   Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc) = (Nop, 0)") | 
|         |    535 apply(simp add: new_tape.simps) | 
|         |    536 apply(rule s_out_fetch, simp, simp) | 
|         |    537 done | 
|         |    538  | 
|         |    539 lemma conf_keep:  | 
|         |    540       "\<lbrakk>t_ncorrect tp;  | 
|         |    541         \<not> (length tp1 div 2 < a \<and> | 
|         |    542         a \<le> length tp1 div 2 + length tp div 2); n > 0\<rbrakk> | 
|         |    543       \<Longrightarrow> t_steps (a, aa, ba) (tp, length tp1 div 2) n = (0, aa, ba)" | 
|         |    544 apply(induct n, simp) | 
|         |    545 apply(case_tac n, simp add: t_steps.simps) | 
|         |    546 apply(rule_tac conf_keep_step, simp+) | 
|         |    547 apply(subgoal_tac " t_steps (a, aa, ba)  | 
|         |    548                (tp, length tp1 div 2) (Suc (Suc nat)) | 
|         |    549          = t_step (t_steps (a, aa, ba)  | 
|         |    550             (tp, length tp1 div 2) (Suc nat)) (tp, length tp1 div 2)") | 
|         |    551 apply(simp) | 
|         |    552 apply(rule_tac conf_keep_step, simp, simp) | 
|         |    553 apply(rule stepn) | 
|         |    554 done | 
|         |    555  | 
|         |    556 lemma state_bef_inside:  | 
|         |    557     "\<lbrakk>t_ncorrect tp1; t_ncorrect tp;  | 
|         |    558       t_steps (s0, l0, r0) (tp, length tp1 div 2) stp = (s, l, r); | 
|         |    559       length tp1 div 2 < s0 \<and>  | 
|         |    560          s0 \<le> length tp1 div 2 + length tp div 2; | 
|         |    561       length tp1 div 2 < s \<and> s \<le> length tp1 div 2 + length tp div 2;  | 
|         |    562       n < stp; t_steps (s0, l0, r0) (tp, length tp1 div 2) n =  | 
|         |    563       (a, aa, ba)\<rbrakk> | 
|         |    564       \<Longrightarrow>  length tp1 div 2 < a \<and>  | 
|         |    565          a \<le> length tp1 div 2 + length tp div 2" | 
|         |    566 apply(subgoal_tac "\<exists> x. stp = n + x", erule exE) | 
|         |    567 apply(simp only: t_step_add) | 
|         |    568 apply(rule classical) | 
|         |    569 apply(subgoal_tac "t_steps (a, aa, ba)  | 
|         |    570           (tp, length tp1 div 2) x = (0, aa, ba)") | 
|         |    571 apply(simp) | 
|         |    572 apply(rule conf_keep, simp, simp, simp) | 
|         |    573 apply(rule_tac x = "stp - n" in exI, simp) | 
|         |    574 done | 
|         |    575  | 
|         |    576 lemma turing_shift_inside:  | 
|         |    577        "\<lbrakk>t_steps (s0, l0, r0) (tp, length tp1 div 2) stp = (s, l, r); | 
|         |    578          length tp1 div 2 < s0 \<and>  | 
|         |    579          s0 \<le> length tp1 div 2 + length tp div 2;  | 
|         |    580          t_ncorrect tp1; t_ncorrect tp; | 
|         |    581          length tp1 div 2 < s \<and>  | 
|         |    582          s \<le> length tp1 div 2 + length tp div 2\<rbrakk> | 
|         |    583        \<Longrightarrow> t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) stp = (s, l, r)" | 
|         |    584 apply(induct stp arbitrary: s l r) | 
|         |    585 apply(simp add: t_steps.simps) | 
|         |    586 apply(subgoal_tac " t_steps (s0, l0, r0)  | 
|         |    587         (tp, length tp1 div 2) (Suc stp) | 
|         |    588                   = t_step (t_steps (s0, l0, r0)  | 
|         |    589         (tp, length tp1 div 2) stp) (tp, length tp1 div 2)") | 
|         |    590 apply(case_tac "t_steps (s0, l0, r0) (tp, length tp1 div 2) stp") | 
|         |    591 apply(subgoal_tac "length tp1 div 2 < a \<and>  | 
|         |    592             a \<le> length tp1 div 2 + length tp div 2") | 
|         |    593 apply(subgoal_tac "t_steps (s0, l0, r0)  | 
|         |    594            (tp1 @ tp @ tp2, 0) stp = (a, b, c)") | 
|         |    595 apply(simp only: stepn, simp) | 
|         |    596 apply(rule_tac t_shift_in_step, simp+) | 
|         |    597 defer | 
|         |    598 apply(rule stepn) | 
|         |    599 apply(rule_tac n = stp and stp = "Suc stp" and a = a  | 
|         |    600                and aa = b and ba = c in state_bef_inside, simp+) | 
|         |    601 done | 
|         |    602  | 
|         |    603 lemma take_Suc_last[elim]: "Suc as \<le> length xs \<Longrightarrow>  | 
|         |    604             take (Suc as) xs = take as xs @ [xs ! as]" | 
|         |    605 apply(induct xs arbitrary: as, simp, simp) | 
|         |    606 apply(case_tac as, simp, simp) | 
|         |    607 done | 
|         |    608  | 
|         |    609 lemma concat_suc: "Suc as \<le> length xs \<Longrightarrow>  | 
|         |    610        concat (take (Suc as) xs) = concat (take as xs) @ xs! as" | 
|         |    611 apply(subgoal_tac "take (Suc as) xs = take as xs @ [xs ! as]", simp) | 
|         |    612 by auto | 
|         |    613  | 
|         |    614 lemma concat_take_suc_iff: "Suc n \<le> length tps \<Longrightarrow>  | 
|         |    615        concat (take n tps) @ (tps ! n) = concat (take (Suc n) tps)" | 
|         |    616 apply(drule_tac concat_suc, simp) | 
|         |    617 done | 
|         |    618  | 
|         |    619 lemma concat_drop_suc_iff:  | 
|         |    620    "Suc n < length tps \<Longrightarrow> concat (drop (Suc n) tps) =  | 
|         |    621            tps ! Suc n @ concat (drop (Suc (Suc n)) tps)" | 
|         |    622 apply(induct tps arbitrary: n, simp, simp) | 
|         |    623 apply(case_tac tps, simp, simp) | 
|         |    624 apply(case_tac n, simp, simp) | 
|         |    625 done | 
|         |    626  | 
|         |    627 declare append_assoc[simp del] | 
|         |    628  | 
|         |    629 lemma  tm_append: "\<lbrakk>n < length tps; tp = tps ! n\<rbrakk> \<Longrightarrow>  | 
|         |    630            \<exists> tp1 tp2. concat tps = tp1 @ tp @ tp2 \<and> tp1 =  | 
|         |    631               concat (take n tps) \<and> tp2 = concat (drop (Suc n) tps)" | 
|         |    632 apply(rule_tac x = "concat (take n tps)" in exI) | 
|         |    633 apply(rule_tac x = "concat (drop (Suc n) tps)" in exI) | 
|         |    634 apply(auto) | 
|         |    635 apply(induct n, simp) | 
|         |    636 apply(case_tac tps, simp, simp, simp) | 
|         |    637 apply(subgoal_tac "concat (take n tps) @ (tps ! n) =  | 
|         |    638                concat (take (Suc n) tps)") | 
|         |    639 apply(simp only: append_assoc[THEN sym], simp only: append_assoc) | 
|         |    640 apply(subgoal_tac " concat (drop (Suc n) tps) = tps ! Suc n @  | 
|         |    641                   concat (drop (Suc (Suc n)) tps)", simp) | 
|         |    642 apply(rule_tac concat_drop_suc_iff, simp) | 
|         |    643 apply(rule_tac concat_take_suc_iff, simp) | 
|         |    644 done | 
|         |    645  | 
|         |    646 declare append_assoc[simp] | 
|         |    647  | 
|         |    648 lemma map_of:  "n < length xs \<Longrightarrow> (map f xs) ! n = f (xs ! n)" | 
|         |    649 by(auto) | 
|         |    650  | 
|         |    651 lemma [simp]: "length (tms_of aprog) = length aprog" | 
|         |    652 apply(auto simp: tms_of.simps tpairs_of.simps) | 
|         |    653 done | 
|         |    654  | 
|         |    655 lemma ci_nth: "\<lbrakk>ly = layout_of aprog; as < length aprog;  | 
|         |    656                 abc_fetch as aprog = Some ins\<rbrakk> | 
|         |    657     \<Longrightarrow> ci ly (start_of ly as) ins = tms_of aprog ! as" | 
|         |    658 apply(simp add: tms_of.simps tpairs_of.simps  | 
|         |    659       abc_fetch.simps  map_of del: map_append) | 
|         |    660 done | 
|         |    661  | 
|         |    662 lemma t_split:"\<lbrakk> | 
|         |    663         ly = layout_of aprog; | 
|         |    664         as < length aprog; abc_fetch as aprog = Some ins\<rbrakk> | 
|         |    665       \<Longrightarrow> \<exists> tp1 tp2. concat (tms_of aprog) =  | 
|         |    666             tp1 @ (ci ly (start_of ly as) ins) @ tp2 | 
|         |    667             \<and> tp1 = concat (take as (tms_of aprog)) \<and>  | 
|         |    668               tp2 = concat (drop (Suc as) (tms_of aprog))" | 
|         |    669 apply(insert tm_append[of "as" "tms_of aprog"  | 
|         |    670                              "ci ly (start_of ly as) ins"], simp) | 
|         |    671 apply(subgoal_tac "ci ly (start_of ly as) ins = (tms_of aprog) ! as") | 
|         |    672 apply(subgoal_tac "length (tms_of aprog) = length aprog", simp, simp) | 
|         |    673 apply(rule_tac ci_nth, auto) | 
|         |    674 done | 
|         |    675  | 
|         |    676 lemma math_sub: "\<lbrakk>x >= Suc 0; x - 1 = z\<rbrakk> \<Longrightarrow> x + y - Suc 0 = z + y" | 
|         |    677 by auto | 
|         |    678  | 
|         |    679 lemma start_more_one: "as \<noteq> 0 \<Longrightarrow> start_of ly as >= Suc 0" | 
|         |    680 apply(induct as, simp add: start_of.simps) | 
|         |    681 apply(case_tac as, auto simp: start_of.simps) | 
|         |    682 done | 
|         |    683  | 
|         |    684 lemma tm_ct: "\<lbrakk>abc2t_correct aprog; tp \<in> set (tms_of aprog)\<rbrakk> \<Longrightarrow>  | 
|         |    685                            t_ncorrect tp" | 
|         |    686 apply(simp add: abc2t_correct.simps tms_of.simps) | 
|         |    687 apply(auto) | 
|         |    688 apply(simp add:list_all_iff, auto) | 
|         |    689 done | 
|         |    690  | 
|         |    691 lemma div_apart: "\<lbrakk>x mod (2::nat) = 0; y mod 2 = 0\<rbrakk>  | 
|         |    692           \<Longrightarrow> (x + y) div 2 = x div 2 + y div 2" | 
|         |    693 apply(drule mod_eqD)+ | 
|         |    694 apply(auto) | 
|         |    695 done | 
|         |    696  | 
|         |    697 lemma div_apart_iff: "\<lbrakk>x mod (2::nat) = 0; y mod 2 = 0\<rbrakk> \<Longrightarrow>  | 
|         |    698            (x + y) mod 2 = 0" | 
|         |    699 apply(auto) | 
|         |    700 done | 
|         |    701  | 
|         |    702 lemma tms_ct: "\<lbrakk>abc2t_correct aprog; n < length aprog\<rbrakk> \<Longrightarrow>  | 
|         |    703          t_ncorrect (concat (take n (tms_of aprog)))" | 
|         |    704 apply(induct n, simp add: t_ncorrect.simps, simp) | 
|         |    705 apply(subgoal_tac "concat (take (Suc n) (tms_of aprog)) =  | 
|         |    706         concat (take n (tms_of aprog)) @ (tms_of aprog ! n)", simp) | 
|         |    707 apply(simp add: t_ncorrect.simps) | 
|         |    708 apply(rule_tac div_apart_iff, simp) | 
|         |    709 apply(subgoal_tac "t_ncorrect (tms_of aprog ! n)",  | 
|         |    710             simp add: t_ncorrect.simps) | 
|         |    711 apply(rule_tac tm_ct, simp) | 
|         |    712 apply(rule_tac nth_mem, simp add: tms_of.simps tpairs_of.simps) | 
|         |    713 apply(rule_tac concat_suc, simp add: tms_of.simps tpairs_of.simps) | 
|         |    714 done | 
|         |    715  | 
|         |    716 lemma tcorrect_div2: "\<lbrakk>abc2t_correct aprog; Suc as < length aprog\<rbrakk> | 
|         |    717   \<Longrightarrow> (length (concat (take as (tms_of aprog))) + length (tms_of aprog | 
|         |    718  ! as)) div 2 = length (concat (take as (tms_of aprog))) div 2 +  | 
|         |    719                  length (tms_of aprog ! as) div 2" | 
|         |    720 apply(subgoal_tac "t_ncorrect (tms_of aprog ! as)") | 
|         |    721 apply(subgoal_tac "t_ncorrect (concat (take as (tms_of aprog)))") | 
|         |    722 apply(rule_tac div_apart) | 
|         |    723 apply(rule tct_div2, simp)+ | 
|         |    724 apply(erule_tac tms_ct, simp) | 
|         |    725 apply(rule_tac tm_ct, simp) | 
|         |    726 apply(rule_tac nth_mem) | 
|         |    727 apply(simp add: tms_of.simps tpairs_of.simps) | 
|         |    728 done | 
|         |    729  | 
|         |    730 lemma [simp]: "length (layout_of aprog) = length aprog" | 
|         |    731 apply(auto simp: layout_of.simps) | 
|         |    732 done | 
|         |    733  | 
|         |    734 lemma start_of_ind: "\<lbrakk>as < length aprog; ly = layout_of aprog\<rbrakk> \<Longrightarrow>  | 
|         |    735        start_of ly (Suc as) = start_of ly as +  | 
|         |    736                           length ((tms_of aprog) ! as) div 2" | 
|         |    737 apply(simp only: start_of.simps, simp) | 
|         |    738 apply(auto simp: start_of.simps tms_of.simps layout_of.simps  | 
|         |    739                  tpairs_of.simps) | 
|         |    740 apply(simp add: ci_length) | 
|         |    741 done | 
|         |    742  | 
|         |    743 lemma concat_take_suc: "Suc n \<le> length xs \<Longrightarrow> | 
|         |    744   concat (take (Suc n) xs) = concat (take n xs) @ (xs ! n)" | 
|         |    745 apply(subgoal_tac "take (Suc n) xs = | 
|         |    746                    take n xs @ [xs ! n]") | 
|         |    747 apply(auto) | 
|         |    748 done | 
|         |    749  | 
|         |    750 lemma ci_length_not0: "Suc 0 <= length (ci ly as i) div 2" | 
|         |    751 apply(subgoal_tac "length (ci ly as i) div 2 = length_of i") | 
|         |    752 apply(simp add: length_of.simps split: abc_inst.splits) | 
|         |    753 apply(rule ci_length) | 
|         |    754 done | 
|         |    755   | 
|         |    756 lemma findnth_length2: "length (findnth n) = 4 * n" | 
|         |    757 apply(induct n, simp) | 
|         |    758 apply(simp) | 
|         |    759 done | 
|         |    760  | 
|         |    761 lemma ci_length2: "length (ci ly as i) = 2 * (length_of i)" | 
|         |    762 apply(simp add: ci.simps length_of.simps tinc_b_def tdec_b_def | 
|         |    763               split: abc_inst.splits, auto) | 
|         |    764 apply(simp add: findnth_length2)+ | 
|         |    765 done | 
|         |    766  | 
|         |    767 lemma tm_mod2: "as < length aprog \<Longrightarrow>  | 
|         |    768              length (tms_of aprog ! as) mod 2 = 0" | 
|         |    769 apply(simp add: tms_of.simps) | 
|         |    770 apply(subgoal_tac "map (\<lambda>(x, y). ci (layout_of aprog) x y)  | 
|         |    771               (tpairs_of aprog) ! as | 
|         |    772                 = (\<lambda>(x, y). ci (layout_of aprog) x y)  | 
|         |    773               ((tpairs_of aprog) ! as)", simp) | 
|         |    774 apply(case_tac "(tpairs_of aprog ! as)", simp) | 
|         |    775 apply(subgoal_tac "length (ci (layout_of aprog) a b) = | 
|         |    776                  2 * (length_of b)", simp) | 
|         |    777 apply(rule ci_length2) | 
|         |    778 apply(rule map_of, simp add: tms_of.simps tpairs_of.simps) | 
|         |    779 done | 
|         |    780  | 
|         |    781 lemma tms_mod2: "as \<le> length aprog \<Longrightarrow>  | 
|         |    782         length (concat (take as (tms_of aprog))) mod 2 = 0" | 
|         |    783 apply(induct as, simp, simp) | 
|         |    784 apply(subgoal_tac "concat (take (Suc as) (tms_of aprog)) | 
|         |    785                   = concat (take as (tms_of aprog)) @  | 
|         |    786                        (tms_of aprog ! as)", auto) | 
|         |    787 apply(rule div_apart_iff, simp, rule tm_mod2, simp) | 
|         |    788 apply(rule concat_take_suc, simp add: tms_of.simps tpairs_of.simps) | 
|         |    789 done | 
|         |    790  | 
|         |    791 lemma [simp]: "\<lbrakk>as < length aprog; (abc_fetch as aprog) = Some ins\<rbrakk> | 
|         |    792        \<Longrightarrow> ci (layout_of aprog)  | 
|         |    793           (start_of (layout_of aprog) as) (ins) \<in> set (tms_of aprog)" | 
|         |    794 apply(insert ci_nth[of "layout_of aprog" aprog as], simp) | 
|         |    795 done | 
|         |    796  | 
|         |    797 lemma startof_not0: "start_of ly as > 0" | 
|         |    798 apply(induct as, simp add: start_of.simps) | 
|         |    799 apply(case_tac as, auto simp: start_of.simps) | 
|         |    800 done | 
|         |    801  | 
|         |    802 declare abc_step_l.simps[simp del] | 
|         |    803 lemma pre_lheq: "\<lbrakk>tp = concat (take as (tms_of aprog)); | 
|         |    804    abc2t_correct aprog; as \<le> length aprog\<rbrakk> \<Longrightarrow>  | 
|         |    805          start_of (layout_of aprog) as - Suc 0 = length tp div 2" | 
|         |    806 apply(induct as arbitrary: tp, simp add: start_of.simps, simp) | 
|         |    807 proof -  | 
|         |    808   fix as tp | 
|         |    809   assume h1: "\<And>tp. tp = concat (take as (tms_of aprog)) \<Longrightarrow>  | 
|         |    810      start_of (layout_of aprog) as - Suc 0 =  | 
|         |    811             length (concat (take as (tms_of aprog))) div 2" | 
|         |    812   and h2: " abc2t_correct aprog" "Suc as \<le> length aprog"  | 
|         |    813   from h2 show "start_of (layout_of aprog) (Suc as) - Suc 0 =  | 
|         |    814           length (concat (take (Suc as) (tms_of aprog))) div 2" | 
|         |    815     apply(insert h1[of "concat (take as (tms_of aprog))"], simp) | 
|         |    816     apply(insert start_of_ind[of as aprog "layout_of aprog"], simp) | 
|         |    817     apply(subgoal_tac "(take (Suc as) (tms_of aprog)) =  | 
|         |    818             take as (tms_of aprog) @ [(tms_of aprog) ! as]", simp) | 
|         |    819     apply(subgoal_tac "(length (concat (take as (tms_of aprog))) +  | 
|         |    820                        length (tms_of aprog ! as)) div 2 | 
|         |    821             = length (concat (take as (tms_of aprog))) div 2 +  | 
|         |    822               length (tms_of aprog ! as) div 2", simp) | 
|         |    823     apply(subgoal_tac "start_of (layout_of aprog) as =  | 
|         |    824        length (concat (take as (tms_of aprog))) div 2 + Suc 0", simp) | 
|         |    825     apply(subgoal_tac "start_of (layout_of aprog) as > 0", simp,  | 
|         |    826            rule_tac startof_not0) | 
|         |    827     apply(insert tm_mod2[of as aprog], simp) | 
|         |    828     apply(insert tms_mod2[of as aprog], simp, arith) | 
|         |    829     apply(rule take_Suc_last, simp) | 
|         |    830     done | 
|         |    831 qed | 
|         |    832  | 
|         |    833 lemma crsp2stateq:  | 
|         |    834   "\<lbrakk>as < length aprog; abc2t_correct aprog; | 
|         |    835        crsp_l (layout_of aprog) (as, am) (a, aa, ba) inres\<rbrakk> \<Longrightarrow>  | 
|         |    836         a = length (concat (take as (tms_of aprog))) div 2 + 1" | 
|         |    837 apply(simp add: crsp_l.simps) | 
|         |    838 apply(insert pre_lheq[of "(concat (take as (tms_of aprog)))" as aprog] | 
|         |    839 , simp)    | 
|         |    840 apply(subgoal_tac "start_of (layout_of aprog) as > 0",  | 
|         |    841       auto intro: startof_not0) | 
|         |    842 done | 
|         |    843  | 
|         |    844 lemma turing_shift_outside:  | 
|         |    845      "\<lbrakk>t_steps (s0, l0, r0) (tp, length tp1 div 2) stp = (s, l, r);  | 
|         |    846        s \<noteq> 0; stp > 0; | 
|         |    847        length tp1 div 2 < s0 \<and>  | 
|         |    848        s0 \<le> length tp1 div 2 + length tp div 2;  | 
|         |    849        t_ncorrect tp1; t_ncorrect tp; | 
|         |    850        \<not> (length tp1 div 2 < s \<and>  | 
|         |    851       s \<le> length tp1 div 2 + length tp div 2)\<rbrakk> | 
|         |    852     \<Longrightarrow> \<exists>stp' > 0. t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) stp'  | 
|         |    853                 = (s, l, r)" | 
|         |    854 apply(rule_tac x = stp in exI) | 
|         |    855 apply(case_tac stp, simp add: t_steps.simps) | 
|         |    856 apply(simp only: stepn) | 
|         |    857 apply(case_tac "t_steps (s0, l0, r0) (tp, length tp1 div 2) nat") | 
|         |    858 apply(subgoal_tac "length tp1 div 2 < a \<and>  | 
|         |    859                    a \<le> length tp1 div 2 + length tp div 2") | 
|         |    860 apply(subgoal_tac "t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) nat  | 
|         |    861                    = (a, b, c)", simp) | 
|         |    862 apply(rule_tac t_shift_in_step, simp+) | 
|         |    863 apply(rule_tac turing_shift_inside, simp+) | 
|         |    864 apply(rule classical) | 
|         |    865 apply(subgoal_tac "t_step (a,b,c)  | 
|         |    866             (tp, length tp1 div 2) = (0, b, c)", simp) | 
|         |    867 apply(rule_tac conf_keep_step, simp+) | 
|         |    868 done | 
|         |    869  | 
|         |    870 lemma turing_shift:  | 
|         |    871   "\<lbrakk>t_steps (s0, (l0, r0)) (tp, (length tp1 div 2)) stp | 
|         |    872    = (s, (l, r)); s \<noteq> 0; stp > 0; | 
|         |    873   (length tp1 div 2 < s0 \<and> s0 <= length tp1 div 2 + length tp div 2); | 
|         |    874   t_ncorrect tp1; t_ncorrect tp\<rbrakk> \<Longrightarrow>  | 
|         |    875          \<exists> stp' > 0. t_steps (s0, (l0, r0)) (tp1 @ tp @ tp2, 0) stp' = | 
|         |    876                     (s, (l, r))" | 
|         |    877 apply(case_tac "s > length tp1 div 2 \<and>  | 
|         |    878               s <= length tp1 div 2 + length tp div 2") | 
|         |    879 apply(subgoal_tac " t_steps (s0, l0, r0) (tp1 @ tp @ tp2, 0) stp =  | 
|         |    880                    (s, l, r)") | 
|         |    881 apply(rule_tac x = stp in exI, simp) | 
|         |    882 apply(rule_tac turing_shift_inside, simp+) | 
|         |    883 apply(rule_tac turing_shift_outside, simp+) | 
|         |    884 done | 
|         |    885  | 
|         |    886 lemma inc_startof_not0:  "start_of ly as \<ge> Suc 0" | 
|         |    887 apply(induct as, simp add: start_of.simps) | 
|         |    888 apply(simp add: start_of.simps) | 
|         |    889 done | 
|         |    890  | 
|         |    891 lemma s_crsp: | 
|         |    892   "\<lbrakk>as < length aprog; abc_fetch as aprog = Some ins; | 
|         |    893   abc2t_correct aprog; | 
|         |    894   crsp_l (layout_of aprog) (as, am) (a, aa, ba) inres\<rbrakk> \<Longrightarrow>   | 
|         |    895   length (concat (take as (tms_of aprog))) div 2 < a  | 
|         |    896       \<and> a \<le> length (concat (take as (tms_of aprog))) div 2 +  | 
|         |    897          length (ci (layout_of aprog) (start_of (layout_of aprog) as) | 
|         |    898          ins) div 2" | 
|         |    899 apply(subgoal_tac "a = length (concat (take as (tms_of aprog))) div  | 
|         |    900                    2 + 1", simp) | 
|         |    901 apply(rule_tac ci_length_not0) | 
|         |    902 apply(rule crsp2stateq, simp+) | 
|         |    903 done | 
|         |    904  | 
|         |    905 lemma tms_out_ex: | 
|         |    906   "\<lbrakk>ly = layout_of aprog; tprog = tm_of aprog; | 
|         |    907   abc2t_correct aprog; | 
|         |    908   crsp_l ly (as, am) tc inres; as < length aprog; | 
|         |    909   abc_fetch as aprog = Some ins; | 
|         |    910   t_steps tc (ci ly (start_of ly as) ins,  | 
|         |    911   (start_of ly as) - 1) n = (s, l, r); | 
|         |    912   n > 0;  | 
|         |    913   abc_step_l (as, am) (abc_fetch as aprog) = (as', am'); | 
|         |    914   s = start_of ly as' | 
|         |    915   \<rbrakk> | 
|         |    916   \<Longrightarrow> \<exists> stp > 0. (t_steps tc (tprog, 0) stp = (s, (l, r)))" | 
|         |    917 apply(simp only: tm_of.simps) | 
|         |    918 apply(subgoal_tac "\<exists> tp1 tp2. concat (tms_of aprog) =  | 
|         |    919       tp1 @ (ci ly (start_of ly as) ins) @ tp2 | 
|         |    920     \<and> tp1 = concat (take as (tms_of aprog)) \<and>  | 
|         |    921       tp2 = concat (drop (Suc as) (tms_of aprog))") | 
|         |    922 apply(erule exE, erule exE, erule conjE, erule conjE, | 
|         |    923       case_tac tc, simp) | 
|         |    924 apply(rule turing_shift) | 
|         |    925 apply(subgoal_tac "start_of (layout_of aprog) as - Suc 0  | 
|         |    926                 = length tp1 div 2", simp) | 
|         |    927 apply(rule_tac pre_lheq, simp, simp, simp) | 
|         |    928 apply(simp add: startof_not0, simp) | 
|         |    929 apply(rule_tac s_crsp, simp, simp, simp, simp) | 
|         |    930 apply(rule tms_ct, simp, simp) | 
|         |    931 apply(rule tm_ct, simp) | 
|         |    932 apply(subgoal_tac "ci (layout_of aprog)  | 
|         |    933                  (start_of (layout_of aprog) as) ins | 
|         |    934                 = (tms_of aprog ! as)", simp) | 
|         |    935 apply(simp add: tms_of.simps tpairs_of.simps) | 
|         |    936 apply(simp add: tms_of.simps tpairs_of.simps abc_fetch.simps) | 
|         |    937 apply(erule_tac t_split, auto simp: tm_of.simps) | 
|         |    938 done | 
|         |    939  | 
|         |    940 (* | 
|         |    941 subsection {* The compilation of @{text "Inc n"} *} | 
|         |    942 *) | 
|         |    943  | 
|         |    944 text {* | 
|         |    945   The lemmas in this section lead to the correctness of  | 
|         |    946   the compilation of @{text "Inc n"} instruction. | 
|         |    947 *} | 
|         |    948  | 
|         |    949 fun at_begin_fst_bwtn :: "inc_inv_t" | 
|         |    950   where | 
|         |    951   "at_begin_fst_bwtn (as, lm) (s, l, r) ires =  | 
|         |    952       (\<exists> lm1 tn rn. lm1 = (lm @ (0\<^bsup>tn\<^esup>)) \<and> length lm1 = s \<and>  | 
|         |    953           (if lm1 = [] then l = Bk # Bk # ires | 
|         |    954            else l = [Bk]@<rev lm1>@Bk#Bk#ires) \<and> r = (Bk\<^bsup>rn\<^esup>))"  | 
|         |    955  | 
|         |    956  | 
|         |    957 fun at_begin_fst_awtn :: "inc_inv_t" | 
|         |    958   where | 
|         |    959   "at_begin_fst_awtn (as, lm) (s, l, r) ires =  | 
|         |    960       (\<exists> lm1 tn rn. lm1 = (lm @ (0\<^bsup>tn\<^esup>)) \<and> length lm1 = s \<and> | 
|         |    961          (if lm1 = []  then l = Bk # Bk # ires | 
|         |    962           else l = [Bk]@<rev lm1>@Bk#Bk#ires) \<and> r = [Oc]@Bk\<^bsup>rn\<^esup> | 
|         |    963   )" | 
|         |    964  | 
|         |    965 fun at_begin_norm :: "inc_inv_t" | 
|         |    966   where | 
|         |    967   "at_begin_norm (as, lm) (s, l, r) ires=  | 
|         |    968       (\<exists> lm1 lm2 rn. lm = lm1 @ lm2 \<and> length lm1 = s \<and>  | 
|         |    969         (if lm1 = [] then l = Bk # Bk # ires | 
|         |    970          else l = Bk # <rev lm1> @ Bk# Bk # ires ) \<and> r = <lm2> @ (Bk\<^bsup>rn\<^esup>))" | 
|         |    971  | 
|         |    972 fun in_middle :: "inc_inv_t" | 
|         |    973   where | 
|         |    974   "in_middle (as, lm) (s, l, r) ires =  | 
|         |    975       (\<exists> lm1 lm2 tn m ml mr rn. lm @ 0\<^bsup>tn\<^esup> = lm1 @ [m] @ lm2 | 
|         |    976        \<and> length lm1 = s \<and> m + 1 = ml + mr \<and>   | 
|         |    977          ml \<noteq> 0 \<and> tn = s + 1 - length lm \<and>  | 
|         |    978        (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires  | 
|         |    979         else l = (Oc\<^bsup>ml\<^esup>)@[Bk]@<rev lm1>@ | 
|         |    980                  Bk # Bk # ires) \<and> (r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ <lm2>@ (Bk\<^bsup>rn\<^esup>) \<or>  | 
|         |    981       (lm2 = [] \<and> r = (Oc\<^bsup>mr\<^esup>))) | 
|         |    982       )" | 
|         |    983  | 
|         |    984 fun inv_locate_a :: "inc_inv_t" | 
|         |    985   where "inv_locate_a (as, lm) (s, l, r) ires =  | 
|         |    986      (at_begin_norm (as, lm) (s, l, r) ires \<or> | 
|         |    987       at_begin_fst_bwtn (as, lm) (s, l, r) ires \<or> | 
|         |    988       at_begin_fst_awtn (as, lm) (s, l, r) ires | 
|         |    989       )" | 
|         |    990  | 
|         |    991 fun inv_locate_b :: "inc_inv_t" | 
|         |    992   where "inv_locate_b (as, lm) (s, l, r) ires =  | 
|         |    993         (in_middle (as, lm) (s, l, r)) ires " | 
|         |    994  | 
|         |    995 fun inv_after_write :: "inc_inv_t" | 
|         |    996   where "inv_after_write (as, lm) (s, l, r) ires =  | 
|         |    997            (\<exists> rn m lm1 lm2. lm = lm1 @ m # lm2 \<and> | 
|         |    998              (if lm1 = [] then l = Oc\<^bsup>m\<^esup> @ Bk # Bk # ires | 
|         |    999               else Oc # l = Oc\<^bsup>Suc m \<^esup>@ Bk # <rev lm1> @  | 
|         |   1000                       Bk # Bk # ires) \<and> r = [Oc] @ <lm2> @ (Bk\<^bsup>rn\<^esup>))" | 
|         |   1001  | 
|         |   1002 fun inv_after_move :: "inc_inv_t" | 
|         |   1003   where "inv_after_move (as, lm) (s, l, r) ires =  | 
|         |   1004       (\<exists> rn m lm1 lm2. lm = lm1 @ m # lm2 \<and> | 
|         |   1005         (if lm1 = [] then l = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # ires | 
|         |   1006          else l = Oc\<^bsup>Suc m\<^esup>@ Bk # <rev lm1> @ Bk # Bk # ires) \<and>  | 
|         |   1007         r = <lm2> @ (Bk\<^bsup>rn\<^esup>))" | 
|         |   1008  | 
|         |   1009 fun inv_after_clear :: "inc_inv_t" | 
|         |   1010   where "inv_after_clear (as, lm) (s, l, r) ires = | 
|         |   1011        (\<exists> rn m lm1 lm2 r'. lm = lm1 @ m # lm2 \<and>  | 
|         |   1012         (if lm1 = [] then l = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # ires | 
|         |   1013          else l = Oc\<^bsup>Suc m\<^esup>@ Bk # <rev lm1> @ Bk # Bk # ires) \<and>  | 
|         |   1014           r = Bk # r' \<and> Oc # r' = <lm2> @ (Bk\<^bsup>rn\<^esup>))" | 
|         |   1015  | 
|         |   1016 fun inv_on_right_moving :: "inc_inv_t" | 
|         |   1017   where "inv_on_right_moving (as, lm) (s, l, r) ires =  | 
|         |   1018        (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and>  | 
|         |   1019             ml + mr = m \<and>  | 
|         |   1020           (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires | 
|         |   1021           else l = (Oc\<^bsup>ml\<^esup>) @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and>  | 
|         |   1022          ((r = (Oc\<^bsup>mr\<^esup>) @ [Bk] @ <lm2> @ (Bk\<^bsup>rn\<^esup>)) \<or>  | 
|         |   1023           (r = (Oc\<^bsup>mr\<^esup>) \<and> lm2 = [])))" | 
|         |   1024  | 
|         |   1025 fun inv_on_left_moving_norm :: "inc_inv_t" | 
|         |   1026   where "inv_on_left_moving_norm (as, lm) (s, l, r) ires = | 
|         |   1027       (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and>   | 
|         |   1028              ml + mr = Suc m \<and> mr > 0 \<and> (if lm1 = [] then l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # ires | 
|         |   1029                                          else l = (Oc\<^bsup>ml\<^esup>) @ Bk # <rev lm1> @ Bk # Bk # ires) | 
|         |   1030         \<and> (r = (Oc\<^bsup>mr\<^esup>) @ Bk # <lm2> @ (Bk\<^bsup>rn\<^esup>) \<or>  | 
|         |   1031            (lm2 = [] \<and> r = Oc\<^bsup>mr\<^esup>)))" | 
|         |   1032  | 
|         |   1033 fun inv_on_left_moving_in_middle_B:: "inc_inv_t" | 
|         |   1034   where "inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires = | 
|         |   1035                 (\<exists> lm1 lm2 rn. lm = lm1 @ lm2 \<and>   | 
|         |   1036                      (if lm1 = [] then l = Bk # ires | 
|         |   1037                       else l = <rev lm1> @ Bk # Bk # ires) \<and>  | 
|         |   1038                       r = Bk # <lm2> @ (Bk\<^bsup>rn\<^esup>))" | 
|         |   1039  | 
|         |   1040 fun inv_on_left_moving :: "inc_inv_t" | 
|         |   1041   where "inv_on_left_moving (as, lm) (s, l, r) ires =  | 
|         |   1042        (inv_on_left_moving_norm  (as, lm) (s, l, r) ires \<or> | 
|         |   1043         inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires)" | 
|         |   1044  | 
|         |   1045  | 
|         |   1046 fun inv_check_left_moving_on_leftmost :: "inc_inv_t" | 
|         |   1047   where "inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires =  | 
|         |   1048                 (\<exists> rn. l = ires \<and> r = [Bk, Bk] @ <lm> @ (Bk\<^bsup>rn\<^esup>))" | 
|         |   1049  | 
|         |   1050 fun inv_check_left_moving_in_middle :: "inc_inv_t" | 
|         |   1051   where "inv_check_left_moving_in_middle (as, lm) (s, l, r) ires =  | 
|         |   1052  | 
|         |   1053               (\<exists> lm1 lm2 r' rn. lm = lm1 @ lm2 \<and> | 
|         |   1054                  (Oc # l = <rev lm1> @ Bk # Bk # ires) \<and> r = Oc # Bk # r' \<and>  | 
|         |   1055                            r' = <lm2> @ (Bk\<^bsup>rn\<^esup>))" | 
|         |   1056  | 
|         |   1057 fun inv_check_left_moving :: "inc_inv_t" | 
|         |   1058   where "inv_check_left_moving (as, lm) (s, l, r) ires =  | 
|         |   1059              (inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires \<or> | 
|         |   1060              inv_check_left_moving_in_middle (as, lm) (s, l, r) ires)" | 
|         |   1061  | 
|         |   1062 fun inv_after_left_moving :: "inc_inv_t" | 
|         |   1063   where "inv_after_left_moving (as, lm) (s, l, r) ires=  | 
|         |   1064               (\<exists> rn. l = Bk # ires \<and> r = Bk # <lm> @ (Bk\<^bsup>rn\<^esup>))" | 
|         |   1065  | 
|         |   1066 fun inv_stop :: "inc_inv_t" | 
|         |   1067   where "inv_stop (as, lm) (s, l, r) ires=  | 
|         |   1068               (\<exists> rn. l = Bk # Bk # ires \<and> r = <lm> @ (Bk\<^bsup>rn\<^esup>))" | 
|         |   1069  | 
|         |   1070  | 
|         |   1071 fun inc_inv :: "layout \<Rightarrow> nat \<Rightarrow> inc_inv_t" | 
|         |   1072   where | 
|         |   1073   "inc_inv ly n (as, lm) (s, l, r) ires = | 
|         |   1074               (let ss = start_of ly as in | 
|         |   1075                let lm' = abc_lm_s lm n ((abc_lm_v lm n)+1) in | 
|         |   1076                 if s = 0 then False | 
|         |   1077                 else if s < ss then False | 
|         |   1078                 else if s < ss + 2 * n then  | 
|         |   1079                    if (s - ss) mod 2 = 0 then  | 
|         |   1080                        inv_locate_a (as, lm) ((s - ss) div 2, l, r) ires | 
|         |   1081                    else inv_locate_b (as, lm) ((s - ss) div 2, l, r) ires | 
|         |   1082                 else if s = ss + 2 * n then  | 
|         |   1083                         inv_locate_a (as, lm) (n, l, r) ires | 
|         |   1084                 else if s = ss + 2 * n + 1 then  | 
|         |   1085                    inv_locate_b (as, lm) (n, l, r) ires | 
|         |   1086                 else if s = ss + 2 * n + 2 then  | 
|         |   1087                    inv_after_write (as, lm') (s - ss, l, r) ires | 
|         |   1088                 else if s = ss + 2 * n + 3 then  | 
|         |   1089                    inv_after_move (as, lm') (s - ss, l, r) ires | 
|         |   1090                 else if s = ss + 2 * n + 4 then  | 
|         |   1091                    inv_after_clear (as, lm') (s - ss, l, r) ires | 
|         |   1092                 else if s = ss + 2 * n + 5 then  | 
|         |   1093                    inv_on_right_moving (as, lm') (s - ss, l, r) ires | 
|         |   1094                 else if s = ss + 2 * n + 6 then  | 
|         |   1095                    inv_on_left_moving (as, lm') (s - ss, l, r) ires | 
|         |   1096                 else if s = ss + 2 * n + 7 then  | 
|         |   1097                    inv_check_left_moving (as, lm') (s - ss, l, r) ires | 
|         |   1098                 else if s = ss + 2 * n + 8 then  | 
|         |   1099                    inv_after_left_moving (as, lm') (s - ss, l, r) ires | 
|         |   1100                 else if s = ss + 2 * n + 9 then  | 
|         |   1101                    inv_stop (as, lm') (s - ss, l, r) ires | 
|         |   1102                 else False) " | 
|         |   1103  | 
|         |   1104 lemma fetch_intro:  | 
|         |   1105   "\<lbrakk>\<And>xs.\<lbrakk>ba = Oc # xs\<rbrakk> \<Longrightarrow> P (fetch prog i Oc); | 
|         |   1106    \<And>xs.\<lbrakk>ba = Bk # xs\<rbrakk> \<Longrightarrow> P (fetch prog i Bk); | 
|         |   1107    ba = [] \<Longrightarrow> P (fetch prog i Bk) | 
|         |   1108    \<rbrakk> \<Longrightarrow> P (fetch prog i  | 
|         |   1109              (case ba of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))" | 
|         |   1110 by (auto split:list.splits block.splits) | 
|         |   1111  | 
|         |   1112 lemma length_findnth[simp]: "length (findnth n) = 4 * n" | 
|         |   1113 apply(induct n, simp) | 
|         |   1114 apply(simp) | 
|         |   1115 done | 
|         |   1116  | 
|         |   1117 declare tshift.simps[simp del] | 
|         |   1118 declare findnth.simps[simp del] | 
|         |   1119  | 
|         |   1120 lemma findnth_nth:  | 
|         |   1121  "\<lbrakk>n > q; x < 4\<rbrakk> \<Longrightarrow>  | 
|         |   1122         (findnth n) ! (4 * q + x) = (findnth (Suc q) ! (4 * q + x))" | 
|         |   1123 apply(induct n, simp) | 
|         |   1124 apply(case_tac "q < n", simp add: findnth.simps, auto) | 
|         |   1125 apply(simp add: nth_append) | 
|         |   1126 apply(subgoal_tac "q = n", simp) | 
|         |   1127 apply(arith) | 
|         |   1128 done | 
|         |   1129  | 
|         |   1130 lemma Suc_pre[simp]: "\<not> a < start_of ly as \<Longrightarrow>  | 
|         |   1131           (Suc a - start_of ly as) = Suc (a - start_of ly as)" | 
|         |   1132 apply(arith) | 
|         |   1133 done | 
|         |   1134  | 
|         |   1135 lemma fetch_locate_a_o: " | 
|         |   1136 \<And>a  q xs. | 
|         |   1137     \<lbrakk>\<not> a < start_of (layout_of aprog) as;  | 
|         |   1138       a < start_of (layout_of aprog) as + 2 * n;  | 
|         |   1139       a - start_of (layout_of aprog) as = 2 * q;  | 
|         |   1140       start_of (layout_of aprog) as > 0\<rbrakk> | 
|         |   1141     \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) | 
|         |   1142          (Inc n)) (Suc (2 * q)) Oc) = (R, a+1)" | 
|         |   1143 apply(auto simp: ci.simps findnth.simps fetch.simps | 
|         |   1144                   nth_of.simps tshift.simps nth_append Suc_pre) | 
|         |   1145 apply(subgoal_tac "(findnth n ! Suc (4 * q)) =  | 
|         |   1146                  findnth (Suc q) ! (4 * q + 1)") | 
|         |   1147 apply(simp add: findnth.simps nth_append) | 
|         |   1148 apply(subgoal_tac " findnth n !(4 * q + 1) =  | 
|         |   1149                  findnth (Suc q) ! (4 * q + 1)", simp) | 
|         |   1150 apply(rule_tac findnth_nth, auto) | 
|         |   1151 done | 
|         |   1152  | 
|         |   1153 lemma fetch_locate_a_b: " | 
|         |   1154 \<And>a  q xs. | 
|         |   1155     \<lbrakk>abc_fetch as aprog = Some (Inc n);   | 
|         |   1156      \<not> a < start_of (layout_of aprog) as;  | 
|         |   1157      a < start_of (layout_of aprog) as + 2 * n;  | 
|         |   1158      a - start_of (layout_of aprog) as = 2 * q;  | 
|         |   1159      start_of (layout_of aprog) as > 0\<rbrakk> | 
|         |   1160     \<Longrightarrow> (fetch (ci (layout_of aprog)  | 
|         |   1161       (start_of (layout_of aprog) as) (Inc n)) (Suc (2 * q)) Bk) | 
|         |   1162        = (W1, a)" | 
|         |   1163 apply(auto simp: ci.simps findnth.simps fetch.simps | 
|         |   1164                  tshift.simps nth_append) | 
|         |   1165 apply(subgoal_tac "(findnth n ! (4 * q)) =  | 
|         |   1166                            findnth (Suc q) ! (4 * q )") | 
|         |   1167 apply(simp add: findnth.simps nth_append) | 
|         |   1168 apply(subgoal_tac " findnth n !(4 * q + 0) = | 
|         |   1169                             findnth (Suc q) ! (4 * q + 0)", simp) | 
|         |   1170 apply(rule_tac findnth_nth, auto) | 
|         |   1171 done | 
|         |   1172  | 
|         |   1173 lemma [intro]: "x mod 2 = Suc 0 \<Longrightarrow> \<exists> q. x = Suc (2 * q)" | 
|         |   1174 apply(drule mod_eqD, auto) | 
|         |   1175 done | 
|         |   1176  | 
|         |   1177 lemma  add3_Suc: "x + 3 = Suc (Suc (Suc x))" | 
|         |   1178 apply(arith) | 
|         |   1179 done | 
|         |   1180  | 
|         |   1181 declare start_of.simps[simp] | 
|         |   1182 (* | 
|         |   1183 lemma layout_not0: "start_of ly as > 0" | 
|         |   1184 by(induct as, auto) | 
|         |   1185 *) | 
|         |   1186 lemma [simp]:  | 
|         |   1187  "\<lbrakk>\<not> a < start_of (layout_of aprog) as;  | 
|         |   1188    a - start_of (layout_of aprog) as = Suc (2 * q);  | 
|         |   1189    abc_fetch as aprog = Some (Inc n);  | 
|         |   1190    start_of (layout_of aprog) as > 0\<rbrakk> | 
|         |   1191     \<Longrightarrow> Suc (Suc (2 * q + start_of (layout_of aprog) as - Suc 0)) = a" | 
|         |   1192 apply(subgoal_tac  | 
|         |   1193 "Suc (Suc (2 * q + start_of (layout_of aprog) as - Suc 0))  | 
|         |   1194               = 2 + 2 * q + start_of (layout_of aprog) as - Suc 0",  | 
|         |   1195   simp, simp add: inc_startof_not0) | 
|         |   1196 done | 
|         |   1197  | 
|         |   1198 lemma fetch_locate_b_o: " | 
|         |   1199 \<And>a  xs. | 
|         |   1200     \<lbrakk>0 < a; \<not> a < start_of (layout_of aprog) as;  | 
|         |   1201   a < start_of (layout_of aprog) as + 2 * n;  | 
|         |   1202  (a - start_of (layout_of aprog) as) mod 2 = Suc 0;  | 
|         |   1203  start_of (layout_of aprog) as > 0\<rbrakk> | 
|         |   1204     \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as)  | 
|         |   1205       (Inc n)) (Suc (a - start_of (layout_of aprog) as)) Oc) = (R, a)" | 
|         |   1206 apply(auto simp: ci.simps findnth.simps fetch.simps | 
|         |   1207                  nth_of.simps tshift.simps nth_append) | 
|         |   1208 apply(subgoal_tac "\<exists> q. (a - start_of (layout_of aprog) as) =  | 
|         |   1209                          2 * q + 1", auto) | 
|         |   1210 apply(subgoal_tac "(findnth n ! Suc (Suc (Suc (4 * q))))  | 
|         |   1211                   = findnth (Suc q) ! (4 * q + 3)") | 
|         |   1212 apply(simp add: findnth.simps nth_append) | 
|         |   1213 apply(subgoal_tac " findnth n ! (4 * q + 3) =  | 
|         |   1214                  findnth (Suc q) ! (4 * q + 3)", simp add: add3_Suc) | 
|         |   1215 apply(rule_tac findnth_nth, auto) | 
|         |   1216 done | 
|         |   1217  | 
|         |   1218 lemma fetch_locate_b_b: " | 
|         |   1219 \<And>a  xs. | 
|         |   1220     \<lbrakk>0 < a;  \<not> a < start_of (layout_of aprog) as;  | 
|         |   1221      a < start_of (layout_of aprog) as + 2 * n;  | 
|         |   1222      (a - start_of (layout_of aprog) as) mod 2 = Suc 0;  | 
|         |   1223      start_of (layout_of aprog) as > 0\<rbrakk> | 
|         |   1224     \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as) | 
|         |   1225         (Inc n)) (Suc (a - start_of (layout_of aprog) as)) Bk)  | 
|         |   1226         = (R, a + 1)" | 
|         |   1227 apply(auto simp: ci.simps findnth.simps fetch.simps | 
|         |   1228                   nth_of.simps tshift.simps nth_append) | 
|         |   1229 apply(subgoal_tac "\<exists> q. (a - start_of (layout_of aprog) as) =  | 
|         |   1230                   2 * q + 1", auto) | 
|         |   1231 apply(subgoal_tac "(findnth n ! Suc ((Suc (4 * q)))) =  | 
|         |   1232                     findnth (Suc q) ! (4 * q + 2)") | 
|         |   1233 apply(simp add: findnth.simps nth_append) | 
|         |   1234 apply(subgoal_tac " findnth n ! (4 * q + 2) =  | 
|         |   1235                     findnth (Suc q) ! (4 * q + 2)", simp) | 
|         |   1236 apply(rule_tac findnth_nth, auto) | 
|         |   1237 done | 
|         |   1238  | 
|         |   1239 lemma fetch_locate_n_a_o:  | 
|         |   1240        "start_of (layout_of aprog) as > 0 | 
|         |   1241        \<Longrightarrow> (fetch (ci (layout_of aprog)  | 
|         |   1242       (start_of (layout_of aprog) as) (Inc n)) (Suc (2 * n)) Oc) =  | 
|         |   1243              (R, start_of (layout_of aprog) as + 2 * n + 1)" | 
|         |   1244 apply(auto simp: ci.simps findnth.simps fetch.simps | 
|         |   1245                   nth_of.simps tshift.simps nth_append tinc_b_def) | 
|         |   1246 done | 
|         |   1247  | 
|         |   1248 lemma fetch_locate_n_a_b: " | 
|         |   1249        start_of (layout_of aprog) as > 0 | 
|         |   1250       \<Longrightarrow> (fetch (ci (layout_of aprog)  | 
|         |   1251     (start_of (layout_of aprog) as) (Inc n)) (Suc (2 * n)) Bk)  | 
|         |   1252    = (W1, start_of (layout_of aprog) as + 2 * n)" | 
|         |   1253 apply(auto simp: ci.simps findnth.simps fetch.simps | 
|         |   1254                   nth_of.simps tshift.simps nth_append tinc_b_def) | 
|         |   1255 done | 
|         |   1256  | 
|         |   1257 lemma fetch_locate_n_b_o: " | 
|         |   1258     start_of (layout_of aprog) as > 0 | 
|         |   1259     \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as)  | 
|         |   1260      (Inc n)) (Suc (Suc (2 * n))) Oc) =  | 
|         |   1261                       (R, start_of (layout_of aprog) as + 2 * n + 1)" | 
|         |   1262 apply(auto simp: ci.simps findnth.simps fetch.simps | 
|         |   1263                   nth_of.simps tshift.simps nth_append tinc_b_def) | 
|         |   1264 done | 
|         |   1265  | 
|         |   1266 lemma fetch_locate_n_b_b: " | 
|         |   1267     start_of (layout_of aprog) as > 0 | 
|         |   1268    \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as)  | 
|         |   1269    (Inc n)) (Suc (Suc (2 * n))) Bk) =  | 
|         |   1270        (W1, start_of (layout_of aprog) as + 2 * n + 2)" | 
|         |   1271 apply(auto simp: ci.simps findnth.simps fetch.simps | 
|         |   1272                   nth_of.simps tshift.simps nth_append tinc_b_def) | 
|         |   1273 done | 
|         |   1274  | 
|         |   1275 lemma fetch_after_write_o: " | 
|         |   1276     start_of (layout_of aprog) as > 0 | 
|         |   1277     \<Longrightarrow> (fetch (ci (layout_of aprog) (start_of (layout_of aprog) as)  | 
|         |   1278             (Inc n)) (Suc (Suc (Suc (2 * n)))) Oc) =  | 
|         |   1279         (R, start_of (layout_of aprog) as + 2*n + 3)" | 
|         |   1280 apply(auto simp: ci.simps findnth.simps fetch.simps | 
|         |   1281                   nth_of.simps tshift.simps nth_append tinc_b_def) | 
|         |   1282 done | 
|         |   1283  | 
|         |   1284 lemma fetch_after_move_o: " | 
|         |   1285       start_of (layout_of aprog) as > 0 | 
|         |   1286       \<Longrightarrow> (fetch (ci (layout_of aprog)  | 
|         |   1287               (start_of (layout_of aprog) as) (Inc n)) (4 + 2 * n) Oc) | 
|         |   1288         = (W0, start_of (layout_of aprog) as + 2 * n + 4)" | 
|         |   1289 apply(auto simp: ci.simps findnth.simps tshift.simps  | 
|         |   1290                  tinc_b_def add3_Suc) | 
|         |   1291 apply(subgoal_tac "4 + 2*n = Suc (2*n + 3)", simp only: fetch.simps) | 
|         |   1292 apply(auto simp: nth_of.simps nth_append) | 
|         |   1293 done | 
|         |   1294  | 
|         |   1295 lemma fetch_after_move_b: " | 
|         |   1296       start_of (layout_of aprog) as > 0 | 
|         |   1297       \<Longrightarrow>(fetch (ci (layout_of aprog)  | 
|         |   1298             (start_of (layout_of aprog) as) (Inc n)) (4 + 2 * n) Bk) | 
|         |   1299        = (L, start_of (layout_of aprog) as + 2 * n + 6)" | 
|         |   1300 apply(auto simp: ci.simps findnth.simps tshift.simps  | 
|         |   1301                  tinc_b_def add3_Suc) | 
|         |   1302 apply(subgoal_tac "4 + 2*n = Suc (2*n + 3)", simp only: fetch.simps) | 
|         |   1303 apply(auto simp: nth_of.simps nth_append) | 
|         |   1304 done | 
|         |   1305  | 
|         |   1306 lemma fetch_clear_b: " | 
|         |   1307      start_of (layout_of aprog) as > 0 | 
|         |   1308       \<Longrightarrow> (fetch (ci (layout_of aprog)  | 
|         |   1309               (start_of (layout_of aprog) as) (Inc n)) (5 + 2 * n) Bk) | 
|         |   1310       = (R, start_of (layout_of aprog) as + 2 * n + 5)" | 
|         |   1311 apply(auto simp: ci.simps findnth.simps  | 
|         |   1312                      tshift.simps tinc_b_def add3_Suc) | 
|         |   1313 apply(subgoal_tac "5 + 2*n = Suc (2*n + 4)", simp only: fetch.simps) | 
|         |   1314 apply(auto simp: nth_of.simps nth_append) | 
|         |   1315 done | 
|         |   1316  | 
|         |   1317 lemma fetch_right_move_o: " | 
|         |   1318       start_of (layout_of aprog) as > 0 | 
|         |   1319       \<Longrightarrow> (fetch (ci (layout_of aprog)  | 
|         |   1320                 (start_of (layout_of aprog) as) (Inc n)) (6 + 2*n) Oc) | 
|         |   1321       = (R, start_of (layout_of aprog) as + 2 * n + 5)" | 
|         |   1322 apply(auto simp: ci.simps findnth.simps tshift.simps  | 
|         |   1323                  tinc_b_def add3_Suc) | 
|         |   1324 apply(subgoal_tac "6 + 2*n = Suc (2*n + 5)", simp only: fetch.simps) | 
|         |   1325 apply(auto simp: nth_of.simps nth_append) | 
|         |   1326 done | 
|         |   1327  | 
|         |   1328 lemma fetch_right_move_b: " | 
|         |   1329       start_of (layout_of aprog) as > 0 | 
|         |   1330       \<Longrightarrow> (fetch (ci (layout_of aprog)  | 
|         |   1331                 (start_of (layout_of aprog) as) (Inc n)) (6 + 2*n) Bk) | 
|         |   1332       = (W1, start_of (layout_of aprog) as + 2 * n + 2)" | 
|         |   1333 apply(auto simp: ci.simps findnth.simps  | 
|         |   1334                  tshift.simps tinc_b_def add3_Suc) | 
|         |   1335 apply(subgoal_tac "6 + 2*n = Suc (2*n + 5)", simp only: fetch.simps) | 
|         |   1336 apply(auto simp: nth_of.simps nth_append) | 
|         |   1337 done | 
|         |   1338  | 
|         |   1339 lemma fetch_left_move_o: " | 
|         |   1340       start_of (layout_of aprog) as > 0 | 
|         |   1341       \<Longrightarrow> (fetch (ci (layout_of aprog)  | 
|         |   1342                (start_of (layout_of aprog) as) (Inc n)) (7 + 2*n) Oc) | 
|         |   1343       = (L, start_of (layout_of aprog) as + 2 * n + 6)" | 
|         |   1344 apply(auto simp: ci.simps findnth.simps tshift.simps  | 
|         |   1345                  tinc_b_def add3_Suc) | 
|         |   1346 apply(subgoal_tac "7 + 2*n = Suc (2*n + 6)", simp only: fetch.simps) | 
|         |   1347 apply(auto simp: nth_of.simps nth_append) | 
|         |   1348 done | 
|         |   1349  | 
|         |   1350 lemma fetch_left_move_b: " | 
|         |   1351       start_of (layout_of aprog) as > 0 | 
|         |   1352       \<Longrightarrow> (fetch (ci (layout_of aprog)  | 
|         |   1353                (start_of (layout_of aprog) as) (Inc n)) (7 + 2*n) Bk) | 
|         |   1354       = (L, start_of (layout_of aprog) as + 2 * n + 7)" | 
|         |   1355 apply(auto simp: ci.simps findnth.simps  | 
|         |   1356                  tshift.simps tinc_b_def add3_Suc) | 
|         |   1357 apply(subgoal_tac "7 + 2*n = Suc (2*n + 6)", simp only: fetch.simps) | 
|         |   1358 apply(auto simp: nth_of.simps nth_append) | 
|         |   1359 done | 
|         |   1360  | 
|         |   1361 lemma fetch_check_left_move_o: " | 
|         |   1362       start_of (layout_of aprog) as > 0 | 
|         |   1363       \<Longrightarrow> (fetch (ci (layout_of aprog)  | 
|         |   1364                (start_of (layout_of aprog) as) (Inc n)) (8 + 2*n) Oc) | 
|         |   1365       = (L, start_of (layout_of aprog) as + 2 * n + 6)" | 
|         |   1366 apply(auto simp: ci.simps findnth.simps tshift.simps tinc_b_def) | 
|         |   1367 apply(subgoal_tac "8 + 2 * n = Suc (2 * n + 7)",  | 
|         |   1368                                   simp only: fetch.simps) | 
|         |   1369 apply(auto simp: nth_of.simps nth_append) | 
|         |   1370 done | 
|         |   1371  | 
|         |   1372 lemma fetch_check_left_move_b: " | 
|         |   1373       start_of (layout_of aprog) as > 0 | 
|         |   1374       \<Longrightarrow> (fetch (ci (layout_of aprog)  | 
|         |   1375               (start_of (layout_of aprog) as) (Inc n)) (8 + 2*n) Bk) | 
|         |   1376       = (R, start_of (layout_of aprog) as + 2 * n + 8)  " | 
|         |   1377 apply(auto simp: ci.simps findnth.simps  | 
|         |   1378                  tshift.simps tinc_b_def add3_Suc) | 
|         |   1379 apply(subgoal_tac "8 + 2*n= Suc (2*n + 7)", simp only: fetch.simps) | 
|         |   1380 apply(auto simp: nth_of.simps nth_append) | 
|         |   1381 done | 
|         |   1382  | 
|         |   1383 lemma fetch_after_left_move: " | 
|         |   1384       start_of (layout_of aprog) as > 0 | 
|         |   1385       \<Longrightarrow> (fetch (ci (layout_of aprog)  | 
|         |   1386               (start_of (layout_of aprog) as) (Inc n)) (9 + 2*n) Bk) | 
|         |   1387      = (R, start_of (layout_of aprog) as + 2 * n + 9)" | 
|         |   1388 apply(auto simp: ci.simps findnth.simps fetch.simps | 
|         |   1389                   nth_of.simps tshift.simps nth_append tinc_b_def) | 
|         |   1390 done | 
|         |   1391  | 
|         |   1392 lemma fetch_stop: " | 
|         |   1393        start_of (layout_of aprog) as > 0 | 
|         |   1394       \<Longrightarrow> (fetch (ci (layout_of aprog)  | 
|         |   1395              (start_of (layout_of aprog) as) (Inc n)) (10 + 2 *n)  b) | 
|         |   1396      = (Nop, 0)" | 
|         |   1397 apply(auto simp: ci.simps findnth.simps fetch.simps | 
|         |   1398                   nth_of.simps tshift.simps nth_append tinc_b_def | 
|         |   1399             split: block.splits) | 
|         |   1400 done | 
|         |   1401  | 
|         |   1402 lemma fetch_state0: " | 
|         |   1403        (fetch (ci (layout_of aprog)  | 
|         |   1404                (start_of (layout_of aprog) as) (Inc n)) 0 b) | 
|         |   1405      = (Nop, 0)" | 
|         |   1406 apply(auto simp: ci.simps findnth.simps fetch.simps | 
|         |   1407                   nth_of.simps tshift.simps nth_append tinc_b_def) | 
|         |   1408 done | 
|         |   1409  | 
|         |   1410 lemmas fetch_simps =  | 
|         |   1411   fetch_locate_a_o fetch_locate_a_b fetch_locate_b_o fetch_locate_b_b  | 
|         |   1412   fetch_locate_n_a_b fetch_locate_n_a_o fetch_locate_n_b_o  | 
|         |   1413   fetch_locate_n_b_b fetch_after_write_o fetch_after_move_o  | 
|         |   1414   fetch_after_move_b fetch_clear_b fetch_right_move_o  | 
|         |   1415   fetch_right_move_b fetch_left_move_o fetch_left_move_b | 
|         |   1416   fetch_after_left_move fetch_check_left_move_o fetch_stop  | 
|         |   1417   fetch_state0 fetch_check_left_move_b | 
|         |   1418  | 
|         |   1419 text {* *} | 
|         |   1420 declare exponent_def[simp del] tape_of_nat_list.simps[simp del]  | 
|         |   1421    at_begin_norm.simps[simp del] at_begin_fst_bwtn.simps[simp del]  | 
|         |   1422    at_begin_fst_awtn.simps[simp del] in_middle.simps[simp del]  | 
|         |   1423    abc_lm_s.simps[simp del] abc_lm_v.simps[simp del]   | 
|         |   1424    ci.simps[simp del] t_step.simps[simp del] | 
|         |   1425    inv_after_move.simps[simp del]  | 
|         |   1426    inv_on_left_moving_norm.simps[simp del]  | 
|         |   1427    inv_on_left_moving_in_middle_B.simps[simp del] | 
|         |   1428    inv_after_clear.simps[simp del]  | 
|         |   1429    inv_after_write.simps[simp del] inv_on_left_moving.simps[simp del] | 
|         |   1430    inv_on_right_moving.simps[simp del]  | 
|         |   1431    inv_check_left_moving.simps[simp del]  | 
|         |   1432    inv_check_left_moving_in_middle.simps[simp del] | 
|         |   1433    inv_check_left_moving_on_leftmost.simps[simp del]  | 
|         |   1434    inv_after_left_moving.simps[simp del] | 
|         |   1435    inv_stop.simps[simp del] inv_locate_a.simps[simp del]  | 
|         |   1436    inv_locate_b.simps[simp del] | 
|         |   1437 declare tms_of.simps[simp del] tm_of.simps[simp del] | 
|         |   1438         layout_of.simps[simp del] abc_fetch.simps [simp del]  | 
|         |   1439         t_step.simps[simp del] t_steps.simps[simp del]  | 
|         |   1440         tpairs_of.simps[simp del] start_of.simps[simp del] | 
|         |   1441         fetch.simps [simp del] new_tape.simps [simp del]  | 
|         |   1442         nth_of.simps [simp del] ci.simps [simp del] | 
|         |   1443         length_of.simps[simp del] | 
|         |   1444  | 
|         |   1445 (*! Start point *) | 
|         |   1446 lemma [simp]: "Suc (2 * q) mod 2 = Suc 0" | 
|         |   1447 by arith | 
|         |   1448  | 
|         |   1449 lemma [simp]: "Suc (2 * q) div 2 = q" | 
|         |   1450 by arith | 
|         |   1451  | 
|         |   1452 lemma [simp]: "\<lbrakk> \<not> a < start_of ly as;  | 
|         |   1453           a < start_of ly as + 2 * n; a - start_of ly as = 2 * q\<rbrakk> | 
|         |   1454              \<Longrightarrow> Suc a < start_of ly as + 2 * n" | 
|         |   1455 apply(arith) | 
|         |   1456 done | 
|         |   1457  | 
|         |   1458 lemma [simp]: "x mod 2 = Suc 0 \<Longrightarrow> (Suc x) mod 2 = 0" | 
|         |   1459 by arith | 
|         |   1460  | 
|         |   1461 lemma [simp]: "x mod 2 = Suc 0 \<Longrightarrow> (Suc x) div 2 = Suc (x div 2)" | 
|         |   1462 by arith | 
|         |   1463 lemma exp_def[simp]: "a\<^bsup>Suc n \<^esup>= a # a\<^bsup>n\<^esup>" | 
|         |   1464 by(simp add: exponent_def) | 
|         |   1465 lemma [intro]: "Bk # r = Oc\<^bsup>mr\<^esup> @ r' \<Longrightarrow> mr = 0" | 
|         |   1466 by(case_tac mr, auto simp: exponent_def) | 
|         |   1467  | 
|         |   1468 lemma [intro]: "Bk # r = replicate mr Oc \<Longrightarrow> mr = 0" | 
|         |   1469 by(case_tac mr, auto) | 
|         |   1470 lemma tape_of_nl_abv_cons[simp]: "xs \<noteq> [] \<Longrightarrow>  | 
|         |   1471                    <x # xs> = Oc\<^bsup>Suc x\<^esup>@ Bk # <xs>" | 
|         |   1472 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) | 
|         |   1473 apply(case_tac xs, simp, simp add: tape_of_nat_list.simps) | 
|         |   1474 done | 
|         |   1475  | 
|         |   1476 lemma [simp]: "<[]::nat list> = []" | 
|         |   1477 by(auto simp: tape_of_nl_abv tape_of_nat_list.simps) | 
|         |   1478 lemma [simp]: "Oc # r = <(lm::nat list)> @ Bk\<^bsup>rn\<^esup>\<Longrightarrow> lm \<noteq> []" | 
|         |   1479 apply(auto simp: tape_of_nl_abv tape_of_nat_list.simps) | 
|         |   1480 apply(case_tac rn, auto simp: exponent_def) | 
|         |   1481 done | 
|         |   1482 lemma BkCons_nil: "Bk # xs = <lm::nat list> @ Bk\<^bsup>rn\<^esup>\<Longrightarrow> lm = []" | 
|         |   1483 apply(case_tac lm, simp) | 
|         |   1484 apply(case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps) | 
|         |   1485 done | 
|         |   1486 lemma BkCons_nil': "Bk # xs = <lm::nat list> @ Bk\<^bsup>ln\<^esup>\<Longrightarrow> lm = []" | 
|         |   1487 by(auto intro: BkCons_nil) | 
|         |   1488  | 
|         |   1489 lemma hd_tl_tape_of_nat_list:   | 
|         |   1490    "tl (lm::nat list) \<noteq> [] \<Longrightarrow> <lm> = <hd lm> @ Bk # <tl lm>" | 
|         |   1491 apply(frule tape_of_nl_abv_cons[of "tl lm" "hd lm"]) | 
|         |   1492 apply(simp add: tape_of_nat_abv Bk_def del: tape_of_nl_abv_cons) | 
|         |   1493 apply(subgoal_tac "lm = hd lm # tl lm", auto) | 
|         |   1494 apply(case_tac lm, auto) | 
|         |   1495 done | 
|         |   1496 lemma [simp]: "Oc # xs = Oc\<^bsup>mr\<^esup> @ Bk # <lm2> @ Bk\<^bsup>rn\<^esup>\<Longrightarrow> mr > 0" | 
|         |   1497 apply(case_tac mr, auto simp: exponent_def) | 
|         |   1498 done | 
|         |   1499  | 
|         |   1500 lemma tape_of_nat_list_cons: "xs \<noteq> [] \<Longrightarrow> tape_of_nat_list (x # xs) = | 
|         |   1501               replicate (Suc x) Oc @ Bk # tape_of_nat_list xs" | 
|         |   1502 apply(drule tape_of_nl_abv_cons[of xs x]) | 
|         |   1503 apply(auto simp: tape_of_nl_abv tape_of_nat_abv Oc_def Bk_def exponent_def) | 
|         |   1504 done | 
|         |   1505  | 
|         |   1506 lemma rev_eq: "rev xs = rev ys \<Longrightarrow> xs = ys" | 
|         |   1507 by simp | 
|         |   1508  | 
|         |   1509 lemma tape_of_nat_list_eq: " xs = ys \<Longrightarrow>  | 
|         |   1510         tape_of_nat_list xs = tape_of_nat_list ys" | 
|         |   1511 by simp | 
|         |   1512  | 
|         |   1513 lemma tape_of_nl_nil_eq: "<(lm::nat list)> = [] = (lm = [])" | 
|         |   1514 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) | 
|         |   1515 apply(case_tac lm, simp add: tape_of_nat_list.simps) | 
|         |   1516 apply(case_tac "list") | 
|         |   1517 apply(auto simp: tape_of_nat_list.simps) | 
|         |   1518 done | 
|         |   1519  | 
|         |   1520 lemma rep_ind: "replicate (Suc n) a = replicate n a @ [a]" | 
|         |   1521 apply(induct n, simp, simp) | 
|         |   1522 done | 
|         |   1523  | 
|         |   1524 lemma [simp]: "Oc # r = <lm::nat list> @ replicate rn Bk \<Longrightarrow> Suc 0 \<le> length lm" | 
|         |   1525 apply(rule_tac classical, auto) | 
|         |   1526 apply(case_tac lm, simp, case_tac rn, auto) | 
|         |   1527 done | 
|         |   1528 lemma Oc_Bk_Cons: "Oc # Bk # list = <lm::nat list> @ Bk\<^bsup>ln\<^esup> \<Longrightarrow>  | 
|         |   1529                    lm \<noteq> [] \<and> hd lm = 0" | 
|         |   1530 apply(case_tac lm, simp, case_tac ln, simp add: exponent_def, simp add: exponent_def, simp) | 
|         |   1531 apply(case_tac lista, auto simp: tape_of_nl_abv tape_of_nat_list.simps) | 
|         |   1532 done  | 
|         |   1533 (*lemma Oc_Oc_Cons: "Oc # Oc # list = <lm::nat list> @ Bk\<^bsup>ln\<^esup> \<Longrightarrow>  | 
|         |   1534                   lm \<noteq> [] \<and> hd lm > 0" | 
|         |   1535 apply(case_tac lm, simp add: exponent_def, case_tac ln, simp, simp) | 
|         |   1536 apply(case_tac lista,  | 
|         |   1537         auto simp: tape_of_nl_abv tape_of_nat_list.simps exponent_def) | 
|         |   1538 apply(case_tac [!] a, auto) | 
|         |   1539 apply(case_tac ln, auto) | 
|         |   1540 done | 
|         |   1541 *) | 
|         |   1542 lemma Oc_nil_zero[simp]: "[Oc] = <lm::nat list> @ Bk\<^bsup>ln\<^esup>  | 
|         |   1543                  \<Longrightarrow> lm = [0] \<and> ln = 0" | 
|         |   1544 apply(case_tac lm, simp) | 
|         |   1545 apply(case_tac ln, auto simp: exponent_def) | 
|         |   1546 apply(case_tac [!] list,  | 
|         |   1547         auto simp: tape_of_nl_abv tape_of_nat_list.simps) | 
|         |   1548 done | 
|         |   1549  | 
|         |   1550 lemma  [simp]: "Oc # r = <lm2> @ replicate rn Bk \<Longrightarrow>  | 
|         |   1551        (\<exists>rn. r = replicate (hd lm2) Oc @ Bk # <tl lm2> @  | 
|         |   1552                       replicate rn Bk) \<or>  | 
|         |   1553           tl lm2 = [] \<and> r = replicate (hd lm2) Oc" | 
|         |   1554 apply(rule_tac disjCI, simp) | 
|         |   1555 apply(case_tac "tl lm2 = []", simp) | 
|         |   1556 apply(case_tac lm2, simp add: tape_of_nl_abv tape_of_nat_list.simps) | 
|         |   1557 apply(case_tac rn, simp, simp, simp) | 
|         |   1558 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps exponent_def) | 
|         |   1559 apply(case_tac rn, simp, simp) | 
|         |   1560 apply(rule_tac x = rn in exI) | 
|         |   1561 apply(simp add: hd_tl_tape_of_nat_list) | 
|         |   1562 apply(simp add: tape_of_nat_abv Oc_def exponent_def) | 
|         |   1563 done | 
|         |   1564  | 
|         |   1565 (*inv: from locate_a to locate_b*) | 
|         |   1566 lemma [simp]:  | 
|         |   1567       "inv_locate_a (as, lm) (q, l, Oc # r) ires | 
|         |   1568        \<Longrightarrow> inv_locate_b (as, lm) (q, Oc # l, r) ires" | 
|         |   1569 apply(simp only: inv_locate_a.simps inv_locate_b.simps in_middle.simps | 
|         |   1570           at_begin_norm.simps at_begin_fst_bwtn.simps | 
|         |   1571           at_begin_fst_awtn.simps) | 
|         |   1572 apply(erule disjE, erule exE, erule exE, erule exE) | 
|         |   1573 apply(rule_tac x = lm1 in exI, rule_tac x = "tl lm2" in exI, simp) | 
|         |   1574 apply(rule_tac x = "0" in exI, rule_tac x = "hd lm2" in exI,  | 
|         |   1575                 auto simp: exponent_def) | 
|         |   1576 apply(rule_tac x = "Suc 0" in exI, simp add:exponent_def) | 
|         |   1577 apply(rule_tac x = "lm @ replicate tn 0" in exI,  | 
|         |   1578       rule_tac x = "[]" in exI,     | 
|         |   1579       rule_tac x = "Suc tn" in exI, rule_tac x = 0 in exI) | 
|         |   1580 apply(simp only: rep_ind, simp) | 
|         |   1581 apply(rule_tac x = "Suc 0" in exI, auto) | 
|         |   1582 apply(case_tac [1-3] rn, simp_all ) | 
|         |   1583 apply(rule_tac x = "lm @ replicate tn 0" in exI,  | 
|         |   1584       rule_tac x = "[]" in exI,  | 
|         |   1585       rule_tac x = "Suc tn" in exI,  | 
|         |   1586       rule_tac x = 0 in exI, simp add: rep_ind del: replicate_Suc split:if_splits) | 
|         |   1587 apply(rule_tac x = "Suc 0" in exI, auto) | 
|         |   1588 apply(case_tac rn, simp, simp) | 
|         |   1589 apply(rule_tac [!] x = "Suc 0" in exI, auto) | 
|         |   1590 apply(case_tac [!] rn, simp_all) | 
|         |   1591 done | 
|         |   1592  | 
|         |   1593 (*inv: from locate_a to _locate_a*) | 
|         |   1594 lemma locate_a_2_locate_a[simp]: "inv_locate_a (as, am) (q, aaa, Bk # xs) ires | 
|         |   1595        \<Longrightarrow> inv_locate_a (as, am) (q, aaa, Oc # xs) ires" | 
|         |   1596 apply(simp only: inv_locate_a.simps at_begin_norm.simps  | 
|         |   1597                  at_begin_fst_bwtn.simps at_begin_fst_awtn.simps) | 
|         |   1598 apply(erule_tac disjE, erule exE, erule exE, erule exE,  | 
|         |   1599       rule disjI2, rule disjI2) | 
|         |   1600 defer | 
|         |   1601 apply(erule_tac disjE, erule exE, erule exE,  | 
|         |   1602       erule exE, rule disjI2, rule disjI2) | 
|         |   1603 prefer 2 | 
|         |   1604 apply(simp) | 
|         |   1605 proof- | 
|         |   1606   fix lm1 tn rn | 
|         |   1607   assume k: "lm1 = am @ 0\<^bsup>tn\<^esup> \<and> length lm1 = q \<and> (if lm1 = [] then aaa = Bk # Bk #  | 
|         |   1608     ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> Bk # xs = Bk\<^bsup>rn\<^esup>" | 
|         |   1609   thus "\<exists>lm1 tn rn. lm1 = am @ 0\<^bsup>tn\<^esup> \<and> length lm1 = q \<and> (if lm1 = [] then  | 
|         |   1610     aaa = Bk # Bk # ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> Oc # xs = [Oc] @ Bk\<^bsup>rn\<^esup>" | 
|         |   1611     (is "\<exists>lm1 tn rn. ?P lm1 tn rn") | 
|         |   1612   proof - | 
|         |   1613     from k have "?P lm1 tn (rn - 1)" | 
|         |   1614       apply(auto simp: Oc_def) | 
|         |   1615       by(case_tac [!] "rn::nat", auto simp: exponent_def) | 
|         |   1616     thus ?thesis by blast | 
|         |   1617   qed | 
|         |   1618 next | 
|         |   1619   fix lm1 lm2 rn | 
|         |   1620   assume h1: "am = lm1 @ lm2 \<and> length lm1 = q \<and> (if lm1 = []  | 
|         |   1621     then aaa = Bk # Bk # ires else aaa = Bk # <rev lm1> @ Bk # Bk # ires) \<and> | 
|         |   1622     Bk # xs = <lm2> @ Bk\<^bsup>rn\<^esup>" | 
|         |   1623   from h1 have h2: "lm2 = []" | 
|         |   1624   proof(rule_tac xs = xs and rn = rn in BkCons_nil, simp) | 
|         |   1625   qed | 
|         |   1626   from h1 and h2 show "\<exists>lm1 tn rn. lm1 = am @ 0\<^bsup>tn\<^esup> \<and> length lm1 = q \<and>  | 
|         |   1627     (if lm1 = [] then aaa = Bk # Bk # ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> | 
|         |   1628     Oc # xs = [Oc] @ Bk\<^bsup>rn\<^esup>"  | 
|         |   1629     (is "\<exists>lm1 tn rn. ?P lm1 tn rn") | 
|         |   1630   proof - | 
|         |   1631     from h1 and h2  have "?P lm1 0 (rn - 1)" | 
|         |   1632       apply(auto simp: Oc_def exponent_def  | 
|         |   1633                       tape_of_nl_abv tape_of_nat_list.simps) | 
|         |   1634       by(case_tac "rn::nat", simp, simp) | 
|         |   1635     thus ?thesis by blast | 
|         |   1636   qed | 
|         |   1637 qed | 
|         |   1638  | 
|         |   1639 lemma [intro]: "\<exists>rn. [a] = a\<^bsup>rn\<^esup>" | 
|         |   1640 by(rule_tac x = "Suc 0" in exI, simp add: exponent_def) | 
|         |   1641  | 
|         |   1642 lemma [intro]: "\<exists>tn. [] = a\<^bsup>tn\<^esup>" | 
|         |   1643 apply(rule_tac x = 0 in exI, simp add: exponent_def) | 
|         |   1644 done | 
|         |   1645  | 
|         |   1646 lemma [intro]:  "at_begin_norm (as, am) (q, aaa, []) ires | 
|         |   1647              \<Longrightarrow> at_begin_norm (as, am) (q, aaa, [Bk]) ires" | 
|         |   1648 apply(simp add: at_begin_norm.simps, erule_tac exE, erule_tac exE) | 
|         |   1649 apply(rule_tac x = lm1 in exI, simp, auto) | 
|         |   1650 done | 
|         |   1651  | 
|         |   1652 lemma [intro]: "at_begin_fst_bwtn (as, am) (q, aaa, []) ires  | 
|         |   1653             \<Longrightarrow> at_begin_fst_bwtn (as, am) (q, aaa, [Bk]) ires" | 
|         |   1654 apply(simp only: at_begin_fst_bwtn.simps, erule_tac exE, erule_tac exE, erule_tac exE) | 
|         |   1655 apply(rule_tac x = "am @ 0\<^bsup>tn\<^esup>" in exI, auto) | 
|         |   1656 done | 
|         |   1657  | 
|         |   1658 lemma [intro]: "at_begin_fst_awtn (as, am) (q, aaa, []) ires | 
|         |   1659            \<Longrightarrow> at_begin_fst_awtn (as, am) (q, aaa, [Bk]) ires" | 
|         |   1660 apply(auto simp: at_begin_fst_awtn.simps) | 
|         |   1661 done  | 
|         |   1662  | 
|         |   1663 lemma [intro]: "inv_locate_a (as, am) (q, aaa, []) ires | 
|         |   1664             \<Longrightarrow> inv_locate_a (as, am) (q, aaa, [Bk]) ires" | 
|         |   1665 apply(simp only: inv_locate_a.simps) | 
|         |   1666 apply(erule disj_forward) | 
|         |   1667 defer | 
|         |   1668 apply(erule disj_forward, auto) | 
|         |   1669 done | 
|         |   1670  | 
|         |   1671 lemma [simp]: "inv_locate_a (as, am) (q, aaa, []) ires \<Longrightarrow>  | 
|         |   1672                inv_locate_a (as, am) (q, aaa, [Oc]) ires" | 
|         |   1673 apply(insert locate_a_2_locate_a [of as am q aaa "[]"]) | 
|         |   1674 apply(subgoal_tac "inv_locate_a (as, am) (q, aaa, [Bk]) ires", auto) | 
|         |   1675 done | 
|         |   1676  | 
|         |   1677 (*inv: from locate_b to locate_b*) | 
|         |   1678 lemma [simp]: "inv_locate_b (as, am) (q, aaa, Oc # xs) ires | 
|         |   1679          \<Longrightarrow> inv_locate_b (as, am) (q, Oc # aaa, xs) ires" | 
|         |   1680 apply(simp only: inv_locate_b.simps in_middle.simps) | 
|         |   1681 apply(erule exE)+ | 
|         |   1682 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,  | 
|         |   1683       rule_tac x = tn in exI, rule_tac x = m in exI) | 
|         |   1684 apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - 1" in exI, | 
|         |   1685       rule_tac x = rn in exI) | 
|         |   1686 apply(case_tac mr, simp_all add: exponent_def, auto) | 
|         |   1687 done | 
|         |   1688 lemma zero_and_nil[intro]: "(Bk # Bk\<^bsup>n\<^esup> = Oc\<^bsup>mr\<^esup> @ Bk # <lm::nat list> @  | 
|         |   1689                              Bk\<^bsup>rn \<^esup>) \<or> (lm2 = [] \<and> Bk # Bk\<^bsup>n\<^esup> = Oc\<^bsup>mr\<^esup>) | 
|         |   1690        \<Longrightarrow> mr = 0 \<and> lm = []" | 
|         |   1691 apply(rule context_conjI) | 
|         |   1692 apply(case_tac mr, auto simp:exponent_def) | 
|         |   1693 apply(insert BkCons_nil[of "replicate (n - 1) Bk" lm rn]) | 
|         |   1694 apply(case_tac n, auto simp: exponent_def Bk_def  tape_of_nl_nil_eq) | 
|         |   1695 done | 
|         |   1696  | 
|         |   1697 lemma tape_of_nat_def: "<[m::nat]> =  Oc # Oc\<^bsup>m\<^esup>" | 
|         |   1698 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) | 
|         |   1699 done | 
|         |   1700 lemma [simp]: "\<lbrakk>inv_locate_b (as, am) (q, aaa, Bk # xs) ires; \<exists>n. xs = Bk\<^bsup>n\<^esup>\<rbrakk> | 
|         |   1701             \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires" | 
|         |   1702 apply(simp add: inv_locate_b.simps inv_locate_a.simps) | 
|         |   1703 apply(rule_tac disjI2, rule_tac disjI1) | 
|         |   1704 apply(simp only: in_middle.simps at_begin_fst_bwtn.simps) | 
|         |   1705 apply(erule_tac exE)+ | 
|         |   1706 apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = tn in exI, simp) | 
|         |   1707 apply(subgoal_tac "mr = 0 \<and> lm2 = []") | 
|         |   1708 defer | 
|         |   1709 apply(rule_tac n = n and mr = mr and lm = "lm2"  | 
|         |   1710                and rn = rn and n = n in zero_and_nil) | 
|         |   1711 apply(auto simp: exponent_def) | 
|         |   1712 apply(case_tac "lm1 = []", auto simp: tape_of_nat_def) | 
|         |   1713 done | 
|         |   1714  | 
|         |   1715 lemma length_equal: "xs = ys \<Longrightarrow> length xs = length ys" | 
|         |   1716 by auto | 
|         |   1717 lemma [simp]: "a\<^bsup>0\<^esup> = []"  | 
|         |   1718 by(simp add: exp_zero) | 
|         |   1719 (*inv: from locate_b to locate_a*) | 
|         |   1720 lemma [simp]: "length (a\<^bsup>b\<^esup>) = b" | 
|         |   1721 apply(simp add: exponent_def) | 
|         |   1722 done | 
|         |   1723  | 
|         |   1724 lemma [simp]: "\<lbrakk>inv_locate_b (as, am) (q, aaa, Bk # xs) ires;  | 
|         |   1725                 \<not> (\<exists>n. xs = Bk\<^bsup>n\<^esup>)\<rbrakk>  | 
|         |   1726        \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires" | 
|         |   1727 apply(simp add: inv_locate_b.simps inv_locate_a.simps) | 
|         |   1728 apply(rule_tac disjI1) | 
|         |   1729 apply(simp only: in_middle.simps at_begin_norm.simps) | 
|         |   1730 apply(erule_tac exE)+ | 
|         |   1731 apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = lm2 in exI, simp) | 
|         |   1732 apply(subgoal_tac "tn = 0", simp add: exponent_def , auto split: if_splits) | 
|         |   1733 apply(case_tac [!] mr, simp_all add: tape_of_nat_def, auto) | 
|         |   1734 apply(case_tac lm2, simp, erule_tac x = rn in allE, simp) | 
|         |   1735 apply(case_tac am, simp, simp) | 
|         |   1736 apply(case_tac lm2, simp, erule_tac x = rn in allE, simp) | 
|         |   1737 apply(drule_tac length_equal, simp) | 
|         |   1738 done | 
|         |   1739  | 
|         |   1740 lemma locate_b_2_a[intro]:  | 
|         |   1741        "inv_locate_b (as, am) (q, aaa, Bk # xs) ires | 
|         |   1742     \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires" | 
|         |   1743 apply(case_tac "\<exists> n. xs = Bk\<^bsup>n\<^esup>", simp, simp) | 
|         |   1744 done | 
|         |   1745  | 
|         |   1746 lemma locate_b_2_locate_a[simp]:  | 
|         |   1747     "\<lbrakk>\<not> a < start_of ly as;  | 
|         |   1748       a < start_of ly as + 2 * n;  | 
|         |   1749       (a - start_of ly as) mod 2 = Suc 0;  | 
|         |   1750      inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, Bk # xs) ires\<rbrakk> | 
|         |   1751    \<Longrightarrow> (Suc a < start_of ly as + 2 * n \<longrightarrow> inv_locate_a (as, am) | 
|         |   1752        (Suc ((a - start_of ly as) div 2), Bk # aaa, xs) ires) \<and> | 
|         |   1753        (\<not> Suc a < start_of ly as + 2 * n \<longrightarrow>  | 
|         |   1754                 inv_locate_a (as, am) (n, Bk # aaa, xs) ires)" | 
|         |   1755 apply(auto) | 
|         |   1756 apply(subgoal_tac "n > 0") | 
|         |   1757 apply(subgoal_tac "(a - start_of ly as) div 2 = n - 1") | 
|         |   1758 apply(insert locate_b_2_a [of as am "n - 1" aaa xs], simp) | 
|         |   1759 apply(arith) | 
|         |   1760 apply(case_tac n, simp, simp) | 
|         |   1761 done | 
|         |   1762  | 
|         |   1763 lemma [simp]:  "inv_locate_b (as, am) (q, l, []) ires  | 
|         |   1764            \<Longrightarrow>  inv_locate_b (as, am) (q, l, [Bk]) ires" | 
|         |   1765 apply(simp only: inv_locate_b.simps in_middle.simps) | 
|         |   1766 apply(erule exE)+ | 
|         |   1767 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,  | 
|         |   1768       rule_tac x = tn in exI, rule_tac x = m in exI,  | 
|         |   1769       rule_tac x = ml in exI, rule_tac x = mr in exI) | 
|         |   1770 apply(auto) | 
|         |   1771 done | 
|         |   1772  | 
|         |   1773 lemma locate_b_2_locate_a_B[simp]:  | 
|         |   1774  "\<lbrakk>\<not> a < start_of ly as;  | 
|         |   1775    a < start_of ly as + 2 * n;  | 
|         |   1776    (a - start_of ly as) mod 2 = Suc 0;  | 
|         |   1777    inv_locate_b (as, am) ((a - start_of ly as) div 2, aaa, []) ires\<rbrakk> | 
|         |   1778    \<Longrightarrow> (Suc a < start_of ly as + 2 * n \<longrightarrow>  | 
|         |   1779      inv_locate_a (as, am)  | 
|         |   1780             (Suc ((a - start_of ly as) div 2), Bk # aaa, []) ires)  | 
|         |   1781     \<and> (\<not> Suc a < start_of ly as + 2 * n \<longrightarrow>  | 
|         |   1782                   inv_locate_a (as, am) (n, Bk # aaa, []) ires)" | 
|         |   1783 apply(insert locate_b_2_locate_a [of a ly as n am aaa "[]"], simp) | 
|         |   1784 done | 
|         |   1785  | 
|         |   1786 (*inv: from locate_b to after_write*) | 
|         |   1787 lemma inv_locate_b_2_after_write[simp]:  | 
|         |   1788       "inv_locate_b (as, am) (n, aaa, Bk # xs) ires | 
|         |   1789       \<Longrightarrow> inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) | 
|         |   1790           (Suc (Suc (2 * n)), aaa, Oc # xs) ires" | 
|         |   1791 apply(auto simp: in_middle.simps inv_after_write.simps  | 
|         |   1792                  abc_lm_v.simps abc_lm_s.simps  inv_locate_b.simps) | 
|         |   1793 apply(subgoal_tac [!] "mr = 0", auto simp: exponent_def split: if_splits) | 
|         |   1794 apply(subgoal_tac "lm2 = []", simp) | 
|         |   1795 apply(rule_tac x = rn in exI, rule_tac x = "Suc m" in exI, | 
|         |   1796       rule_tac x = "lm1" in exI, simp, rule_tac x = "[]" in exI, simp) | 
|         |   1797 apply(case_tac "Suc (length lm1) - length am", simp, simp only: rep_ind, simp) | 
|         |   1798 apply(subgoal_tac "length lm1 - length am = nat", simp, arith) | 
|         |   1799 apply(drule_tac length_equal, simp) | 
|         |   1800 done | 
|         |   1801  | 
|         |   1802 lemma [simp]: "inv_locate_b (as, am) (n, aaa, []) ires \<Longrightarrow>  | 
|         |   1803      inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n)))  | 
|         |   1804                      (Suc (Suc (2 * n)), aaa, [Oc]) ires" | 
|         |   1805 apply(insert inv_locate_b_2_after_write [of as am n aaa "[]"]) | 
|         |   1806 by(simp) | 
|         |   1807  | 
|         |   1808 (*inv: from after_write to after_move*) | 
|         |   1809 lemma [simp]: "inv_after_write (as, lm) (Suc (Suc (2 * n)), l, Oc # r) ires | 
|         |   1810                 \<Longrightarrow> inv_after_move (as, lm) (2 * n + 3, Oc # l, r) ires" | 
|         |   1811 apply(auto simp:inv_after_move.simps inv_after_write.simps split: if_splits) | 
|         |   1812 done | 
|         |   1813  | 
|         |   1814 lemma [simp]: "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n) | 
|         |   1815                 )) (Suc (Suc (2 * n)), aaa, Bk # xs) ires = False" | 
|         |   1816 apply(simp add: inv_after_write.simps ) | 
|         |   1817 done | 
|         |   1818  | 
|         |   1819 lemma [simp]:  | 
|         |   1820  "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n)))  | 
|         |   1821                         (Suc (Suc (2 * n)), aaa, []) ires = False" | 
|         |   1822 apply(simp add: inv_after_write.simps ) | 
|         |   1823 done | 
|         |   1824  | 
|         |   1825 (*inv: from after_move to after_clear*) | 
|         |   1826 lemma [simp]: "inv_after_move (as, lm) (s, l, Oc # r) ires | 
|         |   1827                 \<Longrightarrow> inv_after_clear (as, lm) (s', l, Bk # r) ires" | 
|         |   1828 apply(auto simp: inv_after_move.simps inv_after_clear.simps split: if_splits) | 
|         |   1829 done | 
|         |   1830  | 
|         |   1831 (*inv: from after_move to on_leftmoving*) | 
|         |   1832 lemma inv_after_move_2_inv_on_left_moving[simp]:   | 
|         |   1833    "inv_after_move (as, lm) (s, l, Bk # r) ires | 
|         |   1834    \<Longrightarrow> (l = [] \<longrightarrow>  | 
|         |   1835          inv_on_left_moving (as, lm) (s', [], Bk # Bk # r) ires) \<and> | 
|         |   1836       (l \<noteq> [] \<longrightarrow>  | 
|         |   1837          inv_on_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)" | 
|         |   1838 apply(simp only: inv_after_move.simps inv_on_left_moving.simps) | 
|         |   1839 apply(subgoal_tac "l \<noteq> []", rule conjI, simp, rule impI,  | 
|         |   1840                 rule disjI1, simp only: inv_on_left_moving_norm.simps) | 
|         |   1841 apply(erule exE)+ | 
|         |   1842 apply(subgoal_tac "lm2 = []") | 
|         |   1843 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,   | 
|         |   1844     rule_tac x = m in exI, rule_tac x = m in exI,  | 
|         |   1845     rule_tac x = 1 in exI,   | 
|         |   1846     rule_tac x = "rn - 1" in exI, simp, case_tac rn) | 
|         |   1847 apply(auto simp: exponent_def  intro: BkCons_nil split: if_splits) | 
|         |   1848 done | 
|         |   1849  | 
|         |   1850 lemma [elim]: "[] = <lm::nat list> \<Longrightarrow> lm = []" | 
|         |   1851 using tape_of_nl_nil_eq[of lm] | 
|         |   1852 by simp | 
|         |   1853  | 
|         |   1854 lemma inv_after_move_2_inv_on_left_moving_B[simp]:  | 
|         |   1855     "inv_after_move (as, lm) (s, l, []) ires | 
|         |   1856       \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s', [], [Bk]) ires) \<and> | 
|         |   1857           (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s', tl l, [hd l]) ires)" | 
|         |   1858 apply(simp only: inv_after_move.simps inv_on_left_moving.simps) | 
|         |   1859 apply(subgoal_tac "l \<noteq> []", rule conjI, simp, rule impI, rule disjI1, | 
|         |   1860         simp only: inv_on_left_moving_norm.simps) | 
|         |   1861 apply(erule exE)+ | 
|         |   1862 apply(subgoal_tac "lm2 = []") | 
|         |   1863 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,   | 
|         |   1864       rule_tac x = m in exI, rule_tac x = m in exI,  | 
|         |   1865       rule_tac x = 1 in exI, rule_tac x = "rn - 1" in exI, simp, case_tac rn) | 
|         |   1866 apply(auto simp: exponent_def  tape_of_nl_nil_eq  intro: BkCons_nil  split: if_splits) | 
|         |   1867 done | 
|         |   1868  | 
|         |   1869 (*inv: from after_clear to on_right_moving*) | 
|         |   1870 lemma [simp]: "Oc # r = replicate rn Bk = False" | 
|         |   1871 apply(case_tac rn, simp, simp) | 
|         |   1872 done | 
|         |   1873  | 
|         |   1874 lemma inv_after_clear_2_inv_on_right_moving[simp]:  | 
|         |   1875      "inv_after_clear (as, lm) (2 * n + 4, l, Bk # r) ires | 
|         |   1876       \<Longrightarrow> inv_on_right_moving (as, lm) (2 * n + 5, Bk # l, r) ires" | 
|         |   1877 apply(auto simp: inv_after_clear.simps inv_on_right_moving.simps ) | 
|         |   1878 apply(subgoal_tac "lm2 \<noteq> []") | 
|         |   1879 apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI,  | 
|         |   1880       rule_tac x = "hd lm2" in exI, simp) | 
|         |   1881 apply(rule_tac x = 0 in exI, rule_tac x = "hd lm2" in exI) | 
|         |   1882 apply(simp add: exponent_def, rule conjI) | 
|         |   1883 apply(case_tac [!] "lm2::nat list", auto simp: exponent_def) | 
|         |   1884 apply(case_tac rn, auto split: if_splits simp: tape_of_nat_def) | 
|         |   1885 apply(case_tac list,  | 
|         |   1886      simp add:  tape_of_nl_abv tape_of_nat_list.simps exponent_def) | 
|         |   1887 apply(erule_tac x = "rn - 1" in allE,  | 
|         |   1888       case_tac rn, auto simp: exponent_def) | 
|         |   1889 apply(case_tac list,  | 
|         |   1890      simp add:  tape_of_nl_abv tape_of_nat_list.simps exponent_def) | 
|         |   1891 apply(erule_tac x = "rn - 1" in allE,  | 
|         |   1892       case_tac rn, auto simp: exponent_def) | 
|         |   1893 done | 
|         |   1894  | 
|         |   1895  | 
|         |   1896 lemma [simp]: "inv_after_clear (as, lm) (2 * n + 4, l, []) ires\<Longrightarrow>  | 
|         |   1897                inv_after_clear (as, lm) (2 * n + 4, l, [Bk]) ires"  | 
|         |   1898 by(auto simp: inv_after_clear.simps) | 
|         |   1899  | 
|         |   1900 lemma [simp]: "inv_after_clear (as, lm) (2 * n + 4, l, []) ires | 
|         |   1901              \<Longrightarrow> inv_on_right_moving (as, lm) (2 * n + 5, Bk # l, []) ires" | 
|         |   1902 by(insert  | 
|         |   1903     inv_after_clear_2_inv_on_right_moving[of as lm n l "[]"], simp) | 
|         |   1904  | 
|         |   1905 (*inv: from on_right_moving to on_right_movign*) | 
|         |   1906 lemma [simp]: "inv_on_right_moving (as, lm) (2 * n + 5, l, Oc # r) ires | 
|         |   1907       \<Longrightarrow> inv_on_right_moving (as, lm) (2 * n + 5, Oc # l, r) ires" | 
|         |   1908 apply(auto simp: inv_on_right_moving.simps) | 
|         |   1909 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,  | 
|         |   1910            rule_tac x = "ml + mr" in exI, simp) | 
|         |   1911 apply(rule_tac x = "Suc ml" in exI,  | 
|         |   1912            rule_tac x = "mr - 1" in exI, simp) | 
|         |   1913 apply(case_tac mr, auto simp: exponent_def ) | 
|         |   1914 apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI,  | 
|         |   1915       rule_tac x = "ml + mr" in exI, simp) | 
|         |   1916 apply(rule_tac x = "Suc ml" in exI,  | 
|         |   1917       rule_tac x = "mr - 1" in exI, simp) | 
|         |   1918 apply(case_tac mr, auto split: if_splits simp: exponent_def) | 
|         |   1919 done | 
|         |   1920  | 
|         |   1921 lemma inv_on_right_moving_2_inv_on_right_moving[simp]:  | 
|         |   1922      "inv_on_right_moving (as, lm) (2 * n + 5, l, Bk # r) ires | 
|         |   1923      \<Longrightarrow> inv_after_write (as, lm) (Suc (Suc (2 * n)), l, Oc # r) ires" | 
|         |   1924 apply(auto simp: inv_on_right_moving.simps inv_after_write.simps ) | 
|         |   1925 apply(case_tac mr, auto simp: exponent_def split: if_splits) | 
|         |   1926 apply(case_tac [!] mr, simp_all) | 
|         |   1927 done | 
|         |   1928        | 
|         |   1929 lemma [simp]: "inv_on_right_moving (as, lm) (2 * n + 5, l, []) ires\<Longrightarrow>  | 
|         |   1930              inv_on_right_moving (as, lm) (2 * n + 5, l, [Bk]) ires" | 
|         |   1931 apply(auto simp: inv_on_right_moving.simps exponent_def) | 
|         |   1932 apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, simp) | 
|         |   1933 apply (rule_tac x = m in exI, auto split: if_splits simp: exponent_def) | 
|         |   1934 done | 
|         |   1935  | 
|         |   1936 (*inv: from on_right_moving to after_write*) | 
|         |   1937 lemma [simp]: "inv_on_right_moving (as, lm) (2 * n + 5, l, []) ires | 
|         |   1938        \<Longrightarrow> inv_after_write (as, lm) (Suc (Suc (2 * n)), l, [Oc]) ires" | 
|         |   1939 apply(rule_tac inv_on_right_moving_2_inv_on_right_moving, simp) | 
|         |   1940 done | 
|         |   1941  | 
|         |   1942 (*inv: from on_left_moving to on_left_moving*) | 
|         |   1943 lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm)  | 
|         |   1944                (s, l, Oc # r) ires = False" | 
|         |   1945 apply(auto simp: inv_on_left_moving_in_middle_B.simps ) | 
|         |   1946 done | 
|         |   1947  | 
|         |   1948 lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, l, Bk # r) ires  | 
|         |   1949              = False" | 
|         |   1950 apply(auto simp: inv_on_left_moving_norm.simps) | 
|         |   1951 apply(case_tac [!] mr, auto simp: ) | 
|         |   1952 done | 
|         |   1953  | 
|         |   1954 lemma [intro]: "\<exists>rna. Oc # Oc\<^bsup>m\<^esup> @ Bk # <lm> @ Bk\<^bsup>rn\<^esup> = <m # lm> @ Bk\<^bsup>rna\<^esup>" | 
|         |   1955 apply(case_tac lm, simp add: tape_of_nl_abv tape_of_nat_list.simps) | 
|         |   1956 apply(rule_tac x = "Suc rn" in exI, simp) | 
|         |   1957 apply(case_tac list, simp_all add: tape_of_nl_abv tape_of_nat_list.simps, auto) | 
|         |   1958 done | 
|         |   1959  | 
|         |   1960  | 
|         |   1961 lemma [simp]:  | 
|         |   1962   "\<lbrakk>inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires; | 
|         |   1963     hd l = Bk; l \<noteq> []\<rbrakk> \<Longrightarrow>  | 
|         |   1964      inv_on_left_moving_in_middle_B (as, lm) (s, tl l, Bk # Oc # r) ires" | 
|         |   1965 apply(case_tac l, simp, simp) | 
|         |   1966 apply(simp only: inv_on_left_moving_norm.simps  | 
|         |   1967                  inv_on_left_moving_in_middle_B.simps) | 
|         |   1968 apply(erule_tac exE)+ | 
|         |   1969 apply(rule_tac x = lm1 in exI, rule_tac x = "m # lm2" in exI, auto) | 
|         |   1970 apply(case_tac [!] ml, auto) | 
|         |   1971 apply(rule_tac [!] x = 0 in exI, simp_all add: tape_of_nl_abv tape_of_nat_list.simps) | 
|         |   1972 done | 
|         |   1973  | 
|         |   1974 lemma [simp]: "\<lbrakk>inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires;  | 
|         |   1975                 hd l = Oc; l \<noteq> []\<rbrakk> | 
|         |   1976             \<Longrightarrow> inv_on_left_moving_norm (as, lm)  | 
|         |   1977                                         (s, tl l, Oc # Oc # r) ires" | 
|         |   1978 apply(simp only: inv_on_left_moving_norm.simps) | 
|         |   1979 apply(erule exE)+ | 
|         |   1980 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,  | 
|         |   1981       rule_tac x = m in exI, rule_tac x = "ml - 1" in exI, | 
|         |   1982       rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI, simp) | 
|         |   1983 apply(case_tac ml, auto simp: exponent_def split: if_splits) | 
|         |   1984 done | 
|         |   1985  | 
|         |   1986 lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, [], Oc # r) ires | 
|         |   1987      \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # Oc # r) ires" | 
|         |   1988 apply(auto simp: inv_on_left_moving_norm.simps  | 
|         |   1989                  inv_on_left_moving_in_middle_B.simps split: if_splits) | 
|         |   1990 done | 
|         |   1991  | 
|         |   1992 lemma [simp]:"inv_on_left_moving (as, lm) (s, l, Oc # r) ires | 
|         |   1993     \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s, [], Bk # Oc # r) ires) | 
|         |   1994  \<and>  (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s, tl l, hd l # Oc # r) ires)" | 
|         |   1995 apply(simp add: inv_on_left_moving.simps) | 
|         |   1996 apply(case_tac "l \<noteq> []", rule conjI, simp, simp) | 
|         |   1997 apply(case_tac "hd l", simp, simp, simp) | 
|         |   1998 done | 
|         |   1999  | 
|         |   2000 (*inv: from on_left_moving to check_left_moving*) | 
|         |   2001 lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm)  | 
|         |   2002                                       (s, Bk # list, Bk # r) ires | 
|         |   2003           \<Longrightarrow> inv_check_left_moving_on_leftmost (as, lm)  | 
|         |   2004                                       (s', list, Bk # Bk # r) ires" | 
|         |   2005 apply(auto simp: inv_on_left_moving_in_middle_B.simps  | 
|         |   2006                  inv_check_left_moving_on_leftmost.simps split: if_splits) | 
|         |   2007 apply(case_tac [!] "rev lm1", simp_all) | 
|         |   2008 apply(case_tac [!] lista, simp_all add: tape_of_nl_abv tape_of_nat_list.simps) | 
|         |   2009 done | 
|         |   2010  | 
|         |   2011 lemma [simp]: | 
|         |   2012     "inv_check_left_moving_in_middle (as, lm) (s, l, Bk # r) ires= False" | 
|         |   2013 by(auto simp: inv_check_left_moving_in_middle.simps ) | 
|         |   2014  | 
|         |   2015 lemma [simp]:  | 
|         |   2016  "inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # r) ires\<Longrightarrow>  | 
|         |   2017   inv_check_left_moving_on_leftmost (as, lm) (s', [], Bk # Bk # r) ires" | 
|         |   2018 apply(auto simp: inv_on_left_moving_in_middle_B.simps  | 
|         |   2019                  inv_check_left_moving_on_leftmost.simps split: if_splits) | 
|         |   2020 done | 
|         |   2021  | 
|         |   2022  | 
|         |   2023 lemma [simp]: "inv_check_left_moving_on_leftmost (as, lm)  | 
|         |   2024                                        (s, list, Oc # r) ires= False" | 
|         |   2025 by(auto simp: inv_check_left_moving_on_leftmost.simps split: if_splits) | 
|         |   2026  | 
|         |   2027 lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm)  | 
|         |   2028                                          (s, Oc # list, Bk # r) ires | 
|         |   2029  \<Longrightarrow> inv_check_left_moving_in_middle (as, lm) (s', list, Oc # Bk # r) ires" | 
|         |   2030 apply(auto simp: inv_on_left_moving_in_middle_B.simps  | 
|         |   2031                  inv_check_left_moving_in_middle.simps  split: if_splits) | 
|         |   2032 done | 
|         |   2033  | 
|         |   2034 lemma inv_on_left_moving_2_check_left_moving[simp]: | 
|         |   2035  "inv_on_left_moving (as, lm) (s, l, Bk # r) ires | 
|         |   2036  \<Longrightarrow> (l = [] \<longrightarrow> inv_check_left_moving (as, lm) (s', [], Bk # Bk # r) ires) | 
|         |   2037  \<and> (l \<noteq> [] \<longrightarrow>  | 
|         |   2038       inv_check_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)" | 
|         |   2039 apply(simp add: inv_on_left_moving.simps inv_check_left_moving.simps) | 
|         |   2040 apply(case_tac l, simp, simp) | 
|         |   2041 apply(case_tac a, simp, simp) | 
|         |   2042 done | 
|         |   2043  | 
|         |   2044 lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, l, []) ires = False" | 
|         |   2045 apply(auto simp: inv_on_left_moving_norm.simps) | 
|         |   2046 by(case_tac [!] mr, auto) | 
|         |   2047  | 
|         |   2048 lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires\<Longrightarrow>  | 
|         |   2049      inv_on_left_moving (as, lm) (6 + 2 * n, l, [Bk]) ires" | 
|         |   2050 apply(simp add: inv_on_left_moving.simps) | 
|         |   2051 apply(auto simp: inv_on_left_moving_in_middle_B.simps) | 
|         |   2052 done | 
|         |   2053  | 
|         |   2054 lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires = False" | 
|         |   2055 apply(simp add: inv_on_left_moving.simps) | 
|         |   2056 apply(simp add: inv_on_left_moving_in_middle_B.simps) | 
|         |   2057 done | 
|         |   2058  | 
|         |   2059 lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires | 
|         |   2060  \<Longrightarrow> (l = [] \<longrightarrow> inv_check_left_moving (as, lm) (s', [], [Bk]) ires) \<and> | 
|         |   2061     (l \<noteq> [] \<longrightarrow> inv_check_left_moving (as, lm) (s', tl l, [hd l]) ires)" | 
|         |   2062 by simp | 
|         |   2063  | 
|         |   2064 lemma Oc_Bk_Cons_ex[simp]:  | 
|         |   2065  "Oc # Bk # list = <lm::nat list> @ Bk\<^bsup>ln\<^esup> \<Longrightarrow>  | 
|         |   2066                              \<exists>ln. list = <tl (lm)> @ Bk\<^bsup>ln\<^esup>" | 
|         |   2067 apply(case_tac "lm", simp) | 
|         |   2068 apply(case_tac ln, simp_all add: exponent_def) | 
|         |   2069 apply(case_tac lista,  | 
|         |   2070       auto simp: tape_of_nl_abv tape_of_nat_list.simps exponent_def) | 
|         |   2071 apply(case_tac [!] a, auto simp: ) | 
|         |   2072 apply(case_tac ln, simp, rule_tac x = nat in exI, simp) | 
|         |   2073 done | 
|         |   2074  | 
|         |   2075 lemma [simp]: | 
|         |   2076   "Oc # Bk # list = <rev lm1::nat list> @ Bk\<^bsup>ln\<^esup> \<Longrightarrow>  | 
|         |   2077       \<exists>rna. Oc # Bk # <lm2> @ Bk\<^bsup>rn\<^esup> = <hd (rev lm1) # lm2> @ Bk\<^bsup>rna\<^esup>" | 
|         |   2078 apply(frule Oc_Bk_Cons, simp) | 
|         |   2079 apply(case_tac lm2,  | 
|         |   2080      auto simp: tape_of_nl_abv tape_of_nat_list.simps  exponent_def ) | 
|         |   2081 apply(rule_tac x = "Suc rn" in exI, simp) | 
|         |   2082 done | 
|         |   2083  | 
|         |   2084 (*inv: from check_left_moving to on_left_moving*) | 
|         |   2085 lemma [intro]: "\<exists>rna. a # a\<^bsup>rn\<^esup> = a\<^bsup>rna\<^esup>" | 
|         |   2086 apply(rule_tac x = "Suc rn" in exI, simp) | 
|         |   2087 done | 
|         |   2088  | 
|         |   2089 lemma  | 
|         |   2090 inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[simp]: | 
|         |   2091 "inv_check_left_moving_in_middle (as, lm) (s, Bk # list, Oc # r) ires | 
|         |   2092   \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s', list, Bk # Oc # r) ires" | 
|         |   2093 apply(simp only: inv_check_left_moving_in_middle.simps  | 
|         |   2094                  inv_on_left_moving_in_middle_B.simps) | 
|         |   2095 apply(erule_tac exE)+ | 
|         |   2096 apply(rule_tac x = "rev (tl (rev lm1))" in exI,  | 
|         |   2097       rule_tac x = "[hd (rev lm1)] @ lm2" in exI, auto) | 
|         |   2098 apply(case_tac [!] "rev lm1",simp_all add: tape_of_nl_abv tape_of_nat_list.simps) | 
|         |   2099 apply(case_tac [!] a, simp_all) | 
|         |   2100 apply(case_tac [1] lm2, simp_all add: tape_of_nat_list.simps, auto) | 
|         |   2101 apply(case_tac [3] lm2, simp_all add: tape_of_nat_list.simps, auto) | 
|         |   2102 apply(case_tac [!] lista, simp_all add: tape_of_nat_list.simps) | 
|         |   2103 done | 
|         |   2104  | 
|         |   2105 lemma [simp]:  | 
|         |   2106  "inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires\<Longrightarrow> | 
|         |   2107      inv_check_left_moving_in_middle (as, lm) (s', [Bk], Oc # r) ires" | 
|         |   2108 apply(auto simp: inv_check_left_moving_in_middle.simps ) | 
|         |   2109 done | 
|         |   2110  | 
|         |   2111 lemma [simp]:  | 
|         |   2112  "inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires | 
|         |   2113    \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s', [], Bk # Oc # r) ires" | 
|         |   2114 apply(insert  | 
|         |   2115 inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[of  | 
|         |   2116                   as lm n "[]" r], simp) | 
|         |   2117 done  | 
|         |   2118  | 
|         |   2119 lemma [simp]: "a\<^bsup>0\<^esup> = []" | 
|         |   2120 apply(simp add: exponent_def) | 
|         |   2121 done | 
|         |   2122  | 
|         |   2123 lemma [simp]: "inv_check_left_moving_in_middle (as, lm)  | 
|         |   2124                        (s, Oc # list, Oc # r) ires | 
|         |   2125    \<Longrightarrow> inv_on_left_moving_norm (as, lm) (s', list, Oc # Oc # r) ires" | 
|         |   2126 apply(auto simp: inv_check_left_moving_in_middle.simps  | 
|         |   2127                  inv_on_left_moving_norm.simps) | 
|         |   2128 apply(rule_tac x = "rev (tl (rev lm1))" in exI,  | 
|         |   2129       rule_tac x = lm2 in exI, rule_tac x = "hd (rev lm1)" in exI) | 
|         |   2130 apply(rule_tac conjI) | 
|         |   2131 apply(case_tac "rev lm1", simp, simp) | 
|         |   2132 apply(rule_tac x = "hd (rev lm1) - 1" in exI, auto) | 
|         |   2133 apply(rule_tac [!] x = "Suc (Suc 0)" in exI, simp) | 
|         |   2134 apply(case_tac [!] "rev lm1", simp_all) | 
|         |   2135 apply(case_tac [!] a, simp_all add: tape_of_nl_abv tape_of_nat_list.simps, auto) | 
|         |   2136 done  | 
|         |   2137  | 
|         |   2138 lemma [simp]: "inv_check_left_moving (as, lm) (s, l, Oc # r) ires | 
|         |   2139 \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s', [], Bk # Oc # r) ires) \<and> | 
|         |   2140    (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s', tl l, hd l # Oc # r) ires)" | 
|         |   2141 apply(case_tac l,  | 
|         |   2142       auto simp: inv_check_left_moving.simps inv_on_left_moving.simps) | 
|         |   2143 apply(case_tac a, simp, simp) | 
|         |   2144 done | 
|         |   2145  | 
|         |   2146 (*inv: check_left_moving to after_left_moving*) | 
|         |   2147 lemma [simp]: "inv_check_left_moving (as, lm) (s, l, Bk # r) ires | 
|         |   2148                 \<Longrightarrow> inv_after_left_moving (as, lm) (s', Bk # l, r) ires" | 
|         |   2149 apply(auto simp: inv_check_left_moving.simps  | 
|         |   2150  inv_check_left_moving_on_leftmost.simps inv_after_left_moving.simps) | 
|         |   2151 done | 
|         |   2152  | 
|         |   2153  | 
|         |   2154 lemma [simp]:"inv_check_left_moving (as, lm) (s, l, []) ires | 
|         |   2155       \<Longrightarrow> inv_after_left_moving (as, lm) (s', Bk # l, []) ires" | 
|         |   2156 by(simp add: inv_check_left_moving.simps   | 
|         |   2157 inv_check_left_moving_in_middle.simps  | 
|         |   2158 inv_check_left_moving_on_leftmost.simps) | 
|         |   2159  | 
|         |   2160 (*inv: after_left_moving to inv_stop*) | 
|         |   2161 lemma [simp]: "inv_after_left_moving (as, lm) (s, l, Bk # r) ires | 
|         |   2162        \<Longrightarrow> inv_stop (as, lm) (s', Bk # l, r) ires" | 
|         |   2163 apply(auto simp: inv_after_left_moving.simps inv_stop.simps) | 
|         |   2164 done | 
|         |   2165  | 
|         |   2166 lemma [simp]: "inv_after_left_moving (as, lm) (s, l, []) ires | 
|         |   2167              \<Longrightarrow> inv_stop (as, lm) (s', Bk # l, []) ires" | 
|         |   2168 by(auto simp: inv_after_left_moving.simps) | 
|         |   2169  | 
|         |   2170 (*inv: stop to stop*) | 
|         |   2171 lemma [simp]: "inv_stop (as, lm) (x, l, r) ires \<Longrightarrow>  | 
|         |   2172                inv_stop (as, lm) (y, l, r) ires" | 
|         |   2173 apply(simp add: inv_stop.simps) | 
|         |   2174 done | 
|         |   2175  | 
|         |   2176 lemma [simp]: "inv_after_clear (as, lm) (s, aaa, Oc # xs) ires= False" | 
|         |   2177 apply(auto simp: inv_after_clear.simps ) | 
|         |   2178 done | 
|         |   2179  | 
|         |   2180 lemma [simp]:  | 
|         |   2181   "inv_after_left_moving (as, lm) (s, aaa, Oc # xs) ires = False" | 
|         |   2182 by(auto simp: inv_after_left_moving.simps  ) | 
|         |   2183  | 
|         |   2184 lemma start_of_not0: "as \<noteq> 0 \<Longrightarrow> start_of ly as > 0" | 
|         |   2185 apply(rule startof_not0) | 
|         |   2186 done | 
|         |   2187  | 
|         |   2188 text {* | 
|         |   2189   The single step currectness of the TM complied from Abacus instruction @{text "Inc n"}. | 
|         |   2190   It shows every single step execution of this TM keeps the invariant. | 
|         |   2191 *} | 
|         |   2192  | 
|         |   2193 lemma inc_inv_step:  | 
|         |   2194   assumes  | 
|         |   2195   -- {* Invariant holds on the start *} | 
|         |   2196       h11: "inc_inv ly n (as, am) tc ires"  | 
|         |   2197   -- {* The layout of Abacus program @{text "aprog"} is @{text "ly"} *} | 
|         |   2198   and h12: "ly = layout_of aprog"  | 
|         |   2199   -- {* The instruction at position @{text "as"} is @{text "Inc n"} *} | 
|         |   2200   and h21: "abc_fetch as aprog = Some (Inc n)"  | 
|         |   2201   -- {* TM not yet reach the final state, where @{text "start_of ly as + 2*n + 9"} is the state | 
|         |   2202         where the current TM stops and the next TM starts. *} | 
|         |   2203   and h22: "(\<lambda> (s, l, r). s \<noteq> start_of ly as + 2*n + 9) tc" | 
|         |   2204   shows  | 
|         |   2205   -- {*  | 
|         |   2206   Single step execution of the TM keeps the invaraint, where  | 
|         |   2207   the TM compiled from @{text "Inc n"} is @{text "(ci ly (start_of ly as) (Inc n))"} | 
|         |   2208   @{text "start_of ly as - Suc 0)"} is the offset used to execute this {\em shifted} | 
|         |   2209   TM. | 
|         |   2210   *} | 
|         |   2211   "inc_inv ly n (as, am) (t_step tc (ci ly (start_of ly as) (Inc n), start_of ly as - Suc 0)) ires" | 
|         |   2212 proof - | 
|         |   2213   from h21 h22  have h3 : "start_of (layout_of aprog) as > 0" | 
|         |   2214     apply(case_tac as, simp add: start_of.simps abc_fetch.simps) | 
|         |   2215     apply(insert start_of_not0[of as "layout_of aprog"], simp) | 
|         |   2216     done     | 
|         |   2217   from h11 h12 and h21 h22 and this show ?thesis  | 
|         |   2218     apply(case_tac tc, simp) | 
|         |   2219     apply(case_tac "a = 0",  | 
|         |   2220       auto split:if_splits simp add:t_step.simps, | 
|         |   2221       tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *}) | 
|         |   2222     apply (simp_all add:fetch_simps new_tape.simps) | 
|         |   2223     done | 
|         |   2224 qed | 
|         |   2225  | 
|         |   2226  | 
|         |   2227 lemma t_steps_ind: "t_steps tc (p, off) (Suc n) | 
|         |   2228                  = t_step (t_steps tc (p, off) n) (p, off)" | 
|         |   2229 apply(induct n arbitrary: tc) | 
|         |   2230 apply(simp add: t_steps.simps) | 
|         |   2231 apply(simp add: t_steps.simps) | 
|         |   2232 done | 
|         |   2233  | 
|         |   2234 definition lex_pair :: "((nat \<times> nat) \<times> (nat \<times> nat)) set" | 
|         |   2235   where  | 
|         |   2236   "lex_pair \<equiv> less_than <*lex*> less_than" | 
|         |   2237  | 
|         |   2238 definition lex_triple ::  | 
|         |   2239    "((nat \<times> (nat \<times> nat)) \<times> (nat \<times> (nat \<times> nat))) set" | 
|         |   2240   where "lex_triple \<equiv> less_than <*lex*> lex_pair" | 
|         |   2241  | 
|         |   2242 definition lex_square ::  | 
|         |   2243     "((nat \<times> nat \<times> nat \<times> nat) \<times> (nat \<times> nat \<times> nat \<times> nat)) set" | 
|         |   2244   where "lex_square \<equiv> less_than <*lex*> lex_triple" | 
|         |   2245  | 
|         |   2246 fun abc_inc_stage1 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" | 
|         |   2247   where | 
|         |   2248   "abc_inc_stage1 (s, l, r) ss n =  | 
|         |   2249             (if s = 0 then 0 | 
|         |   2250              else if s \<le> ss+2*n+1 then 5 | 
|         |   2251              else if s\<le> ss+2*n+5 then 4 | 
|         |   2252              else if s \<le> ss+2*n+7 then 3 | 
|         |   2253              else if s = ss+2*n+8 then 2 | 
|         |   2254              else 1)" | 
|         |   2255  | 
|         |   2256 fun abc_inc_stage2 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" | 
|         |   2257   where | 
|         |   2258   "abc_inc_stage2 (s, l, r) ss n = | 
|         |   2259                 (if s \<le> ss + 2*n + 1 then 0 | 
|         |   2260                  else if s = ss + 2*n + 2 then length r | 
|         |   2261                  else if s = ss + 2*n + 3 then length r | 
|         |   2262                  else if s = ss + 2*n + 4 then length r | 
|         |   2263                  else if s = ss + 2*n + 5 then  | 
|         |   2264                                   if r \<noteq> [] then length r | 
|         |   2265                                   else 1 | 
|         |   2266                  else if s = ss+2*n+6 then length l | 
|         |   2267                  else if s = ss+2*n+7 then length l | 
|         |   2268                  else 0)" | 
|         |   2269  | 
|         |   2270 fun abc_inc_stage3 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> block list \<Rightarrow>  nat" | 
|         |   2271   where | 
|         |   2272   "abc_inc_stage3 (s, l, r) ss n ires = ( | 
|         |   2273               if s = ss + 2*n + 3 then 4 | 
|         |   2274               else if s = ss + 2*n + 4 then 3 | 
|         |   2275               else if s = ss + 2*n + 5 then  | 
|         |   2276                    if r \<noteq> [] \<and> hd r = Oc then 2 | 
|         |   2277                    else 1 | 
|         |   2278               else if s = ss + 2*n + 2 then 0 | 
|         |   2279               else if s = ss + 2*n + 6 then  | 
|         |   2280                       if l = Bk # ires \<and> r \<noteq> [] \<and>  hd r = Oc then 2 | 
|         |   2281                       else 1 | 
|         |   2282               else if s = ss + 2*n + 7 then  | 
|         |   2283                       if r \<noteq> [] \<and> hd r = Oc then 3 | 
|         |   2284                       else 0 | 
|         |   2285               else ss+2*n+9 - s)" | 
|         |   2286  | 
|         |   2287 fun abc_inc_stage4 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> block list \<Rightarrow> nat" | 
|         |   2288   where | 
|         |   2289   "abc_inc_stage4 (s, l, r) ss n ires =  | 
|         |   2290             (if s \<le> ss+2*n+1 \<and> (s - ss) mod 2 = 0 then  | 
|         |   2291                 if (r\<noteq>[] \<and> hd r = Oc) then 0 | 
|         |   2292                 else 1 | 
|         |   2293              else if (s \<le> ss+2*n+1 \<and> (s - ss) mod 2 = Suc 0)  | 
|         |   2294                                                  then length r | 
|         |   2295              else if s = ss + 2*n + 6 then  | 
|         |   2296                   if l = Bk # ires \<and> hd r = Bk then 0 | 
|         |   2297                   else Suc (length l) | 
|         |   2298              else 0)" | 
|         |   2299   | 
|         |   2300 fun abc_inc_measure :: "(t_conf \<times> nat \<times> nat \<times> block list) \<Rightarrow>  | 
|         |   2301                         (nat \<times> nat \<times> nat \<times> nat)" | 
|         |   2302   where | 
|         |   2303   "abc_inc_measure (c, ss, n, ires) =  | 
|         |   2304      (abc_inc_stage1 c ss n, abc_inc_stage2 c ss n,  | 
|         |   2305       abc_inc_stage3 c ss n ires, abc_inc_stage4 c ss n ires)" | 
|         |   2306  | 
|         |   2307 definition abc_inc_LE :: "(((nat \<times> block list \<times> block list) \<times> nat \<times>  | 
|         |   2308        nat \<times> block list) \<times> ((nat \<times> block list \<times> block list) \<times> nat \<times> nat \<times> block list)) set" | 
|         |   2309   where "abc_inc_LE \<equiv> (inv_image lex_square abc_inc_measure)" | 
|         |   2310  | 
|         |   2311 lemma wf_lex_triple: "wf lex_triple" | 
|         |   2312 by (auto intro:wf_lex_prod simp:lex_triple_def lex_pair_def) | 
|         |   2313  | 
|         |   2314 lemma wf_lex_square: "wf lex_square" | 
|         |   2315 by (auto intro:wf_lex_triple simp:lex_triple_def lex_square_def lex_pair_def) | 
|         |   2316  | 
|         |   2317 lemma wf_abc_inc_le[intro]: "wf abc_inc_LE" | 
|         |   2318 by(auto intro:wf_inv_image wf_lex_square simp:abc_inc_LE_def) | 
|         |   2319  | 
|         |   2320 (********************************************************************) | 
|         |   2321 declare inc_inv.simps[simp del] | 
|         |   2322  | 
|         |   2323 lemma halt_lemma2':  | 
|         |   2324   "\<lbrakk>wf LE;  \<forall> n. ((\<not> P (f n) \<and> Q (f n)) \<longrightarrow>  | 
|         |   2325     (Q (f (Suc n)) \<and> (f (Suc n), (f n)) \<in> LE)); Q (f 0)\<rbrakk>  | 
|         |   2326       \<Longrightarrow> \<exists> n. P (f n)" | 
|         |   2327 apply(intro exCI, simp) | 
|         |   2328 apply(subgoal_tac "\<forall> n. Q (f n)", simp) | 
|         |   2329 apply(drule_tac f = f in wf_inv_image) | 
|         |   2330 apply(simp add: inv_image_def) | 
|         |   2331 apply(erule wf_induct, simp) | 
|         |   2332 apply(erule_tac x = x in allE) | 
|         |   2333 apply(erule_tac x = n in allE, erule_tac x = n in allE) | 
|         |   2334 apply(erule_tac x = "Suc x" in allE, simp) | 
|         |   2335 apply(rule_tac allI) | 
|         |   2336 apply(induct_tac n, simp) | 
|         |   2337 apply(erule_tac x = na in allE, simp) | 
|         |   2338 done | 
|         |   2339  | 
|         |   2340 lemma halt_lemma2'':  | 
|         |   2341   "\<lbrakk>P (f n); \<not> P (f (0::nat))\<rbrakk> \<Longrightarrow>  | 
|         |   2342          \<exists> n. (P (f n) \<and> (\<forall> i < n. \<not> P (f i)))" | 
|         |   2343 apply(induct n rule: nat_less_induct, auto) | 
|         |   2344 done | 
|         |   2345  | 
|         |   2346 lemma halt_lemma2''': | 
|         |   2347  "\<lbrakk>\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> LE; | 
|         |   2348                  Q (f 0);  \<forall>i<na. \<not> P (f i)\<rbrakk> \<Longrightarrow> Q (f na)" | 
|         |   2349 apply(induct na, simp, simp) | 
|         |   2350 done | 
|         |   2351  | 
|         |   2352 lemma halt_lemma2:  | 
|         |   2353   "\<lbrakk>wf LE;   | 
|         |   2354     \<forall> n. ((\<not> P (f n) \<and> Q (f n)) \<longrightarrow> (Q (f (Suc n)) \<and> (f (Suc n), (f n)) \<in> LE));  | 
|         |   2355     Q (f 0); \<not> P (f 0)\<rbrakk>  | 
|         |   2356   \<Longrightarrow> \<exists> n. P (f n) \<and> Q (f n)" | 
|         |   2357 apply(insert halt_lemma2' [of LE P f Q], simp, erule_tac exE) | 
|         |   2358 apply(subgoal_tac "\<exists> n. (P (f n) \<and> (\<forall> i < n. \<not> P (f i)))") | 
|         |   2359 apply(erule_tac exE)+ | 
|         |   2360 apply(rule_tac x = na in exI, auto) | 
|         |   2361 apply(rule halt_lemma2''', simp, simp, simp) | 
|         |   2362 apply(erule_tac halt_lemma2'', simp) | 
|         |   2363 done | 
|         |   2364  | 
|         |   2365 lemma [simp]:  | 
|         |   2366   "\<lbrakk>ly = layout_of aprog; abc_fetch as aprog = Some (Inc n)\<rbrakk> | 
|         |   2367     \<Longrightarrow> start_of ly (Suc as) = start_of ly as + 2*n +9" | 
|         |   2368 apply(case_tac as, auto simp: abc_fetch.simps start_of.simps  | 
|         |   2369           layout_of.simps length_of.simps split: if_splits) | 
|         |   2370 done | 
|         |   2371  | 
|         |   2372 lemma inc_inv_init:  | 
|         |   2373  "\<lbrakk>abc_fetch as aprog = Some (Inc n);  | 
|         |   2374    crsp_l ly (as, am) (start_of ly as, l, r) ires; ly = layout_of aprog\<rbrakk> | 
|         |   2375   \<Longrightarrow> inc_inv ly n (as, am) (start_of ly as, l, r) ires" | 
|         |   2376 apply(auto simp: crsp_l.simps inc_inv.simps  | 
|         |   2377       inv_locate_a.simps at_begin_fst_bwtn.simps  | 
|         |   2378       at_begin_fst_awtn.simps at_begin_norm.simps ) | 
|         |   2379 apply(auto intro: startof_not0) | 
|         |   2380 done | 
|         |   2381       | 
|         |   2382 lemma inc_inv_stop_pre[simp]:  | 
|         |   2383    "\<lbrakk>ly = layout_of aprog; inc_inv ly n (as, am) (s, l, r) ires;  | 
|         |   2384      s = start_of ly as; abc_fetch as aprog = Some (Inc n)\<rbrakk> | 
|         |   2385     \<Longrightarrow>  (\<forall>na. \<not> (\<lambda>((s, l, r), ss, n', ires'). s = start_of ly (Suc as))  | 
|         |   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) \<and> | 
|         |   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) na, s, n, ires) \<longrightarrow> | 
|         |   2391        (\<lambda>((s, l, r), ss, n', ires'). inc_inv ly n (as, am) (s, l, r) ires')  | 
|         |   2392       (t_steps (s, l, r) (ci ly (start_of ly as)  | 
|         |   2393               (Inc n), start_of ly as - Suc 0) (Suc na), s, n, ires) \<and> | 
|         |   2394      ((t_steps (s, l, r) (ci ly (start_of ly as) (Inc n),  | 
|         |   2395         start_of ly as - Suc 0) (Suc na), s, n, ires),  | 
|         |   2396       t_steps (s, l, r) (ci ly (start_of ly as)  | 
|         |   2397         (Inc n), start_of ly as - Suc 0) na, s, n, ires) \<in> abc_inc_LE)" | 
|         |   2398 apply(rule allI, rule impI, simp add: t_steps_ind, | 
|         |   2399        rule conjI, erule_tac conjE) | 
|         |   2400 apply(rule_tac inc_inv_step, simp, simp, simp) | 
|         |   2401 apply(case_tac "t_steps (start_of (layout_of aprog) as, l, r) (ci (layout_of aprog)  | 
|         |   2402   (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na", simp) | 
|         |   2403 proof - | 
|         |   2404   fix na | 
|         |   2405   assume h1: "abc_fetch as aprog = Some (Inc n)" | 
|         |   2406     "\<not> (\<lambda>(s, l, r) (ss, n', ires'). s = start_of (layout_of aprog) as + 2 * n + 9) | 
|         |   2407     (t_steps (start_of (layout_of aprog) as, l, r) (ci (layout_of aprog)  | 
|         |   2408     (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na)  | 
|         |   2409     (start_of (layout_of aprog) as, n, ires) \<and> | 
|         |   2410     inc_inv (layout_of aprog) n (as, am) (t_steps (start_of (layout_of aprog) as, l, r) | 
|         |   2411     (ci (layout_of aprog) (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na) ires" | 
|         |   2412   from h1 have h2: "start_of (layout_of aprog) as > 0" | 
|         |   2413     apply(rule_tac startof_not0) | 
|         |   2414     done | 
|         |   2415   from h1 and h2 show "((t_step (t_steps (start_of (layout_of aprog) as, l, r) (ci (layout_of aprog) | 
|         |   2416     (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na) | 
|         |   2417     (ci (layout_of aprog) (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0), | 
|         |   2418     start_of (layout_of aprog) as, n, ires), | 
|         |   2419     t_steps (start_of (layout_of aprog) as, l, r)  | 
|         |   2420     (ci (layout_of aprog) (start_of (layout_of aprog) as) (Inc n), start_of (layout_of aprog) as - Suc 0) na,  | 
|         |   2421     start_of (layout_of aprog) as, n, ires) | 
|         |   2422             \<in> abc_inc_LE" | 
|         |   2423     apply(case_tac "(t_steps (start_of (layout_of aprog) as, l, r)  | 
|         |   2424                (ci (layout_of aprog) | 
|         |   2425          (start_of (layout_of aprog) as) (Inc n),  | 
|         |   2426            start_of (layout_of aprog) as - Suc 0) na)", simp) | 
|         |   2427     apply(case_tac "a = 0",  | 
|         |   2428      auto split:if_splits simp add:t_step.simps inc_inv.simps,  | 
|         |   2429        tactic {* ALLGOALS (resolve_tac [@{thm fetch_intro}]) *}) | 
|         |   2430     apply(simp_all add:fetch_simps new_tape.simps) | 
|         |   2431     apply(auto simp add: abc_inc_LE_def   | 
|         |   2432     lex_square_def lex_triple_def lex_pair_def | 
|         |   2433       inv_after_write.simps inv_after_move.simps inv_after_clear.simps | 
|         |   2434       inv_on_left_moving.simps inv_on_left_moving_norm.simps split: if_splits) | 
|         |   2435     done | 
|         |   2436 qed | 
|         |   2437  | 
|         |   2438 lemma inc_inv_stop_pre1:  | 
|         |   2439   "\<lbrakk> | 
|         |   2440   ly = layout_of aprog;  | 
|         |   2441   abc_fetch as aprog = Some (Inc n); | 
|         |   2442   s = start_of ly as;  | 
|         |   2443   inc_inv ly n (as, am) (s, l, r) ires | 
|         |   2444   \<rbrakk> \<Longrightarrow>  | 
|         |   2445   (\<exists> stp > 0. (\<lambda> (s', l', r'). | 
|         |   2446            s' = start_of ly (Suc as) \<and>  | 
|         |   2447            inc_inv ly n (as, am) (s', l', r') ires)  | 
|         |   2448                (t_steps (s, l, r) (ci ly (start_of ly as) (Inc n),  | 
|         |   2449                         start_of ly as - Suc 0) stp))" | 
|         |   2450 apply(insert halt_lemma2[of abc_inc_LE  | 
|         |   2451     "\<lambda> ((s, l, r), ss, n', ires'). s = start_of ly (Suc as)"  | 
|         |   2452     "(\<lambda> stp. (t_steps (s, l, r)  | 
|         |   2453      (ci ly (start_of ly as) (Inc n),  | 
|         |   2454      start_of ly as - Suc 0) stp, s, n, ires))"  | 
|         |   2455     "\<lambda> ((s, l, r), ss, n'). inc_inv ly n (as, am) (s, l, r) ires"]) | 
|         |   2456 apply(insert  wf_abc_inc_le) | 
|         |   2457 apply(insert inc_inv_stop_pre[of ly aprog n as am s l r ires], simp) | 
|         |   2458 apply(simp only: t_steps.simps, auto) | 
|         |   2459 apply(rule_tac x = na in exI) | 
|         |   2460 apply(case_tac "(t_steps (start_of (layout_of aprog) as, l, r)  | 
|         |   2461    (ci (layout_of aprog) (start_of (layout_of aprog) as)  | 
|         |   2462    (Inc n), start_of (layout_of aprog) as - Suc 0) na)", simp) | 
|         |   2463 apply(case_tac na, simp add: t_steps.simps, simp) | 
|         |   2464 done | 
|         |   2465  | 
|         |   2466 lemma inc_inv_stop:  | 
|         |   2467   assumes program_and_layout:  | 
|         |   2468   -- {* There is an Abacus program @{text "aprog"} and its layout is @{text "ly"}: *} | 
|         |   2469   "ly = layout_of aprog" | 
|         |   2470   and an_instruction: | 
|         |   2471   -- {* There is an instruction @{text "Inc n"} at postion @{text "as"} of @{text "aprog"} *} | 
|         |   2472   "abc_fetch as aprog = Some (Inc n)" | 
|         |   2473   and the_start_state: | 
|         |   2474   -- {* According to @{text "ly"} and @{text "as"},  | 
|         |   2475         the start state of the TM compiled from this | 
|         |   2476         @{text "Inc n"} instruction should be @{text "s"}: | 
|         |   2477      *} | 
|         |   2478   "s = start_of ly as" | 
|         |   2479   and inv: | 
|         |   2480   -- {* Invariant holds on configuration @{text "(s, l, r)"} *} | 
|         |   2481   "inc_inv ly n (as, am) (s, l, r) ires" | 
|         |   2482   shows  -- {* After @{text "stp"} steps of execution, the compiled  | 
|         |   2483             TM reaches the start state of next compiled TM and the invariant  | 
|         |   2484             still holds. | 
|         |   2485             *} | 
|         |   2486       "(\<exists> stp > 0. (\<lambda> (s', l', r'). | 
|         |   2487            s' = start_of ly (Suc as) \<and>  | 
|         |   2488            inc_inv ly n (as, am) (s', l', r') ires)  | 
|         |   2489                (t_steps (s, l, r) (ci ly (start_of ly as) (Inc n),  | 
|         |   2490                         start_of ly as - Suc 0) stp))" | 
|         |   2491 proof - | 
|         |   2492   from inc_inv_stop_pre1 [OF  program_and_layout an_instruction the_start_state inv]  | 
|         |   2493   show ?thesis . | 
|         |   2494 qed | 
|         |   2495  | 
|         |   2496 lemma inc_inv_stop_cond:  | 
|         |   2497   "\<lbrakk>ly = layout_of aprog;  | 
|         |   2498     s' = start_of ly (as + 1);  | 
|         |   2499     inc_inv ly n (as, lm) (s', (l', r')) ires;  | 
|         |   2500     abc_fetch as aprog = Some (Inc n)\<rbrakk> \<Longrightarrow> | 
|         |   2501     crsp_l ly (Suc as, abc_lm_s lm n (Suc (abc_lm_v lm n)))  | 
|         |   2502                                                 (s', l', r') ires" | 
|         |   2503 apply(subgoal_tac "s' = start_of ly as + 2*n + 9", simp) | 
|         |   2504 apply(auto simp: inc_inv.simps inv_stop.simps crsp_l.simps ) | 
|         |   2505 done | 
|         |   2506  | 
|         |   2507 lemma inc_crsp_ex_pre: | 
|         |   2508   "\<lbrakk>ly = layout_of aprog;  | 
|         |   2509     crsp_l ly (as, am) tc ires;   | 
|         |   2510     abc_fetch as aprog = Some (Inc n)\<rbrakk> | 
|         |   2511  \<Longrightarrow> \<exists>stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Inc n)))  | 
|         |   2512                 (t_steps tc (ci ly (start_of ly as) (Inc n),  | 
|         |   2513                                 start_of ly as - Suc 0) stp) ires" | 
|         |   2514 proof(case_tac tc, simp add: abc_step_l.simps) | 
|         |   2515   fix a b c | 
|         |   2516   assume h1: "ly = layout_of aprog"  | 
|         |   2517          "crsp_l (layout_of aprog) (as, am) (a, b, c) ires" | 
|         |   2518          "abc_fetch as aprog = Some (Inc n)" | 
|         |   2519   hence h2: "a = start_of ly as" | 
|         |   2520     by(auto simp: crsp_l.simps) | 
|         |   2521   from h1 and h2 have h3:  | 
|         |   2522        "inc_inv ly n (as, am) (start_of ly as, b, c) ires" | 
|         |   2523     by(rule_tac inc_inv_init, simp, simp, simp) | 
|         |   2524   from h1 and h2 and h3 have h4: | 
|         |   2525        "(\<exists> stp > 0. (\<lambda> (s', l', r'). s' =  | 
|         |   2526            start_of ly (Suc as) \<and> inc_inv ly n (as, am) (s', l', r') ires) | 
|         |   2527          (t_steps (a, b, c) (ci ly (start_of ly as)  | 
|         |   2528                  (Inc n), start_of ly as - Suc 0) stp))" | 
|         |   2529     apply(rule_tac inc_inv_stop, auto) | 
|         |   2530     done | 
|         |   2531   from h1 and h2 and h3 and h4 show  | 
|         |   2532      "\<exists>stp > 0. crsp_l (layout_of aprog)  | 
|         |   2533         (Suc as, abc_lm_s am n (Suc (abc_lm_v am n))) | 
|         |   2534        (t_steps (a, b, c) (ci (layout_of aprog)  | 
|         |   2535           (start_of (layout_of aprog) as) (Inc n),  | 
|         |   2536               start_of (layout_of aprog) as - Suc 0) stp) ires" | 
|         |   2537     apply(erule_tac exE) | 
|         |   2538     apply(rule_tac x = stp in exI) | 
|         |   2539     apply(case_tac "(t_steps (a, b, c) (ci (layout_of aprog)  | 
|         |   2540          (start_of (layout_of aprog) as) (Inc n),  | 
|         |   2541              start_of (layout_of aprog) as - Suc 0) stp)", simp) | 
|         |   2542     apply(rule_tac inc_inv_stop_cond, auto) | 
|         |   2543     done | 
|         |   2544 qed | 
|         |   2545  | 
|         |   2546 text {* | 
|         |   2547   The total correctness of the compilaton of @{text "Inc n"} instruction. | 
|         |   2548 *} | 
|         |   2549  | 
|         |   2550 lemma inc_crsp_ex: | 
|         |   2551   assumes layout:  | 
|         |   2552   -- {* For any Abacus program @{text "aprog"}, assuming its layout is @{text "ly"} *} | 
|         |   2553   "ly = layout_of aprog" | 
|         |   2554   and corresponds:  | 
|         |   2555   -- {* Abacus configuration @{text "(as, am)"} is in correspondence with  | 
|         |   2556         TM configuration @{text "tc"} *} | 
|         |   2557   "crsp_l ly (as, am) tc ires" | 
|         |   2558   and inc: | 
|         |   2559   -- {* There is an instruction @{text "Inc n"} at postion @{text "as"} of @{text "aprog"} *} | 
|         |   2560   "abc_fetch as aprog = Some (Inc n)" | 
|         |   2561   shows | 
|         |   2562   -- {*  | 
|         |   2563   After @{text "stp"} steps of execution, the TM compiled from this @{text "Inc n"} | 
|         |   2564   stops with a configuration which corresponds to the Abacus configuration obtained | 
|         |   2565   from the execution of this @{text "Inc n"} instruction. | 
|         |   2566   *}  | 
|         |   2567   "\<exists>stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Inc n)))  | 
|         |   2568                        (t_steps tc (ci ly (start_of ly as) (Inc n),  | 
|         |   2569                                 start_of ly as - Suc 0) stp) ires" | 
|         |   2570 proof - | 
|         |   2571   from inc_crsp_ex_pre [OF layout corresponds inc] show ?thesis . | 
|         |   2572 qed | 
|         |   2573  | 
|         |   2574 (* | 
|         |   2575 subsection {* 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 subsection {* Compilation of @{text "Goto n"}*} | 
|         |   4839 *) | 
|         |   4840  | 
|         |   4841 lemma goto_fetch:  | 
|         |   4842      "fetch (ci (layout_of aprog)  | 
|         |   4843          (start_of (layout_of aprog) as) (Goto n)) (Suc 0)  b | 
|         |   4844      = (Nop, start_of (layout_of aprog) n)" | 
|         |   4845 apply(auto simp: ci.simps fetch.simps nth_of.simps  | 
|         |   4846            split: block.splits) | 
|         |   4847 done | 
|         |   4848  | 
|         |   4849 text {* | 
|         |   4850   Correctness of complied @{text "Goto n"} | 
|         |   4851   *} | 
|         |   4852  | 
|         |   4853 lemma goto_crsp_ex_pre:  | 
|         |   4854   "\<lbrakk>ly = layout_of aprog;  | 
|         |   4855     crsp_l ly (as, am) tc ires; | 
|         |   4856     abc_fetch as aprog = Some (Goto n)\<rbrakk> | 
|         |   4857  \<Longrightarrow> \<exists>stp > 0. crsp_l ly (abc_step_l (as, am) (Some (Goto n)))  | 
|         |   4858       (t_steps tc (ci (layout_of aprog) (start_of ly as) (Goto n),  | 
|         |   4859                                         start_of ly as - Suc 0) stp) ires" | 
|         |   4860 apply(rule_tac x = 1 in exI) | 
|         |   4861 apply(simp add: abc_step_l.simps t_steps.simps t_step.simps) | 
|         |   4862 apply(case_tac tc, simp) | 
|         |   4863 apply(subgoal_tac "a = start_of (layout_of aprog) as", auto) | 
|         |   4864 apply(subgoal_tac "start_of (layout_of aprog) as > 0", simp) | 
|         |   4865 apply(auto simp: goto_fetch new_tape.simps crsp_l.simps) | 
|         |   4866 apply(rule startof_not0) | 
|         |   4867 done | 
|         |   4868  | 
|         |   4869 lemma goto_crsp_ex: | 
|         |   4870   assumes layout: "ly = layout_of aprog" | 
|         |   4871   and goto: "abc_fetch as aprog = Some (Goto n)" | 
|         |   4872   and correspondence: "crsp_l ly (as, am) tc ires" | 
|         |   4873   shows "\<exists>stp>0. crsp_l ly (abc_step_l (as, am) (Some (Goto n)))  | 
|         |   4874               (t_steps tc (ci (layout_of aprog) (start_of ly as) (Goto n), | 
|         |   4875                                            start_of ly as - Suc 0) stp) ires" | 
|         |   4876 proof - | 
|         |   4877   from goto_crsp_ex_pre and layout goto correspondence show "?thesis" by blast | 
|         |   4878 qed | 
|         |   4879  | 
|         |   4880 subsection {* | 
|         |   4881   The correctness of the compiler | 
|         |   4882   *} | 
|         |   4883  | 
|         |   4884 declare abc_step_l.simps[simp del] | 
|         |   4885  | 
|         |   4886 lemma tm_crsp_ex:  | 
|         |   4887          "\<lbrakk>ly = layout_of aprog; | 
|         |   4888            crsp_l ly (as, am) tc ires;  | 
|         |   4889            as < length aprog; | 
|         |   4890            abc_fetch as aprog = Some ins\<rbrakk> | 
|         |   4891       \<Longrightarrow> \<exists> n > 0. crsp_l ly (abc_step_l (as,am) (Some ins)) | 
|         |   4892                (t_steps tc (ci (layout_of aprog) (start_of ly as)  | 
|         |   4893                   (ins), (start_of ly as) - (Suc 0)) n) ires" | 
|         |   4894 apply(case_tac "ins", simp) | 
|         |   4895 apply(auto intro: inc_crsp_ex_pre dec_crsp_ex goto_crsp_ex) | 
|         |   4896 done | 
|         |   4897  | 
|         |   4898 lemma start_of_pre:  | 
|         |   4899   "n < length aprog \<Longrightarrow> start_of (layout_of aprog) n | 
|         |   4900                      = start_of (layout_of (butlast aprog)) n" | 
|         |   4901 apply(induct n, simp add: start_of.simps, simp) | 
|         |   4902 apply(simp add: layout_of.simps start_of.simps) | 
|         |   4903 apply(subgoal_tac "n < length aprog - Suc 0", simp) | 
|         |   4904 apply(subgoal_tac "(aprog ! n) = (butlast aprog ! n)", simp) | 
|         |   4905 proof - | 
|         |   4906   fix n | 
|         |   4907   assume h1: "Suc n < length aprog" | 
|         |   4908   thus "aprog ! n = butlast aprog ! n" | 
|         |   4909     apply(case_tac "length aprog", simp, simp) | 
|         |   4910     apply(insert nth_append[of "butlast aprog" "[last aprog]" n]) | 
|         |   4911     apply(subgoal_tac "(butlast aprog @ [last aprog]) = aprog") | 
|         |   4912     apply(simp split: if_splits) | 
|         |   4913     apply(rule append_butlast_last_id, case_tac aprog, simp, simp) | 
|         |   4914     done | 
|         |   4915 next | 
|         |   4916   fix n | 
|         |   4917   assume "Suc n < length aprog" | 
|         |   4918   thus "n < length aprog - Suc 0" | 
|         |   4919     apply(case_tac aprog, simp, simp) | 
|         |   4920     done | 
|         |   4921 qed | 
|         |   4922      | 
|         |   4923 lemma zip_eq: "xs = ys \<Longrightarrow> zip xs zs = zip ys zs" | 
|         |   4924 by simp | 
|         |   4925  | 
|         |   4926 lemma tpairs_of_append_iff: "length aprog = Suc n \<Longrightarrow>  | 
|         |   4927          tpairs_of aprog = tpairs_of (butlast aprog) @  | 
|         |   4928                      [(start_of (layout_of aprog) n, aprog ! n)]" | 
|         |   4929 apply(simp add: tpairs_of.simps) | 
|         |   4930 apply(insert zip_append[of "map (start_of (layout_of aprog)) [0..<n]" | 
|         |   4931      "butlast aprog" "[start_of (layout_of aprog) n]" "[last aprog]"]) | 
|         |   4932 apply(simp del: zip_append) | 
|         |   4933 apply(subgoal_tac "(butlast aprog @ [last aprog]) = aprog", auto) | 
|         |   4934 apply(rule_tac zip_eq, auto) | 
|         |   4935 apply(rule_tac start_of_pre, simp) | 
|         |   4936 apply(insert last_conv_nth[of aprog], case_tac aprog, simp, simp) | 
|         |   4937 apply(rule append_butlast_last_id, case_tac aprog, simp, simp) | 
|         |   4938 done | 
|         |   4939  | 
|         |   4940 lemma [simp]: "list_all (\<lambda>(n, tm). abacus.t_ncorrect (ci layout n tm)) | 
|         |   4941          (zip (map (start_of layout) [0..<length aprog]) aprog)" | 
|         |   4942 proof(induct "length aprog" arbitrary: aprog, simp) | 
|         |   4943   fix x aprog | 
|         |   4944   assume ind: "\<And>aprog. x = length aprog \<Longrightarrow>  | 
|         |   4945         list_all (\<lambda>(n, tm). abacus.t_ncorrect (ci layout n tm)) | 
|         |   4946            (zip (map (start_of layout) [0..<length aprog]) aprog)" | 
|         |   4947   and h: "Suc x = length (aprog::abc_inst list)" | 
|         |   4948   have g1: "list_all (\<lambda>(n, tm). abacus.t_ncorrect (ci layout n tm))  | 
|         |   4949     (zip (map (start_of layout) [0..<length (butlast aprog)])  | 
|         |   4950                                                  (butlast aprog))" | 
|         |   4951     using h | 
|         |   4952     apply(rule_tac ind, auto) | 
|         |   4953     done | 
|         |   4954   have g2: "(map (start_of layout) [0..<length aprog]) =  | 
|         |   4955                      map (start_of layout) ([0..<length aprog - 1]  | 
|         |   4956          @ [length aprog - 1])" | 
|         |   4957     using h | 
|         |   4958     apply(case_tac aprog, simp, simp) | 
|         |   4959     done | 
|         |   4960   have "\<exists> xs a. aprog = xs @ [a]" | 
|         |   4961     using h | 
|         |   4962     apply(rule_tac x = "butlast aprog" in exI,  | 
|         |   4963           rule_tac x = "last aprog" in exI) | 
|         |   4964     apply(case_tac "aprog = []", simp, simp) | 
|         |   4965     done | 
|         |   4966   from this obtain xs where "\<exists> a. aprog = xs @ [a]" .. | 
|         |   4967   from this obtain a where g3: "aprog = xs @ [a]" .. | 
|         |   4968   from g1 and g2 and g3 show "list_all (\<lambda>(n, tm).  | 
|         |   4969                               abacus.t_ncorrect (ci layout n tm))  | 
|         |   4970               (zip (map (start_of layout) [0..<length aprog]) aprog)" | 
|         |   4971     apply(simp) | 
|         |   4972     apply(auto simp: t_ncorrect.simps ci.simps  tshift.simps  | 
|         |   4973           tinc_b_def tdec_b_def split: abc_inst.splits) | 
|         |   4974     apply arith+ | 
|         |   4975     done | 
|         |   4976 qed | 
|         |   4977  | 
|         |   4978 lemma [intro]: "abc2t_correct aprog" | 
|         |   4979 apply(simp add: abc2t_correct.simps tpairs_of.simps  | 
|         |   4980           split: abc_inst.splits) | 
|         |   4981 done | 
|         |   4982  | 
|         |   4983 lemma as_out: "\<lbrakk>ly = layout_of aprog; tprog = tm_of aprog;  | 
|         |   4984                 crsp_l ly (as, am) tc ires; length aprog \<le> as\<rbrakk>  | 
|         |   4985             \<Longrightarrow> abc_step_l (as, am) (abc_fetch as aprog) = (as, am)" | 
|         |   4986 apply(simp add: abc_fetch.simps abc_step_l.simps) | 
|         |   4987 done | 
|         |   4988  | 
|         |   4989 lemma tm_merge_ex:  | 
|         |   4990   "\<lbrakk>crsp_l (layout_of aprog) (as, am) tc ires;  | 
|         |   4991     as < length aprog;  | 
|         |   4992     abc_fetch as aprog = Some a;  | 
|         |   4993     abc2t_correct aprog; | 
|         |   4994     crsp_l (layout_of aprog) (abc_step_l (as, am) (Some a)) | 
|         |   4995      (t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) | 
|         |   4996          a, start_of (layout_of aprog) as - Suc 0) n) ires;  | 
|         |   4997     n > 0\<rbrakk> | 
|         |   4998    \<Longrightarrow> \<exists>stp > 0. crsp_l (layout_of aprog) (abc_step_l (as, am)  | 
|         |   4999                        (Some a)) (t_steps tc (tm_of aprog, 0) stp) ires" | 
|         |   5000 apply(case_tac "(t_steps tc (ci (layout_of aprog)  | 
|         |   5001            (start_of (layout_of aprog) as) a,  | 
|         |   5002             start_of (layout_of aprog) as - Suc 0) n)",  simp) | 
|         |   5003 apply(case_tac "(abc_step_l (as, am) (Some a))", simp) | 
|         |   5004 proof - | 
|         |   5005   fix aa b c aaa ba  | 
|         |   5006   assume h:  | 
|         |   5007   "crsp_l (layout_of aprog) (as, am) tc ires"  | 
|         |   5008   "as < length aprog"  | 
|         |   5009   "abc_fetch as aprog = Some a"  | 
|         |   5010   "crsp_l (layout_of aprog) (aaa, ba) (aa, b, c) ires"  | 
|         |   5011   "abc2t_correct aprog"  | 
|         |   5012   "n>0" | 
|         |   5013   "t_steps tc (ci (layout_of aprog) (start_of (layout_of aprog) as) a, | 
|         |   5014       start_of (layout_of aprog) as - Suc 0) n = (aa, b, c)"  | 
|         |   5015    "abc_step_l (as, am) (Some a) = (aaa, ba)" | 
|         |   5016   hence "aa = start_of (layout_of aprog) aaa" | 
|         |   5017     apply(simp add: crsp_l.simps) | 
|         |   5018     done | 
|         |   5019   from this and h show  | 
|         |   5020   "\<exists>stp > 0. crsp_l (layout_of aprog) (aaa, ba)  | 
|         |   5021                           (t_steps tc (tm_of aprog, 0) stp) ires" | 
|         |   5022     apply(insert tms_out_ex[of "layout_of aprog" aprog  | 
|         |   5023                 "tm_of aprog" as am tc ires a n aa b c aaa ba], auto) | 
|         |   5024     done | 
|         |   5025 qed | 
|         |   5026   | 
|         |   5027 lemma crsp_inside:  | 
|         |   5028   "\<lbrakk>ly = layout_of aprog;  | 
|         |   5029     tprog = tm_of aprog; | 
|         |   5030     crsp_l ly (as, am) tc ires; | 
|         |   5031     as < length aprog\<rbrakk> \<Longrightarrow>  | 
|         |   5032     (\<exists> stp > 0. crsp_l ly (abc_step_l (as,am) (abc_fetch as aprog))  | 
|         |   5033                                          (t_steps tc (tprog, 0) stp) ires)" | 
|         |   5034 apply(case_tac "abc_fetch as aprog", simp add: abc_fetch.simps) | 
|         |   5035 proof - | 
|         |   5036   fix a | 
|         |   5037   assume "ly = layout_of aprog"  | 
|         |   5038      "tprog = tm_of aprog"  | 
|         |   5039      "crsp_l ly (as, am) tc ires"  | 
|         |   5040      "as < length aprog"  | 
|         |   5041      "abc_fetch as aprog = Some a" | 
|         |   5042   thus "\<exists>stp > 0. crsp_l ly (abc_step_l (as, am)  | 
|         |   5043                  (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires" | 
|         |   5044     proof(insert tm_crsp_ex[of ly aprog as am tc ires a],  | 
|         |   5045           auto intro: tm_merge_ex) | 
|         |   5046   qed | 
|         |   5047 qed | 
|         |   5048  | 
|         |   5049 lemma crsp_outside:  | 
|         |   5050   "\<lbrakk>ly = layout_of aprog; tprog = tm_of aprog; | 
|         |   5051     crsp_l ly (as, am) tc ires; as \<ge> length aprog\<rbrakk> | 
|         |   5052     \<Longrightarrow> (\<exists> stp. crsp_l ly (abc_step_l (as,am) (abc_fetch as aprog))  | 
|         |   5053                                          (t_steps tc (tprog, 0) stp) ires)" | 
|         |   5054 apply(subgoal_tac "abc_step_l (as, am) (abc_fetch as aprog) | 
|         |   5055                 = (as, am)", simp) | 
|         |   5056 apply(rule_tac x = 0 in exI, simp add: t_steps.simps) | 
|         |   5057 apply(rule as_out, simp+) | 
|         |   5058 done | 
|         |   5059  | 
|         |   5060 text {* | 
|         |   5061   Single-step correntess of the compiler. | 
|         |   5062 *} | 
|         |   5063 lemma astep_crsp_pre:  | 
|         |   5064       "\<lbrakk>ly = layout_of aprog;  | 
|         |   5065         tprog = tm_of aprog; | 
|         |   5066         crsp_l ly (as, am) tc ires\<rbrakk> | 
|         |   5067        \<Longrightarrow> (\<exists> stp. crsp_l ly (abc_step_l (as,am)  | 
|         |   5068                   (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires)" | 
|         |   5069 apply(case_tac "as < length aprog") | 
|         |   5070 apply(drule_tac crsp_inside, auto) | 
|         |   5071 apply(rule_tac crsp_outside, simp+) | 
|         |   5072 done | 
|         |   5073  | 
|         |   5074 text {* | 
|         |   5075   Single-step correntess of the compiler. | 
|         |   5076 *} | 
|         |   5077 lemma astep_crsp_pre1:  | 
|         |   5078       "\<lbrakk>ly = layout_of aprog; | 
|         |   5079         tprog = tm_of aprog; | 
|         |   5080         crsp_l ly (as, am) tc ires\<rbrakk> | 
|         |   5081        \<Longrightarrow> (\<exists> stp. crsp_l ly (abc_step_l (as,am)  | 
|         |   5082                   (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires)" | 
|         |   5083 apply(case_tac "as < length aprog") | 
|         |   5084 apply(drule_tac crsp_inside, auto) | 
|         |   5085 apply(rule_tac crsp_outside, simp+) | 
|         |   5086 done | 
|         |   5087  | 
|         |   5088 lemma astep_crsp: | 
|         |   5089   assumes layout:  | 
|         |   5090   -- {* There is a Abacus program @{text "aprog"} with layout @{text "ly"} *} | 
|         |   5091   "ly = layout_of aprog" | 
|         |   5092   and compiled:  | 
|         |   5093   -- {* @{text "tprog"} is the TM compiled from @{text "aprog"} *} | 
|         |   5094   "tprog = tm_of aprog" | 
|         |   5095   and corresponds:  | 
|         |   5096   -- {* Abacus configuration @{text "(as, am)"} is in correspondence with TM configuration | 
|         |   5097    @{text "tc"} *} | 
|         |   5098   "crsp_l ly (as, am) tc ires" | 
|         |   5099   -- {* One step execution of @{text "aprog"} can be simulated by multi-step execution  | 
|         |   5100   of @{text "tprog"} *} | 
|         |   5101   shows "(\<exists> stp. crsp_l ly (abc_step_l (as,am)  | 
|         |   5102                   (abc_fetch as aprog)) (t_steps tc (tprog, 0) stp) ires)" | 
|         |   5103 proof - | 
|         |   5104   from astep_crsp_pre1 [OF layout compiled corresponds] show ?thesis . | 
|         |   5105 qed | 
|         |   5106  | 
|         |   5107 lemma steps_crsp_pre:  | 
|         |   5108     "\<lbrakk>ly = layout_of aprog; tprog = tm_of aprog;  | 
|         |   5109       crsp_l ly ac tc ires; ac' = abc_steps_l ac aprog n\<rbrakk> \<Longrightarrow>  | 
|         |   5110         (\<exists> n'. crsp_l ly ac' (t_steps tc (tprog, 0) n') ires)" | 
|         |   5111 apply(induct n arbitrary: ac' ac tc, simp add: abc_steps_l.simps) | 
|         |   5112 apply(rule_tac x = 0 in exI) | 
|         |   5113 apply(case_tac ac, simp add: abc_steps_l.simps t_steps.simps) | 
|         |   5114 apply(case_tac ac, simp add: abc_steps_l.simps) | 
|         |   5115 apply(subgoal_tac  | 
|         |   5116    "(\<exists> stp. crsp_l ly (abc_step_l (a, b) | 
|         |   5117             (abc_fetch a aprog)) (t_steps tc (tprog, 0) stp) ires)") | 
|         |   5118 apply(erule exE) | 
|         |   5119 apply(subgoal_tac  | 
|         |   5120    "\<exists>n'. crsp_l (layout_of aprog)  | 
|         |   5121     (abc_steps_l (abc_step_l (a, b) (abc_fetch a aprog)) aprog n) | 
|         |   5122          (t_steps ((t_steps tc (tprog, 0) stp)) (tm_of aprog, 0) n') ires") | 
|         |   5123 apply(erule exE) | 
|         |   5124 apply(subgoal_tac  | 
|         |   5125     "t_steps (t_steps tc (tprog, 0) stp) (tm_of aprog, 0) n' = | 
|         |   5126      t_steps tc (tprog, 0) (stp + n')") | 
|         |   5127 apply(rule_tac x = "stp + n'" in exI, simp) | 
|         |   5128 apply(auto intro: astep_crsp simp: t_step_add) | 
|         |   5129 done | 
|         |   5130  | 
|         |   5131 text {* | 
|         |   5132   Multi-step correctess of the compiler. | 
|         |   5133 *} | 
|         |   5134  | 
|         |   5135 lemma steps_crsp:  | 
|         |   5136   assumes layout:  | 
|         |   5137   -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"} *} | 
|         |   5138     "ly = layout_of aprog" | 
|         |   5139   and compiled:  | 
|         |   5140   -- {* @{text "tprog"} is the TM compiled from @{text "aprog"} *} | 
|         |   5141   "tprog = tm_of aprog" | 
|         |   5142   and correspond:  | 
|         |   5143   -- {* Abacus configuration @{text "ac"} is in correspondence with TM configuration @{text "tc"} *} | 
|         |   5144       "crsp_l ly ac tc ires" | 
|         |   5145   and execution:  | 
|         |   5146   -- {* @{text "ac'"} is the configuration obtained from @{text "n"}-step execution  | 
|         |   5147       of @{text "aprog"} starting from configuration @{text "ac"} *} | 
|         |   5148   "ac' = abc_steps_l ac aprog n"  | 
|         |   5149   -- {* There exists steps @{text "n'"} steps, after these steps of execution,  | 
|         |   5150   the Turing configuration such obtained is in correspondence with @{text "ac'"} *} | 
|         |   5151   shows "(\<exists> n'. crsp_l ly ac' (t_steps tc (tprog, 0) n') ires)" | 
|         |   5152 proof - | 
|         |   5153   from steps_crsp_pre [OF layout compiled correspond execution] show ?thesis . | 
|         |   5154 qed | 
|         |   5155  | 
|         |   5156 subsection {* The Mop-up machine *} | 
|         |   5157  | 
|         |   5158 fun mop_bef :: "nat \<Rightarrow> tprog" | 
|         |   5159   where | 
|         |   5160   "mop_bef 0 = []" | | 
|         |   5161   "mop_bef (Suc n) = mop_bef n @  | 
|         |   5162        [(R, 2*n + 3), (W0, 2*n + 2), (R, 2*n + 1), (W1, 2*n + 2)]" | 
|         |   5163  | 
|         |   5164 definition mp_up :: "tprog" | 
|         |   5165   where | 
|         |   5166   "mp_up \<equiv> [(R, 2), (R, 1), (L, 5), (W0, 3), (R, 4), (W0, 3), | 
|         |   5167             (R, 2), (W0, 3), (L, 5), (L, 6), (R, 0), (L, 6)]" | 
|         |   5168  | 
|         |   5169 fun tMp :: "nat \<Rightarrow> nat \<Rightarrow> tprog" | 
|         |   5170   where  | 
|         |   5171   "tMp n off = tshift (mop_bef n @ tshift mp_up (2*n)) off" | 
|         |   5172  | 
|         |   5173 declare  mp_up_def[simp del]  tMp.simps[simp del] mop_bef.simps[simp del] | 
|         |   5174 (**********Begin: equiv among aba and turing***********) | 
|         |   5175  | 
|         |   5176 lemma tm_append_step:  | 
|         |   5177  "\<lbrakk>t_ncorrect tp1; t_step tc (tp1, 0) = (s, l, r); s \<noteq> 0\<rbrakk>  | 
|         |   5178  \<Longrightarrow> t_step tc (tp1 @ tp2, 0) = (s, l, r)" | 
|         |   5179 apply(simp add: t_step.simps) | 
|         |   5180 apply(case_tac tc, simp) | 
|         |   5181 apply(case_tac  | 
|         |   5182        "(fetch tp1 a (case c of [] \<Rightarrow> Bk | | 
|         |   5183                Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))", simp) | 
|         |   5184 apply(case_tac a, simp add: fetch.simps) | 
|         |   5185 apply(simp add: fetch.simps) | 
|         |   5186 apply(case_tac c, simp) | 
|         |   5187 apply(case_tac [!] "ab::block") | 
|         |   5188 apply(auto simp: nth_of.simps nth_append t_ncorrect.simps  | 
|         |   5189            split: if_splits) | 
|         |   5190 done | 
|         |   5191  | 
|         |   5192 lemma state0_ind: "t_steps (0, l, r) (tp, 0) stp = (0, l, r)" | 
|         |   5193 apply(induct stp, simp add: t_steps.simps) | 
|         |   5194 apply(simp add: t_steps.simps t_step.simps fetch.simps new_tape.simps) | 
|         |   5195 done | 
|         |   5196  | 
|         |   5197 lemma tm_append_steps:   | 
|         |   5198  "\<lbrakk>t_ncorrect tp1; t_steps tc (tp1, 0) stp = (s, l ,r); s \<noteq> 0\<rbrakk> | 
|         |   5199   \<Longrightarrow> t_steps tc (tp1 @ tp2, 0) stp = (s, l, r)" | 
|         |   5200 apply(induct stp arbitrary: tc s l r) | 
|         |   5201 apply(case_tac tc,  simp) | 
|         |   5202 apply(simp add: t_steps.simps) | 
|         |   5203 proof - | 
|         |   5204   fix stp tc s l r | 
|         |   5205   assume h1: "\<And>tc s l r. \<lbrakk>t_ncorrect tp1; t_steps tc (tp1, 0) stp =  | 
|         |   5206    (s, l, r); s \<noteq> 0\<rbrakk> \<Longrightarrow> t_steps tc (tp1 @ tp2, 0) stp = (s, l, r)" | 
|         |   5207     and h2: "t_steps tc (tp1, 0) (Suc stp) = (s, l, r)" "s \<noteq> 0"  | 
|         |   5208             "t_ncorrect tp1" | 
|         |   5209   thus "t_steps tc (tp1 @ tp2, 0) (Suc stp) = (s, l, r)" | 
|         |   5210     apply(simp add: t_steps.simps) | 
|         |   5211     apply(case_tac "(t_step tc (tp1, 0))", simp) | 
|         |   5212     proof- | 
|         |   5213       fix a b c  | 
|         |   5214       assume g1: "\<And>tc s l r. \<lbrakk>t_steps tc (tp1, 0) stp = (s, l, r);  | 
|         |   5215                 0 < s\<rbrakk> \<Longrightarrow> t_steps tc (tp1 @ tp2, 0) stp = (s, l, r)" | 
|         |   5216 	and g2: "t_step tc (tp1, 0) = (a, b, c)"  | 
|         |   5217                 "t_steps (a, b, c) (tp1, 0) stp = (s, l, r)"  | 
|         |   5218                 "0 < s"  | 
|         |   5219                 "t_ncorrect tp1" | 
|         |   5220       hence g3: "a > 0" | 
|         |   5221 	apply(case_tac "a::nat", auto simp: t_steps.simps) | 
|         |   5222 	apply(simp add: state0_ind) | 
|         |   5223 	done | 
|         |   5224       from g1 and g2 and this have g4:  | 
|         |   5225                     "(t_step tc (tp1 @ tp2, 0)) = (a, b, c)" | 
|         |   5226 	apply(rule_tac tm_append_step, simp, simp, simp) | 
|         |   5227 	done | 
|         |   5228       from g1 and g2 and g3 and g4 show  | 
|         |   5229           "t_steps (t_step tc (tp1 @ tp2, 0)) (tp1 @ tp2, 0) stp | 
|         |   5230                                                          = (s, l, r)" | 
|         |   5231 	apply(simp) | 
|         |   5232 	done | 
|         |   5233     qed | 
|         |   5234 qed | 
|         |   5235  | 
|         |   5236 lemma shift_fetch:  | 
|         |   5237  "\<lbrakk>n < length tp;  | 
|         |   5238   (tp:: (taction \<times> nat) list) ! n = (aa, ba); | 
|         |   5239    ba \<noteq> 0\<rbrakk>  | 
|         |   5240    \<Longrightarrow> (tshift tp (length tp div 2)) ! n =  | 
|         |   5241                      (aa , ba + length tp div 2)" | 
|         |   5242 apply(simp add: tshift.simps) | 
|         |   5243 done | 
|         |   5244  | 
|         |   5245 lemma tshift_length_equal: "length (tshift tp q) = length tp" | 
|         |   5246 apply(auto simp: tshift.simps) | 
|         |   5247 done | 
|         |   5248  | 
|         |   5249 thm nth_of.simps | 
|         |   5250  | 
|         |   5251 lemma [simp]: "t_ncorrect tp \<Longrightarrow> 2 * (length tp div 2) = length tp" | 
|         |   5252 apply(auto simp: t_ncorrect.simps) | 
|         |   5253 done | 
|         |   5254  | 
|         |   5255 lemma  tm_append_step_equal':  | 
|         |   5256    "\<lbrakk>t_ncorrect tp; t_ncorrect tp'; off = length tp div 2\<rbrakk> \<Longrightarrow>  | 
|         |   5257     (\<lambda> (s, l, r). ((\<lambda> (s', l', r').  | 
|         |   5258       (s'\<noteq> 0 \<longrightarrow> (s = s' + off \<and> l = l' \<and> r = r')))  | 
|         |   5259          (t_step (a, b, c) (tp', 0)))) | 
|         |   5260                (t_step (a + off, b, c) (tp @ tshift tp' off, 0))" | 
|         |   5261 apply(simp add: t_step.simps) | 
|         |   5262 apply(case_tac a, simp add: fetch.simps) | 
|         |   5263 apply(case_tac  | 
|         |   5264 "(fetch tp' a (case c of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))", | 
|         |   5265  simp) | 
|         |   5266 apply(case_tac  | 
|         |   5267 "(fetch (tp @ tshift tp' (length tp div 2)) | 
|         |   5268         (Suc (nat + length tp div 2))  | 
|         |   5269            (case c of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))",  | 
|         |   5270  simp) | 
|         |   5271 apply(case_tac "(new_tape aa (b, c))", | 
|         |   5272       case_tac "(new_tape aaa (b, c))", simp,  | 
|         |   5273       rule impI, simp add: fetch.simps split: block.splits option.splits) | 
|         |   5274 apply (auto simp: nth_of.simps t_ncorrect.simps  | 
|         |   5275                       nth_append tshift_length_equal tshift.simps split: if_splits) | 
|         |   5276 done | 
|         |   5277  | 
|         |   5278  | 
|         |   5279 lemma  tm_append_step_equal:  | 
|         |   5280  "\<lbrakk>t_ncorrect tp; t_ncorrect tp'; off = length tp div 2;  | 
|         |   5281    t_step (a, b, c) (tp', 0) = (aa, ab, bb);  aa \<noteq> 0\<rbrakk> | 
|         |   5282  \<Longrightarrow> t_step (a + length tp div 2, b, c)  | 
|         |   5283         (tp @ tshift tp' (length tp div 2), 0) | 
|         |   5284                           = (aa + length tp div 2, ab, bb)" | 
|         |   5285 apply(insert tm_append_step_equal'[of tp tp' off a b c], simp) | 
|         |   5286 apply(case_tac "(t_step (a + length tp div 2, b, c)  | 
|         |   5287                    (tp @ tshift tp' (length tp div 2), 0))", simp) | 
|         |   5288 done | 
|         |   5289  | 
|         |   5290 lemma tm_append_steps_equal:  | 
|         |   5291  "\<lbrakk>t_ncorrect tp; t_ncorrect tp'; off = length tp div 2\<rbrakk> \<Longrightarrow>  | 
|         |   5292    (\<lambda> (s, l, r). ((\<lambda> (s', l', r'). ((s'\<noteq> 0 \<longrightarrow> s = s' + off \<and> l = l' | 
|         |   5293                      \<and> r = r'))) (t_steps (a, b, c) (tp', 0) stp))) | 
|         |   5294    (t_steps (a + off, b, c) (tp @ tshift tp' off, 0) stp)" | 
|         |   5295 apply(induct stp arbitrary : a b c, simp add: t_steps.simps) | 
|         |   5296 apply(simp add: t_steps.simps) | 
|         |   5297 apply(case_tac "(t_step (a, b, c) (tp', 0))", simp) | 
|         |   5298 apply(case_tac "aa = 0", simp add: state0_ind) | 
|         |   5299 apply(subgoal_tac "(t_step (a + length tp div 2, b, c)  | 
|         |   5300                       (tp @ tshift tp' (length tp div 2), 0))  | 
|         |   5301   = (aa + length tp div 2, ba, ca)", simp) | 
|         |   5302 apply(rule tm_append_step_equal, auto) | 
|         |   5303 done | 
|         |   5304  | 
|         |   5305 (*********Begin: mop_up***************) | 
|         |   5306 type_synonym mopup_type = "t_conf \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> block list \<Rightarrow> bool" | 
|         |   5307  | 
|         |   5308 fun mopup_stop :: "mopup_type" | 
|         |   5309   where | 
|         |   5310   "mopup_stop (s, l, r) lm n ires=  | 
|         |   5311         (\<exists> ln rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \<and> r = <abc_lm_v lm n> @ Bk\<^bsup>rn\<^esup>)" | 
|         |   5312  | 
|         |   5313 fun mopup_bef_erase_a :: "mopup_type" | 
|         |   5314   where | 
|         |   5315   "mopup_bef_erase_a (s, l, r) lm n ires=  | 
|         |   5316          (\<exists> ln m rn. l = Bk \<^bsup>ln\<^esup> @ Bk # Bk # ires \<and>  | 
|         |   5317                   r = Oc\<^bsup>m \<^esup>@ Bk # <(drop ((s + 1) div 2) lm)> @ Bk\<^bsup>rn\<^esup>)" | 
|         |   5318  | 
|         |   5319 fun mopup_bef_erase_b :: "mopup_type" | 
|         |   5320   where | 
|         |   5321   "mopup_bef_erase_b (s, l, r) lm n ires =  | 
|         |   5322       (\<exists> ln m rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \<and> r = Bk # Oc\<^bsup>m\<^esup> @ Bk #  | 
|         |   5323                                       <(drop (s div 2) lm)> @ Bk\<^bsup>rn\<^esup>)" | 
|         |   5324  | 
|         |   5325  | 
|         |   5326 fun mopup_jump_over1 :: "mopup_type" | 
|         |   5327   where | 
|         |   5328   "mopup_jump_over1 (s, l, r) lm n ires =  | 
|         |   5329       (\<exists> ln m1 m2 rn. m1 + m2 = Suc (abc_lm_v lm n) \<and>  | 
|         |   5330         l = Oc\<^bsup>m1\<^esup> @ Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires \<and>  | 
|         |   5331      (r = Oc\<^bsup>m2\<^esup> @ Bk # <(drop (Suc n) lm)> @ Bk\<^bsup>rn \<^esup>\<or>  | 
|         |   5332      (r = Oc\<^bsup>m2\<^esup> \<and> (drop (Suc n) lm) = [])))" | 
|         |   5333  | 
|         |   5334 fun mopup_aft_erase_a :: "mopup_type" | 
|         |   5335   where | 
|         |   5336   "mopup_aft_erase_a (s, l, r) lm n ires =  | 
|         |   5337       (\<exists> lnl lnr rn (ml::nat list) m.  | 
|         |   5338           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>  | 
|         |   5339                                    (r = <ml> @ Bk\<^bsup>rn\<^esup>))" | 
|         |   5340  | 
|         |   5341 fun mopup_aft_erase_b :: "mopup_type" | 
|         |   5342   where | 
|         |   5343   "mopup_aft_erase_b (s, l, r) lm n ires=  | 
|         |   5344    (\<exists> lnl lnr rn (ml::nat list) m.  | 
|         |   5345       m = Suc (abc_lm_v lm n) \<and>  | 
|         |   5346       l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \<and>  | 
|         |   5347      (r = Bk # <ml> @ Bk\<^bsup>rn \<^esup>\<or> | 
|         |   5348       r = Bk # Bk # <ml> @ Bk\<^bsup>rn\<^esup>))" | 
|         |   5349  | 
|         |   5350 fun mopup_aft_erase_c :: "mopup_type" | 
|         |   5351   where | 
|         |   5352   "mopup_aft_erase_c (s, l, r) lm n ires =  | 
|         |   5353  (\<exists> lnl lnr rn (ml::nat list) m.  | 
|         |   5354      m = Suc (abc_lm_v lm n) \<and>  | 
|         |   5355      l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \<and>  | 
|         |   5356     (r = <ml> @ Bk\<^bsup>rn \<^esup>\<or> r = Bk # <ml> @ Bk\<^bsup>rn\<^esup>))" | 
|         |   5357  | 
|         |   5358 fun mopup_left_moving :: "mopup_type" | 
|         |   5359   where | 
|         |   5360   "mopup_left_moving (s, l, r) lm n ires =  | 
|         |   5361   (\<exists> lnl lnr rn m. | 
|         |   5362      m = Suc (abc_lm_v lm n) \<and>  | 
|         |   5363    ((l = Bk\<^bsup>lnr \<^esup>@ Oc\<^bsup>m \<^esup>@ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \<and> r = Bk\<^bsup>rn\<^esup>) \<or> | 
|         |   5364     (l = Oc\<^bsup>m - 1\<^esup> @ Bk\<^bsup>lnl\<^esup> @ Bk # Bk # ires \<and> r = Oc # Bk\<^bsup>rn\<^esup>)))" | 
|         |   5365  | 
|         |   5366 fun mopup_jump_over2 :: "mopup_type" | 
|         |   5367   where | 
|         |   5368   "mopup_jump_over2 (s, l, r) lm n ires =  | 
|         |   5369      (\<exists> ln rn m1 m2. | 
|         |   5370           m1 + m2 = Suc (abc_lm_v lm n)  | 
|         |   5371         \<and> r \<noteq> []  | 
|         |   5372         \<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>))  | 
|         |   5373         \<and> (hd r = Bk \<longrightarrow> (l = Bk\<^bsup>ln\<^esup> @ Bk # ires \<and> r = Bk # Oc\<^bsup>m1 + m2\<^esup> @ Bk\<^bsup>rn\<^esup>)))" | 
|         |   5374  | 
|         |   5375  | 
|         |   5376 fun mopup_inv :: "mopup_type" | 
|         |   5377   where | 
|         |   5378   "mopup_inv (s, l, r) lm n ires =  | 
|         |   5379       (if s = 0 then mopup_stop (s, l, r) lm n ires | 
|         |   5380        else if s \<le> 2*n then | 
|         |   5381                if s mod 2 = 1 then mopup_bef_erase_a (s, l, r) lm n ires | 
|         |   5382                    else mopup_bef_erase_b (s, l, r) lm n ires | 
|         |   5383             else if s = 2*n + 1 then  | 
|         |   5384                 mopup_jump_over1 (s, l, r) lm n ires | 
|         |   5385             else if s = 2*n + 2 then mopup_aft_erase_a (s, l, r) lm n ires | 
|         |   5386             else if s = 2*n + 3 then mopup_aft_erase_b (s, l, r) lm n ires | 
|         |   5387             else if s = 2*n + 4 then mopup_aft_erase_c (s, l, r) lm n ires | 
|         |   5388             else if s = 2*n + 5 then mopup_left_moving (s, l, r) lm n ires | 
|         |   5389             else if s = 2*n + 6 then mopup_jump_over2 (s, l, r) lm n ires | 
|         |   5390             else False)" | 
|         |   5391  | 
|         |   5392 declare  | 
|         |   5393   mopup_jump_over2.simps[simp del] mopup_left_moving.simps[simp del] | 
|         |   5394   mopup_aft_erase_c.simps[simp del] mopup_aft_erase_b.simps[simp del]  | 
|         |   5395   mopup_aft_erase_a.simps[simp del] mopup_jump_over1.simps[simp del] | 
|         |   5396   mopup_bef_erase_a.simps[simp del] mopup_bef_erase_b.simps[simp del] | 
|         |   5397   mopup_stop.simps[simp del] | 
|         |   5398  | 
|         |   5399 lemma mopup_fetch_0[simp]:  | 
|         |   5400      "(fetch (mop_bef n @ tshift mp_up (2 * n)) 0 b) = (Nop, 0)" | 
|         |   5401 by(simp add: fetch.simps) | 
|         |   5402  | 
|         |   5403 lemma mop_bef_length[simp]: "length (mop_bef n) = 4 * n" | 
|         |   5404 apply(induct n, simp add: mop_bef.simps, simp add: mop_bef.simps) | 
|         |   5405 done | 
|         |   5406  | 
|         |   5407 thm findnth_nth | 
|         |   5408 lemma mop_bef_nth:  | 
|         |   5409   "\<lbrakk>q < n; x < 4\<rbrakk> \<Longrightarrow> mop_bef n ! (4 * q + x) =  | 
|         |   5410                              mop_bef (Suc q) ! ((4 * q) + x)" | 
|         |   5411 apply(induct n, simp) | 
|         |   5412 apply(case_tac "q < n", simp add: mop_bef.simps, auto) | 
|         |   5413 apply(simp add: nth_append) | 
|         |   5414 apply(subgoal_tac "q = n", simp) | 
|         |   5415 apply(arith) | 
|         |   5416 done | 
|         |   5417  | 
|         |   5418 lemma fetch_bef_erase_a_o[simp]:  | 
|         |   5419  "\<lbrakk>0 < s; s \<le> 2 * n; s mod 2 = Suc 0\<rbrakk> | 
|         |   5420   \<Longrightarrow> (fetch (mop_bef n @ tshift mp_up (2 * n)) s Oc) = (W0, s + 1)" | 
|         |   5421 apply(subgoal_tac "\<exists> q. s = 2*q + 1", auto) | 
|         |   5422 apply(subgoal_tac "length (mop_bef n) = 4*n") | 
|         |   5423 apply(auto simp: fetch.simps nth_of.simps nth_append) | 
|         |   5424 apply(subgoal_tac "mop_bef n ! (4 * q + 1) =  | 
|         |   5425                       mop_bef (Suc q) ! ((4 * q) + 1)",  | 
|         |   5426       simp add: mop_bef.simps nth_append) | 
|         |   5427 apply(rule mop_bef_nth, auto) | 
|         |   5428 done | 
|         |   5429  | 
|         |   5430 lemma fetch_bef_erase_a_b[simp]: | 
|         |   5431   "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = Suc 0\<rbrakk> | 
|         |   5432    \<Longrightarrow>  (fetch (mop_bef n @ tshift mp_up (2 * n)) s Bk) = (R, s + 2)" | 
|         |   5433 apply(subgoal_tac "\<exists> q. s = 2*q + 1", auto) | 
|         |   5434 apply(subgoal_tac "length (mop_bef n) = 4*n") | 
|         |   5435 apply(auto simp: fetch.simps nth_of.simps nth_append) | 
|         |   5436 apply(subgoal_tac "mop_bef n ! (4 * q + 0) =  | 
|         |   5437                        mop_bef (Suc q) ! ((4 * q + 0))",  | 
|         |   5438       simp add: mop_bef.simps nth_append) | 
|         |   5439 apply(rule mop_bef_nth, auto) | 
|         |   5440 done | 
|         |   5441  | 
|         |   5442 lemma fetch_bef_erase_b_b:  | 
|         |   5443   "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = 0\<rbrakk> \<Longrightarrow>  | 
|         |   5444      (fetch (mop_bef n @ tshift mp_up (2 * n)) s Bk) = (R, s - 1)" | 
|         |   5445 apply(subgoal_tac "\<exists> q. s = 2 * q", auto) | 
|         |   5446 apply(case_tac qa, simp, simp) | 
|         |   5447 apply(auto simp: fetch.simps nth_of.simps nth_append) | 
|         |   5448 apply(subgoal_tac "mop_bef n ! (4 * nat + 2) =  | 
|         |   5449                      mop_bef (Suc nat) ! ((4 * nat) + 2)",  | 
|         |   5450       simp add: mop_bef.simps nth_append) | 
|         |   5451 apply(rule mop_bef_nth, auto) | 
|         |   5452 done | 
|         |   5453  | 
|         |   5454 lemma fetch_jump_over1_o:  | 
|         |   5455  "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (2 * n)) Oc | 
|         |   5456   = (R, Suc (2 * n))" | 
|         |   5457 apply(subgoal_tac "length (mop_bef n) = 4 * n") | 
|         |   5458 apply(auto simp: fetch.simps nth_of.simps mp_up_def nth_append  | 
|         |   5459                  tshift.simps) | 
|         |   5460 done | 
|         |   5461  | 
|         |   5462 lemma fetch_jump_over1_b:  | 
|         |   5463  "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (2 * n)) Bk  | 
|         |   5464  = (R, Suc (Suc (2 * n)))" | 
|         |   5465 apply(subgoal_tac "length (mop_bef n) = 4 * n") | 
|         |   5466 apply(auto simp: fetch.simps nth_of.simps mp_up_def  | 
|         |   5467                  nth_append tshift.simps) | 
|         |   5468 done | 
|         |   5469  | 
|         |   5470 lemma fetch_aft_erase_a_o:  | 
|         |   5471  "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (Suc (2 * n))) Oc  | 
|         |   5472  = (W0, Suc (2 * n + 2))" | 
|         |   5473 apply(subgoal_tac "length (mop_bef n) = 4 * n") | 
|         |   5474 apply(auto simp: fetch.simps nth_of.simps mp_up_def  | 
|         |   5475                  nth_append tshift.simps) | 
|         |   5476 done | 
|         |   5477  | 
|         |   5478 lemma fetch_aft_erase_a_b:  | 
|         |   5479  "fetch (mop_bef n @ tshift mp_up (2 * n)) (Suc (Suc (2 * n))) Bk | 
|         |   5480   = (L, Suc (2 * n + 4))" | 
|         |   5481 apply(subgoal_tac "length (mop_bef n) = 4 * n") | 
|         |   5482 apply(auto simp: fetch.simps nth_of.simps mp_up_def  | 
|         |   5483                  nth_append tshift.simps) | 
|         |   5484 done | 
|         |   5485  | 
|         |   5486 lemma fetch_aft_erase_b_b:  | 
|         |   5487  "fetch (mop_bef n @ tshift mp_up (2 * n)) (2*n + 3) Bk | 
|         |   5488   = (R, Suc (2 * n + 3))" | 
|         |   5489 apply(subgoal_tac "length (mop_bef n) = 4 * n") | 
|         |   5490 apply(subgoal_tac "2*n + 3 = Suc (2*n + 2)", simp only: fetch.simps) | 
|         |   5491 apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) | 
|         |   5492 done | 
|         |   5493  | 
|         |   5494 lemma fetch_aft_erase_c_o:  | 
|         |   5495  "fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 4) Oc  | 
|         |   5496  = (W0, Suc (2 * n + 2))" | 
|         |   5497 apply(subgoal_tac "length (mop_bef n) = 4 * n") | 
|         |   5498 apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps) | 
|         |   5499 apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) | 
|         |   5500 done | 
|         |   5501  | 
|         |   5502 lemma fetch_aft_erase_c_b:  | 
|         |   5503  "fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 4) Bk  | 
|         |   5504  = (R, Suc (2 * n + 1))" | 
|         |   5505 apply(subgoal_tac "length (mop_bef n) = 4 * n") | 
|         |   5506 apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps) | 
|         |   5507 apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) | 
|         |   5508 done | 
|         |   5509  | 
|         |   5510 lemma fetch_left_moving_o:  | 
|         |   5511  "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 5) Oc)  | 
|         |   5512  = (L, 2*n + 6)" | 
|         |   5513 apply(subgoal_tac "length (mop_bef n) = 4 * n") | 
|         |   5514 apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps) | 
|         |   5515 apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) | 
|         |   5516 done | 
|         |   5517  | 
|         |   5518 lemma fetch_left_moving_b:  | 
|         |   5519  "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 5) Bk) | 
|         |   5520   = (L, 2*n + 5)" | 
|         |   5521 apply(subgoal_tac "length (mop_bef n) = 4 * n") | 
|         |   5522 apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps) | 
|         |   5523 apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) | 
|         |   5524 done | 
|         |   5525  | 
|         |   5526 lemma fetch_jump_over2_b: | 
|         |   5527   "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 6) Bk)  | 
|         |   5528  = (R, 0)" | 
|         |   5529 apply(subgoal_tac "length (mop_bef n) = 4 * n") | 
|         |   5530 apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps) | 
|         |   5531 apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) | 
|         |   5532 done | 
|         |   5533  | 
|         |   5534 lemma fetch_jump_over2_o:  | 
|         |   5535 "(fetch (mop_bef n @ tshift mp_up (2 * n)) (2 * n + 6) Oc)  | 
|         |   5536  = (L, 2*n + 6)" | 
|         |   5537 apply(subgoal_tac "length (mop_bef n) = 4 * n") | 
|         |   5538 apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps) | 
|         |   5539 apply(auto simp: nth_of.simps mp_up_def nth_append tshift.simps) | 
|         |   5540 done | 
|         |   5541  | 
|         |   5542 lemmas mopupfetchs =  | 
|         |   5543 fetch_bef_erase_a_o fetch_bef_erase_a_b fetch_bef_erase_b_b  | 
|         |   5544 fetch_jump_over1_o fetch_jump_over1_b fetch_aft_erase_a_o  | 
|         |   5545 fetch_aft_erase_a_b fetch_aft_erase_b_b fetch_aft_erase_c_o  | 
|         |   5546 fetch_aft_erase_c_b fetch_left_moving_o fetch_left_moving_b  | 
|         |   5547 fetch_jump_over2_b fetch_jump_over2_o | 
|         |   5548  | 
|         |   5549 lemma [simp]:  | 
|         |   5550 "\<lbrakk>n < length lm; 0 < s; s mod 2 = Suc 0;  | 
|         |   5551   mopup_bef_erase_a (s, l, Oc # xs) lm n ires;  | 
|         |   5552   Suc s \<le> 2 * n\<rbrakk> \<Longrightarrow>  | 
|         |   5553   mopup_bef_erase_b (Suc s, l, Bk # xs) lm n ires" | 
|         |   5554 apply(auto simp: mopup_bef_erase_a.simps mopup_bef_erase_b.simps ) | 
|         |   5555 apply(rule_tac x = "m - 1" in exI, rule_tac x = rn in exI) | 
|         |   5556 apply(case_tac m, simp, simp) | 
|         |   5557 done | 
|         |   5558  | 
|         |   5559 lemma mopup_false1: | 
|         |   5560   "\<lbrakk>0 < s; s \<le> 2 * n; s mod 2 = Suc 0;  \<not> Suc s \<le> 2 * n\<rbrakk>  | 
|         |   5561   \<Longrightarrow> RR" | 
|         |   5562 apply(arith) | 
|         |   5563 done | 
|         |   5564  | 
|         |   5565 lemma [simp]:  | 
|         |   5566  "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = Suc 0;  | 
|         |   5567    mopup_bef_erase_a (s, l, Oc # xs) lm n ires; r = Oc # xs\<rbrakk> | 
|         |   5568  \<Longrightarrow> (Suc s \<le> 2 * n \<longrightarrow> mopup_bef_erase_b (Suc s, l, Bk # xs) lm n ires)  \<and> | 
|         |   5569      (\<not> Suc s \<le> 2 * n \<longrightarrow> mopup_jump_over1 (Suc s, l, Bk # xs) lm n ires) " | 
|         |   5570 apply(auto elim: mopup_false1) | 
|         |   5571 done | 
|         |   5572  | 
|         |   5573 lemma drop_abc_lm_v_simp[simp]:  | 
|         |   5574    "n < length lm \<Longrightarrow> drop n lm = abc_lm_v lm n # drop (Suc n) lm" | 
|         |   5575 apply(auto simp: abc_lm_v.simps) | 
|         |   5576 apply(drule drop_Suc_conv_tl, simp) | 
|         |   5577 done | 
|         |   5578 lemma [simp]: "(\<exists>rna. Bk\<^bsup>rn\<^esup> = Bk # Bk\<^bsup>rna\<^esup>) \<or> Bk\<^bsup>rn\<^esup> = []" | 
|         |   5579 apply(case_tac rn, simp, auto) | 
|         |   5580 done | 
|         |   5581  | 
|         |   5582 lemma [simp]: "\<exists>lna. Bk # Bk\<^bsup>ln\<^esup> = Bk\<^bsup>lna\<^esup>" | 
|         |   5583 apply(rule_tac x = "Suc ln" in exI, auto) | 
|         |   5584 done | 
|         |   5585  | 
|         |   5586 lemma mopup_bef_erase_a_2_jump_over[simp]:  | 
|         |   5587  "\<lbrakk>n < length lm; 0 < s; s mod 2 = Suc 0;  | 
|         |   5588    mopup_bef_erase_a (s, l, Bk # xs) lm n ires; Suc s = 2 * n\<rbrakk>  | 
|         |   5589 \<Longrightarrow> mopup_jump_over1 (Suc (2 * n), Bk # l, xs) lm n ires" | 
|         |   5590 apply(auto simp: mopup_bef_erase_a.simps mopup_jump_over1.simps) | 
|         |   5591 apply(case_tac m, simp) | 
|         |   5592 apply(rule_tac x = "Suc ln" in exI, rule_tac x = 0 in exI,  | 
|         |   5593       simp add: tape_of_nl_abv) | 
|         |   5594 apply(case_tac "drop (Suc n) lm", auto simp: tape_of_nat_list.simps ) | 
|         |   5595 done | 
|         |   5596  | 
|         |   5597 lemma Suc_Suc_div:  "\<lbrakk>0 < s; s mod 2 = Suc 0; Suc (Suc s) \<le> 2 * n\<rbrakk> | 
|         |   5598            \<Longrightarrow> (Suc (Suc (s div 2))) \<le> n" | 
|         |   5599 apply(arith) | 
|         |   5600 done | 
|         |   5601  | 
|         |   5602 lemma mopup_bef_erase_a_2_a[simp]:  | 
|         |   5603  "\<lbrakk>n < length lm; 0 < s; s mod 2 = Suc 0;  | 
|         |   5604    mopup_bef_erase_a (s, l, Bk # xs) lm n ires;  | 
|         |   5605    Suc (Suc s) \<le> 2 * n\<rbrakk> \<Longrightarrow>  | 
|         |   5606    mopup_bef_erase_a (Suc (Suc s), Bk # l, xs) lm n ires" | 
|         |   5607 apply(auto simp: mopup_bef_erase_a.simps ) | 
|         |   5608 apply(subgoal_tac "drop (Suc (Suc (s div 2))) lm \<noteq> []") | 
|         |   5609 apply(case_tac m, simp) | 
|         |   5610 apply(rule_tac x = "Suc (abc_lm_v lm (Suc (s div 2)))" in exI,  | 
|         |   5611       rule_tac x = rn in exI, simp, simp) | 
|         |   5612 apply(subgoal_tac "(Suc (Suc (s div 2))) \<le> n", simp,  | 
|         |   5613       rule_tac Suc_Suc_div, auto) | 
|         |   5614 done | 
|         |   5615  | 
|         |   5616 lemma mopup_false2:  | 
|         |   5617  "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n;  | 
|         |   5618    s mod 2 = Suc 0; Suc s \<noteq> 2 * n; | 
|         |   5619    \<not> Suc (Suc s) \<le> 2 * n\<rbrakk> \<Longrightarrow> RR" | 
|         |   5620 apply(arith) | 
|         |   5621 done | 
|         |   5622  | 
|         |   5623 lemma [simp]: | 
|         |   5624   "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n;  | 
|         |   5625    s mod 2 = Suc 0;  | 
|         |   5626    mopup_bef_erase_a (s, l, Bk # xs) lm n ires;  | 
|         |   5627    r = Bk # xs\<rbrakk> | 
|         |   5628  \<Longrightarrow> (Suc s = 2 * n \<longrightarrow>  | 
|         |   5629              mopup_jump_over1 (Suc (2 * n), Bk # l, xs) lm n ires) \<and> | 
|         |   5630      (Suc s \<noteq> 2 * n \<longrightarrow>  | 
|         |   5631        (Suc (Suc s) \<le> 2 * n \<longrightarrow>  | 
|         |   5632           mopup_bef_erase_a (Suc (Suc s), Bk # l, xs) lm n ires) \<and>  | 
|         |   5633        (\<not> Suc (Suc s) \<le> 2 * n \<longrightarrow>  | 
|         |   5634           mopup_aft_erase_a (Suc (Suc s), Bk # l, xs) lm n ires))" | 
|         |   5635 apply(auto elim: mopup_false2) | 
|         |   5636 done | 
|         |   5637  | 
|         |   5638 lemma [simp]: "mopup_bef_erase_a (s, l, []) lm n ires \<Longrightarrow>  | 
|         |   5639                         mopup_bef_erase_a (s, l, [Bk]) lm n ires" | 
|         |   5640 apply(auto simp: mopup_bef_erase_a.simps) | 
|         |   5641 done | 
|         |   5642  | 
|         |   5643 lemma [simp]: | 
|         |   5644    "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = Suc 0; | 
|         |   5645      mopup_bef_erase_a (s, l, []) lm n ires; r = []\<rbrakk> | 
|         |   5646     \<Longrightarrow> (Suc s = 2 * n \<longrightarrow>  | 
|         |   5647               mopup_jump_over1 (Suc (2 * n), Bk # l, []) lm n ires) \<and> | 
|         |   5648         (Suc s \<noteq> 2 * n \<longrightarrow>  | 
|         |   5649              (Suc (Suc s) \<le> 2 * n \<longrightarrow>  | 
|         |   5650                  mopup_bef_erase_a (Suc (Suc s), Bk # l, []) lm n ires) \<and> | 
|         |   5651              (\<not> Suc (Suc s) \<le> 2 * n \<longrightarrow>  | 
|         |   5652                  mopup_aft_erase_a (Suc (Suc s), Bk # l, []) lm n ires))" | 
|         |   5653 apply(auto) | 
|         |   5654 done | 
|         |   5655  | 
|         |   5656 lemma "mopup_bef_erase_b (s, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []" | 
|         |   5657 apply(auto simp: mopup_bef_erase_b.simps) | 
|         |   5658 done | 
|         |   5659  | 
|         |   5660 lemma [simp]: "mopup_bef_erase_b (s, l, Oc # xs) lm n ires = False" | 
|         |   5661 apply(auto simp: mopup_bef_erase_b.simps ) | 
|         |   5662 done | 
|         |   5663   | 
|         |   5664 lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow>  | 
|         |   5665                                       (s - Suc 0) mod 2 = Suc 0" | 
|         |   5666 apply(arith) | 
|         |   5667 done | 
|         |   5668  | 
|         |   5669 lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow> | 
|         |   5670                                        s - Suc 0 \<le> 2 * n" | 
|         |   5671 apply(simp) | 
|         |   5672 done | 
|         |   5673  | 
|         |   5674 lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow> \<not> s \<le> Suc 0" | 
|         |   5675 apply(arith) | 
|         |   5676 done | 
|         |   5677  | 
|         |   5678 lemma [simp]: "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n;  | 
|         |   5679                s mod 2 \<noteq> Suc 0;  | 
|         |   5680                mopup_bef_erase_b (s, l, Bk # xs) lm n ires; r = Bk # xs\<rbrakk>  | 
|         |   5681            \<Longrightarrow> mopup_bef_erase_a (s - Suc 0, Bk # l, xs) lm n ires" | 
|         |   5682 apply(auto simp: mopup_bef_erase_b.simps mopup_bef_erase_a.simps) | 
|         |   5683 done | 
|         |   5684  | 
|         |   5685 lemma [simp]: "\<lbrakk>mopup_bef_erase_b (s, l, []) lm n ires\<rbrakk> \<Longrightarrow>  | 
|         |   5686                    mopup_bef_erase_a (s - Suc 0, Bk # l, []) lm n ires" | 
|         |   5687 apply(auto simp: mopup_bef_erase_b.simps mopup_bef_erase_a.simps) | 
|         |   5688 done | 
|         |   5689  | 
|         |   5690 lemma [simp]:  | 
|         |   5691    "\<lbrakk>n < length lm; | 
|         |   5692     mopup_jump_over1 (Suc (2 * n), l, Oc # xs) lm n ires; | 
|         |   5693     r = Oc # xs\<rbrakk> | 
|         |   5694   \<Longrightarrow> mopup_jump_over1 (Suc (2 * n), Oc # l, xs) lm n ires" | 
|         |   5695 apply(auto simp: mopup_jump_over1.simps) | 
|         |   5696 apply(rule_tac x = ln in exI, rule_tac x = "Suc m1" in exI, | 
|         |   5697        rule_tac x = "m2 - 1" in exI) | 
|         |   5698 apply(case_tac "m2", simp, simp, rule_tac x = rn in exI, simp) | 
|         |   5699 apply(rule_tac x = ln in exI, rule_tac x = "Suc m1" in exI,  | 
|         |   5700       rule_tac x = "m2 - 1" in exI) | 
|         |   5701 apply(case_tac m2, simp, simp) | 
|         |   5702 done | 
|         |   5703  | 
|         |   5704 lemma mopup_jump_over1_2_aft_erase_a[simp]:   | 
|         |   5705  "\<lbrakk>n < length lm; mopup_jump_over1 (Suc (2 * n), l, Bk # xs) lm n ires\<rbrakk> | 
|         |   5706   \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, xs) lm n ires" | 
|         |   5707 apply(simp only: mopup_jump_over1.simps mopup_aft_erase_a.simps) | 
|         |   5708 apply(erule_tac exE)+ | 
|         |   5709 apply(rule_tac x = ln in exI, rule_tac x = "Suc 0" in exI) | 
|         |   5710 apply(case_tac m2, simp) | 
|         |   5711 apply(rule_tac x = rn in exI, rule_tac x = "drop (Suc n) lm" in exI,  | 
|         |   5712       simp) | 
|         |   5713 apply(simp) | 
|         |   5714 done | 
|         |   5715  | 
|         |   5716 lemma [simp]:  | 
|         |   5717  "\<lbrakk>n < length lm; mopup_jump_over1 (Suc (2 * n), l, []) lm n ires\<rbrakk> \<Longrightarrow>  | 
|         |   5718     mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, []) lm n ires" | 
|         |   5719 apply(rule mopup_jump_over1_2_aft_erase_a, simp) | 
|         |   5720 apply(auto simp: mopup_jump_over1.simps) | 
|         |   5721 apply(rule_tac x = ln in exI, rule_tac x = m1 in exI,  | 
|         |   5722       rule_tac x = m2 in exI, simp add: ) | 
|         |   5723 apply(rule_tac x = 0 in exI, auto) | 
|         |   5724 done | 
|         |   5725  | 
|         |   5726 lemma [simp]:  | 
|         |   5727  "\<lbrakk>n < length lm;  | 
|         |   5728    mopup_aft_erase_a (Suc (Suc (2 * n)), l, Oc # xs) lm n ires\<rbrakk>  | 
|         |   5729  \<Longrightarrow> mopup_aft_erase_b (Suc (Suc (Suc (2 * n))), l, Bk # xs) lm n ires" | 
|         |   5730 apply(auto simp: mopup_aft_erase_a.simps mopup_aft_erase_b.simps ) | 
|         |   5731 apply(case_tac ml, simp, case_tac rn, simp, simp) | 
|         |   5732 apply(case_tac list, auto simp: tape_of_nl_abv  | 
|         |   5733                                 tape_of_nat_list.simps ) | 
|         |   5734 apply(case_tac a, simp, rule_tac x = rn in exI,  | 
|         |   5735       rule_tac x = "[]" in exI, | 
|         |   5736        simp add: tape_of_nat_list.simps, simp) | 
|         |   5737 apply(rule_tac x = rn in exI, rule_tac x = "[nat]" in exI,  | 
|         |   5738       simp add: tape_of_nat_list.simps ) | 
|         |   5739 apply(case_tac a, simp, rule_tac x = rn in exI,  | 
|         |   5740        rule_tac x = "aa # lista" in exI, simp, simp) | 
|         |   5741 apply(rule_tac x = rn in exI, rule_tac x = "nat # aa # lista" in exI,  | 
|         |   5742        simp add: tape_of_nat_list.simps ) | 
|         |   5743 done | 
|         |   5744  | 
|         |   5745 lemma [simp]: | 
|         |   5746   "mopup_aft_erase_a (Suc (Suc (2 * n)), l, Bk # xs) lm n ires \<Longrightarrow> l \<noteq> []" | 
|         |   5747 apply(auto simp: mopup_aft_erase_a.simps) | 
|         |   5748 done | 
|         |   5749  | 
|         |   5750 lemma [simp]: | 
|         |   5751   "\<lbrakk>n < length lm; | 
|         |   5752     mopup_aft_erase_a (Suc (Suc (2 * n)), l, Bk # xs) lm n ires\<rbrakk> | 
|         |   5753   \<Longrightarrow> mopup_left_moving (5 + 2 * n, tl l, hd l # Bk # xs) lm n ires" | 
|         |   5754 apply(simp only: mopup_aft_erase_a.simps mopup_left_moving.simps) | 
|         |   5755 apply(erule exE)+ | 
|         |   5756 apply(case_tac lnr, simp) | 
|         |   5757 apply(rule_tac x = lnl in exI, simp, rule_tac x = rn in exI, simp) | 
|         |   5758 apply(subgoal_tac "ml = []", simp) | 
|         |   5759 apply(rule_tac xs = xs and rn = rn in BkCons_nil, simp, auto) | 
|         |   5760 apply(subgoal_tac "ml = []", auto) | 
|         |   5761 apply(rule_tac xs = xs and rn = rn in BkCons_nil, simp) | 
|         |   5762 done | 
|         |   5763  | 
|         |   5764 lemma [simp]: | 
|         |   5765   "mopup_aft_erase_a (Suc (Suc (2 * n)), l, []) lm n ires \<Longrightarrow> l \<noteq> []" | 
|         |   5766 apply(simp only: mopup_aft_erase_a.simps) | 
|         |   5767 apply(erule exE)+ | 
|         |   5768 apply(auto) | 
|         |   5769 done | 
|         |   5770  | 
|         |   5771 lemma [simp]: | 
|         |   5772   "\<lbrakk>n < length lm; mopup_aft_erase_a (Suc (Suc (2 * n)), l, []) lm n ires\<rbrakk> | 
|         |   5773   \<Longrightarrow> mopup_left_moving (5 + 2 * n, tl l, [hd l]) lm n ires" | 
|         |   5774 apply(simp only: mopup_aft_erase_a.simps mopup_left_moving.simps) | 
|         |   5775 apply(erule exE)+ | 
|         |   5776 apply(subgoal_tac "ml = [] \<and> rn = 0", erule conjE, erule conjE, simp) | 
|         |   5777 apply(case_tac lnr, simp, rule_tac x = lnl in exI, simp,  | 
|         |   5778       rule_tac x = 0 in exI, simp) | 
|         |   5779 apply(rule_tac x = lnl in exI, rule_tac x = nat in exI,  | 
|         |   5780       rule_tac x = "Suc 0" in exI, simp) | 
|         |   5781 apply(case_tac ml, simp, case_tac rn, simp, simp) | 
|         |   5782 apply(case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps) | 
|         |   5783 done | 
|         |   5784  | 
|         |   5785 lemma [simp]: "mopup_aft_erase_b (2 * n + 3, l, Oc # xs) lm n ires = False" | 
|         |   5786 apply(auto simp: mopup_aft_erase_b.simps ) | 
|         |   5787 done | 
|         |   5788  | 
|         |   5789 lemma [simp]:  | 
|         |   5790  "\<lbrakk>n < length lm;  | 
|         |   5791    mopup_aft_erase_c (2 * n + 4, l, Oc # xs) lm n ires\<rbrakk> | 
|         |   5792   \<Longrightarrow> mopup_aft_erase_b (Suc (Suc (Suc (2 * n))), l, Bk # xs) lm n ires" | 
|         |   5793 apply(auto simp: mopup_aft_erase_c.simps mopup_aft_erase_b.simps ) | 
|         |   5794 apply(case_tac ml, simp, case_tac rn, simp, simp, simp) | 
|         |   5795 apply(case_tac list, auto simp: tape_of_nl_abv  | 
|         |   5796                         tape_of_nat_list.simps tape_of_nat_abv ) | 
|         |   5797 apply(case_tac a, rule_tac x = rn in exI,  | 
|         |   5798       rule_tac x = "[]" in exI, simp add: tape_of_nat_list.simps) | 
|         |   5799 apply(rule_tac x = rn in exI, rule_tac x = "[nat]" in exI,  | 
|         |   5800       simp add: tape_of_nat_list.simps ) | 
|         |   5801 apply(case_tac a, simp, rule_tac x = rn in exI,  | 
|         |   5802       rule_tac x = "aa # lista" in exI, simp) | 
|         |   5803 apply(rule_tac x = rn in exI, rule_tac x = "nat # aa # lista" in exI,  | 
|         |   5804       simp add: tape_of_nat_list.simps ) | 
|         |   5805 done | 
|         |   5806  | 
|         |   5807 lemma mopup_aft_erase_c_aft_erase_a[simp]:  | 
|         |   5808  "\<lbrakk>n < length lm; mopup_aft_erase_c (2 * n + 4, l, Bk # xs) lm n ires\<rbrakk>  | 
|         |   5809  \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, xs) lm n ires" | 
|         |   5810 apply(simp only: mopup_aft_erase_c.simps mopup_aft_erase_a.simps ) | 
|         |   5811 apply(erule_tac exE)+ | 
|         |   5812 apply(erule conjE, erule conjE, erule disjE) | 
|         |   5813 apply(subgoal_tac "ml = []", simp, case_tac rn,  | 
|         |   5814       simp, simp, rule conjI) | 
|         |   5815 apply(rule_tac x = lnl in exI, rule_tac x = "Suc lnr" in exI, simp) | 
|         |   5816 apply(rule_tac x = nat in exI, rule_tac x = "[]" in exI, simp) | 
|         |   5817 apply(rule_tac xs = xs and rn = rn in BkCons_nil, simp, simp,  | 
|         |   5818       rule conjI) | 
|         |   5819 apply(rule_tac x = lnl in exI, rule_tac x = "Suc lnr" in exI, simp) | 
|         |   5820 apply(rule_tac x = rn in exI, rule_tac x = "ml" in exI, simp) | 
|         |   5821 done | 
|         |   5822  | 
|         |   5823 lemma [simp]:  | 
|         |   5824  "\<lbrakk>n < length lm; mopup_aft_erase_c (2 * n + 4, l, []) lm n ires\<rbrakk>  | 
|         |   5825  \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, []) lm n ires" | 
|         |   5826 apply(rule mopup_aft_erase_c_aft_erase_a, simp) | 
|         |   5827 apply(simp only: mopup_aft_erase_c.simps) | 
|         |   5828 apply(erule exE)+ | 
|         |   5829 apply(rule_tac x = lnl in exI, rule_tac x = lnr in exI, simp add: ) | 
|         |   5830 apply(rule_tac x = 0 in exI, rule_tac x = "[]" in exI, simp) | 
|         |   5831 done | 
|         |   5832  | 
|         |   5833 lemma mopup_aft_erase_b_2_aft_erase_c[simp]: | 
|         |   5834   "\<lbrakk>n < length lm; mopup_aft_erase_b (2 * n + 3, l, Bk # xs) lm n ires\<rbrakk>   | 
|         |   5835  \<Longrightarrow> mopup_aft_erase_c (4 + 2 * n, Bk # l, xs) lm n ires" | 
|         |   5836 apply(auto simp: mopup_aft_erase_b.simps mopup_aft_erase_c.simps) | 
|         |   5837 apply(rule_tac x = "lnl" in exI, rule_tac x = "Suc lnr" in exI, simp) | 
|         |   5838 apply(rule_tac x = "lnl" in exI, rule_tac x = "Suc lnr" in exI, simp) | 
|         |   5839 done | 
|         |   5840  | 
|         |   5841 lemma [simp]:  | 
|         |   5842  "\<lbrakk>n < length lm; mopup_aft_erase_b (2 * n + 3, l, []) lm n ires\<rbrakk>  | 
|         |   5843  \<Longrightarrow> mopup_aft_erase_c (4 + 2 * n, Bk # l, []) lm n ires" | 
|         |   5844 apply(rule_tac mopup_aft_erase_b_2_aft_erase_c, simp) | 
|         |   5845 apply(simp add: mopup_aft_erase_b.simps) | 
|         |   5846 done | 
|         |   5847  | 
|         |   5848 lemma [simp]:  | 
|         |   5849     "mopup_left_moving (2 * n + 5, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []" | 
|         |   5850 apply(auto simp: mopup_left_moving.simps) | 
|         |   5851 done | 
|         |   5852  | 
|         |   5853 lemma [simp]:   | 
|         |   5854  "\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, Oc # xs) lm n ires\<rbrakk> | 
|         |   5855   \<Longrightarrow> mopup_jump_over2 (2 * n + 6, tl l, hd l # Oc # xs) lm n ires" | 
|         |   5856 apply(simp only: mopup_left_moving.simps mopup_jump_over2.simps) | 
|         |   5857 apply(erule_tac exE)+ | 
|         |   5858 apply(erule conjE, erule disjE, erule conjE) | 
|         |   5859 apply(case_tac rn, simp, simp add: ) | 
|         |   5860 apply(case_tac "hd l", simp add:  ) | 
|         |   5861 apply(case_tac "abc_lm_v lm n", simp) | 
|         |   5862 apply(rule_tac x = "lnl" in exI, rule_tac x = rn in exI,  | 
|         |   5863       rule_tac x = "Suc 0" in exI, rule_tac x = 0 in exI) | 
|         |   5864 apply(case_tac lnl, simp, simp, simp add: exp_ind[THEN sym], simp) | 
|         |   5865 apply(case_tac "abc_lm_v lm n", simp) | 
|         |   5866 apply(case_tac lnl, simp, simp) | 
|         |   5867 apply(rule_tac x = lnl in exI, rule_tac x = rn in exI) | 
|         |   5868 apply(rule_tac x = nat in exI, rule_tac x = "Suc (Suc 0)" in exI, simp) | 
|         |   5869 done | 
|         |   5870  | 
|         |   5871 lemma [simp]: "mopup_left_moving (2 * n + 5, l, xs) lm n ires \<Longrightarrow> l \<noteq> []" | 
|         |   5872 apply(auto simp: mopup_left_moving.simps) | 
|         |   5873 done | 
|         |   5874  | 
|         |   5875 lemma [simp]: | 
|         |   5876   "\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, Bk # xs) lm n ires\<rbrakk>  | 
|         |   5877  \<Longrightarrow> mopup_left_moving (2 * n + 5, tl l, hd l # Bk # xs) lm n ires" | 
|         |   5878 apply(simp only: mopup_left_moving.simps) | 
|         |   5879 apply(erule exE)+ | 
|         |   5880 apply(case_tac lnr, simp) | 
|         |   5881 apply(rule_tac x = lnl in exI, rule_tac x = 0 in exI,  | 
|         |   5882       rule_tac x = rn in exI, simp, simp) | 
|         |   5883 apply(rule_tac x = lnl in exI, rule_tac x = nat in exI, simp) | 
|         |   5884 done | 
|         |   5885  | 
|         |   5886 lemma [simp]:  | 
|         |   5887 "\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, []) lm n ires\<rbrakk> | 
|         |   5888     \<Longrightarrow> mopup_left_moving (2 * n + 5, tl l, [hd l]) lm n ires" | 
|         |   5889 apply(simp only: mopup_left_moving.simps) | 
|         |   5890 apply(erule exE)+ | 
|         |   5891 apply(case_tac lnr, simp) | 
|         |   5892 apply(rule_tac x = lnl in exI, rule_tac x = 0 in exI,  | 
|         |   5893       rule_tac x = 0 in exI, simp, auto) | 
|         |   5894 done | 
|         |   5895  | 
|         |   5896 lemma [simp]:  | 
|         |   5897  "mopup_jump_over2 (2 * n + 6, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []" | 
|         |   5898 apply(auto simp: mopup_jump_over2.simps ) | 
|         |   5899 done | 
|         |   5900  | 
|         |   5901 lemma [intro]: "\<exists>lna. Bk # Bk\<^bsup>ln\<^esup> = Bk\<^bsup>lna\<^esup> @ [Bk]" | 
|         |   5902 apply(simp only: exp_ind[THEN sym], auto) | 
|         |   5903 done | 
|         |   5904  | 
|         |   5905 lemma [simp]:  | 
|         |   5906 "\<lbrakk>n < length lm; mopup_jump_over2 (2 * n + 6, l, Oc # xs) lm n ires\<rbrakk> | 
|         |   5907  \<Longrightarrow>  mopup_jump_over2 (2 * n + 6, tl l, hd l # Oc # xs) lm n ires" | 
|         |   5908 apply(simp only: mopup_jump_over2.simps) | 
|         |   5909 apply(erule_tac exE)+ | 
|         |   5910 apply(simp add:  , erule conjE, erule_tac conjE) | 
|         |   5911 apply(case_tac m1, simp) | 
|         |   5912 apply(rule_tac x = ln in exI, rule_tac x = rn in exI,  | 
|         |   5913       rule_tac x = 0 in exI, simp) | 
|         |   5914 apply(case_tac ln, simp, simp, simp only: exp_ind[THEN sym], simp) | 
|         |   5915 apply(rule_tac x = ln in exI, rule_tac x = rn in exI,  | 
|         |   5916       rule_tac x = nat in exI, rule_tac x = "Suc m2" in exI, simp) | 
|         |   5917 done | 
|         |   5918  | 
|         |   5919 lemma [simp]: "\<exists>rna. Oc # Oc\<^bsup>a\<^esup> @ Bk\<^bsup>rn\<^esup> = <a> @ Bk\<^bsup>rna\<^esup>" | 
|         |   5920 apply(case_tac a, auto simp: tape_of_nat_abv ) | 
|         |   5921 done | 
|         |   5922  | 
|         |   5923 lemma [simp]:  | 
|         |   5924  "\<lbrakk>n < length lm; mopup_jump_over2 (2 * n + 6, l, Bk # xs) lm n ires\<rbrakk>  | 
|         |   5925   \<Longrightarrow> mopup_stop (0, Bk # l, xs) lm n ires" | 
|         |   5926 apply(auto simp: mopup_jump_over2.simps mopup_stop.simps) | 
|         |   5927 done | 
|         |   5928  | 
|         |   5929 lemma [simp]: "mopup_jump_over2 (2 * n + 6, l, []) lm n ires = False" | 
|         |   5930 apply(simp only: mopup_jump_over2.simps, simp) | 
|         |   5931 done | 
|         |   5932  | 
|         |   5933 lemma mopup_inv_step: | 
|         |   5934   "\<lbrakk>n < length lm; mopup_inv (s, l, r) lm n ires\<rbrakk> | 
|         |   5935   \<Longrightarrow> mopup_inv (t_step (s, l, r)  | 
|         |   5936        ((mop_bef n @ tshift mp_up (2 * n)), 0)) lm n ires" | 
|         |   5937 apply(auto split:if_splits simp add:t_step.simps, | 
|         |   5938       tactic {* ALLGOALS (resolve_tac [@{thm "fetch_intro"}]) *}) | 
|         |   5939 apply(simp_all add: mopupfetchs new_tape.simps) | 
|         |   5940 done | 
|         |   5941  | 
|         |   5942 declare mopup_inv.simps[simp del] | 
|         |   5943  | 
|         |   5944 lemma mopup_inv_steps:  | 
|         |   5945 "\<lbrakk>n < length lm; mopup_inv (s, l, r) lm n ires\<rbrakk> \<Longrightarrow>  | 
|         |   5946      mopup_inv (t_steps (s, l, r)  | 
|         |   5947                    ((mop_bef n @ tshift mp_up (2 * n)), 0) stp) lm n ires" | 
|         |   5948 apply(induct stp, simp add: t_steps.simps) | 
|         |   5949 apply(simp add: t_steps_ind) | 
|         |   5950 apply(case_tac "(t_steps (s, l, r)  | 
|         |   5951                 (mop_bef n @ tshift mp_up (2 * n), 0) stp)", simp) | 
|         |   5952 apply(rule_tac mopup_inv_step, simp, simp) | 
|         |   5953 done | 
|         |   5954  | 
|         |   5955 lemma [simp]:  | 
|         |   5956  "\<lbrakk>n < length lm; Suc 0 \<le> n\<rbrakk> \<Longrightarrow>  | 
|         |   5957             mopup_bef_erase_a (Suc 0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>) lm n ires" | 
|         |   5958 apply(auto simp: mopup_bef_erase_a.simps  abc_lm_v.simps) | 
|         |   5959 apply(case_tac lm, simp, case_tac list, simp, simp) | 
|         |   5960 apply(rule_tac x = "Suc a" in exI, rule_tac x = rn in exI, simp) | 
|         |   5961 done | 
|         |   5962    | 
|         |   5963 lemma [simp]: | 
|         |   5964   "lm \<noteq> [] \<Longrightarrow> mopup_jump_over1 (Suc 0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>) lm 0  ires" | 
|         |   5965 apply(auto simp: mopup_jump_over1.simps) | 
|         |   5966 apply(rule_tac x = ln in exI, rule_tac x = 0 in exI, simp add: ) | 
|         |   5967 apply(case_tac lm, simp, simp add: abc_lm_v.simps) | 
|         |   5968 apply(case_tac rn, simp) | 
|         |   5969 apply(case_tac list, rule_tac disjI2,  | 
|         |   5970       simp add: tape_of_nl_abv tape_of_nat_list.simps) | 
|         |   5971 apply(rule_tac disjI1, | 
|         |   5972       simp add: tape_of_nl_abv tape_of_nat_list.simps ) | 
|         |   5973 apply(rule_tac disjI1, case_tac list,  | 
|         |   5974       simp add: tape_of_nl_abv tape_of_nat_list.simps,  | 
|         |   5975       rule_tac x = nat in exI, simp) | 
|         |   5976 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps ) | 
|         |   5977 done | 
|         |   5978  | 
|         |   5979 lemma mopup_init:  | 
|         |   5980  "\<lbrakk>n < length lm; crsp_l ly (as, lm) (ac, l, r) ires\<rbrakk> \<Longrightarrow>  | 
|         |   5981                                mopup_inv (Suc 0, l, r) lm n ires" | 
|         |   5982 apply(auto simp: crsp_l.simps mopup_inv.simps) | 
|         |   5983 apply(case_tac n, simp, auto simp: mopup_bef_erase_a.simps ) | 
|         |   5984 apply(rule_tac x = "Suc (hd lm)" in exI, rule_tac x = rn in exI, simp) | 
|         |   5985 apply(case_tac lm, simp, case_tac list, simp, case_tac lista, simp add: abc_lm_v.simps) | 
|         |   5986 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps) | 
|         |   5987 apply(simp add: mopup_jump_over1.simps) | 
|         |   5988 apply(rule_tac x = 0 in exI, rule_tac x = 0 in exI, auto) | 
|         |   5989 apply(case_tac [!] n, simp_all) | 
|         |   5990 apply(case_tac [!] lm, simp, case_tac list, simp) | 
|         |   5991 apply(case_tac rn, simp add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps) | 
|         |   5992 apply(erule_tac x = nat in allE, simp add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps) | 
|         |   5993 apply(simp add: abc_lm_v.simps, auto) | 
|         |   5994 apply(case_tac list, simp_all add: tape_of_nl_abv tape_of_nat_list.simps abc_lm_v.simps)  | 
|         |   5995 apply(erule_tac x = rn in allE, simp_all) | 
|         |   5996 done | 
|         |   5997  | 
|         |   5998 fun abc_mopup_stage1 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat" | 
|         |   5999   where | 
|         |   6000   "abc_mopup_stage1 (s, l, r) n =  | 
|         |   6001            (if s > 0 \<and> s \<le> 2*n then 6 | 
|         |   6002             else if s = 2*n + 1 then 4 | 
|         |   6003             else if s \<ge> 2*n + 2 \<and> s \<le> 2*n + 4 then 3 | 
|         |   6004             else if s = 2*n + 5 then 2 | 
|         |   6005             else if s = 2*n + 6 then 1 | 
|         |   6006             else 0)" | 
|         |   6007  | 
|         |   6008 fun abc_mopup_stage2 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat" | 
|         |   6009   where | 
|         |   6010   "abc_mopup_stage2 (s, l, r) n =  | 
|         |   6011            (if s > 0 \<and> s \<le> 2*n then length r | 
|         |   6012             else if s = 2*n + 1 then length r | 
|         |   6013             else if s = 2*n + 5 then length l | 
|         |   6014             else if s = 2*n + 6 then length l | 
|         |   6015             else if s \<ge> 2*n + 2 \<and> s \<le> 2*n + 4 then length r | 
|         |   6016             else 0)" | 
|         |   6017  | 
|         |   6018 fun abc_mopup_stage3 :: "t_conf \<Rightarrow> nat \<Rightarrow> nat" | 
|         |   6019   where | 
|         |   6020   "abc_mopup_stage3 (s, l, r) n =  | 
|         |   6021           (if s > 0 \<and> s \<le> 2*n then  | 
|         |   6022               if hd r = Bk then 0 | 
|         |   6023               else 1 | 
|         |   6024            else if s = 2*n + 2 then 1  | 
|         |   6025            else if s = 2*n + 3 then 0 | 
|         |   6026            else if s = 2*n + 4 then 2 | 
|         |   6027            else 0)" | 
|         |   6028  | 
|         |   6029 fun abc_mopup_measure :: "(t_conf \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)" | 
|         |   6030   where | 
|         |   6031   "abc_mopup_measure (c, n) =  | 
|         |   6032     (abc_mopup_stage1 c n, abc_mopup_stage2 c n,  | 
|         |   6033                                        abc_mopup_stage3 c n)" | 
|         |   6034  | 
|         |   6035 definition abc_mopup_LE :: | 
|         |   6036    "(((nat \<times> block list \<times> block list) \<times> nat) \<times>  | 
|         |   6037     ((nat \<times> block list \<times> block list) \<times> nat)) set" | 
|         |   6038   where | 
|         |   6039   "abc_mopup_LE \<equiv> (inv_image lex_triple abc_mopup_measure)" | 
|         |   6040  | 
|         |   6041 lemma wf_abc_mopup_le[intro]: "wf abc_mopup_LE" | 
|         |   6042 by(auto intro:wf_inv_image wf_lex_triple simp:abc_mopup_LE_def) | 
|         |   6043  | 
|         |   6044 lemma [simp]: "mopup_bef_erase_a (a, aa, []) lm n ires = False" | 
|         |   6045 apply(auto simp: mopup_bef_erase_a.simps) | 
|         |   6046 done | 
|         |   6047  | 
|         |   6048 lemma [simp]: "mopup_bef_erase_b (a, aa, []) lm n ires = False" | 
|         |   6049 apply(auto simp: mopup_bef_erase_b.simps)  | 
|         |   6050 done | 
|         |   6051  | 
|         |   6052 lemma [simp]: "mopup_aft_erase_b (2 * n + 3, aa, []) lm n ires = False" | 
|         |   6053 apply(auto simp: mopup_aft_erase_b.simps) | 
|         |   6054 done | 
|         |   6055  | 
|         |   6056 lemma mopup_halt_pre:  | 
|         |   6057  "\<lbrakk>n < length lm; mopup_inv (Suc 0, l, r) lm n ires; wf abc_mopup_LE\<rbrakk> | 
|         |   6058  \<Longrightarrow>  \<forall>na. \<not> (\<lambda>(s, l, r) n. s = 0) (t_steps (Suc 0, l, r) | 
|         |   6059       (mop_bef n @ tshift mp_up (2 * n), 0) na) n \<longrightarrow> | 
|         |   6060        ((t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0)  | 
|         |   6061         (Suc na), n), | 
|         |   6062        t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) | 
|         |   6063          na, n) \<in> abc_mopup_LE" | 
|         |   6064 apply(rule allI, rule impI, simp add: t_steps_ind) | 
|         |   6065 apply(subgoal_tac "mopup_inv (t_steps (Suc 0, l, r)  | 
|         |   6066                      (mop_bef n @ tshift mp_up (2 * n), 0) na) lm n ires") | 
|         |   6067 apply(case_tac "(t_steps (Suc 0, l, r)  | 
|         |   6068                (mop_bef n @ tshift mp_up (2 * n), 0) na)",  simp) | 
|         |   6069 proof - | 
|         |   6070   fix na a b c | 
|         |   6071   assume  "n < length lm" "mopup_inv (a, b, c) lm n ires" "0 < a" | 
|         |   6072   thus "((t_step (a, b, c) (mop_bef n @ tshift mp_up (2 * n), 0), n), | 
|         |   6073          (a, b, c), n) \<in> abc_mopup_LE" | 
|         |   6074     apply(auto split:if_splits simp add:t_step.simps mopup_inv.simps, | 
|         |   6075       tactic {* ALLGOALS (resolve_tac [@{thm "fetch_intro"}]) *}) | 
|         |   6076     apply(simp_all add: mopupfetchs new_tape.simps abc_mopup_LE_def  | 
|         |   6077                    lex_triple_def lex_pair_def ) | 
|         |   6078     done | 
|         |   6079 next | 
|         |   6080   fix na | 
|         |   6081   assume "n < length lm" "mopup_inv (Suc 0, l, r) lm n ires" | 
|         |   6082   thus "mopup_inv (t_steps (Suc 0, l, r)  | 
|         |   6083        (mop_bef n @ tshift mp_up (2 * n), 0) na) lm n ires" | 
|         |   6084     apply(rule mopup_inv_steps) | 
|         |   6085     done | 
|         |   6086 qed | 
|         |   6087  | 
|         |   6088 lemma mopup_halt: "\<lbrakk>n < length lm; crsp_l ly (as, lm) (s, l, r) ires\<rbrakk> \<Longrightarrow>  | 
|         |   6089   \<exists> stp. (\<lambda> (s, l, r). s = 0) (t_steps (Suc 0, l, r)  | 
|         |   6090         ((mop_bef n @ tshift mp_up (2 * n)), 0) stp)" | 
|         |   6091 apply(subgoal_tac "mopup_inv (Suc 0, l, r) lm n ires") | 
|         |   6092 apply(insert wf_abc_mopup_le) | 
|         |   6093 apply(insert halt_lemma[of abc_mopup_LE  | 
|         |   6094     "\<lambda> ((s, l, r), n). s = 0"  | 
|         |   6095     "\<lambda> stp. (t_steps (Suc 0, l, r) ((mop_bef n @ tshift mp_up (2 * n)) | 
|         |   6096            , 0) stp, n)"], auto) | 
|         |   6097 apply(insert mopup_halt_pre[of n lm l r], simp, erule exE) | 
|         |   6098 apply(rule_tac x = na in exI, case_tac "(t_steps (Suc 0, l, r)  | 
|         |   6099           (mop_bef n @ tshift mp_up (2 * n), 0) na)", simp) | 
|         |   6100 apply(rule_tac mopup_init, auto) | 
|         |   6101 done | 
|         |   6102 (***End: mopup stop****) | 
|         |   6103  | 
|         |   6104 lemma mopup_halt_conf_pre:  | 
|         |   6105  "\<lbrakk>n < length lm; crsp_l ly (as, lm) (s, l, r) ires\<rbrakk>  | 
|         |   6106   \<Longrightarrow> \<exists> na. (\<lambda> (s', l', r').  s' = 0 \<and> mopup_stop (s', l', r') lm n ires) | 
|         |   6107       (t_steps (Suc 0, l, r)  | 
|         |   6108             ((mop_bef n @ tshift mp_up (2 * n)), 0) na)" | 
|         |   6109 apply(subgoal_tac "\<exists> stp. (\<lambda> (s, l, r). s = 0)  | 
|         |   6110  (t_steps (Suc 0, l, r) ((mop_bef n @ tshift mp_up (2 * n)), 0) stp)", | 
|         |   6111        erule exE) | 
|         |   6112 apply(rule_tac x = stp in exI,  | 
|         |   6113       case_tac "(t_steps (Suc 0, l, r)  | 
|         |   6114           (mop_bef n @ tshift mp_up (2 * n), 0) stp)", simp) | 
|         |   6115 apply(subgoal_tac " mopup_inv (Suc 0, l, r) lm n ires") | 
|         |   6116 apply(subgoal_tac "mopup_inv (t_steps (Suc 0, l, r)  | 
|         |   6117             (mop_bef n @ tshift mp_up (2 * n), 0) stp) lm n ires", simp) | 
|         |   6118 apply(simp only: mopup_inv.simps) | 
|         |   6119 apply(rule_tac mopup_inv_steps, simp, simp) | 
|         |   6120 apply(rule_tac mopup_init, simp, simp) | 
|         |   6121 apply(rule_tac mopup_halt, simp, simp) | 
|         |   6122 done | 
|         |   6123  | 
|         |   6124 lemma  mopup_halt_conf: | 
|         |   6125   assumes len: "n < length lm" | 
|         |   6126   and correspond: "crsp_l ly (as, lm) (s, l, r) ires" | 
|         |   6127   shows  | 
|         |   6128   "\<exists> na. (\<lambda> (s', l', r'). ((\<exists>ln rn. s' = 0 \<and> l' = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires | 
|         |   6129                            \<and> r' = Oc\<^bsup>Suc (abc_lm_v lm n)\<^esup> @ Bk\<^bsup>rn\<^esup>))) | 
|         |   6130              (t_steps (Suc 0, l, r)  | 
|         |   6131                   ((mop_bef n @ tshift mp_up (2 * n)), 0) na)" | 
|         |   6132 using len correspond mopup_halt_conf_pre[of n lm ly as s l r ires] | 
|         |   6133 apply(simp add: mopup_stop.simps tape_of_nat_abv tape_of_nat_list.simps) | 
|         |   6134 done | 
|         |   6135  | 
|         |   6136 subsection {* Final results about Abacus machine *} | 
|         |   6137  | 
|         |   6138 lemma mopup_halt_bef: "\<lbrakk>n < length lm; crsp_l ly (as, lm) (s, l, r) ires\<rbrakk>  | 
|         |   6139     \<Longrightarrow> \<exists>stp. (\<lambda>(s, l, r). s \<noteq> 0 \<and> ((\<lambda> (s', l', r'). s' = 0) | 
|         |   6140    (t_step (s, l, r) (mop_bef n @ tshift mp_up (2 * n), 0)))) | 
|         |   6141     (t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) stp)" | 
|         |   6142 apply(insert mopup_halt[of n lm ly as s l r ires], simp, erule_tac exE) | 
|         |   6143 proof - | 
|         |   6144   fix stp | 
|         |   6145   assume "n < length lm"  | 
|         |   6146          "crsp_l ly (as, lm) (s, l, r) ires"  | 
|         |   6147          "(\<lambda>(s, l, r). s = 0)  | 
|         |   6148             (t_steps (Suc 0, l, r)  | 
|         |   6149               (mop_bef n @ tshift mp_up (2 * n), 0) stp)" | 
|         |   6150   thus "\<exists>stp. (\<lambda>(s, ab). 0 < s \<and> (\<lambda>(s', l', r'). s' = 0)  | 
|         |   6151    (t_step (s, ab) (mop_bef n @ tshift mp_up (2 * n), 0)))  | 
|         |   6152     (t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) stp)" | 
|         |   6153   proof(induct stp, simp add: t_steps.simps, simp) | 
|         |   6154     fix stpa | 
|         |   6155     assume h1:  | 
|         |   6156       "(\<lambda>(s, l, r). s = 0) (t_steps (Suc 0, l, r)  | 
|         |   6157            (mop_bef n @ tshift mp_up (2 * n), 0) stpa) \<Longrightarrow> | 
|         |   6158        \<exists>stp. (\<lambda>(s, ab). 0 < s \<and> (\<lambda>(s', l', r'). s' = 0)  | 
|         |   6159          (t_step (s, ab) (mop_bef n @ tshift mp_up (2 * n), 0)))  | 
|         |   6160             (t_steps (Suc 0, l, r)  | 
|         |   6161               (mop_bef n @ tshift mp_up (2 * n), 0) stp)" | 
|         |   6162       and h2:  | 
|         |   6163         "(\<lambda>(s, l, r). s = 0) (t_steps (Suc 0, l, r)  | 
|         |   6164                     (mop_bef n @ tshift mp_up (2 * n), 0) (Suc stpa))" | 
|         |   6165          "n < length lm"  | 
|         |   6166          "crsp_l ly (as, lm) (s, l, r) ires" | 
|         |   6167     thus "\<exists>stp. (\<lambda>(s, ab). 0 < s \<and> (\<lambda>(s', l', r'). s' = 0)  | 
|         |   6168              (t_step (s, ab) (mop_bef n @ tshift mp_up (2 * n), 0))) ( | 
|         |   6169                 t_steps (Suc 0, l, r)  | 
|         |   6170                   (mop_bef n @ tshift mp_up (2 * n), 0) stp)" | 
|         |   6171       apply(case_tac "(\<lambda>(s, l, r). s = 0) (t_steps (Suc 0, l, r)  | 
|         |   6172                      (mop_bef n @ tshift mp_up (2 * n), 0) stpa)",  | 
|         |   6173             simp) | 
|         |   6174       apply(rule_tac x = "stpa" in exI) | 
|         |   6175       apply(case_tac "(t_steps (Suc 0, l, r)  | 
|         |   6176                          (mop_bef n @ tshift mp_up (2 * n), 0) stpa)", | 
|         |   6177             simp add: t_steps_ind) | 
|         |   6178       done | 
|         |   6179   qed | 
|         |   6180 qed | 
|         |   6181  | 
|         |   6182 lemma tshift_nth_state0: "\<lbrakk>n < length tp; tp ! n = (a, 0)\<rbrakk> | 
|         |   6183     \<Longrightarrow> tshift tp off ! n = (a, 0)" | 
|         |   6184 apply(induct n, case_tac tp, simp) | 
|         |   6185 apply(auto simp: tshift.simps) | 
|         |   6186 done | 
|         |   6187  | 
|         |   6188 lemma shift_length: "length (tshift tp n) = length tp" | 
|         |   6189 apply(auto simp: tshift.simps) | 
|         |   6190 done | 
|         |   6191  | 
|         |   6192 lemma even_Suc_le: "\<lbrakk>y mod 2 = 0; 2 * x < y\<rbrakk> \<Longrightarrow> Suc (2 * x) < y" | 
|         |   6193 by arith | 
|         |   6194  | 
|         |   6195 lemma [simp]: "(4::nat) * n mod 2 = 0" | 
|         |   6196 by arith | 
|         |   6197  | 
|         |   6198 lemma tm_append_fetch_equal:  | 
|         |   6199   "\<lbrakk>t_ncorrect (tm_of aprog); s'> 0; | 
|         |   6200     fetch (mop_bef n @ tshift mp_up (2 * n)) s' b = (a, 0)\<rbrakk> | 
|         |   6201 \<Longrightarrow> fetch (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n))  | 
|         |   6202     (length (tm_of aprog) div 2)) (s' + length (tm_of aprog) div 2) b  | 
|         |   6203    = (a, 0)" | 
|         |   6204 apply(case_tac s', simp) | 
|         |   6205 apply(auto simp: fetch.simps nth_of.simps t_ncorrect.simps shift_length nth_append | 
|         |   6206                  tshift.simps split: list.splits block.splits split: if_splits) | 
|         |   6207 done | 
|         |   6208  | 
|         |   6209 lemma [simp]: | 
|         |   6210   "\<lbrakk>t_ncorrect (tm_of aprog); | 
|         |   6211     t_step (s', l', r') (mop_bef n @ tshift mp_up (2 * n), 0) =  | 
|         |   6212                                                (0, l'', r''); s' > 0\<rbrakk> | 
|         |   6213   \<Longrightarrow> t_step (s' + length (tm_of aprog) div 2, l', r')  | 
|         |   6214         (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) | 
|         |   6215            (length (tm_of aprog) div 2), 0) = (0, l'', r'')" | 
|         |   6216 apply(simp add: t_step.simps) | 
|         |   6217 apply(subgoal_tac  | 
|         |   6218    "(fetch (mop_bef n @ tshift mp_up (2 * n)) s'  | 
|         |   6219               (case r' of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc)) | 
|         |   6220   = (fetch (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n))  | 
|         |   6221        (length (tm_of aprog) div 2)) (s' + length (tm_of aprog) div 2) | 
|         |   6222     (case r' of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))", simp) | 
|         |   6223 apply(case_tac "(fetch (mop_bef n @ tshift mp_up (2 * n)) s'  | 
|         |   6224        (case r' of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc))", simp) | 
|         |   6225 apply(drule_tac tm_append_fetch_equal, auto) | 
|         |   6226 done | 
|         |   6227  | 
|         |   6228 lemma [intro]:  | 
|         |   6229   "start_of (layout_of aprog) (length aprog) - Suc 0 =  | 
|         |   6230                                       length (tm_of aprog) div 2" | 
|         |   6231 apply(subgoal_tac  "abc2t_correct aprog") | 
|         |   6232 apply(insert pre_lheq[of "concat (take (length aprog)  | 
|         |   6233        (tms_of aprog))" "length aprog" aprog], simp add: tm_of.simps) | 
|         |   6234 by auto | 
|         |   6235  | 
|         |   6236 lemma tm_append_stop_step:  | 
|         |   6237   "\<lbrakk>t_ncorrect (tm_of aprog);  | 
|         |   6238     t_ncorrect (mop_bef n @ tshift mp_up (2 * n)); n < length lm;  | 
|         |   6239    (t_steps (Suc 0, l, r) (mop_bef n @ tshift mp_up (2 * n), 0) stp) = | 
|         |   6240                          (s', l', r'); | 
|         |   6241     s' \<noteq> 0; | 
|         |   6242     t_step (s', l', r') (mop_bef n @ tshift mp_up (2 * n), 0)  | 
|         |   6243                                                      = (0, l'', r'')\<rbrakk> | 
|         |   6244      \<Longrightarrow> | 
|         |   6245 (t_steps ((start_of (layout_of aprog) (length aprog), l, r))  | 
|         |   6246   (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) | 
|         |   6247    (start_of (layout_of aprog) (length aprog) - Suc 0), 0) (Suc stp)) | 
|         |   6248   = (0, l'', r'')" | 
|         |   6249 apply(insert tm_append_steps_equal[of "tm_of aprog"  | 
|         |   6250       "(mop_bef n @ tshift mp_up (2 * n))" | 
|         |   6251       "(start_of (layout_of aprog) (length aprog) - Suc 0)"  | 
|         |   6252       "Suc 0" l r stp], simp) | 
|         |   6253 apply(subgoal_tac "(start_of (layout_of aprog) (length aprog) - Suc 0) | 
|         |   6254               = (length (tm_of aprog) div 2)", simp add: t_steps_ind) | 
|         |   6255 apply(case_tac  | 
|         |   6256  "(t_steps (start_of (layout_of aprog) (length aprog), l, r)  | 
|         |   6257       (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) | 
|         |   6258            (length (tm_of aprog) div 2), 0) stp)", simp) | 
|         |   6259 apply(subgoal_tac "start_of (layout_of aprog) (length aprog) > 0",  | 
|         |   6260       case_tac "start_of (layout_of aprog) (length aprog)",  | 
|         |   6261       simp, simp) | 
|         |   6262 apply(rule startof_not0, auto) | 
|         |   6263 done | 
|         |   6264  | 
|         |   6265 lemma start_of_out_range:  | 
|         |   6266 "as \<ge> length aprog \<Longrightarrow>  | 
|         |   6267    start_of (layout_of aprog) as =  | 
|         |   6268              start_of (layout_of aprog) (length aprog)" | 
|         |   6269 apply(induct as, simp) | 
|         |   6270 apply(case_tac "length aprog = Suc as", simp) | 
|         |   6271 apply(simp add: start_of.simps) | 
|         |   6272 done | 
|         |   6273  | 
|         |   6274 lemma [intro]: "t_ncorrect (tm_of aprog)" | 
|         |   6275 apply(simp add: tm_of.simps) | 
|         |   6276 apply(insert tms_mod2[of "length aprog" aprog],  | 
|         |   6277                                 simp add: t_ncorrect.simps) | 
|         |   6278 done | 
|         |   6279  | 
|         |   6280 lemma abacus_turing_eq_halt_case_pre:  | 
|         |   6281    "\<lbrakk>ly = layout_of aprog;  | 
|         |   6282      tprog = tm_of aprog;  | 
|         |   6283      crsp_l ly ac tc ires; | 
|         |   6284      n < length am; | 
|         |   6285      abc_steps_l ac aprog stp = (as, am);  | 
|         |   6286      mop_ss = start_of ly (length aprog); | 
|         |   6287      as \<ge> length aprog\<rbrakk> | 
|         |   6288      \<Longrightarrow> \<exists> stp. (\<lambda> (s, l, r). s = 0 \<and> mopup_inv (0, l, r) am n ires) | 
|         |   6289                 (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp)" | 
|         |   6290 apply(insert steps_crsp[of ly aprog tprog ac tc ires "(as, am)" stp], auto) | 
|         |   6291 apply(case_tac "(t_steps tc (tm_of aprog, 0) n')",   | 
|         |   6292       simp add: tMp.simps) | 
|         |   6293 apply(subgoal_tac "t_ncorrect (mop_bef n @ tshift mp_up (2 * n))") | 
|         |   6294 proof - | 
|         |   6295   fix n' a b c | 
|         |   6296   assume h1:  | 
|         |   6297     "crsp_l (layout_of aprog) ac tc ires"  | 
|         |   6298     "abc_steps_l ac aprog stp = (as, am)"  | 
|         |   6299     "length aprog \<le> as" | 
|         |   6300     "crsp_l (layout_of aprog) (as, am) (a, b, c) ires" | 
|         |   6301     "t_steps tc (tm_of aprog, 0) n' = (a, b, c)" | 
|         |   6302     "n < length am" | 
|         |   6303     "t_ncorrect (mop_bef n @ tshift mp_up (2 * n))" | 
|         |   6304   hence h2: | 
|         |   6305   "t_steps tc (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) | 
|         |   6306     (start_of (layout_of aprog) (length aprog) - Suc 0), 0) n'  | 
|         |   6307                                     = (a, b, c)"  | 
|         |   6308     apply(rule_tac tm_append_steps, simp) | 
|         |   6309     apply(simp add: crsp_l.simps, auto) | 
|         |   6310     apply(simp add: crsp_l.simps) | 
|         |   6311     apply(rule startof_not0) | 
|         |   6312     done | 
|         |   6313   from h1 have h3:  | 
|         |   6314   "\<exists>stp. (\<lambda>(s, l, r). s \<noteq> 0 \<and> ((\<lambda> (s', l', r'). s' = 0)  | 
|         |   6315            (t_step (s, l, r) (mop_bef n @ tshift mp_up (2 * n), 0)))) | 
|         |   6316          (t_steps (Suc 0, b, c)  | 
|         |   6317                (mop_bef n @ tshift mp_up (2 * n), 0) stp)" | 
|         |   6318     apply(rule_tac mopup_halt_bef, auto) | 
|         |   6319     done | 
|         |   6320   from h1 and h2 and h3 show  | 
|         |   6321     "\<exists>stp. case t_steps tc (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n)) | 
|         |   6322     (start_of (layout_of aprog) (length aprog) - Suc 0), 0) stp of (s, ab) | 
|         |   6323     \<Rightarrow> s = 0 \<and> mopup_inv (0, ab) am n ires" | 
|         |   6324   proof(erule_tac exE,  | 
|         |   6325     case_tac "(t_steps (Suc 0, b, c)  | 
|         |   6326               (mop_bef n @ tshift mp_up (2 * n), 0) stpa)", simp, | 
|         |   6327     case_tac "(t_step (aa, ba, ca)  | 
|         |   6328               (mop_bef n @ tshift mp_up (2 * n), 0))", simp) | 
|         |   6329     fix stpa aa ba ca aaa baa caa | 
|         |   6330     assume g1: "0 < aa \<and> aaa = 0"  | 
|         |   6331       "t_steps (Suc 0, b, c)  | 
|         |   6332       (mop_bef n @ tshift mp_up (2 * n), 0) stpa = (aa, ba,ca)"  | 
|         |   6333       "t_step (aa, ba, ca) (mop_bef n @ tshift mp_up (2 * n), 0) | 
|         |   6334       = (0, baa, caa)" | 
|         |   6335     from h1 and this have g2:  | 
|         |   6336       "t_steps (start_of (layout_of aprog) (length aprog), b, c)  | 
|         |   6337          (tm_of aprog @ tshift (mop_bef n @ tshift mp_up (2 * n))  | 
|         |   6338            (start_of (layout_of aprog) (length aprog) - Suc 0), 0)  | 
|         |   6339                 (Suc stpa) = (0, baa, caa)" | 
|         |   6340       apply(rule_tac tm_append_stop_step, auto) | 
|         |   6341       done | 
|         |   6342     from h1 and h2 and g1 and this show "?thesis" | 
|         |   6343       apply(rule_tac x = "n' + Suc stpa" in exI) | 
|         |   6344       apply(simp add: t_steps_ind del: t_steps.simps) | 
|         |   6345       apply(subgoal_tac "a = start_of (layout_of aprog)  | 
|         |   6346                                           (length aprog)", simp) | 
|         |   6347       apply(insert mopup_inv_steps[of n am "Suc 0" b c ires "Suc stpa"], | 
|         |   6348             simp add: t_steps_ind) | 
|         |   6349       apply(subgoal_tac "mopup_inv (Suc 0, b, c) am n ires", simp) | 
|         |   6350       apply(rule_tac mopup_init, simp, simp) | 
|         |   6351       apply(simp add: crsp_l.simps) | 
|         |   6352       apply(erule_tac start_of_out_range) | 
|         |   6353       done | 
|         |   6354   qed | 
|         |   6355 next | 
|         |   6356   show " t_ncorrect (mop_bef n @ tshift mp_up (2 * n))" | 
|         |   6357     apply(auto simp: t_ncorrect.simps tshift.simps mp_up_def) | 
|         |   6358     done    | 
|         |   6359 qed | 
|         |   6360  | 
|         |   6361 text {* | 
|         |   6362   One of the main theorems about Abacus compilation. | 
|         |   6363 *} | 
|         |   6364  | 
|         |   6365 lemma abacus_turing_eq_halt_case:  | 
|         |   6366   assumes layout:  | 
|         |   6367   -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *} | 
|         |   6368   "ly = layout_of aprog" | 
|         |   6369   and complied:  | 
|         |   6370   -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *} | 
|         |   6371   "tprog = tm_of aprog" | 
|         |   6372   and correspond:  | 
|         |   6373   -- {* TM configuration @{text "tc"} and Abacus configuration @{text "ac"} | 
|         |   6374   are in correspondence: *} | 
|         |   6375   "crsp_l ly ac tc ires" | 
|         |   6376   and halt_state:  | 
|         |   6377   -- {* @{text "as"} is a program label outside the range of @{text "aprog"}. So  | 
|         |   6378   if Abacus is in such a state, it is in halt state: *} | 
|         |   6379   "as \<ge> length aprog" | 
|         |   6380   and abc_exec:  | 
|         |   6381   -- {* Supposing after @{text "stp"} step of execution, Abacus program @{text "aprog"} | 
|         |   6382   reaches such a halt state: *} | 
|         |   6383   "abc_steps_l ac aprog stp = (as, am)" | 
|         |   6384   and rs_len:  | 
|         |   6385   -- {* @{text "n"} is a memory address in the range of Abacus memory @{text "am"}: *} | 
|         |   6386   "n < length am" | 
|         |   6387   and mopup_start: | 
|         |   6388   -- {* The startling label for mopup mahines, according to the layout and Abacus program  | 
|         |   6389    should be @{text "mop_ss"}: *} | 
|         |   6390   "mop_ss = start_of ly (length aprog)" | 
|         |   6391   shows  | 
|         |   6392   -- {*  | 
|         |   6393   After @{text "stp"} steps of execution of the TM composed of @{text "tprog"} and the mopup  | 
|         |   6394   TM @{text "(tMp n (mop_ss - 1))"} will halt and gives rise to a configuration which  | 
|         |   6395   only hold the content of memory cell @{text "n"}: | 
|         |   6396   *} | 
|         |   6397   "\<exists> stp. (\<lambda> (s, l, r). \<exists> ln rn. s = 0 \<and>  l = Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires | 
|         |   6398      \<and> r = Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>rn\<^esup>) | 
|         |   6399            (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp)" | 
|         |   6400 proof - | 
|         |   6401   from layout complied correspond halt_state abc_exec rs_len mopup_start | 
|         |   6402        and abacus_turing_eq_halt_case_pre [of ly aprog tprog ac tc ires n am stp as mop_ss] | 
|         |   6403   show "?thesis"  | 
|         |   6404     apply(simp add: mopup_inv.simps mopup_stop.simps tape_of_nat_abv) | 
|         |   6405     done | 
|         |   6406 qed | 
|         |   6407  | 
|         |   6408 lemma abc_unhalt_case_zero:  | 
|         |   6409 "\<lbrakk>crsp_l (layout_of aprog) ac tc ires; | 
|         |   6410   n < length am;  | 
|         |   6411   \<forall>stp. (\<lambda>(as, am). as < length aprog) (abc_steps_l ac aprog stp)\<rbrakk> | 
|         |   6412  \<Longrightarrow> \<exists>astp bstp. 0 \<le> bstp \<and>  | 
|         |   6413           crsp_l (layout_of aprog) (abc_steps_l ac aprog astp)  | 
|         |   6414                 (t_steps tc (tm_of aprog, 0) bstp) ires" | 
|         |   6415 apply(rule_tac x = "Suc 0" in exI) | 
|         |   6416 apply(case_tac " abc_steps_l ac aprog (Suc 0)", simp) | 
|         |   6417 proof - | 
|         |   6418   fix a b | 
|         |   6419   assume "crsp_l (layout_of aprog) ac tc ires"  | 
|         |   6420          "abc_steps_l ac aprog (Suc 0) = (a, b)" | 
|         |   6421   thus "\<exists>bstp. crsp_l (layout_of aprog) (a, b)  | 
|         |   6422                (t_steps tc (tm_of aprog, 0) bstp) ires" | 
|         |   6423     apply(insert steps_crsp[of "layout_of aprog" aprog  | 
|         |   6424                   "tm_of aprog" ac tc ires "(a, b)" "Suc 0"], auto) | 
|         |   6425     done | 
|         |   6426 qed | 
|         |   6427  | 
|         |   6428 declare abc_steps_l.simps[simp del] | 
|         |   6429  | 
|         |   6430 lemma abc_steps_ind:  | 
|         |   6431  "let (as, am) = abc_steps_l ac aprog stp in  | 
|         |   6432    abc_steps_l ac aprog (Suc stp) = | 
|         |   6433               abc_step_l (as, am) (abc_fetch as aprog) " | 
|         |   6434 proof(simp) | 
|         |   6435   show "(\<lambda>(as, am). abc_steps_l ac aprog (Suc stp) =  | 
|         |   6436         abc_step_l (as, am) (abc_fetch as aprog))  | 
|         |   6437               (abc_steps_l ac aprog stp)" | 
|         |   6438   proof(induct stp arbitrary: ac) | 
|         |   6439     fix ac | 
|         |   6440     show "(\<lambda>(as, am). abc_steps_l ac aprog (Suc 0) =  | 
|         |   6441             abc_step_l (as, am) (abc_fetch as aprog))   | 
|         |   6442                     (abc_steps_l ac aprog 0)" | 
|         |   6443       apply(case_tac "ac:: nat \<times> nat list",  | 
|         |   6444             simp add: abc_steps_l.simps) | 
|         |   6445       apply(case_tac "(abc_step_l (a, b) (abc_fetch a aprog))", | 
|         |   6446             simp add: abc_steps_l.simps) | 
|         |   6447       done | 
|         |   6448   next | 
|         |   6449     fix stp ac | 
|         |   6450     assume h1: | 
|         |   6451       "(\<And>ac. (\<lambda>(as, am). abc_steps_l ac aprog (Suc stp) = | 
|         |   6452                             abc_step_l (as, am) (abc_fetch as aprog))  | 
|         |   6453              (abc_steps_l ac aprog stp))" | 
|         |   6454     thus  | 
|         |   6455       "(\<lambda>(as, am). abc_steps_l ac aprog (Suc (Suc stp)) = | 
|         |   6456               abc_step_l (as, am) (abc_fetch as aprog))  | 
|         |   6457                              (abc_steps_l ac aprog (Suc stp))" | 
|         |   6458       apply(case_tac "ac::nat \<times> nat list", simp) | 
|         |   6459       apply(subgoal_tac  | 
|         |   6460            "abc_steps_l (a, b) aprog (Suc (Suc stp)) = | 
|         |   6461             abc_steps_l (abc_step_l (a, b) (abc_fetch a aprog))  | 
|         |   6462                                               aprog (Suc stp)", simp) | 
|         |   6463       apply(case_tac "(abc_step_l (a, b) (abc_fetch a aprog))", simp) | 
|         |   6464     proof - | 
|         |   6465       fix a b aa ba | 
|         |   6466       assume h2: "abc_step_l (a, b) (abc_fetch a aprog) = (aa, ba)" | 
|         |   6467       from h1 and h2  show  | 
|         |   6468       "(\<lambda>(as, am). abc_steps_l (aa, ba) aprog (Suc stp) =  | 
|         |   6469           abc_step_l (as, am) (abc_fetch as aprog))  | 
|         |   6470                     (abc_steps_l (a, b) aprog (Suc stp))" | 
|         |   6471 	apply(insert h1[of "(aa, ba)"]) | 
|         |   6472 	apply(simp add: abc_steps_l.simps) | 
|         |   6473 	apply(insert h2, simp) | 
|         |   6474 	done | 
|         |   6475     next | 
|         |   6476       fix a b | 
|         |   6477       show  | 
|         |   6478         "abc_steps_l (a, b) aprog (Suc (Suc stp)) =  | 
|         |   6479          abc_steps_l (abc_step_l (a, b) (abc_fetch a aprog))  | 
|         |   6480                                                    aprog (Suc stp)" | 
|         |   6481 	apply(simp only: abc_steps_l.simps) | 
|         |   6482 	done | 
|         |   6483     qed | 
|         |   6484   qed | 
|         |   6485 qed | 
|         |   6486  | 
|         |   6487 lemma abc_unhalt_case_induct:  | 
|         |   6488   "\<lbrakk>crsp_l (layout_of aprog) ac tc ires; | 
|         |   6489     n < length am;  | 
|         |   6490     \<forall>stp. (\<lambda>(as, am). as < length aprog) (abc_steps_l ac aprog stp);  | 
|         |   6491     stp \<le> bstp; | 
|         |   6492     crsp_l (layout_of aprog) (abc_steps_l ac aprog astp)  | 
|         |   6493                            (t_steps tc (tm_of aprog, 0) bstp) ires\<rbrakk> | 
|         |   6494  \<Longrightarrow> \<exists>astp bstp. Suc stp \<le> bstp \<and> crsp_l (layout_of aprog)  | 
|         |   6495        (abc_steps_l ac aprog astp) (t_steps tc (tm_of aprog, 0) bstp) ires" | 
|         |   6496 apply(rule_tac x = "Suc astp" in exI) | 
|         |   6497 apply(case_tac "abc_steps_l ac aprog astp") | 
|         |   6498 proof - | 
|         |   6499   fix a b | 
|         |   6500   assume  | 
|         |   6501     "\<forall>stp. (\<lambda>(as, am). as < length aprog)   | 
|         |   6502                  (abc_steps_l ac aprog stp)"  | 
|         |   6503     "stp \<le> bstp" | 
|         |   6504     "crsp_l (layout_of aprog) (abc_steps_l ac aprog astp)  | 
|         |   6505       (t_steps tc (tm_of aprog, 0) bstp) ires"  | 
|         |   6506     "abc_steps_l ac aprog astp = (a, b)" | 
|         |   6507   thus  | 
|         |   6508  "\<exists>bstp\<ge>Suc stp. crsp_l (layout_of aprog) | 
|         |   6509        (abc_steps_l ac aprog (Suc astp))  | 
|         |   6510    (t_steps tc (tm_of aprog, 0) bstp) ires" | 
|         |   6511     apply(insert crsp_inside[of "layout_of aprog" aprog  | 
|         |   6512       "tm_of aprog" a b "(t_steps tc (tm_of aprog, 0) bstp)" "ires"], auto) | 
|         |   6513     apply(erule_tac x = astp in allE, auto) | 
|         |   6514     apply(rule_tac x = "bstp + stpa" in exI, simp) | 
|         |   6515     apply(insert abc_steps_ind[of ac aprog "astp"], simp) | 
|         |   6516     done | 
|         |   6517 qed    | 
|         |   6518  | 
|         |   6519 lemma abc_unhalt_case:  | 
|         |   6520   "\<lbrakk>crsp_l (layout_of aprog) ac tc ires;   | 
|         |   6521     \<forall>stp. (\<lambda>(as, am). as < length aprog) (abc_steps_l ac aprog stp)\<rbrakk> | 
|         |   6522  \<Longrightarrow> (\<exists> astp bstp. bstp \<ge> stp \<and>  | 
|         |   6523          crsp_l (layout_of aprog) (abc_steps_l ac aprog astp)  | 
|         |   6524                                 (t_steps tc (tm_of aprog, 0) bstp) ires)" | 
|         |   6525 apply(induct stp) | 
|         |   6526 apply(rule_tac abc_unhalt_case_zero, auto) | 
|         |   6527 apply(rule_tac abc_unhalt_case_induct, auto) | 
|         |   6528 done | 
|         |   6529    | 
|         |   6530 lemma abacus_turing_eq_unhalt_case_pre:  | 
|         |   6531   "\<lbrakk>ly = layout_of aprog;  | 
|         |   6532     tprog = tm_of aprog; | 
|         |   6533     crsp_l ly ac tc ires; | 
|         |   6534     \<forall> stp. ((\<lambda> (as, am). as < length aprog) | 
|         |   6535                        (abc_steps_l ac aprog stp)); | 
|         |   6536     mop_ss = start_of ly (length aprog)\<rbrakk> | 
|         |   6537   \<Longrightarrow> (\<not> (\<exists> stp. (\<lambda> (s, l, r). s = 0) | 
|         |   6538               (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp)))" | 
|         |   6539   apply(auto) | 
|         |   6540 proof - | 
|         |   6541   fix stp a b | 
|         |   6542   assume h1:  | 
|         |   6543     "crsp_l (layout_of aprog) ac tc ires"  | 
|         |   6544     "\<forall>stp. (\<lambda>(as, am). as < length aprog) (abc_steps_l ac aprog stp)" | 
|         |   6545     "t_steps tc (tm_of aprog @ tMp n (start_of (layout_of aprog)  | 
|         |   6546     (length aprog) - Suc 0), 0) stp = (0, a, b)" | 
|         |   6547   thus "False" | 
|         |   6548   proof(insert abc_unhalt_case[of aprog ac tc ires stp], auto,  | 
|         |   6549         case_tac "(abc_steps_l ac aprog astp)",  | 
|         |   6550         case_tac "(t_steps tc (tm_of aprog, 0) bstp)", simp) | 
|         |   6551     fix astp bstp aa ba aaa baa c | 
|         |   6552     assume h2:  | 
|         |   6553       "abc_steps_l ac aprog astp = (aa, ba)" "stp \<le> bstp" | 
|         |   6554       "t_steps tc (tm_of aprog, 0) bstp = (aaa, baa, c)"  | 
|         |   6555       "crsp_l (layout_of aprog) (aa, ba) (aaa, baa, c) ires" | 
|         |   6556     hence h3:  | 
|         |   6557       "t_steps tc (tm_of aprog @ tMp n  | 
|         |   6558        (start_of (layout_of aprog) (length aprog) - Suc 0), 0) bstp  | 
|         |   6559                     = (aaa, baa, c)" | 
|         |   6560       apply(intro tm_append_steps, auto) | 
|         |   6561       apply(simp add: crsp_l.simps, rule startof_not0) | 
|         |   6562       done | 
|         |   6563     from h2 have h4: "\<exists> diff. bstp = stp + diff" | 
|         |   6564       apply(rule_tac x = "bstp - stp" in exI, simp) | 
|         |   6565       done | 
|         |   6566     from h4 and h3 and h2  and h1 show "?thesis" | 
|         |   6567       apply(auto) | 
|         |   6568       apply(simp add: state0_ind crsp_l.simps) | 
|         |   6569       apply(subgoal_tac "start_of (layout_of aprog) aa > 0", simp) | 
|         |   6570       apply(rule startof_not0) | 
|         |   6571       done | 
|         |   6572   qed | 
|         |   6573 qed | 
|         |   6574  | 
|         |   6575 lemma abacus_turing_eq_unhalt_case: | 
|         |   6576   assumes layout:  | 
|         |   6577   -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *} | 
|         |   6578   "ly = layout_of aprog" | 
|         |   6579   and compiled:  | 
|         |   6580   -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *} | 
|         |   6581   "tprog = tm_of aprog" | 
|         |   6582   and correspond:  | 
|         |   6583   -- {*  | 
|         |   6584   TM configuration @{text "tc"} and Abacus configuration @{text "ac"} | 
|         |   6585   are in correspondence:  | 
|         |   6586   *} | 
|         |   6587   "crsp_l ly ac tc ires" | 
|         |   6588   and abc_unhalt:  | 
|         |   6589   -- {* | 
|         |   6590   If, no matter how many steps the Abacus program @{text "aprog"} executes, it | 
|         |   6591   may never reach a halt state.  | 
|         |   6592   *} | 
|         |   6593   "\<forall> stp. ((\<lambda> (as, am). as < length aprog) | 
|         |   6594                        (abc_steps_l ac aprog stp))" | 
|         |   6595   and mopup_start: "mop_ss = start_of ly (length aprog)" | 
|         |   6596   shows | 
|         |   6597   -- {* | 
|         |   6598   The the TM composed of TM @{text "tprog"} and the moupup TM may never reach a halt state as well. | 
|         |   6599   *} | 
|         |   6600   "\<not> (\<exists> stp. (\<lambda> (s, l, r). s = 0) | 
|         |   6601               (t_steps tc (tprog @ (tMp n (mop_ss - 1)), 0) stp))" | 
|         |   6602   using layout compiled correspond abc_unhalt mopup_start | 
|         |   6603   apply(rule_tac abacus_turing_eq_unhalt_case_pre, auto) | 
|         |   6604   done | 
|         |   6605  | 
|         |   6606  | 
|         |   6607 definition abc_list_crsp:: "nat list \<Rightarrow> nat list \<Rightarrow> bool" | 
|         |   6608   where | 
|         |   6609   "abc_list_crsp xs ys = (\<exists> n. xs = ys @ 0\<^bsup>n\<^esup> \<or> ys = xs @ 0\<^bsup>n\<^esup>)" | 
|         |   6610 lemma [intro]: "abc_list_crsp (lm @ 0\<^bsup>m\<^esup>) lm" | 
|         |   6611 apply(auto simp: abc_list_crsp_def) | 
|         |   6612 done | 
|         |   6613  | 
|         |   6614 lemma abc_list_crsp_lm_v:  | 
|         |   6615   "abc_list_crsp lma lmb \<Longrightarrow> abc_lm_v lma n = abc_lm_v lmb n" | 
|         |   6616 apply(auto simp: abc_list_crsp_def abc_lm_v.simps  | 
|         |   6617                  nth_append exponent_def) | 
|         |   6618 done | 
|         |   6619  | 
|         |   6620 lemma  rep_app_cons_iff:  | 
|         |   6621   "k < n \<Longrightarrow> replicate n a[k:=b] =  | 
|         |   6622           replicate k a @ b # replicate (n - k - 1) a" | 
|         |   6623 apply(induct n arbitrary: k, simp) | 
|         |   6624 apply(simp split:nat.splits) | 
|         |   6625 done | 
|         |   6626  | 
|         |   6627 lemma abc_list_crsp_lm_s:  | 
|         |   6628   "abc_list_crsp lma lmb \<Longrightarrow>  | 
|         |   6629       abc_list_crsp (abc_lm_s lma m n) (abc_lm_s lmb m n)" | 
|         |   6630 apply(auto simp: abc_list_crsp_def abc_lm_v.simps abc_lm_s.simps) | 
|         |   6631 apply(simp_all add: list_update_append, auto simp: exponent_def) | 
|         |   6632 proof - | 
|         |   6633   fix na | 
|         |   6634   assume h: "m < length lmb + na" " \<not> m < length lmb" | 
|         |   6635   hence "m - length lmb < na" by simp | 
|         |   6636   hence "replicate na 0[(m- length lmb):= n] =  | 
|         |   6637            replicate (m - length lmb) 0 @ n #  | 
|         |   6638               replicate (na - (m - length lmb) - 1) 0" | 
|         |   6639     apply(erule_tac rep_app_cons_iff) | 
|         |   6640     done | 
|         |   6641   thus "\<exists>nb. replicate na 0[m - length lmb := n] = | 
|         |   6642                  replicate (m - length lmb) 0 @ n # replicate nb 0 \<or> | 
|         |   6643                  replicate (m - length lmb) 0 @ [n] = | 
|         |   6644                  replicate na 0[m - length lmb := n] @ replicate nb 0" | 
|         |   6645     apply(auto) | 
|         |   6646     done | 
|         |   6647 next | 
|         |   6648   fix na | 
|         |   6649   assume h: "\<not> m < length lmb + na" | 
|         |   6650   show  | 
|         |   6651     "\<exists>nb. replicate na 0 @ replicate (m - (length lmb + na)) 0 @ [n] = | 
|         |   6652            replicate (m - length lmb) 0 @ n # replicate nb 0 \<or> | 
|         |   6653           replicate (m - length lmb) 0 @ [n] = | 
|         |   6654             replicate na 0 @ | 
|         |   6655             replicate (m - (length lmb + na)) 0 @ n # replicate nb 0" | 
|         |   6656     apply(rule_tac x = 0 in exI, simp, auto) | 
|         |   6657     using h | 
|         |   6658     apply(simp add: replicate_add[THEN sym]) | 
|         |   6659     done | 
|         |   6660 next | 
|         |   6661   fix na | 
|         |   6662   assume h: "\<not> m < length lma" "m < length lma + na" | 
|         |   6663   hence "m - length lma < na" by simp | 
|         |   6664   hence  | 
|         |   6665     "replicate na 0[(m- length lma):= n] = replicate (m - length lma)  | 
|         |   6666                   0 @ n # replicate (na - (m - length lma) - 1) 0" | 
|         |   6667     apply(erule_tac rep_app_cons_iff) | 
|         |   6668     done | 
|         |   6669   thus "\<exists>nb. replicate (m - length lma) 0 @ [n] = | 
|         |   6670                  replicate na 0[m - length lma := n] @ replicate nb 0  | 
|         |   6671            \<or> replicate na 0[m - length lma := n] = | 
|         |   6672                  replicate (m - length lma) 0 @ n # replicate nb 0" | 
|         |   6673     apply(auto) | 
|         |   6674     done | 
|         |   6675 next | 
|         |   6676   fix na | 
|         |   6677   assume "\<not> m < length lma + na" | 
|         |   6678   thus " \<exists>nb. replicate (m - length lma) 0 @ [n] = | 
|         |   6679             replicate na 0 @ | 
|         |   6680             replicate (m - (length lma + na)) 0 @ n # replicate nb 0  | 
|         |   6681         \<or>   replicate na 0 @  | 
|         |   6682                replicate (m - (length lma + na)) 0 @ [n] = | 
|         |   6683             replicate (m - length lma) 0 @ n # replicate nb 0" | 
|         |   6684     apply(rule_tac x = 0 in exI, simp, auto) | 
|         |   6685     apply(simp add: replicate_add[THEN sym]) | 
|         |   6686     done | 
|         |   6687 qed | 
|         |   6688  | 
|         |   6689 lemma abc_list_crsp_step:  | 
|         |   6690   "\<lbrakk>abc_list_crsp lma lmb; abc_step_l (aa, lma) i = (a, lma');  | 
|         |   6691     abc_step_l (aa, lmb) i = (a', lmb')\<rbrakk> | 
|         |   6692     \<Longrightarrow> a' = a \<and> abc_list_crsp lma' lmb'" | 
|         |   6693 apply(case_tac i, auto simp: abc_step_l.simps  | 
|         |   6694        abc_list_crsp_lm_s abc_list_crsp_lm_v Let_def  | 
|         |   6695                        split: abc_inst.splits if_splits) | 
|         |   6696 done | 
|         |   6697  | 
|         |   6698 lemma abc_steps_red:  | 
|         |   6699   "abc_steps_l ac aprog stp = (as, am) \<Longrightarrow> | 
|         |   6700      abc_steps_l ac aprog (Suc stp) =  | 
|         |   6701            abc_step_l (as, am) (abc_fetch as aprog)" | 
|         |   6702 using abc_steps_ind[of ac aprog stp] | 
|         |   6703 apply(simp) | 
|         |   6704 done | 
|         |   6705  | 
|         |   6706 lemma abc_list_crsp_steps:  | 
|         |   6707   "\<lbrakk>abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog stp = (a, lm'); aprog \<noteq> []\<rbrakk>  | 
|         |   6708       \<Longrightarrow> \<exists> lma. abc_steps_l (0, lm) aprog stp = (a, lma) \<and>  | 
|         |   6709                                           abc_list_crsp lm' lma" | 
|         |   6710 apply(induct stp arbitrary: a lm', simp add: abc_steps_l.simps, auto) | 
|         |   6711 apply(case_tac "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog stp",  | 
|         |   6712       simp add: abc_steps_ind) | 
|         |   6713 proof - | 
|         |   6714   fix stp a lm' aa b | 
|         |   6715   assume ind: | 
|         |   6716     "\<And>a lm'. aa = a \<and> b = lm' \<Longrightarrow>  | 
|         |   6717      \<exists>lma. abc_steps_l (0, lm) aprog stp = (a, lma) \<and> | 
|         |   6718                                           abc_list_crsp lm' lma" | 
|         |   6719     and h: "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog (Suc stp) = (a, lm')"  | 
|         |   6720            "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog stp = (aa, b)"  | 
|         |   6721            "aprog \<noteq> []" | 
|         |   6722   hence g1: "abc_steps_l (0, lm @ 0\<^bsup>m\<^esup>) aprog (Suc stp) | 
|         |   6723           = abc_step_l (aa, b) (abc_fetch aa aprog)" | 
|         |   6724     apply(rule_tac abc_steps_red, simp) | 
|         |   6725     done | 
|         |   6726   have "\<exists>lma. abc_steps_l (0, lm) aprog stp = (aa, lma) \<and>  | 
|         |   6727               abc_list_crsp b lma" | 
|         |   6728     apply(rule_tac ind, simp) | 
|         |   6729     done | 
|         |   6730   from this obtain lma where g2:  | 
|         |   6731     "abc_steps_l (0, lm) aprog stp = (aa, lma) \<and>  | 
|         |   6732      abc_list_crsp b lma"   .. | 
|         |   6733   hence g3: "abc_steps_l (0, lm) aprog (Suc stp) | 
|         |   6734           = abc_step_l (aa, lma) (abc_fetch aa aprog)" | 
|         |   6735     apply(rule_tac abc_steps_red, simp) | 
|         |   6736     done | 
|         |   6737   show "\<exists>lma. abc_steps_l (0, lm) aprog (Suc stp) = (a, lma) \<and>   | 
|         |   6738               abc_list_crsp lm' lma" | 
|         |   6739     using g1 g2 g3 h | 
|         |   6740     apply(auto) | 
|         |   6741     apply(case_tac "abc_step_l (aa, b) (abc_fetch aa aprog)", | 
|         |   6742           case_tac "abc_step_l (aa, lma) (abc_fetch aa aprog)", simp) | 
|         |   6743     apply(rule_tac abc_list_crsp_step, auto) | 
|         |   6744     done | 
|         |   6745 qed | 
|         |   6746  | 
|         |   6747 lemma [simp]: "(case ca of [] \<Rightarrow> Bk | Bk # xs \<Rightarrow> Bk | Oc # xs \<Rightarrow> Oc) = | 
|         |   6748                 (case ca of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)" | 
|         |   6749 by(case_tac ca, simp_all, case_tac a, simp, simp) | 
|         |   6750  | 
|         |   6751 lemma steps_eq: "length t mod 2 = 0 \<Longrightarrow>  | 
|         |   6752                     t_steps c (t, 0) stp = steps c t stp" | 
|         |   6753 apply(induct stp) | 
|         |   6754 apply(simp add: steps.simps t_steps.simps) | 
|         |   6755 apply(simp add:tstep_red t_steps_ind) | 
|         |   6756 apply(case_tac "steps c t stp", simp) | 
|         |   6757 apply(auto simp: t_step.simps tstep.simps) | 
|         |   6758 done | 
|         |   6759  | 
|         |   6760 lemma crsp_l_start: "crsp_l ly (0, lm) (Suc 0, Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>) ires" | 
|         |   6761 apply(simp add: crsp_l.simps, auto simp: start_of.simps) | 
|         |   6762 done | 
|         |   6763  | 
|         |   6764 lemma t_ncorrect_app: "\<lbrakk>t_ncorrect t1; t_ncorrect t2\<rbrakk> \<Longrightarrow>  | 
|         |   6765                                           t_ncorrect (t1 @ t2)" | 
|         |   6766 apply(simp add: t_ncorrect.simps, auto) | 
|         |   6767 done | 
|         |   6768  | 
|         |   6769 lemma [simp]:  | 
|         |   6770   "(length (tm_of aprog) +  | 
|         |   6771     length (tMp n (start_of ly (length aprog) - Suc 0))) mod 2 = 0" | 
|         |   6772 apply(subgoal_tac  | 
|         |   6773  "t_ncorrect (tm_of aprog @ tMp n  | 
|         |   6774              (start_of ly (length aprog) - Suc 0))") | 
|         |   6775 apply(simp add: t_ncorrect.simps) | 
|         |   6776 apply(rule_tac t_ncorrect_app,  | 
|         |   6777       auto simp: tMp.simps t_ncorrect.simps tshift.simps mp_up_def) | 
|         |   6778 apply(subgoal_tac | 
|         |   6779        "t_ncorrect (tm_of aprog)", simp add: t_ncorrect.simps) | 
|         |   6780 apply(auto) | 
|         |   6781 done | 
|         |   6782  | 
|         |   6783 lemma [simp]: "takeWhile (\<lambda>a. a = Oc)  | 
|         |   6784               (replicate rs Oc @ replicate rn Bk) = replicate rs Oc" | 
|         |   6785 apply(induct rs, auto) | 
|         |   6786 apply(induct rn, auto) | 
|         |   6787 done | 
|         |   6788  | 
|         |   6789 lemma abacus_turing_eq_halt':  | 
|         |   6790   "\<lbrakk>ly = layout_of aprog;  | 
|         |   6791     tprog = tm_of aprog;  | 
|         |   6792     n < length am; | 
|         |   6793     abc_steps_l (0, lm) aprog stp = (as, am);  | 
|         |   6794     mop_ss = start_of ly (length aprog); | 
|         |   6795     as \<ge> length aprog\<rbrakk> | 
|         |   6796     \<Longrightarrow> \<exists> stp m l. steps (Suc 0, Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>)  | 
|         |   6797                 (tprog @ (tMp n (mop_ss - 1))) stp | 
|         |   6798                   = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>l\<^esup>)" | 
|         |   6799 apply(drule_tac tc = "(Suc 0, Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>)" in  | 
|         |   6800                abacus_turing_eq_halt_case, auto intro: crsp_l_start) | 
|         |   6801 apply(subgoal_tac  | 
|         |   6802          "length (tm_of aprog @ tMp n  | 
|         |   6803                   (start_of ly (length aprog) - Suc 0)) mod 2 = 0") | 
|         |   6804 apply(simp add: steps_eq) | 
|         |   6805 apply(rule_tac x = stpa in exI,  | 
|         |   6806        simp add:  exponent_def, auto) | 
|         |   6807 done | 
|         |   6808  | 
|         |   6809  | 
|         |   6810 lemma list_length: "xs = ys \<Longrightarrow> length xs = length ys" | 
|         |   6811 by simp | 
|         |   6812 lemma [elim]: "tinres (Bk\<^bsup>m\<^esup>) b \<Longrightarrow> \<exists>m. b = Bk\<^bsup>m\<^esup>" | 
|         |   6813 apply(auto simp: tinres_def) | 
|         |   6814 apply(rule_tac x = "m-n" in exI,  | 
|         |   6815              auto simp: exponent_def replicate_add[THEN sym])  | 
|         |   6816 apply(case_tac "m < n", auto) | 
|         |   6817 apply(drule_tac list_length, auto) | 
|         |   6818 apply(subgoal_tac "\<exists> d. m = d + n", auto simp: replicate_add) | 
|         |   6819 apply(rule_tac x = "m - n" in exI, simp) | 
|         |   6820 done | 
|         |   6821 lemma [intro]: "tinres [Bk] (Bk\<^bsup>k\<^esup>) " | 
|         |   6822 apply(auto simp: tinres_def exponent_def) | 
|         |   6823 apply(case_tac k, auto) | 
|         |   6824 apply(rule_tac x = "Suc 0" in exI, simp) | 
|         |   6825 done | 
|         |   6826  | 
|         |   6827 lemma abacus_turing_eq_halt_pre:  | 
|         |   6828  "\<lbrakk>ly = layout_of aprog;  | 
|         |   6829    tprog = tm_of aprog;  | 
|         |   6830    n < length am;      | 
|         |   6831    abc_steps_l (0, lm) aprog stp = (as, am);   | 
|         |   6832    mop_ss = start_of ly (length aprog); | 
|         |   6833    as \<ge> length aprog\<rbrakk> | 
|         |   6834   \<Longrightarrow> \<exists> stp m l. steps  (Suc 0, Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>) | 
|         |   6835                (tprog @ (tMp n (mop_ss - 1))) stp | 
|         |   6836                  = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>l\<^esup>)" | 
|         |   6837 using abacus_turing_eq_halt' | 
|         |   6838 apply(simp) | 
|         |   6839 done | 
|         |   6840  | 
|         |   6841  | 
|         |   6842 text {* | 
|         |   6843   Main theorem for the case when the original Abacus program does halt. | 
|         |   6844 *} | 
|         |   6845  | 
|         |   6846 lemma abacus_turing_eq_halt:  | 
|         |   6847   assumes layout: | 
|         |   6848   "ly = layout_of aprog" | 
|         |   6849   -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *} | 
|         |   6850   and compiled: "tprog = tm_of aprog" | 
|         |   6851   -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *} | 
|         |   6852   and halt_state:  | 
|         |   6853    -- {* @{text "as"} is a program label outside the range of @{text "aprog"}. So  | 
|         |   6854   if Abacus is in such a state, it is in halt state: *} | 
|         |   6855   "as \<ge> length aprog" | 
|         |   6856   and abc_exec:  | 
|         |   6857   -- {* Supposing after @{text "stp"} step of execution, Abacus program @{text "aprog"} | 
|         |   6858   reaches such a halt state: *} | 
|         |   6859   "abc_steps_l (0, lm) aprog stp = (as, am)" | 
|         |   6860   and rs_locate:  | 
|         |   6861    -- {* @{text "n"} is a memory address in the range of Abacus memory @{text "am"}: *} | 
|         |   6862   "n < length am"   | 
|         |   6863   and mopup_start:  | 
|         |   6864    -- {* The startling label for mopup mahines, according to the layout and Abacus program  | 
|         |   6865    should be @{text "mop_ss"}: *} | 
|         |   6866   "mop_ss = start_of ly (length aprog)" | 
|         |   6867   shows  | 
|         |   6868   -- {*  | 
|         |   6869   After @{text "stp"} steps of execution of the TM composed of @{text "tprog"} and the mopup  | 
|         |   6870   TM @{text "(tMp n (mop_ss - 1))"} will halt and gives rise to a configuration which  | 
|         |   6871   only hold the content of memory cell @{text "n"}: | 
|         |   6872   *} | 
|         |   6873   "\<exists> stp m l. steps (Suc 0, Bk # Bk # ires, <lm> @ Bk\<^bsup>rn\<^esup>) (tprog @ (tMp n (mop_ss - 1))) stp | 
|         |   6874                       = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (abc_lm_v am n)\<^esup> @ Bk\<^bsup>l\<^esup>)" | 
|         |   6875   using layout compiled halt_state abc_exec rs_locate mopup_start | 
|         |   6876   by(rule_tac abacus_turing_eq_halt_pre, auto) | 
|         |   6877  | 
|         |   6878 lemma abacus_turing_eq_uhalt':  | 
|         |   6879  "\<lbrakk>ly = layout_of aprog;  | 
|         |   6880    tprog = tm_of aprog;  | 
|         |   6881    \<forall> stp. ((\<lambda> (as, am). as < length aprog)  | 
|         |   6882                    (abc_steps_l (0, lm) aprog stp)); | 
|         |   6883    mop_ss = start_of ly (length aprog)\<rbrakk> | 
|         |   6884   \<Longrightarrow> (\<not> (\<exists> stp. isS0 (steps (Suc 0, [Bk, Bk], <lm>)  | 
|         |   6885                       (tprog @ (tMp n (mop_ss - 1))) stp)))" | 
|         |   6886 apply(drule_tac tc = "(Suc 0, [Bk, Bk], <lm>)" and n = n and ires = "[]" in  | 
|         |   6887          abacus_turing_eq_unhalt_case, auto intro: crsp_l_start) | 
|         |   6888 apply(simp add: crsp_l.simps start_of.simps) | 
|         |   6889 apply(erule_tac x = stp in allE, erule_tac x = stp in allE) | 
|         |   6890 apply(subgoal_tac  | 
|         |   6891    "length (tm_of aprog @ tMp n  | 
|         |   6892          (start_of ly (length aprog) - Suc 0)) mod 2 = 0") | 
|         |   6893 apply(simp add: steps_eq, auto simp: isS0_def) | 
|         |   6894 done | 
|         |   6895  | 
|         |   6896 text {* | 
|         |   6897   Main theorem for the case when the original Abacus program does not halt. | 
|         |   6898   *} | 
|         |   6899 lemma abacus_turing_eq_uhalt: | 
|         |   6900   assumes layout:  | 
|         |   6901   -- {* There is an Abacus program @{text "aprog"} with layout @{text "ly"}: *} | 
|         |   6902   "ly = layout_of aprog" | 
|         |   6903   and compiled: | 
|         |   6904    -- {* The TM compiled from @{text "aprog"} is @{text "tprog"}: *} | 
|         |   6905   "tprog = tm_of aprog" | 
|         |   6906   and abc_unhalt: | 
|         |   6907   -- {* | 
|         |   6908   If, no matter how many steps the Abacus program @{text "aprog"} executes, it | 
|         |   6909   may never reach a halt state.  | 
|         |   6910   *} | 
|         |   6911   "\<forall> stp. ((\<lambda> (as, am). as < length aprog)  | 
|         |   6912                       (abc_steps_l (0, lm) aprog stp))" | 
|         |   6913   and mop_start: "mop_ss = start_of ly (length aprog)" | 
|         |   6914   shows  | 
|         |   6915    -- {* | 
|         |   6916   The the TM composed of TM @{text "tprog"} and the moupup TM may never reach a halt state as well. | 
|         |   6917   *} | 
|         |   6918   "\<not> (\<exists> stp. isS0 (steps (Suc 0, [Bk, Bk], <lm>)  | 
|         |   6919                     (tprog @ (tMp n (mop_ss - 1))) stp))" | 
|         |   6920   using abacus_turing_eq_uhalt' | 
|         |   6921         layout compiled abc_unhalt mop_start | 
|         |   6922   by(auto) | 
|         |   6923  | 
|         |   6924 end | 
|         |   6925  |