thys/UTM.thy
changeset 131 e995ae949731
parent 130 1e89c65f844b
child 133 ca7fb6848715
equal deleted inserted replaced
130:1e89c65f844b 131:e995ae949731
     1 theory UTM
     1 theory UTM
     2 imports Main uncomputable recursive abacus UF GCD 
     2 imports Main recursive abacus UF GCD turing_hoare
     3 begin
     3 begin
     4 
     4 
     5 section {* Wang coding of input arguments *}
     5 section {* Wang coding of input arguments *}
     6 
     6 
     7 text {*
     7 text {*
    22 \newlength{\baseheight}
    22 \newlength{\baseheight}
    23 \settoheight{\baseheight}{$B:R$}
    23 \settoheight{\baseheight}{$B:R$}
    24 \newcommand{\vsep}{5\baseheight}
    24 \newcommand{\vsep}{5\baseheight}
    25 
    25 
    26 The TM used to generate the Wang's code of input arguments is divided into three TMs
    26 The TM used to generate the Wang's code of input arguments is divided into three TMs
    27  executed sequentially, namely $prepare$, $mainwork$ and $adjust$¡£According to the
    27  executed sequentially, namely $prepare$, $mainwork$ and $adjust$\<exclamdown>\<pounds>According to the
    28  convention, start state of ever TM is fixed to state $1$ while the final state is
    28  convention, start state of ever TM is fixed to state $1$ while the final state is
    29  fixed to $0$.
    29  fixed to $0$.
    30 
    30 
    31 The input and output of $prepare$ are illustrated respectively by Figure
    31 The input and output of $prepare$ are illustrated respectively by Figure
    32 \ref{prepare_input} and \ref{prepare_output}.
    32 \ref{prepare_input} and \ref{prepare_output}.
   507 
   507 
   508 definition fourtimes_ly :: "nat list"
   508 definition fourtimes_ly :: "nat list"
   509   where
   509   where
   510   "fourtimes_ly = layout_of abc_fourtimes"
   510   "fourtimes_ly = layout_of abc_fourtimes"
   511 
   511 
   512 definition t_twice :: "tprog"
   512 definition t_twice_compile :: "instr list"
   513   where
   513 where
   514   "t_twice = change_termi_state (tm_of (abc_twice) @ (tMp 1 (start_of twice_ly (length abc_twice) - Suc 0)))"
   514   "t_twice_compile= (tm_of abc_twice @ (shift (mopup 1) (length (tm_of abc_twice) div 2)))"
   515 
   515 
   516 definition t_fourtimes :: "tprog"
   516 definition t_twice :: "instr list"
   517   where
   517   where
   518   "t_fourtimes = change_termi_state (tm_of (abc_fourtimes) @ 
   518   "t_twice = adjust t_twice_compile"
   519              (tMp 1 (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)))"
   519 
   520 
   520 definition t_fourtimes_compile :: "instr list"
       
   521 where
       
   522   "t_fourtimes_compile= (tm_of abc_fourtimes @ (shift (mopup 1) (length (tm_of abc_fourtimes) div 2)))"
       
   523 
       
   524 definition t_fourtimes :: "instr list"
       
   525   where
       
   526   "t_fourtimes = adjust t_fourtimes_compile"
   521 
   527 
   522 definition t_twice_len :: "nat"
   528 definition t_twice_len :: "nat"
   523   where
   529   where
   524   "t_twice_len = length t_twice div 2"
   530   "t_twice_len = length t_twice div 2"
   525 
   531 
   526 definition t_wcode_main_first_part:: "tprog"
   532 definition t_wcode_main_first_part:: "instr list"
   527   where
   533   where
   528   "t_wcode_main_first_part \<equiv> 
   534   "t_wcode_main_first_part \<equiv> 
   529                    [(L, 1), (L, 2), (L, 7), (R, 3),
   535                    [(L, 1), (L, 2), (L, 7), (R, 3),
   530                     (R, 4), (W0, 3), (R, 4), (R, 5),
   536                     (R, 4), (W0, 3), (R, 4), (R, 5),
   531                     (W1, 6), (R, 5), (R, 13), (L, 6),
   537                     (W1, 6), (R, 5), (R, 13), (L, 6),
   532                     (R, 0), (R, 8), (R, 9), (Nop, 8),
   538                     (R, 0), (R, 8), (R, 9), (Nop, 8),
   533                     (R, 10), (W0, 9), (R, 10), (R, 11), 
   539                     (R, 10), (W0, 9), (R, 10), (R, 11), 
   534                     (W1, 12), (R, 11), (R, t_twice_len + 14), (L, 12)]"
   540                     (W1, 12), (R, 11), (R, t_twice_len + 14), (L, 12)]"
   535 
   541 
   536 definition t_wcode_main :: "tprog"
   542 definition t_wcode_main :: "instr list"
   537   where
   543   where
   538   "t_wcode_main = (t_wcode_main_first_part @ tshift t_twice 12 @ [(L, 1), (L, 1)]
   544   "t_wcode_main = (t_wcode_main_first_part @ shift t_twice 12 @ [(L, 1), (L, 1)]
   539                     @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])"
   545                     @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])"
   540 
   546 
   541 fun bl_bin :: "block list \<Rightarrow> nat"
   547 fun bl_bin :: "cell list \<Rightarrow> nat"
   542   where
   548   where
   543   "bl_bin [] = 0" 
   549   "bl_bin [] = 0" 
   544 | "bl_bin (Bk # xs) = 2 * bl_bin xs"
   550 | "bl_bin (Bk # xs) = 2 * bl_bin xs"
   545 | "bl_bin (Oc # xs) = Suc (2 * bl_bin xs)"
   551 | "bl_bin (Oc # xs) = Suc (2 * bl_bin xs)"
   546 
   552 
   547 declare bl_bin.simps[simp del]
   553 declare bl_bin.simps[simp del]
   548 
   554 
   549 type_synonym bin_inv_t = "block list \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
   555 type_synonym bin_inv_t = "cell list \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
   550 
   556 
   551 fun wcode_before_double :: "bin_inv_t"
   557 fun wcode_before_double :: "bin_inv_t"
   552   where
   558   where
   553   "wcode_before_double ires rs (l, r) =
   559   "wcode_before_double ires rs (l, r) =
   554      (\<exists> ln rn. l = Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
   560      (\<exists> ln rn. l = Bk # Bk # Bk\<up>(ln) @ Oc # ires \<and> 
   555                r = Oc\<^bsup>(Suc (Suc rs))\<^esup> @ Bk\<^bsup>rn \<^esup>)"
   561                r = Oc\<up>((Suc (Suc rs))) @ Bk\<up>(rn ))"
   556 
   562 
   557 declare wcode_before_double.simps[simp del]
   563 declare wcode_before_double.simps[simp del]
   558 
   564 
   559 fun wcode_after_double :: "bin_inv_t"
   565 fun wcode_after_double :: "bin_inv_t"
   560   where
   566   where
   561   "wcode_after_double ires rs (l, r) = 
   567   "wcode_after_double ires rs (l, r) = 
   562      (\<exists> ln rn. l = Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and>
   568      (\<exists> ln rn. l = Bk # Bk # Bk\<up>(ln) @ Oc # ires \<and>
   563          r = Oc\<^bsup>Suc (Suc (Suc 2*rs))\<^esup> @ Bk\<^bsup>rn\<^esup>)"
   569          r = Oc\<up>(Suc (Suc (Suc 2*rs))) @ Bk\<up>(rn))"
   564 
   570 
   565 declare wcode_after_double.simps[simp del]
   571 declare wcode_after_double.simps[simp del]
   566 
   572 
   567 fun wcode_on_left_moving_1_B :: "bin_inv_t"
   573 fun wcode_on_left_moving_1_B :: "bin_inv_t"
   568   where
   574   where
   569   "wcode_on_left_moving_1_B ires rs (l, r) = 
   575   "wcode_on_left_moving_1_B ires rs (l, r) = 
   570      (\<exists> ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # Oc # ires \<and> 
   576      (\<exists> ml mr rn. l = Bk\<up>(ml) @ Oc # Oc # ires \<and> 
   571                r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
   577                r = Bk\<up>(mr) @ Oc\<up>(Suc rs) @ Bk\<up>(rn) \<and>
   572                ml + mr > Suc 0 \<and> mr > 0)"
   578                ml + mr > Suc 0 \<and> mr > 0)"
   573 
   579 
   574 declare wcode_on_left_moving_1_B.simps[simp del]
   580 declare wcode_on_left_moving_1_B.simps[simp del]
   575 
   581 
   576 fun wcode_on_left_moving_1_O :: "bin_inv_t"
   582 fun wcode_on_left_moving_1_O :: "bin_inv_t"
   577   where
   583   where
   578   "wcode_on_left_moving_1_O ires rs (l, r) = 
   584   "wcode_on_left_moving_1_O ires rs (l, r) = 
   579      (\<exists> ln rn.
   585      (\<exists> ln rn.
   580                l = Oc # ires \<and> 
   586                l = Oc # ires \<and> 
   581                r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
   587                r = Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
   582 
   588 
   583 declare wcode_on_left_moving_1_O.simps[simp del]
   589 declare wcode_on_left_moving_1_O.simps[simp del]
   584 
   590 
   585 fun wcode_on_left_moving_1 :: "bin_inv_t"
   591 fun wcode_on_left_moving_1 :: "bin_inv_t"
   586   where
   592   where
   591 
   597 
   592 fun wcode_on_checking_1 :: "bin_inv_t"
   598 fun wcode_on_checking_1 :: "bin_inv_t"
   593   where
   599   where
   594    "wcode_on_checking_1 ires rs (l, r) = 
   600    "wcode_on_checking_1 ires rs (l, r) = 
   595     (\<exists> ln rn. l = ires \<and>
   601     (\<exists> ln rn. l = ires \<and>
   596               r = Oc # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
   602               r = Oc # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
   597 
   603 
   598 fun wcode_erase1 :: "bin_inv_t"
   604 fun wcode_erase1 :: "bin_inv_t"
   599   where
   605   where
   600 "wcode_erase1 ires rs (l, r) = 
   606 "wcode_erase1 ires rs (l, r) = 
   601        (\<exists> ln rn. l = Oc # ires \<and> 
   607        (\<exists> ln rn. l = Oc # ires \<and> 
   602                  tl r = Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
   608                  tl r = Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
   603 
   609 
   604 declare wcode_erase1.simps [simp del]
   610 declare wcode_erase1.simps [simp del]
   605 
   611 
   606 fun wcode_on_right_moving_1 :: "bin_inv_t"
   612 fun wcode_on_right_moving_1 :: "bin_inv_t"
   607   where
   613   where
   608   "wcode_on_right_moving_1 ires rs (l, r) = 
   614   "wcode_on_right_moving_1 ires rs (l, r) = 
   609        (\<exists> ml mr rn.        
   615        (\<exists> ml mr rn.        
   610              l = Bk\<^bsup>ml\<^esup> @ Oc # ires \<and> 
   616              l = Bk\<up>(ml) @ Oc # ires \<and> 
   611              r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
   617              r = Bk\<up>(mr) @ Oc\<up>(Suc rs) @ Bk\<up>(rn) \<and>
   612              ml + mr > Suc 0)"
   618              ml + mr > Suc 0)"
   613 
   619 
   614 declare wcode_on_right_moving_1.simps [simp del] 
   620 declare wcode_on_right_moving_1.simps [simp del] 
   615 
   621 
   616 declare wcode_on_right_moving_1.simps[simp del]
   622 declare wcode_on_right_moving_1.simps[simp del]
   617 
   623 
   618 fun wcode_goon_right_moving_1 :: "bin_inv_t"
   624 fun wcode_goon_right_moving_1 :: "bin_inv_t"
   619   where
   625   where
   620   "wcode_goon_right_moving_1 ires rs (l, r) = 
   626   "wcode_goon_right_moving_1 ires rs (l, r) = 
   621       (\<exists> ml mr ln rn. 
   627       (\<exists> ml mr ln rn. 
   622             l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
   628             l = Oc\<up>(ml) @ Bk # Bk # Bk\<up>(ln) @ Oc # ires \<and> 
   623             r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
   629             r = Oc\<up>(mr) @ Bk\<up>(rn) \<and>
   624             ml + mr = Suc rs)"
   630             ml + mr = Suc rs)"
   625 
   631 
   626 declare wcode_goon_right_moving_1.simps[simp del]
   632 declare wcode_goon_right_moving_1.simps[simp del]
   627 
   633 
   628 fun wcode_backto_standard_pos_B :: "bin_inv_t"
   634 fun wcode_backto_standard_pos_B :: "bin_inv_t"
   629   where
   635   where
   630   "wcode_backto_standard_pos_B ires rs (l, r) = 
   636   "wcode_backto_standard_pos_B ires rs (l, r) = 
   631           (\<exists> ln rn. l =  Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
   637           (\<exists> ln rn. l =  Bk # Bk\<up>(ln) @ Oc # ires \<and> 
   632                r =  Bk # Oc\<^bsup>(Suc (Suc rs))\<^esup> @ Bk\<^bsup>rn \<^esup>)"
   638                r =  Bk # Oc\<up>((Suc (Suc rs))) @ Bk\<up>(rn ))"
   633 
   639 
   634 declare wcode_backto_standard_pos_B.simps[simp del]
   640 declare wcode_backto_standard_pos_B.simps[simp del]
   635 
   641 
   636 fun wcode_backto_standard_pos_O :: "bin_inv_t"
   642 fun wcode_backto_standard_pos_O :: "bin_inv_t"
   637   where
   643   where
   638    "wcode_backto_standard_pos_O ires rs (l, r) = 
   644    "wcode_backto_standard_pos_O ires rs (l, r) = 
   639         (\<exists> ml mr ln rn. 
   645         (\<exists> ml mr ln rn. 
   640             l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and>
   646             l = Oc\<up>(ml) @ Bk # Bk # Bk\<up>(ln) @ Oc # ires \<and>
   641             r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
   647             r = Oc\<up>(mr) @ Bk\<up>(rn) \<and>
   642             ml + mr = Suc (Suc rs) \<and> mr > 0)"
   648             ml + mr = Suc (Suc rs) \<and> mr > 0)"
   643 
   649 
   644 declare wcode_backto_standard_pos_O.simps[simp del]
   650 declare wcode_backto_standard_pos_O.simps[simp del]
   645 
   651 
   646 fun wcode_backto_standard_pos :: "bin_inv_t"
   652 fun wcode_backto_standard_pos :: "bin_inv_t"
   649                                             wcode_backto_standard_pos_O ires rs (l, r))"
   655                                             wcode_backto_standard_pos_O ires rs (l, r))"
   650 
   656 
   651 declare wcode_backto_standard_pos.simps[simp del]
   657 declare wcode_backto_standard_pos.simps[simp del]
   652 
   658 
   653 lemma [simp]: "<0::nat> = [Oc]"
   659 lemma [simp]: "<0::nat> = [Oc]"
   654 apply(simp add: tape_of_nat_abv exponent_def tape_of_nat_list.simps)
   660 apply(simp add: tape_of_nat_abv tape_of_nat_list.simps)
   655 done
   661 done
   656 
   662 
   657 lemma tape_of_Suc_nat: "<Suc (a ::nat)> = replicate a Oc @ [Oc, Oc]"
   663 lemma tape_of_Suc_nat: "<Suc (a ::nat)> = replicate a Oc @ [Oc, Oc]"
   658 apply(simp add: tape_of_nat_abv exp_ind tape_of_nat_list.simps)
   664 apply(simp only: tape_of_nat_abv exp_ind, simp)
   659 apply(simp only: exp_ind_def[THEN sym])
       
   660 apply(simp only: exp_ind, simp, simp add: exponent_def)
       
   661 done
   665 done
   662 
   666 
   663 lemma [simp]: "length (<a::nat>) = Suc a"
   667 lemma [simp]: "length (<a::nat>) = Suc a"
   664 apply(simp add: tape_of_nat_abv tape_of_nat_list.simps)
   668 apply(simp add: tape_of_nat_abv tape_of_nat_list.simps)
   665 done
   669 done
   666 
   670 
   667 lemma [simp]: "<[a::nat]> = <a>"
   671 lemma [simp]: "<[a::nat]> = <a>"
   668 apply(simp add: tape_of_nat_abv tape_of_nl_abv exponent_def
   672 apply(simp add: tape_of_nat_abv tape_of_nl_abv
   669                 tape_of_nat_list.simps)
   673   tape_of_nat_list.simps)
   670 done
   674 done
   671 
   675 
   672 lemma bin_wc_eq: "bl_bin xs = bl2wc xs"
   676 lemma bin_wc_eq: "bl_bin xs = bl2wc xs"
   673 proof(induct xs)
   677 proof(induct xs)
   674   show " bl_bin [] = bl2wc []" 
   678   show " bl_bin [] = bl2wc []" 
   681     apply(case_tac a, simp_all add: bl_bin.simps bl2wc.simps)
   685     apply(case_tac a, simp_all add: bl_bin.simps bl2wc.simps)
   682     apply(simp_all add: bl2nat.simps bl2nat_double)
   686     apply(simp_all add: bl2nat.simps bl2nat_double)
   683     done
   687     done
   684 qed
   688 qed
   685 
   689 
   686 declare exp_def[simp del]
       
   687 
       
   688 lemma bl_bin_nat_Suc:  
   690 lemma bl_bin_nat_Suc:  
   689   "bl_bin (<Suc a>) = bl_bin (<a>) + 2^(Suc a)"
   691   "bl_bin (<Suc a>) = bl_bin (<a>) + 2^(Suc a)"
   690 apply(simp add: tape_of_nat_abv bin_wc_eq)
   692 apply(simp add: tape_of_nat_abv bl_bin.simps)
   691 apply(simp add: bl2wc.simps)
   693 apply(induct a, auto simp: bl_bin.simps)
   692 done
   694 done
   693 lemma [simp]: " rev (a\<^bsup>aa\<^esup>) = a\<^bsup>aa\<^esup>"
   695 
   694 apply(simp add: exponent_def)
   696 lemma [simp]: " rev (a\<up>(aa)) = a\<up>(aa)"
   695 done
   697 apply(simp)
   696  
   698 done
   697 declare tape_of_nl_abv_cons[simp del]
   699 
       
   700 lemma tape_of_nl_append_one: "lm \<noteq> [] \<Longrightarrow>  <lm @ [a]> = <lm> @ Bk # Oc\<up>Suc a"
       
   701 apply(induct lm, auto simp: tape_of_nl_cons split:if_splits)
       
   702 done
   698 
   703 
   699 lemma tape_of_nl_rev: "rev (<lm::nat list>) = (<rev lm>)"
   704 lemma tape_of_nl_rev: "rev (<lm::nat list>) = (<rev lm>)"
   700 apply(induct lm rule: list_tl_induct, simp)
   705 apply(induct lm, simp, auto)
   701 apply(case_tac "list = []", simp add: tape_of_nl_abv tape_of_nat_list.simps)
   706 apply(auto simp: tape_of_nl_cons tape_of_nl_append_one split: if_splits)
   702 apply(simp add: tape_of_nat_list_butlast_last tape_of_nl_abv_cons)
   707 apply(simp add: exp_ind[THEN sym])
   703 done
   708 done
   704 lemma [simp]: "a\<^bsup>Suc 0\<^esup> = [a]" 
   709 
   705 by(simp add: exp_def)
   710 lemma [simp]: "a\<up>(Suc 0) = [a]" 
   706 lemma tape_of_nl_cons_app1: "(<a # xs @ [b]>) = (Oc\<^bsup>Suc a\<^esup> @ Bk # (<xs@ [b]>))"
   711 by(simp)
       
   712 
       
   713 lemma tape_of_nl_cons_app1: "(<a # xs @ [b]>) = (Oc\<up>(Suc a) @ Bk # (<xs@ [b]>))"
   707 apply(case_tac xs, simp add: tape_of_nl_abv tape_of_nat_list.simps)
   714 apply(case_tac xs, simp add: tape_of_nl_abv tape_of_nat_list.simps)
   708 apply(simp add: tape_of_nl_abv  tape_of_nat_list.simps)
   715 apply(simp add: tape_of_nl_abv  tape_of_nat_list.simps)
   709 done
   716 done
   710 
   717 
   711 lemma bl_bin_bk_oc[simp]:
   718 lemma bl_bin_bk_oc[simp]:
   714 apply(simp add: bin_wc_eq)
   721 apply(simp add: bin_wc_eq)
   715 using bl2nat_cons_oc[of "xs @ [Bk]"]
   722 using bl2nat_cons_oc[of "xs @ [Bk]"]
   716 apply(simp add: bl2nat_cons_bk bl2wc.simps)
   723 apply(simp add: bl2nat_cons_bk bl2wc.simps)
   717 done
   724 done
   718 
   725 
   719 lemma tape_of_nat[simp]: "(<a::nat>) = Oc\<^bsup>Suc a\<^esup>"
   726 lemma tape_of_nat[simp]: "(<a::nat>) = Oc\<up>(Suc a)"
   720 apply(simp add: tape_of_nat_abv)
   727 apply(simp add: tape_of_nat_abv)
   721 done
   728 done
   722 lemma tape_of_nl_cons_app2: "(<c # xs @ [b]>) = (<c # xs> @ Bk # Oc\<^bsup>Suc b\<^esup>)"
   729 
       
   730 lemma tape_of_nl_cons_app2: "(<c # xs @ [b]>) = (<c # xs> @ Bk # Oc\<up>(Suc b))"
   723 proof(induct "length xs" arbitrary: xs c,
   731 proof(induct "length xs" arbitrary: xs c,
   724   simp add: tape_of_nl_abv  tape_of_nat_list.simps)
   732   simp add: tape_of_nl_abv  tape_of_nat_list.simps)
   725   fix x xs c
   733   fix x xs c
   726   assume ind: "\<And>xs c. x = length xs \<Longrightarrow> <c # xs @ [b]> = 
   734   assume ind: "\<And>xs c. x = length xs \<Longrightarrow> <c # xs @ [b]> = 
   727     <c # xs> @ Bk # Oc\<^bsup>Suc b\<^esup>"
   735     <c # xs> @ Bk # Oc\<up>(Suc b)"
   728     and h: "Suc x = length (xs::nat list)" 
   736     and h: "Suc x = length (xs::nat list)" 
   729   show "<c # xs @ [b]> = <c # xs> @ Bk # Oc\<^bsup>Suc b\<^esup>"
   737   show "<c # xs @ [b]> = <c # xs> @ Bk # Oc\<up>(Suc b)"
   730   proof(case_tac xs, simp add: tape_of_nl_abv  tape_of_nat_list.simps)
   738   proof(case_tac xs, simp add: tape_of_nl_abv  tape_of_nat_list.simps)
   731     fix a list
   739     fix a list
   732     assume g: "xs = a # list"
   740     assume g: "xs = a # list"
   733     hence k: "<a # list @ [b]> =  <a # list> @ Bk # Oc\<^bsup>Suc b\<^esup>"
   741     hence k: "<a # list @ [b]> =  <a # list> @ Bk # Oc\<up>(Suc b)"
   734       apply(rule_tac ind)
   742       apply(rule_tac ind)
   735       using h
   743       using h
   736       apply(simp)
   744       apply(simp)
   737       done
   745       done
   738     from g and k show "<c # xs @ [b]> = <c # xs> @ Bk # Oc\<^bsup>Suc b\<^esup>"
   746     from g and k show "<c # xs @ [b]> = <c # xs> @ Bk # Oc\<up>(Suc b)"
   739       apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
   747       apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
   740       done
   748       done
   741   qed
   749   qed
   742 qed
   750 qed
   743 
   751 
   744 lemma [simp]: "length (<aa # a # list>) = Suc (Suc aa) + length (<a # list>)"
   752 lemma [simp]: "length (<aa # a # list>) = Suc (Suc aa) + length (<a # list>)"
   745 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
   753 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
   746 done
   754 done
   747 
   755 
   748 lemma [simp]: "bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # tape_of_nat_list (a # lista) @ [Bk, Oc]) =
   756 lemma [simp]: "bl_bin (Oc\<up>(Suc aa) @ Bk # tape_of_nat_list (a # lista) @ [Bk, Oc]) =
   749               bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # tape_of_nat_list (a # lista)) + 
   757               bl_bin (Oc\<up>(Suc aa) @ Bk # tape_of_nat_list (a # lista)) + 
   750               2* 2^(length (Oc\<^bsup>Suc aa\<^esup> @ Bk # tape_of_nat_list (a # lista)))"
   758               2* 2^(length (Oc\<up>(Suc aa) @ Bk # tape_of_nat_list (a # lista)))"
   751 using bl_bin_bk_oc[of "Oc\<^bsup>Suc aa\<^esup> @ Bk # tape_of_nat_list (a # lista)"]
   759 using bl_bin_bk_oc[of "Oc\<up>(Suc aa) @ Bk # tape_of_nat_list (a # lista)"]
   752 apply(simp)
   760 apply(simp)
   753 done
   761 done
       
   762 
       
   763 declare replicate_Suc[simp del]
   754 
   764 
   755 lemma [simp]: 
   765 lemma [simp]: 
   756   "bl_bin (<aa # list>) + (4 * rs + 4) * 2 ^ (length (<aa # list>) - Suc 0)
   766   "bl_bin (<aa # list>) + (4 * rs + 4) * 2 ^ (length (<aa # list>) - Suc 0)
   757   = bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # <list @ [0]>) + rs * (2 * 2 ^ (aa + length (<list @ [0]>)))"
   767   = bl_bin (Oc\<up>(Suc aa) @ Bk # <list @ [0]>) + rs * (2 * 2 ^ (aa + length (<list @ [0]>)))"
   758 apply(case_tac "list", simp add: add_mult_distrib, simp)
   768 
       
   769 apply(case_tac "list", simp add: add_mult_distrib)
   759 apply(simp add: tape_of_nl_cons_app2 add_mult_distrib)
   770 apply(simp add: tape_of_nl_cons_app2 add_mult_distrib)
   760 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
   771 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
   761 done
   772 done
   762   
   773 
   763 lemma tape_of_nl_app_Suc: "((<list @ [Suc ab]>)) = (<list @ [ab]>) @ [Oc]"
   774 lemma tape_of_nl_app_Suc: "((<list @ [Suc ab]>)) = (<list @ [ab]>) @ [Oc]"
   764 apply(induct list)
   775 apply(induct list)
   765 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps exp_ind)
   776 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps exp_ind)
   766 apply(case_tac list)
   777 apply(case_tac list)
   767 apply(simp_all add:tape_of_nl_abv tape_of_nat_list.simps exp_ind)
   778 apply(simp_all add:tape_of_nl_abv tape_of_nat_list.simps exp_ind)
   768 done
   779 done
   769 
   780 
   770 lemma [simp]: "bl_bin (Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [ab]> @ [Oc])
   781 lemma [simp]: "bl_bin (Oc # Oc\<up>(aa) @ Bk # <list @ [ab]> @ [Oc])
   771               = bl_bin (Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [ab]>) +
   782               = bl_bin (Oc # Oc\<up>(aa) @ Bk # <list @ [ab]>) +
   772               2^(length (Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [ab]>))"
   783               2^(length (Oc # Oc\<up>(aa) @ Bk # <list @ [ab]>))"
   773 apply(simp add: bin_wc_eq)
   784 apply(simp add: bin_wc_eq)
   774 apply(simp add: bl2nat_cons_oc bl2wc.simps)
   785 apply(simp add: bl2nat_cons_oc bl2wc.simps)
   775 using bl2nat_cons_oc[of "Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [ab]>"]
   786 using bl2nat_cons_oc[of "Oc # Oc\<up>(aa) @ Bk # <list @ [ab]>"]
   776 apply(simp)
   787 apply(simp)
   777 done
   788 done
   778 lemma [simp]: "bl_bin (Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [ab]>) + (4 * 2 ^ (aa + length (<list @ [ab]>)) +
   789 lemma [simp]: "bl_bin (Oc # Oc\<up>(aa) @ Bk # <list @ [ab]>) + (4 * 2 ^ (aa + length (<list @ [ab]>)) +
   779          4 * (rs * 2 ^ (aa + length (<list @ [ab]>)))) =
   790          4 * (rs * 2 ^ (aa + length (<list @ [ab]>)))) =
   780        bl_bin (Oc # Oc\<^bsup>aa\<^esup> @ Bk # <list @ [Suc ab]>) +
   791        bl_bin (Oc # Oc\<up>(aa) @ Bk # <list @ [Suc ab]>) +
   781          rs * (2 * 2 ^ (aa + length (<list @ [Suc ab]>)))"
   792          rs * (2 * 2 ^ (aa + length (<list @ [Suc ab]>)))"
   782 apply(simp add: tape_of_nl_app_Suc)
   793 apply(simp add: tape_of_nl_app_Suc)
   783 done
   794 done
   784 
   795 
   785 declare tape_of_nat[simp del]
   796 declare tape_of_nat[simp del]
   796           else if st = 13 then wcode_before_double ires rs (l, r)
   807           else if st = 13 then wcode_before_double ires rs (l, r)
   797           else False)"
   808           else False)"
   798 
   809 
   799 declare wcode_double_case_inv.simps[simp del]
   810 declare wcode_double_case_inv.simps[simp del]
   800 
   811 
   801 fun wcode_double_case_state :: "t_conf \<Rightarrow> nat"
   812 fun wcode_double_case_state :: "config \<Rightarrow> nat"
   802   where
   813   where
   803   "wcode_double_case_state (st, l, r) = 
   814   "wcode_double_case_state (st, l, r) = 
   804    13 - st"
   815    13 - st"
   805 
   816 
   806 fun wcode_double_case_step :: "t_conf \<Rightarrow> nat"
   817 fun wcode_double_case_step :: "config \<Rightarrow> nat"
   807   where
   818   where
   808   "wcode_double_case_step (st, l, r) = 
   819   "wcode_double_case_step (st, l, r) = 
   809       (if st = Suc 0 then (length l)
   820       (if st = Suc 0 then (length l)
   810       else if st = Suc (Suc 0) then (length r)
   821       else if st = Suc (Suc 0) then (length r)
   811       else if st = 3 then 
   822       else if st = 3 then 
   813       else if st = 4 then (length r)
   824       else if st = 4 then (length r)
   814       else if st = 5 then (length r)
   825       else if st = 5 then (length r)
   815       else if st = 6 then (length l)
   826       else if st = 6 then (length l)
   816       else 0)"
   827       else 0)"
   817 
   828 
   818 fun wcode_double_case_measure :: "t_conf \<Rightarrow> nat \<times> nat"
   829 fun wcode_double_case_measure :: "config \<Rightarrow> nat \<times> nat"
   819   where
   830   where
   820   "wcode_double_case_measure (st, l, r) = 
   831   "wcode_double_case_measure (st, l, r) = 
   821      (wcode_double_case_state (st, l, r), 
   832      (wcode_double_case_state (st, l, r), 
   822       wcode_double_case_step (st, l, r))"
   833       wcode_double_case_step (st, l, r))"
   823 
   834 
   824 definition wcode_double_case_le :: "(t_conf \<times> t_conf) set"
   835 definition wcode_double_case_le :: "(config \<times> config) set"
   825   where "wcode_double_case_le \<equiv> (inv_image lex_pair wcode_double_case_measure)"
   836   where "wcode_double_case_le \<equiv> (inv_image lex_pair wcode_double_case_measure)"
   826 
   837 
   827 lemma [intro]: "wf lex_pair"
   838 lemma [intro]: "wf lex_pair"
   828 by(auto intro:wf_lex_prod simp:lex_pair_def)
   839 by(auto intro:wf_lex_prod simp:lex_pair_def)
   829 
   840 
   855 apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
   866 apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
   856                 fetch.simps nth_of.simps)
   867                 fetch.simps nth_of.simps)
   857 done
   868 done
   858 
   869 
   859 lemma [simp]: "fetch t_wcode_main 4 Bk = (R, 4)"
   870 lemma [simp]: "fetch t_wcode_main 4 Bk = (R, 4)"
   860 apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
   871 apply(subgoal_tac "4 = Suc 3")
   861                 fetch.simps nth_of.simps)
   872 apply(simp only: t_wcode_main_def t_wcode_main_first_part_def
       
   873                 fetch.simps nth_of.simps, auto)
   862 done
   874 done
   863 
   875 
   864 lemma [simp]: "fetch t_wcode_main 4 Oc = (R, 5)"
   876 lemma [simp]: "fetch t_wcode_main 4 Oc = (R, 5)"
   865 apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
   877 apply(subgoal_tac "4 = Suc 3")
   866                 fetch.simps nth_of.simps)
   878 apply(simp only: t_wcode_main_def t_wcode_main_first_part_def
       
   879                 fetch.simps nth_of.simps, auto)
   867 done
   880 done
   868 
   881 
   869 lemma [simp]: "fetch t_wcode_main 5 Oc = (R, 5)"
   882 lemma [simp]: "fetch t_wcode_main 5 Oc = (R, 5)"
   870 apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
   883 apply(subgoal_tac "5 = Suc 4")
   871                 fetch.simps nth_of.simps)
   884 apply(simp only: t_wcode_main_def t_wcode_main_first_part_def
       
   885                 fetch.simps nth_of.simps, auto)
   872 done
   886 done
   873 
   887 
   874 lemma [simp]: "fetch t_wcode_main 5 Bk = (W1, 6)"
   888 lemma [simp]: "fetch t_wcode_main 5 Bk = (W1, 6)"
   875 apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
   889 apply(subgoal_tac "5 = Suc 4")
   876                 fetch.simps nth_of.simps)
   890 apply(simp only: t_wcode_main_def t_wcode_main_first_part_def
       
   891                 fetch.simps nth_of.simps, auto)
   877 done
   892 done
   878 
   893 
   879 lemma [simp]: "fetch t_wcode_main 6 Bk = (R, 13)"
   894 lemma [simp]: "fetch t_wcode_main 6 Bk = (R, 13)"
   880 apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
   895 apply(subgoal_tac "6 = Suc 5")
   881                 fetch.simps nth_of.simps)
   896 apply(simp only: t_wcode_main_def t_wcode_main_first_part_def
   882 done
   897                 fetch.simps nth_of.simps, auto)
   883 
   898 done
       
   899   
   884 lemma [simp]: "fetch t_wcode_main 6 Oc = (L, 6)"
   900 lemma [simp]: "fetch t_wcode_main 6 Oc = (L, 6)"
   885 apply(simp add: t_wcode_main_def t_wcode_main_first_part_def
   901 apply(subgoal_tac "6 = Suc 5")
   886                 fetch.simps nth_of.simps)
   902 apply(simp only: t_wcode_main_def t_wcode_main_first_part_def
   887 done
   903                 fetch.simps nth_of.simps, auto)
   888 lemma [elim]: "Bk\<^bsup>mr\<^esup> = [] \<Longrightarrow> mr = 0"
   904 done
   889 apply(case_tac mr, auto simp: exponent_def)
   905 
       
   906 lemma [elim]: "Bk\<up>(mr) = [] \<Longrightarrow> mr = 0"
       
   907 apply(case_tac mr, auto)
   890 done
   908 done
   891 
   909 
   892 lemma [simp]: "wcode_on_left_moving_1 ires rs (b, []) = False"
   910 lemma [simp]: "wcode_on_left_moving_1 ires rs (b, []) = False"
   893 apply(simp add: wcode_on_left_moving_1.simps wcode_on_left_moving_1_B.simps
   911 apply(simp add: wcode_on_left_moving_1.simps wcode_on_left_moving_1_B.simps
   894                 wcode_on_left_moving_1_O.simps, auto)
   912                 wcode_on_left_moving_1_O.simps) 
   895 done
   913 done                                           
   896 
   914 
   897 
   915 
   898 declare wcode_on_checking_1.simps[simp del]
   916 declare wcode_on_checking_1.simps[simp del]
   899 
   917 
   900 lemmas wcode_double_case_inv_simps = 
   918 lemmas wcode_double_case_inv_simps = 
   919 apply(case_tac ml, simp)
   937 apply(case_tac ml, simp)
   920 apply(rule_tac x = "mr - Suc (Suc 0)" in exI, rule_tac x = rn in exI)
   938 apply(rule_tac x = "mr - Suc (Suc 0)" in exI, rule_tac x = rn in exI)
   921 apply(case_tac mr, simp, case_tac nat, simp, simp add: exp_ind)
   939 apply(case_tac mr, simp, case_tac nat, simp, simp add: exp_ind)
   922 apply(rule_tac disjI1)
   940 apply(rule_tac disjI1)
   923 apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI, 
   941 apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI, 
   924       simp add: exp_ind_def)
   942       simp, simp add: replicate_Suc)
   925 apply(erule_tac exE)+
   943 apply(erule_tac exE)+
   926 apply(simp)
   944 apply(simp)
   927 done
   945 done
   928 
   946 
       
   947 declare replicate_Suc[simp]
   929 
   948 
   930 lemma [elim]: 
   949 lemma [elim]: 
   931   "\<lbrakk>wcode_on_left_moving_1 ires rs (b, Oc # list); tl b = aa \<and> hd b # Oc # list = ba\<rbrakk> 
   950   "\<lbrakk>wcode_on_left_moving_1 ires rs (b, Oc # list); tl b = aa \<and> hd b # Oc # list = ba\<rbrakk> 
   932     \<Longrightarrow> wcode_on_checking_1 ires rs (aa, ba)"
   951     \<Longrightarrow> wcode_on_checking_1 ires rs (aa, ba)"
   933 apply(simp only: wcode_double_case_inv_simps)
   952 apply(simp only: wcode_double_case_inv_simps)
   934 apply(erule_tac disjE)
   953 apply(erule_tac disjE)
   935 apply(erule_tac [!] exE)+
   954 apply(erule_tac [!] exE)+
   936 apply(case_tac mr, simp, simp add: exp_ind_def)
   955 apply(case_tac mr, simp, auto)
   937 apply(rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
   956 done
   938 done
       
   939 
       
   940 
   957 
   941 lemma [simp]: "wcode_on_checking_1 ires rs (b, []) = False" 
   958 lemma [simp]: "wcode_on_checking_1 ires rs (b, []) = False" 
   942 apply(auto simp: wcode_double_case_inv_simps)
   959 apply(auto simp: wcode_double_case_inv_simps)
   943 done         
   960 done         
   944  
   961  
   965 lemma [simp]: "wcode_erase1 ires rs (b, []) = False"
   982 lemma [simp]: "wcode_erase1 ires rs (b, []) = False"
   966 apply(simp add: wcode_double_case_inv_simps)
   983 apply(simp add: wcode_double_case_inv_simps)
   967 done
   984 done
   968 
   985 
   969 lemma [simp]: "wcode_on_right_moving_1 ires rs (b, []) = False"
   986 lemma [simp]: "wcode_on_right_moving_1 ires rs (b, []) = False"
   970 apply(simp add: wcode_double_case_inv_simps exp_ind_def)
   987 apply(simp add: wcode_double_case_inv_simps)
   971 done
   988 done
   972 
   989 
   973 lemma [simp]: "wcode_on_right_moving_1 ires rs (b, []) = False"
   990 lemma [simp]: "wcode_on_right_moving_1 ires rs (b, []) = False"
   974 apply(simp add: wcode_double_case_inv_simps exp_ind_def)
   991 apply(simp add: wcode_double_case_inv_simps)
   975 done
   992 done
   976 
   993 
   977 lemma [elim]: "\<lbrakk>wcode_on_right_moving_1 ires rs (b, Bk # ba);  Bk # b = aa \<and> list = b\<rbrakk> \<Longrightarrow> 
   994 lemma [elim]: "\<lbrakk>wcode_on_right_moving_1 ires rs (b, Bk # ba);  Bk # b = aa \<and> list = b\<rbrakk> \<Longrightarrow> 
   978   wcode_on_right_moving_1 ires rs (aa, ba)"
   995   wcode_on_right_moving_1 ires rs (aa, ba)"
   979 apply(simp only: wcode_double_case_inv_simps)
   996 apply(simp only: wcode_double_case_inv_simps)
   980 apply(erule_tac exE)+
   997 apply(erule_tac exE)+
   981 apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - Suc 0" in exI,
   998 apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - Suc 0" in exI,
   982       rule_tac x = rn in exI)
   999       rule_tac x = rn in exI)
   983 apply(simp add: exp_ind_def)
  1000 apply(simp)
   984 apply(case_tac mr, simp, simp add: exp_ind_def)
  1001 apply(case_tac mr, simp, simp)
   985 done
  1002 done
   986 
  1003 
   987 lemma [elim]: 
  1004 lemma [elim]: 
   988   "\<lbrakk>wcode_on_right_moving_1 ires rs (b, Oc # ba); Oc # b = aa \<and> list = ba\<rbrakk> 
  1005   "\<lbrakk>wcode_on_right_moving_1 ires rs (b, Oc # ba); Oc # b = aa \<and> list = ba\<rbrakk> 
   989   \<Longrightarrow> wcode_goon_right_moving_1 ires rs (aa, ba)"
  1006   \<Longrightarrow> wcode_goon_right_moving_1 ires rs (aa, ba)"
   990 apply(simp only: wcode_double_case_inv_simps)
  1007 apply(simp only: wcode_double_case_inv_simps)
   991 apply(erule_tac exE)+
  1008 apply(erule_tac exE)+
   992 apply(rule_tac x = "Suc 0" in exI, rule_tac x = "rs" in exI,
  1009 apply(rule_tac x = "Suc 0" in exI, rule_tac x = "rs" in exI,
   993       rule_tac x = "ml - Suc (Suc 0)" in exI, rule_tac x = rn in exI)
  1010       rule_tac x = "ml - Suc (Suc 0)" in exI, rule_tac x = rn in exI)
   994 apply(case_tac mr, simp_all add: exp_ind_def)
  1011 apply(case_tac mr, simp_all)
   995 apply(case_tac ml, simp, case_tac nat, simp, simp)
  1012 apply(case_tac ml, simp, case_tac nat, simp, simp)
   996 apply(simp add: exp_ind_def)
       
   997 done
  1013 done
   998 
  1014 
   999 lemma [simp]: 
  1015 lemma [simp]: 
  1000   "wcode_on_right_moving_1 ires rs (b, []) \<Longrightarrow> False"
  1016   "wcode_on_right_moving_1 ires rs (b, []) \<Longrightarrow> False"
  1001 apply(simp add: wcode_double_case_inv_simps exponent_def)
  1017 apply(simp add: wcode_double_case_inv_simps)
  1002 done
  1018 done
  1003 
  1019 
  1004 lemma [elim]: "\<lbrakk>wcode_erase1 ires rs (b, Bk # ba); Bk # b = aa \<and> list = ba; c = Bk # ba\<rbrakk> 
  1020 lemma [elim]: "\<lbrakk>wcode_erase1 ires rs (b, Bk # ba); Bk # b = aa \<and> list = ba; c = Bk # ba\<rbrakk> 
  1005   \<Longrightarrow> wcode_on_right_moving_1 ires rs (aa, ba)"
  1021   \<Longrightarrow> wcode_on_right_moving_1 ires rs (aa, ba)"
  1006 apply(simp only: wcode_double_case_inv_simps)
  1022 apply(simp only: wcode_double_case_inv_simps)
  1007 apply(erule_tac exE)+
  1023 apply(erule_tac exE)+
  1008 apply(rule_tac x = "Suc 0" in exI, rule_tac x = "Suc (Suc ln)" in exI, 
  1024 apply(rule_tac x = "Suc 0" in exI, rule_tac x = "Suc (Suc ln)" in exI, 
  1009       rule_tac x = rn in exI, simp add: exp_ind)
  1025       rule_tac x = rn in exI, simp add: exp_ind del: replicate_Suc)
  1010 done
  1026 done
  1011 
  1027 
  1012 lemma [elim]: "\<lbrakk>wcode_erase1 ires rs (aa, Oc # list);  b = aa \<and> Bk # list = ba\<rbrakk> \<Longrightarrow> 
  1028 lemma [elim]: "\<lbrakk>wcode_erase1 ires rs (aa, Oc # list);  b = aa \<and> Bk # list = ba\<rbrakk> \<Longrightarrow> 
  1013   wcode_erase1 ires rs (aa, ba)"
  1029   wcode_erase1 ires rs (aa, ba)"
  1014 apply(simp only: wcode_double_case_inv_simps)
  1030 apply(simp only: wcode_double_case_inv_simps)
  1022 apply(erule_tac exE)+
  1038 apply(erule_tac exE)+
  1023 apply(rule_tac disjI2)
  1039 apply(rule_tac disjI2)
  1024 apply(simp only:wcode_backto_standard_pos_O.simps)
  1040 apply(simp only:wcode_backto_standard_pos_O.simps)
  1025 apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, rule_tac x = ln in exI,
  1041 apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, rule_tac x = ln in exI,
  1026       rule_tac x = rn in exI, simp)
  1042       rule_tac x = rn in exI, simp)
  1027 apply(case_tac mr, simp_all add: exponent_def)
       
  1028 done
  1043 done
  1029 
  1044 
  1030 lemma [elim]: 
  1045 lemma [elim]: 
  1031   "\<lbrakk>wcode_goon_right_moving_1 ires rs (aa, Bk # list);  b = aa \<and> Oc # list = ba\<rbrakk>
  1046   "\<lbrakk>wcode_goon_right_moving_1 ires rs (aa, Bk # list);  b = aa \<and> Oc # list = ba\<rbrakk>
  1032   \<Longrightarrow> wcode_backto_standard_pos ires rs (aa, ba)"
  1047   \<Longrightarrow> wcode_backto_standard_pos ires rs (aa, ba)"
  1034 apply(erule_tac exE)+
  1049 apply(erule_tac exE)+
  1035 apply(rule_tac disjI2)
  1050 apply(rule_tac disjI2)
  1036 apply(simp only:wcode_backto_standard_pos_O.simps)
  1051 apply(simp only:wcode_backto_standard_pos_O.simps)
  1037 apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, rule_tac x = ln in exI,
  1052 apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, rule_tac x = ln in exI,
  1038       rule_tac x = "rn - Suc 0" in exI, simp)
  1053       rule_tac x = "rn - Suc 0" in exI, simp)
  1039 apply(case_tac mr, simp, case_tac rn, simp, simp_all add: exp_ind_def)
  1054 apply(case_tac mr, simp, case_tac rn, simp, simp_all)
  1040 done
  1055 done
  1041 
  1056 
  1042 lemma [elim]: "\<lbrakk>wcode_goon_right_moving_1 ires rs (b, Oc # ba);  Oc # b = aa \<and> list = ba\<rbrakk> 
  1057 lemma [elim]: "\<lbrakk>wcode_goon_right_moving_1 ires rs (b, Oc # ba);  Oc # b = aa \<and> list = ba\<rbrakk> 
  1043   \<Longrightarrow> wcode_goon_right_moving_1 ires rs (aa, ba)"
  1058   \<Longrightarrow> wcode_goon_right_moving_1 ires rs (aa, ba)"
  1044 apply(simp only: wcode_double_case_inv_simps)
  1059 apply(simp only: wcode_double_case_inv_simps)
  1045 apply(erule_tac exE)+
  1060 apply(erule_tac exE)+
  1046 apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - Suc 0" in exI, 
  1061 apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - Suc 0" in exI, 
  1047       rule_tac x = ln in exI, rule_tac x = rn in exI)
  1062       rule_tac x = ln in exI, rule_tac x = rn in exI)
  1048 apply(simp add: exp_ind_def)
  1063 apply(simp)
  1049 apply(case_tac mr, simp, case_tac rn, simp_all add: exp_ind_def)
  1064 apply(case_tac mr, simp, case_tac rn, simp_all)
  1050 done
  1065 done
  1051 
  1066 
  1052 lemma [elim]: "\<lbrakk>wcode_backto_standard_pos ires rs (b, []);  Bk # b = aa\<rbrakk> \<Longrightarrow> False"
  1067 lemma [elim]: "\<lbrakk>wcode_backto_standard_pos ires rs (b, []);  Bk # b = aa\<rbrakk> \<Longrightarrow> False"
  1053 apply(auto simp: wcode_double_case_inv_simps wcode_backto_standard_pos_O.simps
  1068 apply(auto simp: wcode_double_case_inv_simps wcode_backto_standard_pos_O.simps
  1054                  wcode_backto_standard_pos_B.simps)
  1069                  wcode_backto_standard_pos_B.simps)
  1055 apply(case_tac mr, simp_all add: exp_ind_def)
       
  1056 done
  1070 done
  1057 
  1071 
  1058 lemma [elim]: "\<lbrakk>wcode_backto_standard_pos ires rs (b, Bk # ba); Bk # b = aa \<and> list = ba\<rbrakk> 
  1072 lemma [elim]: "\<lbrakk>wcode_backto_standard_pos ires rs (b, Bk # ba); Bk # b = aa \<and> list = ba\<rbrakk> 
  1059   \<Longrightarrow> wcode_before_double ires rs (aa, ba)"
  1073   \<Longrightarrow> wcode_before_double ires rs (aa, ba)"
  1060 apply(simp only: wcode_double_case_inv_simps wcode_backto_standard_pos_B.simps
  1074 apply(simp only: wcode_double_case_inv_simps wcode_backto_standard_pos_B.simps
  1061                  wcode_backto_standard_pos_O.simps wcode_before_double.simps)
  1075                  wcode_backto_standard_pos_O.simps wcode_before_double.simps)
  1062 apply(erule_tac disjE)
  1076 apply(erule_tac disjE)
  1063 apply(erule_tac exE)+
  1077 apply(erule_tac exE)+
  1064 apply(rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
  1078 apply(rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
  1065 apply(auto)
  1079 apply(auto)
  1066 apply(case_tac [!] mr, simp_all add: exp_ind_def)
  1080 apply(case_tac [!] mr, simp_all)
  1067 done
  1081 done
  1068 
  1082 
  1069 lemma [simp]: "wcode_backto_standard_pos ires rs ([], Oc # list) = False"
  1083 lemma [simp]: "wcode_backto_standard_pos ires rs ([], Oc # list) = False"
  1070 apply(auto simp: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps
  1084 apply(auto simp: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps
  1071                  wcode_backto_standard_pos_O.simps)
  1085                  wcode_backto_standard_pos_O.simps)
  1072 done
  1086 done
  1073 
  1087 
  1074 lemma [simp]: "wcode_backto_standard_pos ires rs (b, []) = False"
  1088 lemma [simp]: "wcode_backto_standard_pos ires rs (b, []) = False"
  1075 apply(auto simp: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps
  1089 apply(auto simp: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps
  1076                  wcode_backto_standard_pos_O.simps)
  1090                  wcode_backto_standard_pos_O.simps)
  1077 apply(case_tac mr, simp, simp add: exp_ind_def)
       
  1078 done
  1091 done
  1079 
  1092 
  1080 lemma [elim]: "\<lbrakk>wcode_backto_standard_pos ires rs (b, Oc # list); tl b = aa; hd b # Oc # list =  ba\<rbrakk>
  1093 lemma [elim]: "\<lbrakk>wcode_backto_standard_pos ires rs (b, Oc # list); tl b = aa; hd b # Oc # list =  ba\<rbrakk>
  1081        \<Longrightarrow> wcode_backto_standard_pos ires rs (aa, ba)"
  1094        \<Longrightarrow> wcode_backto_standard_pos ires rs (aa, ba)"
  1082 apply(simp only:  wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps
  1095 apply(simp only:  wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps
  1088 apply(rule_tac disjI1, rule_tac conjI)
  1101 apply(rule_tac disjI1, rule_tac conjI)
  1089 apply(rule_tac x = ln  in exI, simp, rule_tac x = rn in exI, simp)
  1102 apply(rule_tac x = ln  in exI, simp, rule_tac x = rn in exI, simp)
  1090 apply(rule_tac disjI2)
  1103 apply(rule_tac disjI2)
  1091 apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = ln in exI, 
  1104 apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = ln in exI, 
  1092       rule_tac x = rn in exI, simp)
  1105       rule_tac x = rn in exI, simp)
  1093 apply(simp add: exp_ind_def)
  1106 done
  1094 done
  1107 
  1095 
  1108 declare nth_of.simps[simp del] fetch.simps[simp del]
  1096 declare new_tape.simps[simp del] nth_of.simps[simp del] fetch.simps[simp del]
       
  1097 lemma wcode_double_case_first_correctness:
  1109 lemma wcode_double_case_first_correctness:
  1098   "let P = (\<lambda> (st, l, r). st = 13) in 
  1110   "let P = (\<lambda> (st, l, r). st = 13) in 
  1099        let Q = (\<lambda> (st, l, r). wcode_double_case_inv st ires rs (l, r)) in 
  1111        let Q = (\<lambda> (st, l, r). wcode_double_case_inv st ires rs (l, r)) in 
  1100        let f = (\<lambda> stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp) in
  1112        let f = (\<lambda> stp. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp) in
  1101        \<exists> n .P (f n) \<and> Q (f (n::nat))"
  1113        \<exists> n .P (f n) \<and> Q (f (n::nat))"
  1102 proof -
  1114 proof -
  1103   let ?P = "(\<lambda> (st, l, r). st = 13)"
  1115   let ?P = "(\<lambda> (st, l, r). st = 13)"
  1104   let ?Q = "(\<lambda> (st, l, r). wcode_double_case_inv st ires rs (l, r))"
  1116   let ?Q = "(\<lambda> (st, l, r). wcode_double_case_inv st ires rs (l, r))"
  1105   let ?f = "(\<lambda> stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp)"
  1117   let ?f = "(\<lambda> stp. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp)"
  1106   have "\<exists> n. ?P (?f n) \<and> ?Q (?f (n::nat))"
  1118   have "\<exists> n. ?P (?f n) \<and> ?Q (?f (n::nat))"
  1107   proof(rule_tac halt_lemma2)
  1119   proof(rule_tac halt_lemma2)
  1108     show "wf wcode_double_case_le"
  1120     show "wf wcode_double_case_le"
  1109       by auto
  1121       by auto
  1110   next
  1122   next
  1111     show "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow>
  1123     show "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow>
  1112                    ?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_double_case_le"
  1124                    ?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_double_case_le"
  1113     proof(rule_tac allI, case_tac "?f na", simp add: tstep_red)
  1125     proof(rule_tac allI, case_tac "?f na", simp add: step_red)
  1114       fix na a b c
  1126       fix na a b c
  1115       show "a \<noteq> 13 \<and> wcode_double_case_inv a ires rs (b, c) \<longrightarrow>
  1127       show "a \<noteq> 13 \<and> wcode_double_case_inv a ires rs (b, c) \<longrightarrow>
  1116                (case tstep (a, b, c) t_wcode_main of (st, x) \<Rightarrow> 
  1128                (case step0 (a, b, c) t_wcode_main of (st, x) \<Rightarrow> 
  1117                    wcode_double_case_inv st ires rs x) \<and> 
  1129                    wcode_double_case_inv st ires rs x) \<and> 
  1118                 (tstep (a, b, c) t_wcode_main, a, b, c) \<in> wcode_double_case_le"
  1130                 (step0 (a, b, c) t_wcode_main, a, b, c) \<in> wcode_double_case_le"
  1119         apply(rule_tac impI, simp add: wcode_double_case_inv.simps)
  1131         apply(rule_tac impI, simp add: wcode_double_case_inv.simps)
  1120         apply(auto split: if_splits simp: tstep.simps, 
  1132         apply(auto split: if_splits simp: step.simps, 
  1121               case_tac [!] c, simp_all, case_tac [!] "(c::block list)!0")
  1133               case_tac [!] c, simp_all, case_tac [!] "(c::cell list)!0")
  1122         apply(simp_all add: new_tape.simps wcode_double_case_inv.simps wcode_double_case_le_def
  1134         apply(simp_all add: wcode_double_case_inv.simps wcode_double_case_le_def
  1123                                         lex_pair_def)
  1135                                         lex_pair_def)
  1124         apply(auto split: if_splits)
  1136         apply(auto split: if_splits)
  1125         done
  1137         done
  1126     qed
  1138     qed
  1127   next
  1139   next
  1128     show "?Q (?f 0)"
  1140     show "?Q (?f 0)"
  1129       apply(simp add: steps.simps wcode_double_case_inv.simps 
  1141       apply(simp add: steps.simps wcode_double_case_inv.simps 
  1130                                   wcode_on_left_moving_1.simps
  1142                                   wcode_on_left_moving_1.simps
  1131                                   wcode_on_left_moving_1_B.simps)
  1143                                   wcode_on_left_moving_1_B.simps)
  1132       apply(rule_tac disjI1)
  1144       apply(rule_tac disjI1)
  1133       apply(rule_tac x = "Suc m" in exI, simp add: exp_ind_def)
  1145       apply(rule_tac x = "Suc m" in exI, simp)
  1134       apply(rule_tac x = "Suc 0" in exI, simp add: exp_ind_def)
  1146       apply(rule_tac x = "Suc 0" in exI, simp)
  1135       apply(auto)
       
  1136       done
  1147       done
  1137   next
  1148   next
  1138     show "\<not> ?P (?f 0)"
  1149     show "\<not> ?P (?f 0)"
  1139       apply(simp add: steps.simps)
  1150       apply(simp add: steps.simps)
  1140       done
  1151       done
  1141   qed
  1152   qed
  1142   thus "let P = \<lambda>(st, l, r). st = 13;
  1153   thus "let P = \<lambda>(st, l, r). st = 13;
  1143     Q = \<lambda>(st, l, r). wcode_double_case_inv st ires rs (l, r);
  1154     Q = \<lambda>(st, l, r). wcode_double_case_inv st ires rs (l, r);
  1144     f = steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main
  1155     f = steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main
  1145     in \<exists>n. P (f n) \<and> Q (f n)"
  1156     in \<exists>n. P (f n) \<and> Q (f n)"
  1146     apply(simp add: Let_def)
  1157     apply(simp add: Let_def)
  1147     done
  1158     done
  1148 qed
  1159 qed
  1149     
  1160     
  1150 lemma [elim]: "t_ncorrect tp
  1161 lemma tm_append_shift_append_steps: 
  1151     \<Longrightarrow> t_ncorrect (tshift tp a)"
  1162 "\<lbrakk>steps0 (st, l, r) tp stp = (st', l', r'); 
  1152 apply(simp add: t_ncorrect.simps shift_length)
  1163   0 < st';
  1153 done
  1164   length tp1 mod 2 = 0
  1154 
  1165   \<rbrakk>
  1155 lemma tshift_fetch: "\<lbrakk> fetch tp a b = (aa, st'); 0 < st'\<rbrakk>
  1166   \<Longrightarrow> steps0 (st + length tp1 div 2, l, r) (tp1 @ shift tp (length tp1 div 2) @ tp2) stp 
  1156        \<Longrightarrow> fetch (tshift tp (length tp1 div 2)) a b 
  1167   = (st' + length tp1 div 2, l', r')"
  1157           = (aa, st' + length tp1 div 2)"
       
  1158 apply(subgoal_tac "a > 0")
       
  1159 apply(auto simp: fetch.simps nth_of.simps shift_length nth_map
       
  1160                  tshift.simps split: block.splits if_splits)
       
  1161 done
       
  1162 
       
  1163 lemma t_steps_steps_eq: "\<lbrakk>steps (st, l, r) tp stp = (st', l', r');
       
  1164          0 < st';  
       
  1165          0 < st \<and> st \<le> length tp div 2; 
       
  1166          t_ncorrect tp1;
       
  1167           t_ncorrect tp\<rbrakk>
       
  1168     \<Longrightarrow> t_steps (st + length tp1 div 2, l, r) (tshift tp (length tp1 div 2), 
       
  1169                                                       length tp1 div 2) stp
       
  1170        = (st' + length tp1 div 2, l', r')"
       
  1171 apply(induct stp arbitrary: st' l' r', simp add: steps.simps t_steps.simps,
       
  1172       simp add: tstep_red stepn)
       
  1173 apply(case_tac "(steps (st, l, r) tp stp)", simp)
       
  1174 proof -
  1168 proof -
  1175   fix stp st' l' r' a b c
  1169   assume h: 
  1176   assume ind: "\<And>st' l' r'.
  1170     "steps0 (st, l, r) tp stp = (st', l', r')"
  1177     \<lbrakk>a = st' \<and> b = l' \<and> c = r'; 0 < st'\<rbrakk>
  1171     "0 < st'"
  1178     \<Longrightarrow> t_steps (st + length tp1 div 2, l, r) 
  1172     "length tp1 mod 2 = 0 "
  1179     (tshift tp (length tp1 div 2), length tp1 div 2) stp = 
  1173   from h have 
  1180      (st' + length tp1 div 2, l', r')"
  1174     "steps (st + length tp1 div 2, l, r) (tp1 @ shift tp (length tp1 div 2), 0) stp = 
  1181   and h: "tstep (a, b, c) tp = (st', l', r')" "0 < st'" "t_ncorrect tp1"  "t_ncorrect tp"
  1175                             (st' + length tp1 div 2, l', r')"
  1182   have k: "t_steps (st + length tp1 div 2, l, r) (tshift tp (length tp1 div 2),
  1176     by(rule_tac tm_append_second_steps_eq, simp_all)
  1183          length tp1 div 2) stp = (a + length tp1 div 2, b, c)"
  1177   then have "steps (st + length tp1 div 2, l, r) ((tp1 @ shift tp (length tp1 div 2)) @ tp2, 0) stp = 
  1184     apply(rule_tac ind, simp)
  1178                             (st' + length tp1 div 2, l', r')"
  1185     using h
  1179     using h
  1186     apply(case_tac a, simp_all add: tstep.simps fetch.simps)
  1180     apply(rule_tac tm_append_first_steps_eq, simp_all)
  1187     done
  1181     done
  1188   from h and this show "t_step (t_steps (st + length tp1 div 2, l, r) (tshift tp (length tp1 div 2), length tp1 div 2) stp)
  1182   thus "?thesis"
  1189            (tshift tp (length tp1 div 2), length tp1 div 2) =
  1183     by simp
  1190           (st' + length tp1 div 2, l', r')"
       
  1191     apply(simp add: k)
       
  1192     apply(simp add: tstep.simps t_step.simps)
       
  1193     apply(case_tac "fetch tp a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp)
       
  1194     apply(subgoal_tac "fetch (tshift tp (length tp1 div 2)) a
       
  1195                        (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) = (aa, st' + length tp1 div 2)", simp)
       
  1196     apply(simp add: tshift_fetch)
       
  1197     done
       
  1198 qed 
  1184 qed 
  1199 
  1185 
  1200 lemma t_tshift_lemma: "\<lbrakk> steps (st, l, r) tp stp = (st', l', r'); 
       
  1201                          st' \<noteq> 0; 
       
  1202                          stp > 0;
       
  1203                          0 < st \<and> st \<le> length tp div 2;
       
  1204                          t_ncorrect tp1;
       
  1205                          t_ncorrect tp;
       
  1206                          t_ncorrect tp2
       
  1207                          \<rbrakk>
       
  1208          \<Longrightarrow> \<exists> stp>0. steps (st + length tp1 div 2, l, r) (tp1 @ tshift tp (length tp1 div 2) @ tp2) stp 
       
  1209                   = (st' + length tp1 div 2, l', r')"
       
  1210 proof -
       
  1211   assume h: "steps (st, l, r) tp stp = (st', l', r')"
       
  1212     "st' \<noteq> 0" "stp > 0"
       
  1213     "0 < st \<and> st \<le> length tp div 2"
       
  1214     "t_ncorrect tp1"
       
  1215     "t_ncorrect tp"
       
  1216     "t_ncorrect tp2"
       
  1217   from h have 
       
  1218     "\<exists>stp>0. t_steps (st + length tp1 div 2, l, r) (tp1 @ tshift tp (length tp1 div 2) @ tp2, 0) stp = 
       
  1219                             (st' + length tp1 div 2, l', r')"
       
  1220     apply(rule_tac stp = stp in turing_shift, simp_all add: shift_length)
       
  1221     apply(simp add: t_steps_steps_eq)
       
  1222     apply(simp add: t_ncorrect.simps shift_length)
       
  1223     done
       
  1224   thus "\<exists> stp>0. steps (st + length tp1 div 2, l, r) (tp1 @ tshift tp (length tp1 div 2) @ tp2) stp 
       
  1225                   = (st' + length tp1 div 2, l', r')"
       
  1226     apply(erule_tac exE)
       
  1227     apply(rule_tac x = stp in exI, simp)
       
  1228     apply(subgoal_tac "length (tp1 @ tshift tp (length tp1 div 2) @ tp2) mod 2 = 0")
       
  1229     apply(simp only: steps_eq)
       
  1230     using h
       
  1231     apply(auto simp: t_ncorrect.simps shift_length)
       
  1232     apply arith
       
  1233     done
       
  1234 qed  
       
  1235   
       
  1236 
       
  1237 lemma t_twice_len_ge: "Suc 0 \<le> length t_twice div 2"
  1186 lemma t_twice_len_ge: "Suc 0 \<le> length t_twice div 2"
  1238 apply(simp add: t_twice_def tMp.simps shift_length)
  1187 apply(simp add: t_twice_def mopup.simps t_twice_compile_def)
  1239 done
  1188 done
  1240 
  1189 
  1241 lemma [intro]: "rec_calc_rel (recf.id (Suc 0) 0) [rs] rs"
  1190 lemma [intro]: "rec_calc_rel (recf.id (Suc 0) 0) [rs] rs"
  1242   apply(rule_tac calc_id, simp_all)
  1191   apply(rule_tac calc_id, simp_all)
  1243   done
  1192   done
  1249 
  1198 
  1250 lemma  [intro]: "rec_calc_rel rec_mult [rs, 2] (2 * rs)"
  1199 lemma  [intro]: "rec_calc_rel rec_mult [rs, 2] (2 * rs)"
  1251 using prime_rel_exec_eq[of "rec_mult" "[rs, 2]"  "2*rs"]
  1200 using prime_rel_exec_eq[of "rec_mult" "[rs, 2]"  "2*rs"]
  1252 apply(subgoal_tac "primerec rec_mult (Suc (Suc 0))", auto)
  1201 apply(subgoal_tac "primerec rec_mult (Suc (Suc 0))", auto)
  1253 done
  1202 done
  1254 lemma t_twice_correct: "\<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) 
  1203 
  1255             (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0)) stp =
  1204 declare start_of.simps[simp del]
  1256        (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1205 
       
  1206 lemma t_twice_correct: 
       
  1207   "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) 
       
  1208   (tm_of abc_twice @ shift (mopup (Suc 0)) ((length (tm_of abc_twice) div 2))) stp =
       
  1209   (0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))"
  1257 proof(case_tac "rec_ci rec_twice")
  1210 proof(case_tac "rec_ci rec_twice")
  1258   fix a b c
  1211   fix a b c
  1259   assume h: "rec_ci rec_twice = (a, b, c)"
  1212   assume h: "rec_ci rec_twice = (a, b, c)"
  1260   have "\<exists>stp m l. steps (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\<^bsup>n\<^esup>) (tm_of abc_twice @ tMp (Suc 0) 
  1213   have "\<exists>stp m l. steps0 (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\<up>(n)) (tm_of abc_twice @ shift (mopup 1) 
  1261     (start_of twice_ly (length abc_twice) - 1)) stp = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2*rs)\<^esup> @ Bk\<^bsup>l\<^esup>)"
  1214     (length (tm_of abc_twice) div 2)) stp = (0, Bk\<up>(m) @ Bk # Bk # ires, Oc\<up>(Suc (2*rs)) @ Bk\<up>(l))"
  1262   proof(rule_tac t_compiled_by_rec)
  1215   proof(rule_tac recursive_compile_to_tm_correct)
  1263     show "rec_ci rec_twice = (a, b, c)" by (simp add: h)
  1216     show "rec_ci rec_twice = (a, b, c)" by (simp add: h)
  1264   next
  1217   next
  1265     show "rec_calc_rel rec_twice [rs] (2 * rs)"
  1218     show "rec_calc_rel rec_twice [rs] (2 * rs)"
  1266       apply(simp add: rec_twice_def)
  1219       apply(simp add: rec_twice_def)
  1267       apply(rule_tac rs =  "[rs, 2]" in calc_cn, simp_all)
  1220       apply(rule_tac rs =  "[rs, 2]" in calc_cn, simp_all)
  1268       apply(rule_tac allI, case_tac k, auto)
  1221       apply(rule_tac allI, case_tac k, auto)
  1269       done
  1222       done
  1270   next
  1223   next
  1271     show "length [rs] = Suc 0" by simp
  1224     show "length [rs] = 1" by simp
       
  1225   next	
       
  1226     show "layout_of (a [+] dummy_abc 1) = layout_of (a [+] dummy_abc 1)" by simp
  1272   next
  1227   next
  1273     show "layout_of (a [+] dummy_abc (Suc 0)) = layout_of (a [+] dummy_abc (Suc 0))"
  1228     show "tm_of abc_twice = tm_of (a [+] dummy_abc 1)"
  1274       by simp
       
  1275   next
       
  1276     show "start_of twice_ly (length abc_twice) = 
       
  1277       start_of (layout_of (a [+] dummy_abc (Suc 0))) (length (a [+] dummy_abc (Suc 0)))"
       
  1278       using h
       
  1279       apply(simp add: twice_ly_def abc_twice_def)
       
  1280       done
       
  1281   next
       
  1282     show "tm_of abc_twice = tm_of (a [+] dummy_abc (Suc 0))"
       
  1283       using h
  1229       using h
  1284       apply(simp add: abc_twice_def)
  1230       apply(simp add: abc_twice_def)
  1285       done
  1231       done
  1286   qed
  1232   qed
  1287   thus "\<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) 
  1233   thus "?thesis"
  1288             (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0)) stp =
       
  1289        (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
       
  1290     apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
  1234     apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
  1291     done
  1235     done
  1292 qed
  1236 qed
  1293 
  1237 
  1294 lemma change_termi_state_fetch: "\<lbrakk>fetch ap a b = (aa, st); st > 0\<rbrakk>
  1238 
  1295        \<Longrightarrow> fetch (change_termi_state ap) a b = (aa, st)"
  1239 lemma adjust_fetch0: 
  1296 apply(case_tac b, auto simp: fetch.simps nth_of.simps change_termi_state.simps nth_map
  1240   "\<lbrakk>0 < a; a \<le> length ap div 2;  fetch ap a b = (aa, 0)\<rbrakk>
  1297                        split: if_splits block.splits)
  1241   \<Longrightarrow> fetch (adjust ap) a b = (aa, Suc (length ap div 2))"
  1298 done
  1242 apply(case_tac b, auto simp: fetch.simps nth_of.simps nth_map
  1299 
  1243                        split: if_splits)
  1300 lemma change_termi_state_exec_in_range:
  1244 apply(case_tac [!] a, auto simp: fetch.simps nth_of.simps)
  1301      "\<lbrakk>steps (st, l, r) ap stp = (st', l', r'); st' \<noteq> 0\<rbrakk>
  1245 done
  1302     \<Longrightarrow> steps (st, l, r) (change_termi_state ap) stp = (st', l', r')"
  1246 
  1303 proof(induct stp arbitrary: st l r st' l' r', simp add: steps.simps)
  1247 lemma adjust_fetch_norm: 
  1304   fix stp st l r st' l' r'
  1248   "\<lbrakk>st > 0;  st \<le> length tp div 2; fetch ap st b = (aa, ns); ns \<noteq> 0\<rbrakk>
  1305   assume ind: "\<And>st l r st' l' r'. 
  1249  \<Longrightarrow>  fetch (turing_basic.adjust ap) st b = (aa, ns)"
  1306     \<lbrakk>steps (st, l, r) ap stp = (st', l', r'); st' \<noteq> 0\<rbrakk> \<Longrightarrow>
  1250  apply(case_tac b, auto simp: fetch.simps nth_of.simps nth_map
  1307     steps (st, l, r) (change_termi_state ap) stp = (st', l', r')"
  1251                        split: if_splits)
  1308   and h: "steps (st, l, r) ap (Suc stp) = (st', l', r')" "st' \<noteq> 0"
  1252 apply(case_tac [!] st, auto simp: fetch.simps nth_of.simps)
  1309   from h show "steps (st, l, r) (change_termi_state ap) (Suc stp) = (st', l', r')"
  1253 done
  1310   proof(simp add: tstep_red, case_tac "steps (st, l, r) ap stp", simp)
  1254 
  1311     fix a b c
  1255 lemma adjust_step_eq: 
  1312     assume g: "steps (st, l, r) ap stp = (a, b, c)"
  1256   assumes exec: "step0 (st,l,r) ap = (st', l', r')"
  1313               "tstep (a, b, c) ap = (st', l', r')" "0 < st'"
  1257   and wf_tm: "tm_wf (ap, 0)"
  1314     hence "steps (st, l, r) (change_termi_state ap) stp = (a, b, c)"
  1258   and notfinal: "st' > 0"
  1315       apply(rule_tac ind, simp)
  1259   shows "step0 (st, l, r) (adjust ap) = (st', l', r')"
  1316       apply(case_tac a, simp_all add: tstep_0)
  1260   using assms
  1317       done
  1261 proof -
  1318     from g and this show "tstep (steps (st, l, r) (change_termi_state ap) stp)
  1262   have "st > 0"
  1319       (change_termi_state ap) = (st', l', r')"
  1263     using assms
  1320       apply(simp add: tstep.simps)
  1264     by(case_tac st, simp_all add: step.simps fetch.simps)
  1321       apply(case_tac "fetch ap a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp)
  1265   moreover hence "st \<le> (length ap) div 2"
  1322       apply(subgoal_tac "fetch (change_termi_state ap) a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)
  1266     using assms
  1323                    = (aa, st')", simp)
  1267     apply(case_tac "st \<le> (length ap) div 2", simp)
  1324       apply(simp add: change_termi_state_fetch)
  1268     apply(case_tac st, auto simp: step.simps fetch.simps)
  1325       done
  1269     apply(case_tac "read r", simp_all add: fetch.simps nth_of.simps)
  1326   qed
  1270     done   
       
  1271   ultimately have "fetch (adjust ap) st (read r) = fetch ap st (read r)"
       
  1272     using assms
       
  1273     apply(case_tac "fetch ap st (read r)")
       
  1274     apply(drule_tac adjust_fetch_norm, simp_all)
       
  1275     apply(simp add: step.simps)
       
  1276     done
       
  1277   thus "?thesis"
       
  1278     using exec
       
  1279     by(simp add: step.simps)
  1327 qed
  1280 qed
  1328 
  1281 
  1329 lemma change_termi_state_fetch0: 
  1282 declare adjust.simps[simp del]
  1330   "\<lbrakk>0 < a; a \<le> length ap div 2; t_correct ap; fetch ap a b = (aa, 0)\<rbrakk>
  1283 
  1331   \<Longrightarrow> fetch (change_termi_state ap) a b = (aa, Suc (length ap div 2))"
  1284 lemma adjust_steps_eq: 
  1332 apply(case_tac b, auto simp: fetch.simps nth_of.simps change_termi_state.simps nth_map
  1285   assumes exec: "steps0 (st,l,r) ap stp = (st', l', r')"
  1333                        split: if_splits block.splits)
  1286   and wf_tm: "tm_wf (ap, 0)"
  1334 done
  1287   and notfinal: "st' > 0"
  1335 
  1288   shows "steps0 (st, l, r) (adjust ap) stp = (st', l', r')"
  1336 lemma turing_change_termi_state: 
  1289   using exec notfinal
  1337   "\<lbrakk>steps (Suc 0, l, r) ap stp = (0, l', r'); t_correct ap\<rbrakk>
  1290 proof(induct stp arbitrary: st' l' r')
  1338      \<Longrightarrow> \<exists> stp. steps (Suc 0, l, r) (change_termi_state ap) stp = 
  1291   case 0
       
  1292   thus "?case"
       
  1293     by(simp add: steps.simps)
       
  1294 next
       
  1295   case (Suc stp st' l' r')
       
  1296   have ind: "\<And>st' l' r'. \<lbrakk>steps0 (st, l, r) ap stp = (st', l', r'); 0 < st'\<rbrakk> 
       
  1297     \<Longrightarrow> steps0 (st, l, r) (turing_basic.adjust ap) stp = (st', l', r')" by fact
       
  1298   have h: "steps0 (st, l, r) ap (Suc stp) = (st', l', r')" by fact
       
  1299   have g:   "0 < st'" by fact
       
  1300   obtain st'' l'' r'' where a: "steps0 (st, l, r) ap stp = (st'', l'', r'')"
       
  1301     by (metis prod_cases3)
       
  1302   hence c:"0 < st''"
       
  1303     using h g
       
  1304     apply(simp add: step_red)
       
  1305     apply(case_tac st'', auto)
       
  1306     done
       
  1307   hence b: "steps0 (st, l, r) (turing_basic.adjust ap) stp = (st'', l'', r'')"
       
  1308     using a
       
  1309     by(rule_tac ind, simp_all)
       
  1310   thus "?case"
       
  1311     using assms a b h g
       
  1312     apply(simp add: step_red) 
       
  1313     apply(rule_tac adjust_step_eq, simp_all)
       
  1314     done
       
  1315 qed 
       
  1316 
       
  1317 lemma adjust_halt_eq:
       
  1318   assumes exec: "steps0 (1, l, r) ap stp = (0, l', r')"
       
  1319   and tm_wf: "tm_wf (ap, 0)" 
       
  1320   shows "\<exists> stp. steps0 (Suc 0, l, r) (adjust ap) stp = 
  1339         (Suc (length ap div 2), l', r')"
  1321         (Suc (length ap div 2), l', r')"
  1340 apply(drule first_halt_point)
  1322 proof -
  1341 apply(erule_tac exE)
  1323   have "\<exists> stp. \<not> is_final (steps0 (1, l, r) ap stp) \<and> (steps0 (1, l, r) ap (Suc stp) = (0, l', r'))"
  1342 apply(rule_tac x = "Suc stp" in exI, simp add: tstep_red)
  1324     thm before_final using exec
  1343 apply(case_tac "steps (Suc 0, l, r) ap stp")
  1325     by(erule_tac before_final)
  1344 apply(simp add: isS0_def change_termi_state_exec_in_range)
  1326   then obtain stpa where a: 
  1345 apply(subgoal_tac "steps (Suc 0, l, r) (change_termi_state ap) stp = (a, b, c)", simp)
  1327     "\<not> is_final (steps0 (1, l, r) ap stpa) \<and> (steps0 (1, l, r) ap (Suc stpa) = (0, l', r'))" ..
  1346 apply(simp add: tstep.simps)
  1328   obtain sa la ra where b:"steps0 (1, l, r) ap stpa = (sa, la, ra)"  by (metis prod_cases3)
  1347 apply(case_tac "fetch ap a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp)
  1329   hence c: "steps0 (Suc 0, l, r) (adjust ap) stpa = (sa, la, ra)"
  1348 apply(subgoal_tac "fetch (change_termi_state ap) a 
  1330     using assms a
  1349   (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) = (aa, Suc (length ap div 2))", simp)
  1331     apply(rule_tac adjust_steps_eq, simp_all)
  1350 apply(rule_tac ap = ap in change_termi_state_fetch0, simp_all)
  1332     done
  1351 apply(rule_tac tp = "(l, r)" and l = b and r = c  and stp = stp and A = ap in s_keep, simp_all)
  1333   have d: "sa \<le> length ap div 2"
  1352 apply(simp add: change_termi_state_exec_in_range)
  1334     using steps_in_range[of  "(l, r)" ap stpa] a tm_wf b
       
  1335     by(simp)
       
  1336   obtain ac ns where e: "fetch ap sa (read ra) = (ac, ns)"
       
  1337     by (metis prod.exhaust)
       
  1338   hence f: "ns = 0"
       
  1339     using b a
       
  1340     apply(simp add: step_red step.simps)
       
  1341     done
       
  1342   have k: "fetch (adjust ap) sa (read ra) = (ac, Suc (length ap div 2))"
       
  1343     using a b c d e f
       
  1344     apply(rule_tac adjust_fetch0, simp_all)
       
  1345     done
       
  1346   from a b e f k and c show "?thesis"
       
  1347     apply(rule_tac x = "Suc stpa" in exI)
       
  1348     apply(simp add: step_red, auto)
       
  1349     apply(simp add: step.simps)
       
  1350     done
       
  1351 qed    
       
  1352    
       
  1353 declare tm_wf.simps[simp del]
       
  1354 
       
  1355 lemma [simp]: " tm_wf (t_twice_compile, 0)"
       
  1356 apply(simp only: t_twice_compile_def)
       
  1357 apply(rule_tac t_compiled_correct)
       
  1358 apply(simp_all add: abc_twice_def)
  1353 done
  1359 done
  1354 
  1360 
  1355 lemma t_twice_change_term_state:
  1361 lemma t_twice_change_term_state:
  1356   "\<exists> stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp
  1362   "\<exists> stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) t_twice stp
  1357      = (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1363      = (Suc t_twice_len, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))"
  1358 using t_twice_correct[of ires rs n]
  1364 proof -
  1359 apply(erule_tac exE)
  1365   have "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) 
  1360 apply(erule_tac exE)
  1366     (tm_of abc_twice @ shift (mopup (Suc 0)) ((length (tm_of abc_twice) div 2))) stp =
  1361 apply(erule_tac exE)
  1367     (0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))"
  1362 proof(drule_tac turing_change_termi_state)
  1368     by(rule_tac t_twice_correct)
  1363   fix stp ln rn
  1369   then obtain stp ln rn where " steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) 
  1364   show "t_correct (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0))"
  1370     (tm_of abc_twice @ shift (mopup (Suc 0)) ((length (tm_of abc_twice) div 2))) stp =
  1365     apply(rule_tac t_compiled_correct, simp_all)
  1371     (0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))" by blast
  1366     apply(simp add: twice_ly_def)
  1372   hence "\<exists> stp. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
       
  1373     (adjust t_twice_compile) stp
       
  1374      = (Suc (length t_twice_compile div 2), Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))"
       
  1375     thm adjust_halt_eq
       
  1376     apply(rule_tac stp = stp in adjust_halt_eq)
       
  1377     apply(simp add: t_twice_compile_def, auto)
  1367     done
  1378     done
  1368 next
  1379   then obtain stpb where 
  1369   fix stp ln rn
  1380     "steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
  1370   show "\<exists>stp. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
  1381     (adjust t_twice_compile) stpb
  1371     (change_termi_state (tm_of abc_twice @ tMp (Suc 0) 
  1382      = (Suc (length t_twice_compile div 2), Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))" ..
  1372     (start_of twice_ly (length abc_twice) - Suc 0))) stp =
  1383   thus "?thesis"
  1373     (Suc (length (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0)) div 2),
  1384     apply(simp add: t_twice_def t_twice_len_def)
  1374     Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>) \<Longrightarrow>
  1385     by metis
  1375     \<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp = 
       
  1376     (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
       
  1377     apply(erule_tac exE)
       
  1378     apply(simp add: t_twice_len_def t_twice_def)
       
  1379     apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
       
  1380     done
       
  1381 qed
  1386 qed
  1382 
  1387 
       
  1388 lemma [intro]: "length t_wcode_main_first_part mod 2 = 0"
       
  1389 apply(auto simp: t_wcode_main_first_part_def)
       
  1390 done
       
  1391 
  1383 lemma t_twice_append_pre:
  1392 lemma t_twice_append_pre:
  1384   "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp
  1393   "steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) t_twice stp
  1385   = (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)
  1394   = (Suc t_twice_len, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))
  1386    \<Longrightarrow> \<exists> stp>0. steps (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
  1395    \<Longrightarrow> steps0 (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
  1387      (t_wcode_main_first_part @ tshift t_twice (length t_wcode_main_first_part div 2) @
  1396      (t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @
  1388       ([(L, 1), (L, 1)] @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp 
  1397       ([(L, 1), (L, 1)] @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp 
  1389     = (Suc (t_twice_len) + length t_wcode_main_first_part div 2, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1398     = (Suc (t_twice_len) + length t_wcode_main_first_part div 2, 
  1390 proof(rule_tac t_tshift_lemma, simp_all add: t_twice_len_ge)
  1399              Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))"
  1391   assume "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp = 
  1400 by(rule_tac tm_append_shift_append_steps, auto)
  1392     (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1401 
  1393   thus "0 < stp"
       
  1394     apply(case_tac stp, simp add: steps.simps t_twice_len_ge t_twice_len_def)
       
  1395     using t_twice_len_ge
       
  1396     apply(simp, simp)
       
  1397     done
       
  1398 next
       
  1399   show "t_ncorrect t_wcode_main_first_part"
       
  1400     apply(simp add: t_ncorrect.simps t_wcode_main_first_part_def)
       
  1401     done
       
  1402 next
       
  1403   show "t_ncorrect t_twice"
       
  1404     using length_tm_even[of abc_twice]
       
  1405     apply(auto simp: t_ncorrect.simps t_twice_def)
       
  1406     apply(arith)
       
  1407     done
       
  1408 next
       
  1409   show "t_ncorrect ((L, Suc 0) # (L, Suc 0) #
       
  1410        tshift t_fourtimes (t_twice_len + 13) @ [(L, Suc 0), (L, Suc 0)])"
       
  1411     using length_tm_even[of abc_fourtimes]
       
  1412     apply(simp add: t_ncorrect.simps shift_length t_fourtimes_def)
       
  1413     apply arith
       
  1414     done
       
  1415 qed
       
  1416   
       
  1417 lemma t_twice_append:
  1402 lemma t_twice_append:
  1418   "\<exists> stp ln rn. steps (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
  1403   "\<exists> stp ln rn. steps0 (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
  1419      (t_wcode_main_first_part @ tshift t_twice (length t_wcode_main_first_part div 2) @
  1404      (t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @
  1420       ([(L, 1), (L, 1)] @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp 
  1405       ([(L, 1), (L, 1)] @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp 
  1421     = (Suc (t_twice_len) + length t_wcode_main_first_part div 2, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1406     = (Suc (t_twice_len) + length t_wcode_main_first_part div 2, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))"
  1422   using t_twice_change_term_state[of ires rs n]
  1407   using t_twice_change_term_state[of ires rs n]
  1423   apply(erule_tac exE)
  1408   apply(erule_tac exE)
  1424   apply(erule_tac exE)
  1409   apply(erule_tac exE)
  1425   apply(erule_tac exE)
  1410   apply(erule_tac exE)
  1426   apply(drule_tac t_twice_append_pre)
  1411   apply(drule_tac t_twice_append_pre)
  1427   apply(erule_tac exE)
  1412   apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI)
  1428   apply(rule_tac x = stpa in exI, rule_tac x = ln in exI, rule_tac x = rn in exI)
       
  1429   apply(simp)
  1413   apply(simp)
  1430   done
  1414   done
  1431   
  1415   
       
  1416 lemma mopup_mod2: "length (mopup k) mod 2  = 0"
       
  1417 apply(auto simp: mopup.simps)
       
  1418 by arith
       
  1419 
  1432 lemma [simp]: "fetch t_wcode_main (Suc (t_twice_len + length t_wcode_main_first_part div 2)) Oc
  1420 lemma [simp]: "fetch t_wcode_main (Suc (t_twice_len + length t_wcode_main_first_part div 2)) Oc
  1433      = (L, Suc 0)"
  1421      = (L, Suc 0)"
  1434 apply(subgoal_tac "length (t_twice) mod 2 = 0")
  1422 apply(subgoal_tac "length (t_twice) mod 2 = 0")
  1435 apply(simp add: t_wcode_main_def nth_append fetch.simps t_wcode_main_first_part_def 
  1423 apply(simp add: t_wcode_main_def nth_append fetch.simps t_wcode_main_first_part_def 
  1436   nth_of.simps shift_length t_twice_len_def, auto)
  1424   nth_of.simps t_twice_len_def, auto)
  1437 apply(simp add: t_twice_def)
  1425 apply(simp add: t_twice_def t_twice_compile_def)
  1438 apply(subgoal_tac "length (tm_of abc_twice) mod 2 = 0")
  1426 using mopup_mod2[of 1]
  1439 apply arith
  1427 apply(simp)
  1440 apply(rule_tac tm_even)
  1428 by arith
  1441 done
       
  1442 
  1429 
  1443 lemma wcode_jump1: 
  1430 lemma wcode_jump1: 
  1444   "\<exists> stp ln rn. steps (Suc (t_twice_len) + length t_wcode_main_first_part div 2,
  1431   "\<exists> stp ln rn. steps0 (Suc (t_twice_len) + length t_wcode_main_first_part div 2,
  1445                        Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>n\<^esup>)
  1432                        Bk\<up>(m) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(n))
  1446      t_wcode_main stp 
  1433      t_wcode_main stp 
  1447     = (Suc 0, Bk\<^bsup>ln\<^esup> @ Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1434     = (Suc 0, Bk\<up>(ln) @ Bk # ires, Bk # Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))"
  1448 apply(rule_tac x = "Suc 0" in exI, rule_tac x = "m" in exI, rule_tac x = n in exI)
  1435 apply(rule_tac x = "Suc 0" in exI, rule_tac x = "m" in exI, rule_tac x = n in exI)
  1449 apply(simp add: steps.simps tstep.simps exp_ind_def new_tape.simps)
  1436 apply(simp add: steps.simps step.simps exp_ind)
  1450 apply(case_tac m, simp, simp add: exp_ind_def)
  1437 apply(case_tac m, simp_all)
  1451 apply(simp add: exp_ind_def[THEN sym] exp_ind[THEN sym])
  1438 apply(simp add: exp_ind[THEN sym])
  1452 done
  1439 done
  1453 
  1440 
  1454 lemma wcode_main_first_part_len:
  1441 lemma wcode_main_first_part_len:
  1455   "length t_wcode_main_first_part = 24"
  1442   "length t_wcode_main_first_part = 24"
  1456   apply(simp add: t_wcode_main_first_part_def)
  1443   apply(simp add: t_wcode_main_first_part_def)
  1457   done
  1444   done
  1458 
  1445 
  1459 lemma wcode_double_case: 
  1446 lemma wcode_double_case: 
  1460   shows "\<exists>stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  1447   shows "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
  1461           (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1448           (Suc 0, Bk # Bk\<up>(ln) @ Oc # ires, Bk # Oc\<up>(Suc (2 * rs + 2)) @ Bk\<up>(rn))"
  1462 proof -
  1449 proof -
  1463   have "\<exists>stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  1450   have "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
  1464           (13,  Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1451           (13,  Bk # Bk # Bk\<up>(ln) @ Oc # ires, Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn))"
  1465     using wcode_double_case_first_correctness[of ires rs m n]
  1452     using wcode_double_case_first_correctness[of ires rs m n]
  1466     apply(simp)
  1453     apply(simp)
  1467     apply(erule_tac exE)
  1454     apply(erule_tac exE)
  1468     apply(case_tac "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, 
  1455     apply(case_tac "steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, 
  1469            Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main na",
  1456            Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main na",
  1470           auto simp: wcode_double_case_inv.simps
  1457           auto simp: wcode_double_case_inv.simps
  1471                      wcode_before_double.simps)
  1458                      wcode_before_double.simps)
  1472     apply(rule_tac x = na in exI, rule_tac x = ln in exI, rule_tac x = rn in exI)
  1459     apply(rule_tac x = na in exI, rule_tac x = ln in exI, rule_tac x = rn in exI)
  1473     apply(simp)
  1460     apply(simp)
  1474     done    
  1461     done    
  1475   from this obtain stpa lna rna where stp1: 
  1462   from this obtain stpa lna rna where stp1: 
  1476     "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa = 
  1463     "steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stpa = 
  1477     (13, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast
  1464     (13, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rna))" by blast
  1478   have "\<exists> stp ln rn. steps (13, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp =
  1465   have "\<exists> stp ln rn. steps0 (13, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rna)) t_wcode_main stp =
  1479     (13 + t_twice_len, Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1466     (13 + t_twice_len, Bk # Bk # Bk\<up>(ln) @ Oc # ires, Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rn))"
  1480     using t_twice_append[of "Bk\<^bsup>lna\<^esup> @ Oc # ires" "Suc rs" rna]
  1467     using t_twice_append[of "Bk\<up>(lna) @ Oc # ires" "Suc rs" rna]
  1481     apply(erule_tac exE)
  1468     apply(erule_tac exE)
  1482     apply(erule_tac exE)
  1469     apply(erule_tac exE)
  1483     apply(erule_tac exE)
  1470     apply(erule_tac exE)
  1484     apply(simp add: wcode_main_first_part_len)
  1471     apply(simp add: wcode_main_first_part_len)
  1485     apply(rule_tac x = stp in exI, rule_tac x = "ln + lna" in exI, 
  1472     apply(rule_tac x = stp in exI, rule_tac x = "ln + lna" in exI, 
  1486           rule_tac x = rn in exI)
  1473           rule_tac x = rn in exI)
  1487     apply(simp add: t_wcode_main_def)
  1474     apply(simp add: t_wcode_main_def)
  1488     apply(simp add: exp_ind_def[THEN sym] exp_add[THEN sym])
  1475     apply(simp add: replicate_Suc[THEN sym] exp_add[THEN sym] del: replicate_Suc)
  1489     done
  1476     done
  1490   from this obtain stpb lnb rnb where stp2: 
  1477   from this obtain stpb lnb rnb where stp2: 
  1491     "steps (13, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb =
  1478     "steps0 (13, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rna)) t_wcode_main stpb =
  1492     (13 + t_twice_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnb\<^esup>)" by blast
  1479     (13 + t_twice_len, Bk # Bk # Bk\<up>(lnb) @ Oc # ires, Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rnb))" by blast
  1493   have "\<exists>stp ln rn. steps (13 + t_twice_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires,
  1480   have "\<exists>stp ln rn. steps0 (13 + t_twice_len, Bk # Bk # Bk\<up>(lnb) @ Oc # ires,
  1494     Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnb\<^esup>) t_wcode_main stp = 
  1481     Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rnb)) t_wcode_main stp = 
  1495        (Suc 0,  Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1482        (Suc 0,  Bk # Bk\<up>(ln) @ Oc # ires, Bk # Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rn))"
  1496     using wcode_jump1[of lnb "Oc # ires" "Suc rs" rnb]
  1483     using wcode_jump1[of lnb "Oc # ires" "Suc rs" rnb]
  1497     apply(erule_tac exE)
  1484     apply(erule_tac exE)
  1498     apply(erule_tac exE)
  1485     apply(erule_tac exE)
  1499     apply(erule_tac exE)
  1486     apply(erule_tac exE)
  1500     apply(rule_tac x = stp in exI, 
  1487     apply(rule_tac x = stp in exI, 
  1501           rule_tac x = ln in exI, 
  1488           rule_tac x = ln in exI, 
  1502           rule_tac x = rn in exI, simp add:wcode_main_first_part_len t_wcode_main_def)
  1489           rule_tac x = rn in exI, simp add:wcode_main_first_part_len t_wcode_main_def)
  1503     apply(subgoal_tac "Bk\<^bsup>lnb\<^esup> @ Bk # Bk # Oc # ires = Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires", simp)
  1490     apply(subgoal_tac "Bk\<up>(lnb) @ Bk # Bk # Oc # ires = Bk # Bk # Bk\<up>(lnb) @ Oc # ires", simp)
  1504     apply(simp add: exp_ind_def[THEN sym] exp_ind[THEN sym])
  1491     apply(simp add: replicate_Suc[THEN sym] exp_ind[THEN sym] del: replicate_Suc)
  1505     apply(simp)
  1492     apply(simp)
  1506     apply(case_tac lnb, simp, simp add: exp_ind_def[THEN sym] exp_ind)
  1493     apply(simp add: replicate_Suc[THEN sym] exp_ind del: replicate_Suc)
  1507     done               
  1494     done               
  1508   from this obtain stpc lnc rnc where stp3: 
  1495   from this obtain stpc lnc rnc where stp3: 
  1509     "steps (13 + t_twice_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires,
  1496     "steps0 (13 + t_twice_len, Bk # Bk # Bk\<up>(lnb) @ Oc # ires,
  1510     Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnb\<^esup>) t_wcode_main stpc = 
  1497     Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rnb)) t_wcode_main stpc = 
  1511        (Suc 0,  Bk # Bk\<^bsup>lnc\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnc\<^esup>)"
  1498        (Suc 0,  Bk # Bk\<up>(lnc) @ Oc # ires, Bk # Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rnc))"
  1512     by blast
  1499     by blast
  1513   from stp1 stp2 stp3 show "?thesis"
  1500   from stp1 stp2 stp3 show "?thesis"
  1514     apply(rule_tac x = "stpa + stpb + stpc" in exI, rule_tac x = lnc in exI,
  1501     apply(rule_tac x = "stpa + stpb + stpc" in exI, rule_tac x = lnc in exI,
  1515          rule_tac x = rnc in exI)
  1502          rule_tac x = rnc in exI)
  1516     apply(simp add: steps_add)
  1503     apply(simp add: steps_add)
  1520 
  1507 
  1521 (* Begin: fourtime_case*)
  1508 (* Begin: fourtime_case*)
  1522 fun wcode_on_left_moving_2_B :: "bin_inv_t"
  1509 fun wcode_on_left_moving_2_B :: "bin_inv_t"
  1523   where
  1510   where
  1524   "wcode_on_left_moving_2_B ires rs (l, r) =
  1511   "wcode_on_left_moving_2_B ires rs (l, r) =
  1525      (\<exists> ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # Bk # Oc # ires \<and>
  1512      (\<exists> ml mr rn. l = Bk\<up>(ml) @ Oc # Bk # Oc # ires \<and>
  1526                  r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
  1513                  r = Bk\<up>(mr) @ Oc\<up>(Suc rs) @ Bk\<up>(rn) \<and> 
  1527                  ml + mr > Suc 0 \<and> mr > 0)"
  1514                  ml + mr > Suc 0 \<and> mr > 0)"
  1528 
  1515 
  1529 fun wcode_on_left_moving_2_O :: "bin_inv_t"
  1516 fun wcode_on_left_moving_2_O :: "bin_inv_t"
  1530   where
  1517   where
  1531   "wcode_on_left_moving_2_O ires rs (l, r) =
  1518   "wcode_on_left_moving_2_O ires rs (l, r) =
  1532      (\<exists> ln rn. l = Bk # Oc # ires \<and>
  1519      (\<exists> ln rn. l = Bk # Oc # ires \<and>
  1533                r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1520                r = Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
  1534 
  1521 
  1535 fun wcode_on_left_moving_2 :: "bin_inv_t"
  1522 fun wcode_on_left_moving_2 :: "bin_inv_t"
  1536   where
  1523   where
  1537   "wcode_on_left_moving_2 ires rs (l, r) = 
  1524   "wcode_on_left_moving_2 ires rs (l, r) = 
  1538       (wcode_on_left_moving_2_B ires rs (l, r) \<or> 
  1525       (wcode_on_left_moving_2_B ires rs (l, r) \<or> 
  1540 
  1527 
  1541 fun wcode_on_checking_2 :: "bin_inv_t"
  1528 fun wcode_on_checking_2 :: "bin_inv_t"
  1542   where
  1529   where
  1543   "wcode_on_checking_2 ires rs (l, r) =
  1530   "wcode_on_checking_2 ires rs (l, r) =
  1544        (\<exists> ln rn. l = Oc#ires \<and> 
  1531        (\<exists> ln rn. l = Oc#ires \<and> 
  1545                  r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1532                  r = Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
  1546 
  1533 
  1547 fun wcode_goon_checking :: "bin_inv_t"
  1534 fun wcode_goon_checking :: "bin_inv_t"
  1548   where
  1535   where
  1549   "wcode_goon_checking ires rs (l, r) =
  1536   "wcode_goon_checking ires rs (l, r) =
  1550        (\<exists> ln rn. l = ires \<and>
  1537        (\<exists> ln rn. l = ires \<and>
  1551                  r = Oc # Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1538                  r = Oc # Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
  1552 
  1539 
  1553 fun wcode_right_move :: "bin_inv_t"
  1540 fun wcode_right_move :: "bin_inv_t"
  1554   where
  1541   where
  1555   "wcode_right_move ires rs (l, r) = 
  1542   "wcode_right_move ires rs (l, r) = 
  1556      (\<exists> ln rn. l = Oc # ires \<and>
  1543      (\<exists> ln rn. l = Oc # ires \<and>
  1557                  r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1544                  r = Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
  1558 
  1545 
  1559 fun wcode_erase2 :: "bin_inv_t"
  1546 fun wcode_erase2 :: "bin_inv_t"
  1560   where
  1547   where
  1561   "wcode_erase2 ires rs (l, r) = 
  1548   "wcode_erase2 ires rs (l, r) = 
  1562         (\<exists> ln rn. l = Bk # Oc # ires \<and>
  1549         (\<exists> ln rn. l = Bk # Oc # ires \<and>
  1563                  tl r = Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1550                  tl r = Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
  1564 
  1551 
  1565 fun wcode_on_right_moving_2 :: "bin_inv_t"
  1552 fun wcode_on_right_moving_2 :: "bin_inv_t"
  1566   where
  1553   where
  1567   "wcode_on_right_moving_2 ires rs (l, r) = 
  1554   "wcode_on_right_moving_2 ires rs (l, r) = 
  1568         (\<exists> ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # ires \<and> 
  1555         (\<exists> ml mr rn. l = Bk\<up>(ml) @ Oc # ires \<and> 
  1569                      r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> ml + mr > Suc 0)"
  1556                      r = Bk\<up>(mr) @ Oc\<up>(Suc rs) @ Bk\<up>(rn) \<and> ml + mr > Suc 0)"
  1570 
  1557 
  1571 fun wcode_goon_right_moving_2 :: "bin_inv_t"
  1558 fun wcode_goon_right_moving_2 :: "bin_inv_t"
  1572   where
  1559   where
  1573   "wcode_goon_right_moving_2 ires rs (l, r) = 
  1560   "wcode_goon_right_moving_2 ires rs (l, r) = 
  1574         (\<exists> ml mr ln rn. l = Oc\<^bsup>ml\<^esup> @ Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and>
  1561         (\<exists> ml mr ln rn. l = Oc\<up>(ml) @ Bk # Bk # Bk\<up>(ln) @ Oc # ires \<and>
  1575                         r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> ml + mr = Suc rs)"
  1562                         r = Oc\<up>(mr) @ Bk\<up>(rn) \<and> ml + mr = Suc rs)"
  1576 
  1563 
  1577 fun wcode_backto_standard_pos_2_B :: "bin_inv_t"
  1564 fun wcode_backto_standard_pos_2_B :: "bin_inv_t"
  1578   where
  1565   where
  1579   "wcode_backto_standard_pos_2_B ires rs (l, r) = 
  1566   "wcode_backto_standard_pos_2_B ires rs (l, r) = 
  1580            (\<exists> ln rn. l = Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
  1567            (\<exists> ln rn. l = Bk # Bk\<up>(ln) @ Oc # ires \<and> 
  1581                      r = Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1568                      r = Bk # Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn))"
  1582 
  1569 
  1583 fun wcode_backto_standard_pos_2_O :: "bin_inv_t"
  1570 fun wcode_backto_standard_pos_2_O :: "bin_inv_t"
  1584   where
  1571   where
  1585   "wcode_backto_standard_pos_2_O ires rs (l, r) = 
  1572   "wcode_backto_standard_pos_2_O ires rs (l, r) = 
  1586           (\<exists> ml mr ln rn. l = Oc\<^bsup>ml \<^esup>@ Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
  1573           (\<exists> ml mr ln rn. l = Oc\<up>(ml )@ Bk # Bk # Bk\<up>(ln) @ Oc # ires \<and> 
  1587                           r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
  1574                           r = Oc\<up>(mr) @ Bk\<up>(rn) \<and> 
  1588                           ml + mr = (Suc (Suc rs)) \<and> mr > 0)"
  1575                           ml + mr = (Suc (Suc rs)) \<and> mr > 0)"
  1589 
  1576 
  1590 fun wcode_backto_standard_pos_2 :: "bin_inv_t"
  1577 fun wcode_backto_standard_pos_2 :: "bin_inv_t"
  1591   where
  1578   where
  1592   "wcode_backto_standard_pos_2 ires rs (l, r) = 
  1579   "wcode_backto_standard_pos_2 ires rs (l, r) = 
  1594            wcode_backto_standard_pos_2_B ires rs (l, r))"
  1581            wcode_backto_standard_pos_2_B ires rs (l, r))"
  1595 
  1582 
  1596 fun wcode_before_fourtimes :: "bin_inv_t"
  1583 fun wcode_before_fourtimes :: "bin_inv_t"
  1597   where
  1584   where
  1598   "wcode_before_fourtimes ires rs (l, r) = 
  1585   "wcode_before_fourtimes ires rs (l, r) = 
  1599           (\<exists> ln rn. l = Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> 
  1586           (\<exists> ln rn. l = Bk # Bk # Bk\<up>(ln) @ Oc # ires \<and> 
  1600                     r = Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1587                     r = Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn))"
  1601 
  1588 
  1602 declare wcode_on_left_moving_2_B.simps[simp del] wcode_on_left_moving_2.simps[simp del]
  1589 declare wcode_on_left_moving_2_B.simps[simp del] wcode_on_left_moving_2.simps[simp del]
  1603         wcode_on_left_moving_2_O.simps[simp del] wcode_on_checking_2.simps[simp del]
  1590         wcode_on_left_moving_2_O.simps[simp del] wcode_on_checking_2.simps[simp del]
  1604         wcode_goon_checking.simps[simp del] wcode_right_move.simps[simp del]
  1591         wcode_goon_checking.simps[simp del] wcode_right_move.simps[simp del]
  1605         wcode_erase2.simps[simp del]
  1592         wcode_erase2.simps[simp del]
  1630             else if st = t_twice_len + 14 then wcode_before_fourtimes ires rs (l, r)
  1617             else if st = t_twice_len + 14 then wcode_before_fourtimes ires rs (l, r)
  1631             else False)"
  1618             else False)"
  1632 
  1619 
  1633 declare wcode_fourtimes_case_inv.simps[simp del]
  1620 declare wcode_fourtimes_case_inv.simps[simp del]
  1634 
  1621 
  1635 fun wcode_fourtimes_case_state :: "t_conf \<Rightarrow> nat"
  1622 fun wcode_fourtimes_case_state :: "config \<Rightarrow> nat"
  1636   where
  1623   where
  1637   "wcode_fourtimes_case_state (st, l, r) = 13 - st"
  1624   "wcode_fourtimes_case_state (st, l, r) = 13 - st"
  1638 
  1625 
  1639 fun wcode_fourtimes_case_step :: "t_conf \<Rightarrow> nat"
  1626 fun wcode_fourtimes_case_step :: "config \<Rightarrow> nat"
  1640   where
  1627   where
  1641   "wcode_fourtimes_case_step (st, l, r) = 
  1628   "wcode_fourtimes_case_step (st, l, r) = 
  1642          (if st = Suc 0 then length l
  1629          (if st = Suc 0 then length l
  1643           else if st = 9 then 
  1630           else if st = 9 then 
  1644            (if hd r = Oc then 1
  1631            (if hd r = Oc then 1
  1646           else if st = 10 then length r
  1633           else if st = 10 then length r
  1647           else if st = 11 then length r
  1634           else if st = 11 then length r
  1648           else if st = 12 then length l
  1635           else if st = 12 then length l
  1649           else 0)"
  1636           else 0)"
  1650 
  1637 
  1651 fun wcode_fourtimes_case_measure :: "t_conf \<Rightarrow> nat \<times> nat"
  1638 fun wcode_fourtimes_case_measure :: "config \<Rightarrow> nat \<times> nat"
  1652   where
  1639   where
  1653   "wcode_fourtimes_case_measure (st, l, r) = 
  1640   "wcode_fourtimes_case_measure (st, l, r) = 
  1654      (wcode_fourtimes_case_state (st, l, r), 
  1641      (wcode_fourtimes_case_state (st, l, r), 
  1655       wcode_fourtimes_case_step (st, l, r))"
  1642       wcode_fourtimes_case_step (st, l, r))"
  1656 
  1643 
  1657 definition wcode_fourtimes_case_le :: "(t_conf \<times> t_conf) set"
  1644 definition wcode_fourtimes_case_le :: "(config \<times> config) set"
  1658   where "wcode_fourtimes_case_le \<equiv> (inv_image lex_pair wcode_fourtimes_case_measure)"
  1645   where "wcode_fourtimes_case_le \<equiv> (inv_image lex_pair wcode_fourtimes_case_measure)"
  1659 
  1646 
  1660 lemma wf_wcode_fourtimes_case_le[intro]: "wf wcode_fourtimes_case_le"
  1647 lemma wf_wcode_fourtimes_case_le[intro]: "wf wcode_fourtimes_case_le"
  1661 by(auto intro:wf_inv_image simp: wcode_fourtimes_case_le_def)
  1648 by(auto intro:wf_inv_image simp: wcode_fourtimes_case_le_def)
  1662 
  1649 
  1664 apply(simp add: t_wcode_main_def fetch.simps 
  1651 apply(simp add: t_wcode_main_def fetch.simps 
  1665   t_wcode_main_first_part_def nth_of.simps)
  1652   t_wcode_main_first_part_def nth_of.simps)
  1666 done
  1653 done
  1667 
  1654 
  1668 lemma [simp]: "fetch t_wcode_main 7 Oc = (R, 8)"
  1655 lemma [simp]: "fetch t_wcode_main 7 Oc = (R, 8)"
  1669 apply(simp add: t_wcode_main_def fetch.simps 
  1656 apply(subgoal_tac "7 = Suc 6")
       
  1657 apply(simp only: t_wcode_main_def fetch.simps 
  1670   t_wcode_main_first_part_def nth_of.simps)
  1658   t_wcode_main_first_part_def nth_of.simps)
       
  1659 apply(auto)
  1671 done
  1660 done
  1672  
  1661  
  1673 lemma [simp]: "fetch t_wcode_main 8 Bk = (R, 9)"
  1662 lemma [simp]: "fetch t_wcode_main 8 Bk = (R, 9)"
  1674 apply(simp add: t_wcode_main_def fetch.simps 
  1663 apply(subgoal_tac "8 = Suc 7")
       
  1664 apply(simp only: t_wcode_main_def fetch.simps 
  1675   t_wcode_main_first_part_def nth_of.simps)
  1665   t_wcode_main_first_part_def nth_of.simps)
  1676 done
  1666 apply(auto)
       
  1667 done
       
  1668 
  1677 
  1669 
  1678 lemma [simp]: "fetch t_wcode_main 9 Bk = (R, 10)"
  1670 lemma [simp]: "fetch t_wcode_main 9 Bk = (R, 10)"
  1679 apply(simp add: t_wcode_main_def fetch.simps 
  1671 apply(subgoal_tac "9 = Suc 8")
       
  1672 apply(simp only: t_wcode_main_def fetch.simps 
  1680   t_wcode_main_first_part_def nth_of.simps)
  1673   t_wcode_main_first_part_def nth_of.simps)
       
  1674 apply(auto)
  1681 done
  1675 done
  1682 
  1676 
  1683 lemma [simp]: "fetch t_wcode_main 9 Oc = (W0, 9)"
  1677 lemma [simp]: "fetch t_wcode_main 9 Oc = (W0, 9)"
  1684 apply(simp add: t_wcode_main_def fetch.simps 
  1678 apply(subgoal_tac "9 = Suc 8")
       
  1679 apply(simp only: t_wcode_main_def fetch.simps 
  1685   t_wcode_main_first_part_def nth_of.simps)
  1680   t_wcode_main_first_part_def nth_of.simps)
       
  1681 apply(auto)
  1686 done
  1682 done
  1687 
  1683 
  1688 lemma [simp]: "fetch t_wcode_main 10 Bk = (R, 10)"
  1684 lemma [simp]: "fetch t_wcode_main 10 Bk = (R, 10)"
  1689 apply(simp add: t_wcode_main_def fetch.simps 
  1685 apply(subgoal_tac "10 = Suc 9")
       
  1686 apply(simp only: t_wcode_main_def fetch.simps 
  1690   t_wcode_main_first_part_def nth_of.simps)
  1687   t_wcode_main_first_part_def nth_of.simps)
       
  1688 apply(auto)
  1691 done
  1689 done
  1692 
  1690 
  1693 lemma [simp]: "fetch t_wcode_main 10 Oc = (R, 11)"
  1691 lemma [simp]: "fetch t_wcode_main 10 Oc = (R, 11)"
  1694 apply(simp add: t_wcode_main_def fetch.simps 
  1692 apply(subgoal_tac "10 = Suc 9")
       
  1693 apply(simp only: t_wcode_main_def fetch.simps 
  1695   t_wcode_main_first_part_def nth_of.simps)
  1694   t_wcode_main_first_part_def nth_of.simps)
  1696 done 
  1695 apply(auto)
       
  1696 done
  1697 
  1697 
  1698 lemma [simp]: "fetch t_wcode_main 11 Bk = (W1, 12)"
  1698 lemma [simp]: "fetch t_wcode_main 11 Bk = (W1, 12)"
  1699 apply(simp add: t_wcode_main_def fetch.simps 
  1699 apply(subgoal_tac "11 = Suc 10")
       
  1700 apply(simp only: t_wcode_main_def fetch.simps 
  1700   t_wcode_main_first_part_def nth_of.simps)
  1701   t_wcode_main_first_part_def nth_of.simps)
       
  1702 apply(auto)
  1701 done
  1703 done
  1702 
  1704 
  1703 lemma [simp]: "fetch t_wcode_main 11 Oc = (R, 11)"
  1705 lemma [simp]: "fetch t_wcode_main 11 Oc = (R, 11)"
  1704 apply(simp add: t_wcode_main_def fetch.simps 
  1706 apply(subgoal_tac "11 = Suc 10")
       
  1707 apply(simp only: t_wcode_main_def fetch.simps 
  1705   t_wcode_main_first_part_def nth_of.simps)
  1708   t_wcode_main_first_part_def nth_of.simps)
  1706 done 
  1709 apply(auto)
       
  1710 done
  1707 
  1711 
  1708 lemma [simp]: "fetch t_wcode_main 12 Oc = (L, 12)"
  1712 lemma [simp]: "fetch t_wcode_main 12 Oc = (L, 12)"
  1709 apply(simp add: t_wcode_main_def fetch.simps 
  1713 apply(subgoal_tac "12 = Suc 11")
       
  1714 apply(simp only: t_wcode_main_def fetch.simps 
  1710   t_wcode_main_first_part_def nth_of.simps)
  1715   t_wcode_main_first_part_def nth_of.simps)
  1711 done 
  1716 apply(auto)
       
  1717 done
  1712 
  1718 
  1713 lemma [simp]: "fetch t_wcode_main 12 Bk = (R, t_twice_len + 14)"
  1719 lemma [simp]: "fetch t_wcode_main 12 Bk = (R, t_twice_len + 14)"
  1714 apply(simp add: t_wcode_main_def fetch.simps 
  1720 apply(subgoal_tac "12 = Suc 11")
       
  1721 apply(simp only: t_wcode_main_def fetch.simps 
  1715   t_wcode_main_first_part_def nth_of.simps)
  1722   t_wcode_main_first_part_def nth_of.simps)
  1716 done
  1723 apply(auto)
  1717 
  1724 done
  1718 
  1725 
  1719 lemma [simp]: "wcode_on_left_moving_2 ires rs (b, []) = False"
  1726 lemma [simp]: "wcode_on_left_moving_2 ires rs (b, []) = False"
  1720 apply(auto simp: wcode_fourtimes_invs)
  1727 apply(auto simp: wcode_fourtimes_invs)
  1721 done
  1728 done
  1722 
  1729 
  1735 lemma [simp]: "wcode_erase2 ires rs (b, []) = False"
  1742 lemma [simp]: "wcode_erase2 ires rs (b, []) = False"
  1736 apply(auto simp: wcode_fourtimes_invs)
  1743 apply(auto simp: wcode_fourtimes_invs)
  1737 done
  1744 done
  1738 
  1745 
  1739 lemma [simp]: "wcode_on_right_moving_2 ires rs (b, []) = False"
  1746 lemma [simp]: "wcode_on_right_moving_2 ires rs (b, []) = False"
  1740 apply(auto simp: wcode_fourtimes_invs exponent_def)
  1747 apply(auto simp: wcode_fourtimes_invs)
  1741 done
  1748 done
  1742 
  1749 
  1743 lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, []) = False"
  1750 lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, []) = False"
  1744 apply(auto simp: wcode_fourtimes_invs exponent_def)
  1751 apply(auto simp: wcode_fourtimes_invs)
  1745 done
  1752 done
  1746     
  1753     
  1747 lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  1754 lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  1748 apply(simp add: wcode_fourtimes_invs, auto)
  1755 apply(simp add: wcode_fourtimes_invs, auto)
  1749 done
  1756 done     
  1750         
  1757 
  1751 lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Bk # list) \<Longrightarrow>  wcode_on_left_moving_2 ires rs (tl b, hd b # Bk # list)"
  1758 lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Bk # list) \<Longrightarrow>  wcode_on_left_moving_2 ires rs (tl b, hd b # Bk # list)"
  1752 apply(simp only: wcode_fourtimes_invs)
  1759 apply(simp only: wcode_fourtimes_invs)
  1753 apply(erule_tac disjE)
  1760 apply(erule_tac disjE)
  1754 apply(erule_tac exE)+
  1761 apply(erule_tac exE)+
  1755 apply(case_tac ml, simp)
  1762 apply(case_tac ml, simp)
  1756 apply(rule_tac x = "mr - (Suc (Suc 0))" in exI, rule_tac x = rn in exI, simp)
  1763 apply(rule_tac x = "mr - (Suc (Suc 0))" in exI, rule_tac x = rn in exI, simp)
  1757 apply(case_tac mr, simp, case_tac nat, simp, simp add: exp_ind)
  1764 apply(case_tac mr, simp, case_tac nat, simp, simp add: exp_ind del: replicate_Suc)
  1758 apply(rule_tac disjI1)
  1765 apply(rule_tac disjI1)
  1759 apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI,
  1766 apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI,
  1760       simp add: exp_ind_def)
  1767       simp add: replicate_Suc)
  1761 apply(simp)
  1768 apply(simp)
  1762 done
  1769 done
  1763 
  1770 
  1764 lemma [simp]: "wcode_on_checking_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  1771 lemma [simp]: "wcode_on_checking_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  1765 apply(auto simp: wcode_fourtimes_invs)
  1772 apply(auto simp: wcode_fourtimes_invs)
  1789 done
  1796 done
  1790 
  1797 
  1791 lemma [simp]: "wcode_erase2 ires rs (b, Bk # list) \<Longrightarrow> wcode_on_right_moving_2 ires rs (Bk # b, list)"
  1798 lemma [simp]: "wcode_erase2 ires rs (b, Bk # list) \<Longrightarrow> wcode_on_right_moving_2 ires rs (Bk # b, list)"
  1792 apply(auto simp:wcode_fourtimes_invs )
  1799 apply(auto simp:wcode_fourtimes_invs )
  1793 apply(rule_tac x = "Suc (Suc 0)" in exI, simp add: exp_ind)
  1800 apply(rule_tac x = "Suc (Suc 0)" in exI, simp add: exp_ind)
  1794 apply(rule_tac x =  "Suc (Suc ln)" in exI, simp add: exp_ind, auto)
  1801 apply(rule_tac x =  "Suc (Suc ln)" in exI, simp add: exp_ind del: replicate_Suc)
  1795 done
  1802 done
  1796 
  1803 
  1797 lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  1804 lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  1798 apply(auto simp:wcode_fourtimes_invs )
  1805 apply(auto simp:wcode_fourtimes_invs )
  1799 done
  1806 done
  1800 
  1807 
  1801 lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Bk # list)
  1808 lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Bk # list)
  1802        \<Longrightarrow> wcode_on_right_moving_2 ires rs (Bk # b, list)"
  1809        \<Longrightarrow> wcode_on_right_moving_2 ires rs (Bk # b, list)"
  1803 apply(auto simp: wcode_fourtimes_invs)
  1810 apply(auto simp: wcode_fourtimes_invs)
  1804 apply(rule_tac x = "Suc ml" in exI, simp add: exp_ind_def)
  1811 apply(rule_tac x = "Suc ml" in exI, simp)
  1805 apply(rule_tac x = "mr - 1" in exI, case_tac mr, auto simp: exp_ind_def)
  1812 apply(rule_tac x = "mr - 1" in exI, case_tac mr,auto)
  1806 done
  1813 done
  1807 
  1814 
  1808 lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  1815 lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  1809 apply(auto simp: wcode_fourtimes_invs)
  1816 apply(auto simp: wcode_fourtimes_invs)
  1810 done
  1817 done
  1812 lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Bk # list) \<Longrightarrow> 
  1819 lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Bk # list) \<Longrightarrow> 
  1813                  wcode_backto_standard_pos_2 ires rs (b, Oc # list)"
  1820                  wcode_backto_standard_pos_2 ires rs (b, Oc # list)"
  1814 apply(simp add: wcode_fourtimes_invs, auto)
  1821 apply(simp add: wcode_fourtimes_invs, auto)
  1815 apply(rule_tac x = ml in exI, auto)
  1822 apply(rule_tac x = ml in exI, auto)
  1816 apply(rule_tac x = "Suc 0" in exI, simp)
  1823 apply(rule_tac x = "Suc 0" in exI, simp)
  1817 apply(case_tac mr, simp_all add: exp_ind_def)
  1824 apply(case_tac mr, simp_all)
  1818 apply(rule_tac x = "rn - 1" in exI, simp)
  1825 apply(rule_tac x = "rn - 1" in exI, simp)
  1819 apply(case_tac rn, simp, simp add: exp_ind_def)
  1826 apply(case_tac rn, simp, simp)
  1820 done
  1827 done
  1821    
  1828    
  1822 lemma  [simp]: "wcode_backto_standard_pos_2 ires rs (b, Bk # list) \<Longrightarrow>  b \<noteq> []"
  1829 lemma  [simp]: "wcode_backto_standard_pos_2 ires rs (b, Bk # list) \<Longrightarrow>  b \<noteq> []"
  1823 apply(simp add: wcode_fourtimes_invs, auto)
  1830 apply(simp add: wcode_fourtimes_invs, auto)
  1824 done
  1831 done
  1828 done
  1835 done
  1829 
  1836 
  1830 lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Oc # list) \<Longrightarrow> 
  1837 lemma [simp]: "wcode_on_left_moving_2 ires rs (b, Oc # list) \<Longrightarrow> 
  1831                      wcode_on_checking_2 ires rs (tl b, hd b # Oc # list)"
  1838                      wcode_on_checking_2 ires rs (tl b, hd b # Oc # list)"
  1832 apply(auto simp: wcode_fourtimes_invs)
  1839 apply(auto simp: wcode_fourtimes_invs)
  1833 apply(case_tac [!] mr, simp_all add: exp_ind_def)
  1840 apply(case_tac [!] mr, simp_all)
  1834 done
  1841 done
  1835 
  1842 
  1836 lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, []) \<Longrightarrow> b \<noteq> []"
  1843 lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, []) \<Longrightarrow> b \<noteq> []"
  1837 apply(auto simp: wcode_fourtimes_invs)
  1844 apply(auto simp: wcode_fourtimes_invs)
  1838 done
  1845 done
  1842 apply(simp only: wcode_fourtimes_invs)
  1849 apply(simp only: wcode_fourtimes_invs)
  1843 apply(erule_tac exE)+
  1850 apply(erule_tac exE)+
  1844 apply(rule_tac disjI1)
  1851 apply(rule_tac disjI1)
  1845 apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, 
  1852 apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, 
  1846       rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
  1853       rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
  1847 apply(case_tac mr, simp, simp add: exp_ind_def)
       
  1848 done
  1854 done
  1849 
  1855 
  1850 lemma "wcode_backto_standard_pos_2 ires rs (b, Bk # list)
  1856 lemma "wcode_backto_standard_pos_2 ires rs (b, Bk # list)
  1851        \<Longrightarrow> (\<exists>ln. b = Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires) \<and> (\<exists>rn. list = Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1857        \<Longrightarrow> (\<exists>ln. b = Bk # Bk\<up>(ln) @ Oc # ires) \<and> (\<exists>rn. list = Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn))"
  1852 apply(auto simp: wcode_fourtimes_invs)
  1858 apply(auto simp: wcode_fourtimes_invs)
  1853 apply(case_tac [!] mr, auto simp: exp_ind_def)
  1859 apply(case_tac [!] mr, auto)
  1854 done
  1860 done
  1855 
  1861 
  1856 
  1862 
  1857 lemma [simp]: "wcode_on_checking_2 ires rs (b, Oc # list) \<Longrightarrow> False"
  1863 lemma [simp]: "wcode_on_checking_2 ires rs (b, Oc # list) \<Longrightarrow> False"
  1858 apply(simp add: wcode_fourtimes_invs)
  1864 apply(simp add: wcode_fourtimes_invs)
  1885 done
  1891 done
  1886 
  1892 
  1887 lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Oc # list)
  1893 lemma [simp]: "wcode_on_right_moving_2 ires rs (b, Oc # list)
  1888        \<Longrightarrow> wcode_goon_right_moving_2 ires rs (Oc # b, list)"
  1894        \<Longrightarrow> wcode_goon_right_moving_2 ires rs (Oc # b, list)"
  1889 apply(auto simp: wcode_fourtimes_invs)
  1895 apply(auto simp: wcode_fourtimes_invs)
  1890 apply(case_tac mr, simp_all add: exp_ind_def)
  1896 apply(case_tac mr, simp_all)
  1891 apply(rule_tac x = "Suc 0" in exI, auto)
  1897 apply(rule_tac x = "Suc 0" in exI, auto)
  1892 apply(rule_tac x = "ml - 2" in exI)
  1898 apply(rule_tac x = "ml - 2" in exI)
  1893 apply(case_tac ml, simp, case_tac nat, simp_all add: exp_ind_def)
  1899 apply(case_tac ml, simp, case_tac nat, simp_all)
  1894 done
  1900 done
  1895 
  1901 
  1896 lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
  1902 lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
  1897 apply(simp only:wcode_fourtimes_invs, auto)
  1903 apply(simp only:wcode_fourtimes_invs, auto)
  1898 done
  1904 done
  1899 
  1905 
  1900 lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, Bk # list)
  1906 lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, Bk # list)
  1901        \<Longrightarrow> (\<exists>ln. b = Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires) \<and> (\<exists>rn. list = Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  1907        \<Longrightarrow> (\<exists>ln. b = Bk # Bk\<up>(ln) @ Oc # ires) \<and> (\<exists>rn. list = Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn))"
  1902 apply(simp add: wcode_fourtimes_invs, auto)
  1908 apply(simp add: wcode_fourtimes_invs, auto)
  1903 apply(case_tac [!] mr, simp_all add: exp_ind_def)
  1909 apply(case_tac [!] mr, simp_all)
  1904 done
  1910 done
  1905 
  1911 
  1906 lemma [simp]: "wcode_on_checking_2 ires rs (b, Oc # list) = False"
  1912 lemma [simp]: "wcode_on_checking_2 ires rs (b, Oc # list) = False"
  1907 apply(simp add: wcode_fourtimes_invs)
  1913 apply(simp add: wcode_fourtimes_invs)
  1908 done
  1914 done
  1909 
  1915 
  1910 lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Oc # list) \<Longrightarrow>
  1916 lemma [simp]: "wcode_goon_right_moving_2 ires rs (b, Oc # list) \<Longrightarrow>
  1911        wcode_goon_right_moving_2 ires rs (Oc # b, list)"
  1917        wcode_goon_right_moving_2 ires rs (Oc # b, list)"
  1912 apply(simp only:wcode_fourtimes_invs, auto)
  1918 apply(simp only:wcode_fourtimes_invs, auto)
  1913 apply(rule_tac x = "Suc ml" in exI, auto simp: exp_ind_def)
  1919 apply(rule_tac x = "Suc ml" in exI, simp)
  1914 apply(rule_tac x = "mr - 1" in exI)
  1920 apply(rule_tac x = "mr - 1" in exI)
  1915 apply(case_tac mr, case_tac rn, auto simp: exp_ind_def)
  1921 apply(case_tac mr, case_tac rn, auto)
  1916 done
  1922 done
  1917 
  1923 
  1918 lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
  1924 lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
  1919 apply(simp only: wcode_fourtimes_invs, auto)
  1925 apply(simp only: wcode_fourtimes_invs, auto)
  1920 done
  1926 done
  1922 lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, Oc # list)    
  1928 lemma [simp]: "wcode_backto_standard_pos_2 ires rs (b, Oc # list)    
  1923             \<Longrightarrow> wcode_backto_standard_pos_2 ires rs (tl b, hd b # Oc # list)"
  1929             \<Longrightarrow> wcode_backto_standard_pos_2 ires rs (tl b, hd b # Oc # list)"
  1924 apply(simp only: wcode_fourtimes_invs)
  1930 apply(simp only: wcode_fourtimes_invs)
  1925 apply(erule_tac disjE)
  1931 apply(erule_tac disjE)
  1926 apply(erule_tac exE)+
  1932 apply(erule_tac exE)+
  1927 apply(case_tac ml, simp)
  1933 apply(case_tac ml, auto)
  1928 apply(rule_tac disjI2)
  1934 apply(rule_tac x = nat in exI, auto)
  1929 apply(rule_tac conjI, rule_tac x = ln in exI, simp)
  1935 apply(rule_tac x = "Suc mr" in exI, simp)
  1930 apply(rule_tac x = rn in exI, simp)
       
  1931 apply(rule_tac disjI1)
       
  1932 apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, 
       
  1933       rule_tac x = ln in exI, rule_tac x = rn in exI, simp add: exp_ind_def)
       
  1934 apply(simp)
       
  1935 done
  1936 done
  1936 
  1937 
  1937 lemma wcode_fourtimes_case_first_correctness:
  1938 lemma wcode_fourtimes_case_first_correctness:
  1938  shows "let P = (\<lambda> (st, l, r). st = t_twice_len + 14) in 
  1939  shows "let P = (\<lambda> (st, l, r). st = t_twice_len + 14) in 
  1939   let Q = (\<lambda> (st, l, r). wcode_fourtimes_case_inv st ires rs (l, r)) in 
  1940   let Q = (\<lambda> (st, l, r). wcode_fourtimes_case_inv st ires rs (l, r)) in 
  1940   let f = (\<lambda> stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp) in
  1941   let f = (\<lambda> stp. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp) in
  1941   \<exists> n .P (f n) \<and> Q (f (n::nat))"
  1942   \<exists> n .P (f n) \<and> Q (f (n::nat))"
  1942 proof -
  1943 proof -
  1943   let ?P = "(\<lambda> (st, l, r). st = t_twice_len + 14)"
  1944   let ?P = "(\<lambda> (st, l, r). st = t_twice_len + 14)"
  1944   let ?Q = "(\<lambda> (st, l, r). wcode_fourtimes_case_inv st ires rs (l, r))"
  1945   let ?Q = "(\<lambda> (st, l, r). wcode_fourtimes_case_inv st ires rs (l, r))"
  1945   let ?f = "(\<lambda> stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp)"
  1946   let ?f = "(\<lambda> stp. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp)"
  1946   have "\<exists> n . ?P (?f n) \<and> ?Q (?f (n::nat))"
  1947   have "\<exists> n . ?P (?f n) \<and> ?Q (?f (n::nat))"
  1947   proof(rule_tac halt_lemma2)
  1948   proof(rule_tac halt_lemma2)
  1948     show "wf wcode_fourtimes_case_le"
  1949     show "wf wcode_fourtimes_case_le"
  1949       by auto
  1950       by auto
  1950   next
  1951   next
  1951     show "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow>
  1952     show "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow>
  1952                   ?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_fourtimes_case_le"
  1953                   ?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_fourtimes_case_le"
  1953     apply(rule_tac allI,
  1954     apply(rule_tac allI,
  1954      case_tac "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main na", simp,
  1955      case_tac "steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main na", simp,
  1955      rule_tac impI)
  1956      rule_tac impI)
  1956     apply(simp add: tstep_red tstep.simps, case_tac c, simp, case_tac [2] aa, simp_all)
  1957     apply(simp add: step_red step.simps, case_tac c, simp, case_tac [2] aa, simp_all)
  1957     
  1958     apply(simp_all add: wcode_fourtimes_case_inv.simps
  1958     apply(simp_all add: wcode_fourtimes_case_inv.simps new_tape.simps 
       
  1959                         wcode_fourtimes_case_le_def lex_pair_def split: if_splits)
  1959                         wcode_fourtimes_case_le_def lex_pair_def split: if_splits)
       
  1960     apply(auto simp: wcode_backto_standard_pos_2.simps wcode_backto_standard_pos_2_O.simps
       
  1961       wcode_backto_standard_pos_2_B.simps)
       
  1962     apply(case_tac mr, simp_all)
  1960     done
  1963     done
  1961   next
  1964   next
  1962     show "?Q (?f 0)"
  1965     show "?Q (?f 0)"
  1963       apply(simp add: steps.simps wcode_fourtimes_case_inv.simps)
  1966       apply(simp add: steps.simps wcode_fourtimes_case_inv.simps)
  1964       apply(simp add: wcode_on_left_moving_2.simps wcode_on_left_moving_2_B.simps 
  1967       apply(simp add: wcode_on_left_moving_2.simps wcode_on_left_moving_2_B.simps 
  1965                       wcode_on_left_moving_2_O.simps)
  1968                       wcode_on_left_moving_2_O.simps)
  1966       apply(rule_tac x = "Suc m" in exI, simp add: exp_ind_def)
  1969       apply(rule_tac x = "Suc m" in exI, simp )
  1967       apply(rule_tac x ="Suc 0" in exI, auto)
  1970       apply(rule_tac x ="Suc 0" in exI, auto)
  1968       done
  1971       done
  1969   next
  1972   next
  1970     show "\<not> ?P (?f 0)"
  1973     show "\<not> ?P (?f 0)"
  1971       apply(simp add: steps.simps)
  1974       apply(simp add: steps.simps)
  1979 definition t_fourtimes_len :: "nat"
  1982 definition t_fourtimes_len :: "nat"
  1980   where
  1983   where
  1981   "t_fourtimes_len = (length t_fourtimes div 2)"
  1984   "t_fourtimes_len = (length t_fourtimes div 2)"
  1982 
  1985 
  1983 lemma t_fourtimes_len_gr:  "t_fourtimes_len > 0"
  1986 lemma t_fourtimes_len_gr:  "t_fourtimes_len > 0"
  1984 apply(simp add: t_fourtimes_len_def t_fourtimes_def)
  1987 apply(simp add: t_fourtimes_len_def t_fourtimes_def mopup.simps t_fourtimes_compile_def)
       
  1988 done
       
  1989 
       
  1990 lemma [intro]: "rec_calc_rel (constn 4) [rs] 4"
       
  1991 using prime_rel_exec_eq[of "constn 4" "[rs]" 4]
       
  1992 apply(subgoal_tac "primerec (constn 4) 1", auto)
       
  1993 done
       
  1994 
       
  1995 lemma [intro]: "rec_calc_rel rec_mult [rs, 4] (4 * rs)"
       
  1996 using prime_rel_exec_eq[of "rec_mult" "[rs, 4]"  "4*rs"]
       
  1997 apply(subgoal_tac "primerec rec_mult 2", auto simp: numeral_2_eq_2)
  1985 done
  1998 done
  1986 
  1999 
  1987 lemma t_fourtimes_correct: 
  2000 lemma t_fourtimes_correct: 
  1988   "\<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) 
  2001   "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) 
  1989     (tm_of abc_fourtimes @ tMp (Suc 0) (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) stp =
  2002     (tm_of abc_fourtimes @ shift (mopup 1) (length (tm_of abc_fourtimes) div 2)) stp =
  1990        (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2003        (0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))"
  1991 proof(case_tac "rec_ci rec_fourtimes")
  2004 proof(case_tac "rec_ci rec_fourtimes")
  1992   fix a b c
  2005   fix a b c
  1993   assume h: "rec_ci rec_fourtimes = (a, b, c)"
  2006   assume h: "rec_ci rec_fourtimes = (a, b, c)"
  1994   have "\<exists>stp m l. steps (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\<^bsup>n\<^esup>) (tm_of abc_fourtimes @ tMp (Suc 0) 
  2007   have "\<exists>stp m l. steps0 (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\<up>(n)) (tm_of abc_fourtimes @ shift (mopup 1) 
  1995     (start_of fourtimes_ly (length abc_fourtimes) - 1)) stp = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4*rs)\<^esup> @ Bk\<^bsup>l\<^esup>)"
  2008     (length (tm_of abc_fourtimes) div 2)) stp = (0, Bk\<up>(m) @ Bk # Bk # ires, Oc\<up>(Suc (4*rs)) @ Bk\<up>(l))"
  1996   proof(rule_tac t_compiled_by_rec)
  2009   proof(rule_tac recursive_compile_to_tm_correct)
  1997     show "rec_ci rec_fourtimes = (a, b, c)" by (simp add: h)
  2010     show "rec_ci rec_fourtimes = (a, b, c)" by (simp add: h)
  1998   next
  2011   next
  1999     show "rec_calc_rel rec_fourtimes [rs] (4 * rs)"
  2012     show "rec_calc_rel rec_fourtimes [rs] (4 * rs)"
  2000       using prime_rel_exec_eq [of rec_fourtimes "[rs]" "4 * rs"]
  2013       apply(simp add: rec_fourtimes_def)
  2001       apply(subgoal_tac "primerec rec_fourtimes (length [rs])")
  2014       apply(rule_tac rs =  "[rs, 4]" in calc_cn, simp_all)
  2002       apply(simp_all add: rec_fourtimes_def rec_exec.simps)
  2015       apply(rule_tac allI, case_tac k, auto simp: mult_lemma)
  2003       apply(auto)
       
  2004       apply(simp only: Nat.One_nat_def[THEN sym], auto)
       
  2005       done
  2016       done
  2006   next
  2017   next
  2007     show "length [rs] = Suc 0" by simp
  2018     show "length [rs] = 1" by simp
       
  2019   next	
       
  2020     show "layout_of (a [+] dummy_abc 1) = layout_of (a [+] dummy_abc 1)" by simp
  2008   next
  2021   next
  2009     show "layout_of (a [+] dummy_abc (Suc 0)) = layout_of (a [+] dummy_abc (Suc 0))"
  2022     show "tm_of abc_fourtimes = tm_of (a [+] dummy_abc 1)"
  2010       by simp
       
  2011   next
       
  2012     show "start_of fourtimes_ly (length abc_fourtimes) = 
       
  2013       start_of (layout_of (a [+] dummy_abc (Suc 0))) (length (a [+] dummy_abc (Suc 0)))"
       
  2014       using h
       
  2015       apply(simp add: fourtimes_ly_def abc_fourtimes_def)
       
  2016       done
       
  2017   next
       
  2018     show "tm_of abc_fourtimes = tm_of (a [+] dummy_abc (Suc 0))"
       
  2019       using h
  2023       using h
  2020       apply(simp add: abc_fourtimes_def)
  2024       apply(simp add: abc_fourtimes_def)
  2021       done
  2025       done
  2022   qed
  2026   qed
  2023   thus "\<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) 
  2027   thus "?thesis"
  2024             (tm_of abc_fourtimes @ tMp (Suc 0) (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) stp =
       
  2025        (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
       
  2026     apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
  2028     apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
  2027     done
  2029     done
  2028 qed
  2030 qed
  2029 
  2031 
       
  2032 lemma wf_fourtimes[intro]: "tm_wf (t_fourtimes_compile, 0)"
       
  2033 apply(simp only: t_fourtimes_compile_def)
       
  2034 apply(rule_tac t_compiled_correct)
       
  2035 apply(simp_all add: abc_twice_def)
       
  2036 done
       
  2037 
  2030 lemma t_fourtimes_change_term_state:
  2038 lemma t_fourtimes_change_term_state:
  2031   "\<exists> stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp
  2039   "\<exists> stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) t_fourtimes stp
  2032      = (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2040      = (Suc t_fourtimes_len, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))"
  2033 using t_fourtimes_correct[of ires rs n]
  2041 proof -
  2034 apply(erule_tac exE)
  2042   have "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) 
  2035 apply(erule_tac exE)
  2043     (tm_of abc_fourtimes @ shift (mopup 1) ((length (tm_of abc_fourtimes) div 2))) stp =
  2036 apply(erule_tac exE)
  2044     (0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))"
  2037 proof(drule_tac turing_change_termi_state)
  2045     by(rule_tac t_fourtimes_correct)
  2038   fix stp ln rn
  2046   then obtain stp ln rn where 
  2039   show "t_correct (tm_of abc_fourtimes @ tMp (Suc 0) 
  2047     "steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) 
  2040     (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))"
  2048     (tm_of abc_fourtimes @ shift (mopup 1) ((length (tm_of abc_fourtimes) div 2))) stp =
  2041     apply(rule_tac t_compiled_correct, auto simp: fourtimes_ly_def)
  2049     (0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))" by blast
       
  2050   hence "\<exists> stp. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
       
  2051     (adjust t_fourtimes_compile) stp
       
  2052      = (Suc (length t_fourtimes_compile div 2), Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))"
       
  2053     apply(rule_tac stp = stp in adjust_halt_eq)
       
  2054     apply(simp add: t_fourtimes_compile_def, auto)
  2042     done
  2055     done
  2043 next
  2056   then obtain stpb where 
  2044   fix stp ln rn
  2057     "steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
  2045   show "\<exists>stp. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
  2058     (adjust t_fourtimes_compile) stpb
  2046     (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) 
  2059      = (Suc (length t_fourtimes_compile div 2), Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))" ..
  2047         (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) stp =
  2060   thus "?thesis"
  2048     (Suc (length (tm_of abc_fourtimes @ tMp (Suc 0) (start_of fourtimes_ly 
  2061     apply(simp add: t_fourtimes_def t_fourtimes_len_def)
  2049     (length abc_fourtimes) - Suc 0)) div 2), Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>) \<Longrightarrow>
  2062     by metis
  2050     \<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp =
       
  2051     (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
       
  2052     apply(erule_tac exE)
       
  2053     apply(simp add: t_fourtimes_len_def t_fourtimes_def)
       
  2054     apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI, simp)
       
  2055     done
       
  2056 qed
  2063 qed
  2057 
  2064 
       
  2065 lemma [intro]: "length t_twice mod 2 = 0"
       
  2066 apply(auto simp: t_twice_def t_twice_compile_def)
       
  2067 done
       
  2068 
  2058 lemma t_fourtimes_append_pre:
  2069 lemma t_fourtimes_append_pre:
  2059   "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp
  2070   "steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) t_fourtimes stp
  2060   = (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)
  2071   = (Suc t_fourtimes_len, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))
  2061    \<Longrightarrow> \<exists> stp>0. steps (Suc 0 + length (t_wcode_main_first_part @ 
  2072    \<Longrightarrow> steps0 (Suc 0 + length (t_wcode_main_first_part @ 
  2062               tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2,
  2073               shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2,
  2063        Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
  2074        Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
  2064      ((t_wcode_main_first_part @ 
  2075      ((t_wcode_main_first_part @ 
  2065   tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) @ 
  2076   shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) @ 
  2066   tshift t_fourtimes (length (t_wcode_main_first_part @ 
  2077   shift t_fourtimes (length (t_wcode_main_first_part @ 
  2067   tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2) @ ([(L, 1), (L, 1)])) stp 
  2078   shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2) @ ([(L, 1), (L, 1)])) stp 
  2068   = (Suc t_fourtimes_len + length (t_wcode_main_first_part @ 
  2079   = ((Suc t_fourtimes_len) + length (t_wcode_main_first_part @ 
  2069   tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2,
  2080   shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2,
  2070   Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2081   Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))"
  2071 proof(rule_tac t_tshift_lemma, auto)
  2082 apply(rule_tac tm_append_shift_append_steps, simp_all)
  2072   assume "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp =
  2083 apply(auto simp: t_wcode_main_first_part_def)
  2073     (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2084 done
  2074   thus "0 < stp"
  2085 
  2075     using t_fourtimes_len_gr
       
  2076     apply(case_tac stp, simp_all add: steps.simps)
       
  2077     done
       
  2078 next
       
  2079   show "Suc 0 \<le> length t_fourtimes div 2"
       
  2080     apply(simp add: t_fourtimes_def shift_length tMp.simps)
       
  2081     done
       
  2082 next
       
  2083   show "t_ncorrect (t_wcode_main_first_part @ 
       
  2084     tshift t_twice (length t_wcode_main_first_part div 2) @ 
       
  2085     [(L, Suc 0), (L, Suc 0)])"
       
  2086     apply(simp add: t_ncorrect.simps t_wcode_main_first_part_def shift_length
       
  2087                     t_twice_def)
       
  2088     using tm_even[of abc_twice]
       
  2089     by arith
       
  2090 next
       
  2091   show "t_ncorrect t_fourtimes"
       
  2092     apply(simp add: t_fourtimes_def steps.simps t_ncorrect.simps)
       
  2093     using tm_even[of abc_fourtimes]
       
  2094     by arith
       
  2095 next
       
  2096   show "t_ncorrect [(L, Suc 0), (L, Suc 0)]"
       
  2097     apply(simp add: t_ncorrect.simps)
       
  2098     done
       
  2099 qed
       
  2100 
  2086 
  2101 lemma [simp]: "length t_wcode_main_first_part = 24"
  2087 lemma [simp]: "length t_wcode_main_first_part = 24"
  2102 apply(simp add: t_wcode_main_first_part_def)
  2088 apply(simp add: t_wcode_main_first_part_def)
  2103 done
  2089 done
  2104 
  2090 
  2105 lemma [simp]: "(26 + length t_twice) div 2 = (length t_twice) div 2 + 13"
  2091 lemma [simp]: "(26 + length t_twice) div 2 = (length t_twice) div 2 + 13"
  2106 using tm_even[of abc_twice]
  2092 apply(simp add: t_twice_def t_twice_def)
  2107 apply(simp add: t_twice_def)
  2093 done
  2108 done
  2094 
  2109 
  2095 lemma [simp]: "((26 + length (shift t_twice 12)) div 2)
  2110 lemma [simp]: "((26 + length (tshift t_twice 12)) div 2)
  2096              = (length (shift t_twice 12) div 2 + 13)"
  2111              = (length (tshift t_twice 12) div 2 + 13)"
       
  2112 using tm_even[of abc_twice]
       
  2113 apply(simp add: t_twice_def)
  2097 apply(simp add: t_twice_def)
  2114 done 
  2098 done 
  2115 
  2099 
  2116 lemma [simp]: "t_twice_len + 14 =  14 + length (tshift t_twice 12) div 2"
  2100 lemma [simp]: "t_twice_len + 14 =  14 + length (shift t_twice 12) div 2"
  2117 using tm_even[of abc_twice]
  2101 apply(simp add: t_twice_def t_twice_len_def)
  2118 apply(simp add: t_twice_def t_twice_len_def shift_length)
       
  2119 done
  2102 done
  2120 
  2103 
  2121 lemma t_fourtimes_append:
  2104 lemma t_fourtimes_append:
  2122   "\<exists> stp ln rn. 
  2105   "\<exists> stp ln rn. 
  2123   steps (Suc 0 + length (t_wcode_main_first_part @ tshift t_twice
  2106   steps0 (Suc 0 + length (t_wcode_main_first_part @ shift t_twice
  2124   (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, 
  2107   (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, 
  2125   Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
  2108   Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
  2126   ((t_wcode_main_first_part @ tshift t_twice (length t_wcode_main_first_part div 2) @
  2109   ((t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @
  2127   [(L, 1), (L, 1)]) @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)]) stp 
  2110   [(L, 1), (L, 1)]) @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)]) stp 
  2128   = (Suc t_fourtimes_len + length (t_wcode_main_first_part @ tshift t_twice
  2111   = (Suc t_fourtimes_len + length (t_wcode_main_first_part @ shift t_twice
  2129   (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires,
  2112   (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, Bk\<up>(ln) @ Bk # Bk # ires,
  2130                                                                  Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2113                                                                  Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))"
  2131   using t_fourtimes_change_term_state[of ires rs n]
  2114   using t_fourtimes_change_term_state[of ires rs n]
  2132   apply(erule_tac exE)
  2115   apply(erule_tac exE)
  2133   apply(erule_tac exE)
  2116   apply(erule_tac exE)
  2134   apply(erule_tac exE)
  2117   apply(erule_tac exE)
  2135   apply(drule_tac t_fourtimes_append_pre)
  2118   apply(drule_tac t_fourtimes_append_pre)
  2136   apply(erule_tac exE)
  2119   apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI)
  2137   apply(rule_tac x = stpa in exI, rule_tac x = ln in exI, rule_tac x = rn in exI)
  2120   apply(simp add: t_twice_len_def)
  2138   apply(simp add: t_twice_len_def shift_length)
       
  2139   done
  2121   done
  2140 
  2122 
  2141 lemma t_wcode_main_len: "length t_wcode_main = length t_twice + length t_fourtimes + 28"
  2123 lemma t_wcode_main_len: "length t_wcode_main = length t_twice + length t_fourtimes + 28"
  2142 apply(simp add: t_wcode_main_def shift_length)
  2124 apply(simp add: t_wcode_main_def)
  2143 done
  2125 done
  2144  
  2126 
       
  2127 lemma even_twice_len: "length t_twice mod 2 = 0"
       
  2128 apply(auto simp: t_twice_def t_twice_compile_def)
       
  2129 done
       
  2130 
       
  2131 lemma even_fourtimes_len: "length t_fourtimes mod 2 = 0"
       
  2132 apply(auto simp: t_fourtimes_def t_fourtimes_compile_def)
       
  2133 done
       
  2134 
       
  2135 lemma [simp]: "2 * (length t_twice div 2) = length t_twice"
       
  2136 using even_twice_len
       
  2137 by arith
       
  2138 
       
  2139 lemma [simp]: "2 * (length t_fourtimes div 2) = length t_fourtimes"
       
  2140 using even_fourtimes_len
       
  2141 by arith
       
  2142 
       
  2143 lemma [simp]: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) Oc
       
  2144              = (L, Suc 0)" 
       
  2145 apply(subgoal_tac "14 = Suc 13")
       
  2146 apply(simp only: fetch.simps add_Suc nth_of.simps t_wcode_main_def)
       
  2147 apply(simp add:length_append length_shift Parity.two_times_even_div_two even_twice_len t_fourtimes_len_def)
       
  2148 by arith
       
  2149 
       
  2150 lemma [simp]: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) Bk
       
  2151              = (L, Suc 0)"
       
  2152 apply(subgoal_tac "14 = Suc 13")
       
  2153 apply(simp only: fetch.simps add_Suc nth_of.simps t_wcode_main_def)
       
  2154 apply(simp add:length_append length_shift Parity.two_times_even_div_two even_twice_len t_fourtimes_len_def nth_append)
       
  2155 by arith
       
  2156 
  2145 lemma [simp]: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) b
  2157 lemma [simp]: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) b
  2146              = (L, Suc 0)"
  2158              = (L, Suc 0)"
  2147 using tm_even[of "abc_twice"] tm_even[of "abc_fourtimes"]
  2159 apply(case_tac b, simp_all)
  2148 apply(case_tac b)
       
  2149 apply(simp_all only: fetch.simps)
       
  2150 apply(auto simp: nth_of.simps t_wcode_main_len t_twice_len_def
       
  2151                  t_fourtimes_def t_twice_def t_fourtimes_def t_fourtimes_len_def)
       
  2152 apply(auto simp: t_wcode_main_def t_wcode_main_first_part_def shift_length t_twice_def nth_append 
       
  2153                     t_fourtimes_def)
       
  2154 done
  2160 done
  2155 
  2161 
  2156 lemma wcode_jump2: 
  2162 lemma wcode_jump2: 
  2157   "\<exists> stp ln rn. steps (t_twice_len + 14 + t_fourtimes_len
  2163   "\<exists> stp ln rn. steps0 (t_twice_len + 14 + t_fourtimes_len
  2158   , Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, Oc\<^bsup>Suc (4 * rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>) t_wcode_main stp =
  2164   , Bk # Bk # Bk\<up>(lnb) @ Oc # ires, Oc\<up>(Suc (4 * rs + 4)) @ Bk\<up>(rnb)) t_wcode_main stp =
  2159   (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4 * rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2165   (Suc 0, Bk # Bk\<up>(ln) @ Oc # ires, Bk # Oc\<up>(Suc (4 * rs + 4)) @ Bk\<up>(rn))"
  2160 apply(rule_tac x = "Suc 0" in exI)
  2166 apply(rule_tac x = "Suc 0" in exI)
  2161 apply(simp add: steps.simps shift_length)
  2167 apply(simp add: steps.simps)
  2162 apply(rule_tac x = lnb in exI, rule_tac x = rnb in exI)
  2168 apply(rule_tac x = lnb in exI, rule_tac x = rnb in exI)
  2163 apply(simp add: tstep.simps new_tape.simps)
  2169 apply(simp add: step.simps)
  2164 done
  2170 done
  2165 
  2171 
  2166 lemma wcode_fourtimes_case:
  2172 lemma wcode_fourtimes_case:
  2167   shows "\<exists>stp ln rn.
  2173   shows "\<exists>stp ln rn.
  2168   steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  2174   steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
  2169   (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2175   (Suc 0, Bk # Bk\<up>(ln) @ Oc # ires, Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rn))"
  2170 proof -
  2176 proof -
  2171   have "\<exists>stp ln rn.
  2177   have "\<exists>stp ln rn.
  2172   steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  2178   steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
  2173   (t_twice_len + 14, Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2179   (t_twice_len + 14, Bk # Bk # Bk\<up>(ln) @ Oc # ires, Oc\<up>(Suc (rs + 1)) @ Bk\<up>(rn))"
  2174     using wcode_fourtimes_case_first_correctness[of ires rs m n]
  2180     using wcode_fourtimes_case_first_correctness[of ires rs m n]
  2175     apply(simp add: wcode_fourtimes_case_inv.simps, auto)
  2181     apply(simp add: wcode_fourtimes_case_inv.simps, auto)
  2176     apply(rule_tac x = na in exI, rule_tac x = ln in exI,
  2182     apply(rule_tac x = na in exI, rule_tac x = ln in exI,
  2177           rule_tac x = rn in exI)
  2183           rule_tac x = rn in exI)
  2178     apply(simp)
  2184     apply(simp)
  2179     done
  2185     done
  2180   from this obtain stpa lna rna where stp1:
  2186   from this obtain stpa lna rna where stp1:
  2181     "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa =
  2187     "steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stpa =
  2182   (t_twice_len + 14, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast
  2188   (t_twice_len + 14, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (rs + 1)) @ Bk\<up>(rna))" by blast
  2183   have "\<exists>stp ln rn. steps (t_twice_len + 14, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rna\<^esup>)
  2189   have "\<exists>stp ln rn. steps0 (t_twice_len + 14, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (rs + 1)) @ Bk\<up>(rna))
  2184                      t_wcode_main stp =
  2190                      t_wcode_main stp =
  2185           (t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires,  Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2191           (t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\<up>(ln) @ Oc # ires,  Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rn))"
  2186     using t_fourtimes_append[of " Bk\<^bsup>lna\<^esup> @ Oc # ires" "rs + 1" rna]
  2192     using t_fourtimes_append[of " Bk\<up>(lna) @ Oc # ires" "rs + 1" rna]
  2187     apply(erule_tac exE)
  2193     apply(erule_tac exE)
  2188     apply(erule_tac exE)
  2194     apply(erule_tac exE)
  2189     apply(erule_tac exE)
  2195     apply(erule_tac exE)
  2190     apply(simp add: t_wcode_main_def)
  2196     apply(simp add: t_wcode_main_def)
  2191     apply(rule_tac x = stp in exI, 
  2197     apply(rule_tac x = stp in exI, 
  2192           rule_tac x = "ln + lna" in exI,
  2198           rule_tac x = "ln + lna" in exI,
  2193           rule_tac x = rn in exI, simp)
  2199           rule_tac x = rn in exI, simp)
  2194     apply(simp add: exp_ind_def[THEN sym] exp_add[THEN sym])
  2200     apply(simp add: replicate_Suc[THEN sym] exp_add[THEN sym] del: replicate_Suc)
  2195     done
  2201     done
  2196   from this obtain stpb lnb rnb where stp2:
  2202   from this obtain stpb lnb rnb where stp2:
  2197     "steps (t_twice_len + 14, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rna\<^esup>)
  2203     "steps0 (t_twice_len + 14, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (rs + 1)) @ Bk\<up>(rna))
  2198                      t_wcode_main stpb =
  2204                      t_wcode_main stpb =
  2199        (t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires,  Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>)"
  2205        (t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\<up>(lnb) @ Oc # ires,  Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rnb))"
  2200     by blast
  2206     by blast
  2201   have "\<exists>stp ln rn. steps (t_twice_len + 14 + t_fourtimes_len,
  2207   have "\<exists>stp ln rn. steps0 (t_twice_len + 14 + t_fourtimes_len,
  2202     Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires,  Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>)
  2208     Bk # Bk # Bk\<up>(lnb) @ Oc # ires,  Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rnb))
  2203     t_wcode_main stp =
  2209     t_wcode_main stp =
  2204     (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2210     (Suc 0, Bk # Bk\<up>(ln) @ Oc # ires, Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rn))"
  2205     apply(rule wcode_jump2)
  2211     apply(rule wcode_jump2)
  2206     done
  2212     done
  2207   from this obtain stpc lnc rnc where stp3: 
  2213   from this obtain stpc lnc rnc where stp3: 
  2208     "steps (t_twice_len + 14 + t_fourtimes_len,
  2214     "steps0 (t_twice_len + 14 + t_fourtimes_len,
  2209     Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires,  Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>)
  2215     Bk # Bk # Bk\<up>(lnb) @ Oc # ires,  Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rnb))
  2210     t_wcode_main stpc =
  2216     t_wcode_main stpc =
  2211     (Suc 0, Bk # Bk\<^bsup>lnc\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnc\<^esup>)"
  2217     (Suc 0, Bk # Bk\<up>(lnc) @ Oc # ires, Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rnc))"
  2212     by blast
  2218     by blast
  2213   from stp1 stp2 stp3 show "?thesis"
  2219   from stp1 stp2 stp3 show "?thesis"
  2214     apply(rule_tac x = "stpa + stpb + stpc" in exI,
  2220     apply(rule_tac x = "stpa + stpb + stpc" in exI,
  2215           rule_tac x = lnc in exI, rule_tac x = rnc in exI)
  2221           rule_tac x = lnc in exI, rule_tac x = rnc in exI)
  2216     apply(simp add: steps_add)
  2222     apply(simp add: steps_add)
  2220 (**********************************************************)
  2226 (**********************************************************)
  2221 
  2227 
  2222 fun wcode_on_left_moving_3_B :: "bin_inv_t"
  2228 fun wcode_on_left_moving_3_B :: "bin_inv_t"
  2223   where
  2229   where
  2224   "wcode_on_left_moving_3_B ires rs (l, r) = 
  2230   "wcode_on_left_moving_3_B ires rs (l, r) = 
  2225        (\<exists> ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # Bk # Bk # ires \<and>
  2231        (\<exists> ml mr rn. l = Bk\<up>(ml) @ Oc # Bk # Bk # ires \<and>
  2226                     r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
  2232                     r = Bk\<up>(mr) @ Oc\<up>(Suc rs) @ Bk\<up>(rn) \<and> 
  2227                     ml + mr > Suc 0 \<and> mr > 0 )"
  2233                     ml + mr > Suc 0 \<and> mr > 0 )"
  2228 
  2234 
  2229 fun wcode_on_left_moving_3_O :: "bin_inv_t"
  2235 fun wcode_on_left_moving_3_O :: "bin_inv_t"
  2230   where
  2236   where
  2231   "wcode_on_left_moving_3_O ires rs (l, r) = 
  2237   "wcode_on_left_moving_3_O ires rs (l, r) = 
  2232          (\<exists> ln rn. l = Bk # Bk # ires \<and>
  2238          (\<exists> ln rn. l = Bk # Bk # ires \<and>
  2233                    r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2239                    r = Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
  2234 
  2240 
  2235 fun wcode_on_left_moving_3 :: "bin_inv_t"
  2241 fun wcode_on_left_moving_3 :: "bin_inv_t"
  2236   where
  2242   where
  2237   "wcode_on_left_moving_3 ires rs (l, r) = 
  2243   "wcode_on_left_moving_3 ires rs (l, r) = 
  2238        (wcode_on_left_moving_3_B ires rs (l, r) \<or>  
  2244        (wcode_on_left_moving_3_B ires rs (l, r) \<or>  
  2240 
  2246 
  2241 fun wcode_on_checking_3 :: "bin_inv_t"
  2247 fun wcode_on_checking_3 :: "bin_inv_t"
  2242   where
  2248   where
  2243   "wcode_on_checking_3 ires rs (l, r) = 
  2249   "wcode_on_checking_3 ires rs (l, r) = 
  2244          (\<exists> ln rn. l = Bk # ires \<and>
  2250          (\<exists> ln rn. l = Bk # ires \<and>
  2245              r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2251              r = Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
  2246 
  2252 
  2247 fun wcode_goon_checking_3 :: "bin_inv_t"
  2253 fun wcode_goon_checking_3 :: "bin_inv_t"
  2248   where
  2254   where
  2249   "wcode_goon_checking_3 ires rs (l, r) = 
  2255   "wcode_goon_checking_3 ires rs (l, r) = 
  2250          (\<exists> ln rn. l = ires \<and>
  2256          (\<exists> ln rn. l = ires \<and>
  2251              r = Bk # Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2257              r = Bk # Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
  2252 
  2258 
  2253 fun wcode_stop :: "bin_inv_t"
  2259 fun wcode_stop :: "bin_inv_t"
  2254   where
  2260   where
  2255   "wcode_stop ires rs (l, r) = 
  2261   "wcode_stop ires rs (l, r) = 
  2256           (\<exists> ln rn. l = Bk # ires \<and>
  2262           (\<exists> ln rn. l = Bk # ires \<and>
  2257              r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2263              r = Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
  2258 
  2264 
  2259 fun wcode_halt_case_inv :: "nat \<Rightarrow> bin_inv_t"
  2265 fun wcode_halt_case_inv :: "nat \<Rightarrow> bin_inv_t"
  2260   where
  2266   where
  2261   "wcode_halt_case_inv st ires rs (l, r) = 
  2267   "wcode_halt_case_inv st ires rs (l, r) = 
  2262           (if st = 0 then wcode_stop ires rs (l, r)
  2268           (if st = 0 then wcode_stop ires rs (l, r)
  2263            else if st = Suc 0 then wcode_on_left_moving_3 ires rs (l, r)
  2269            else if st = Suc 0 then wcode_on_left_moving_3 ires rs (l, r)
  2264            else if st = Suc (Suc 0) then wcode_on_checking_3 ires rs (l, r)
  2270            else if st = Suc (Suc 0) then wcode_on_checking_3 ires rs (l, r)
  2265            else if st = 7 then wcode_goon_checking_3 ires rs (l, r)
  2271            else if st = 7 then wcode_goon_checking_3 ires rs (l, r)
  2266            else False)"
  2272            else False)"
  2267 
  2273 
  2268 fun wcode_halt_case_state :: "t_conf \<Rightarrow> nat"
  2274 fun wcode_halt_case_state :: "config \<Rightarrow> nat"
  2269   where
  2275   where
  2270   "wcode_halt_case_state (st, l, r) = 
  2276   "wcode_halt_case_state (st, l, r) = 
  2271            (if st = 1 then 5
  2277            (if st = 1 then 5
  2272             else if st = Suc (Suc 0) then 4
  2278             else if st = Suc (Suc 0) then 4
  2273             else if st = 7 then 3
  2279             else if st = 7 then 3
  2274             else 0)"
  2280             else 0)"
  2275 
  2281 
  2276 fun wcode_halt_case_step :: "t_conf \<Rightarrow> nat"
  2282 fun wcode_halt_case_step :: "config \<Rightarrow> nat"
  2277   where
  2283   where
  2278   "wcode_halt_case_step (st, l, r) = 
  2284   "wcode_halt_case_step (st, l, r) = 
  2279          (if st = 1 then length l
  2285          (if st = 1 then length l
  2280          else 0)"
  2286          else 0)"
  2281 
  2287 
  2282 fun wcode_halt_case_measure :: "t_conf \<Rightarrow> nat \<times> nat"
  2288 fun wcode_halt_case_measure :: "config \<Rightarrow> nat \<times> nat"
  2283   where
  2289   where
  2284   "wcode_halt_case_measure (st, l, r) = 
  2290   "wcode_halt_case_measure (st, l, r) = 
  2285      (wcode_halt_case_state (st, l, r), 
  2291      (wcode_halt_case_state (st, l, r), 
  2286       wcode_halt_case_step (st, l, r))"
  2292       wcode_halt_case_step (st, l, r))"
  2287 
  2293 
  2288 definition wcode_halt_case_le :: "(t_conf \<times> t_conf) set"
  2294 definition wcode_halt_case_le :: "(config \<times> config) set"
  2289   where "wcode_halt_case_le \<equiv> (inv_image lex_pair wcode_halt_case_measure)"
  2295   where "wcode_halt_case_le \<equiv> (inv_image lex_pair wcode_halt_case_measure)"
  2290 
  2296 
  2291 lemma wf_wcode_halt_case_le[intro]: "wf wcode_halt_case_le"
  2297 lemma wf_wcode_halt_case_le[intro]: "wf wcode_halt_case_le"
  2292 by(auto intro:wf_inv_image simp: wcode_halt_case_le_def)
  2298 by(auto intro:wf_inv_image simp: wcode_halt_case_le_def)
  2293 
  2299 
  2299   wcode_on_left_moving_3_B.simps wcode_on_left_moving_3_O.simps
  2305   wcode_on_left_moving_3_B.simps wcode_on_left_moving_3_O.simps
  2300   wcode_on_checking_3.simps wcode_goon_checking_3.simps 
  2306   wcode_on_checking_3.simps wcode_goon_checking_3.simps 
  2301   wcode_on_left_moving_3.simps wcode_stop.simps
  2307   wcode_on_left_moving_3.simps wcode_stop.simps
  2302 
  2308 
  2303 lemma [simp]: "fetch t_wcode_main 7 Bk = (R, 0)"
  2309 lemma [simp]: "fetch t_wcode_main 7 Bk = (R, 0)"
  2304 apply(simp add: fetch.simps t_wcode_main_def nth_append nth_of.simps
  2310 apply(subgoal_tac "7 = Suc 6")
       
  2311 apply(simp only: fetch.simps t_wcode_main_def nth_append nth_of.simps
  2305                 t_wcode_main_first_part_def)
  2312                 t_wcode_main_first_part_def)
       
  2313 apply(auto)
  2306 done
  2314 done
  2307 
  2315 
  2308 lemma [simp]: "wcode_on_left_moving_3 ires rs (b, [])  = False"
  2316 lemma [simp]: "wcode_on_left_moving_3 ires rs (b, [])  = False"
  2309 apply(simp only: wcode_halt_invs)
  2317 apply(simp only: wcode_halt_invs)
  2310 apply(simp add: exp_ind_def)
  2318 apply(simp)
  2311 done    
  2319 done    
  2312 
  2320 
  2313 lemma [simp]: "wcode_on_checking_3 ires rs (b, []) = False"
  2321 lemma [simp]: "wcode_on_checking_3 ires rs (b, []) = False"
  2314 apply(simp add: wcode_halt_invs)
  2322 apply(simp add: wcode_halt_invs)
  2315 done
  2323 done
  2323 apply(simp only: wcode_halt_invs)
  2331 apply(simp only: wcode_halt_invs)
  2324 apply(erule_tac disjE)
  2332 apply(erule_tac disjE)
  2325 apply(erule_tac exE)+
  2333 apply(erule_tac exE)+
  2326 apply(case_tac ml, simp)
  2334 apply(case_tac ml, simp)
  2327 apply(rule_tac x = "mr - 2" in exI, rule_tac x = rn in exI)
  2335 apply(rule_tac x = "mr - 2" in exI, rule_tac x = rn in exI)
  2328 apply(case_tac mr, simp, simp add: exp_ind, simp add: exp_ind[THEN sym])
  2336 apply(case_tac mr, simp, simp add: exp_ind del: replicate_Suc)
       
  2337 apply(case_tac nat, simp, simp add: exp_ind del: replicate_Suc)
  2329 apply(rule_tac disjI1)
  2338 apply(rule_tac disjI1)
  2330 apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, 
  2339 apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, 
  2331       rule_tac x = rn in exI, simp add: exp_ind_def)
  2340       rule_tac x = rn in exI, simp)
  2332 apply(simp)
  2341 apply(simp)
  2333 done
  2342 done
  2334 
  2343 
  2335 lemma [simp]: "wcode_goon_checking_3 ires rs (b, Bk # list) \<Longrightarrow> 
  2344 lemma [simp]: "wcode_goon_checking_3 ires rs (b, Bk # list) \<Longrightarrow> 
  2336   (b = [] \<longrightarrow> wcode_stop ires rs ([Bk], list)) \<and>
  2345   (b = [] \<longrightarrow> wcode_stop ires rs ([Bk], list)) \<and>
  2343 done
  2352 done
  2344 
  2353 
  2345 lemma [simp]: "wcode_on_left_moving_3 ires rs (b, Oc # list) \<Longrightarrow> 
  2354 lemma [simp]: "wcode_on_left_moving_3 ires rs (b, Oc # list) \<Longrightarrow> 
  2346                wcode_on_checking_3 ires rs (tl b, hd b # Oc # list)"
  2355                wcode_on_checking_3 ires rs (tl b, hd b # Oc # list)"
  2347 apply(simp add:wcode_halt_invs, auto)
  2356 apply(simp add:wcode_halt_invs, auto)
  2348 apply(case_tac [!] mr, simp_all add: exp_ind_def)
  2357 apply(case_tac [!] mr, simp_all)
  2349 done     
  2358 done     
  2350 
  2359 
  2351 lemma [simp]: "wcode_on_checking_3 ires rs (b, Oc # list) = False"
  2360 lemma [simp]: "wcode_on_checking_3 ires rs (b, Oc # list) = False"
  2352 apply(auto simp: wcode_halt_invs)
  2361 apply(auto simp: wcode_halt_invs)
  2353 done
  2362 done
  2354 
  2363 
  2355 lemma [simp]: "wcode_on_left_moving_3 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  2364 lemma [simp]: "wcode_on_left_moving_3 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  2356 apply(simp add: wcode_halt_invs, auto)
  2365 apply(simp add: wcode_halt_invs, auto)
  2357 done
  2366 done
  2358 
       
  2359 
  2367 
  2360 lemma [simp]: "wcode_on_checking_3 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  2368 lemma [simp]: "wcode_on_checking_3 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  2361 apply(auto simp: wcode_halt_invs)
  2369 apply(auto simp: wcode_halt_invs)
  2362 done
  2370 done
  2363 
  2371 
  2371 done
  2379 done
  2372 
  2380 
  2373 lemma t_halt_case_correctness: 
  2381 lemma t_halt_case_correctness: 
  2374 shows "let P = (\<lambda> (st, l, r). st = 0) in 
  2382 shows "let P = (\<lambda> (st, l, r). st = 0) in 
  2375        let Q = (\<lambda> (st, l, r). wcode_halt_case_inv st ires rs (l, r)) in 
  2383        let Q = (\<lambda> (st, l, r). wcode_halt_case_inv st ires rs (l, r)) in 
  2376        let f = (\<lambda> stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp) in
  2384        let f = (\<lambda> stp. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp) in
  2377        \<exists> n .P (f n) \<and> Q (f (n::nat))"
  2385        \<exists> n .P (f n) \<and> Q (f (n::nat))"
  2378 proof -
  2386 proof -
  2379   let ?P = "(\<lambda> (st, l, r). st = 0)"
  2387   let ?P = "(\<lambda> (st, l, r). st = 0)"
  2380   let ?Q = "(\<lambda> (st, l, r). wcode_halt_case_inv st ires rs (l, r))"
  2388   let ?Q = "(\<lambda> (st, l, r). wcode_halt_case_inv st ires rs (l, r))"
  2381   let ?f = "(\<lambda> stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp)"
  2389   let ?f = "(\<lambda> stp. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp)"
  2382   have "\<exists> n. ?P (?f n) \<and> ?Q (?f (n::nat))"
  2390   have "\<exists> n. ?P (?f n) \<and> ?Q (?f (n::nat))"
  2383   proof(rule_tac halt_lemma2)
  2391   proof(rule_tac halt_lemma2)
  2384     show "wf wcode_halt_case_le" by auto
  2392     show "wf wcode_halt_case_le" by auto
  2385   next
  2393   next
  2386     show "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow> 
  2394     show "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow> 
  2387                     ?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_halt_case_le"
  2395                     ?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_halt_case_le"
  2388       apply(rule_tac allI, rule_tac impI, case_tac "?f na")
  2396       apply(rule_tac allI, rule_tac impI, case_tac "?f na")
  2389       apply(simp add: tstep_red tstep.simps)
  2397       apply(simp add: step_red step.simps)
  2390       apply(case_tac c, simp, case_tac [2] aa)
  2398       apply(case_tac c, simp, case_tac [2] aa)
  2391       apply(simp_all split: if_splits add: new_tape.simps wcode_halt_case_le_def lex_pair_def)
  2399       apply(simp_all split: if_splits add: wcode_halt_case_le_def lex_pair_def)
  2392       done      
  2400       done      
  2393   next 
  2401   next 
  2394     show "?Q (?f 0)"
  2402     show "?Q (?f 0)"
  2395       apply(simp add: steps.simps wcode_halt_invs)
  2403       apply(simp add: steps.simps wcode_halt_invs)
  2396       apply(rule_tac x = "Suc m" in exI, simp add: exp_ind_def)
  2404       apply(rule_tac x = "Suc m" in exI, simp)
  2397       apply(rule_tac x = "Suc 0" in exI, auto)
  2405       apply(rule_tac x = "Suc 0" in exI, auto)
  2398       done
  2406       done
  2399   next
  2407   next
  2400     show "\<not> ?P (?f 0)"
  2408     show "\<not> ?P (?f 0)"
  2401       apply(simp add: steps.simps)
  2409       apply(simp add: steps.simps)
  2405     apply(auto)
  2413     apply(auto)
  2406     done
  2414     done
  2407 qed
  2415 qed
  2408 
  2416 
  2409 declare wcode_halt_case_inv.simps[simp del]
  2417 declare wcode_halt_case_inv.simps[simp del]
  2410 lemma [intro]: "\<exists> xs. (<rev list @ [aa::nat]> :: block list) = Oc # xs"
  2418 lemma [intro]: "\<exists> xs. (<rev list @ [aa::nat]> :: cell list) = Oc # xs"
  2411 apply(case_tac "rev list", simp)
  2419 apply(case_tac "rev list", simp)
  2412 apply(simp add: tape_of_nat_abv tape_of_nat_list.simps exp_ind_def)
  2420 apply(simp add: tape_of_nl_cons)
  2413 apply(case_tac list, simp, simp)
       
  2414 done
  2421 done
  2415 
  2422 
  2416 lemma wcode_halt_case:
  2423 lemma wcode_halt_case:
  2417   "\<exists>stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
  2424   "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n))
  2418   t_wcode_main stp  = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2425   t_wcode_main stp  = (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
  2419   using t_halt_case_correctness[of ires rs m n]
  2426   using t_halt_case_correctness[of ires rs m n]
  2420 apply(simp)
  2427 apply(simp)
  2421 apply(erule_tac exE)
  2428 apply(erule_tac exE)
  2422 apply(case_tac "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires,
  2429 apply(case_tac "steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Bk # ires,
  2423                 Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main na")
  2430                 Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main na")
  2424 apply(auto simp: wcode_halt_case_inv.simps wcode_stop.simps)
  2431 apply(auto simp: wcode_halt_case_inv.simps wcode_stop.simps)
  2425 apply(rule_tac x = na in exI, rule_tac x = ln in exI, 
  2432 apply(rule_tac x = na in exI, rule_tac x = ln in exI, 
  2426       rule_tac x = rn in exI, simp)
  2433       rule_tac x = rn in exI, simp)
  2427 done
  2434 done
  2428 
  2435 
  2429 lemma bl_bin_one: "bl_bin [Oc] =  Suc 0"
  2436 lemma bl_bin_one: "bl_bin [Oc] =  Suc 0"
  2430 apply(simp add: bl_bin.simps)
  2437 apply(simp add: bl_bin.simps)
  2431 done
  2438 done
  2432 
  2439 
       
  2440 lemma [simp]: "bl_bin [Oc] = 1"
       
  2441 apply(simp add: bl_bin.simps)
       
  2442 done
       
  2443 
       
  2444 lemma [intro]: "2 * 2 ^ a = Suc (Suc (2 * bl_bin (Oc \<up> a)))"
       
  2445 apply(induct a, auto simp: bl_bin.simps)
       
  2446 done
       
  2447 declare replicate_Suc[simp del]
       
  2448 
  2433 lemma t_wcode_main_lemma_pre:
  2449 lemma t_wcode_main_lemma_pre:
  2434   "\<lbrakk>args \<noteq> []; lm = <args::nat list>\<rbrakk> \<Longrightarrow> 
  2450   "\<lbrakk>args \<noteq> []; lm = <args::nat list>\<rbrakk> \<Longrightarrow> 
  2435        \<exists> stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev lm @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main
  2451        \<exists> stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main
  2436                     stp
  2452                     stp
  2437       = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin lm + rs * 2^(length lm - 1) \<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2453       = (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin lm + rs * 2^(length lm - 1) ) @ Bk\<up>(rn))"
  2438 proof(induct "length args" arbitrary: args lm rs m n, simp)
  2454 proof(induct "length args" arbitrary: args lm rs m n, simp)
  2439   fix x args lm rs m n
  2455   fix x args lm rs m n
  2440   assume ind:
  2456   assume ind:
  2441     "\<And>args lm rs m n.
  2457     "\<And>args lm rs m n.
  2442     \<lbrakk>x = length args; (args::nat list) \<noteq> []; lm = <args>\<rbrakk>
  2458     \<lbrakk>x = length args; (args::nat list) \<noteq> []; lm = <args>\<rbrakk>
  2443     \<Longrightarrow> \<exists>stp ln rn.
  2459     \<Longrightarrow> \<exists>stp ln rn.
  2444     steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev lm @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  2460     steps0 (Suc 0, Bk # Bk\<up>(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
  2445     (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin lm + rs * 2 ^ (length lm - 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2461     (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin lm + rs * 2 ^ (length lm - 1)) @ Bk\<up>(rn))"
  2446   
       
  2447     and h: "Suc x = length args" "(args::nat list) \<noteq> []" "lm = <args>"
  2462     and h: "Suc x = length args" "(args::nat list) \<noteq> []" "lm = <args>"
  2448   from h have "\<exists> (a::nat) xs. args = xs @ [a]"
  2463   from h have "\<exists> (a::nat) xs. args = xs @ [a]"
  2449     apply(rule_tac x = "last args" in exI)
  2464     apply(rule_tac x = "last args" in exI)
  2450     apply(rule_tac x = "butlast args" in exI, auto)
  2465     apply(rule_tac x = "butlast args" in exI, auto)
  2451     done    
  2466     done    
  2452   from this obtain a xs where "args = xs @ [a]" by blast
  2467   from this obtain a xs where "args = xs @ [a]" by blast
  2453   from h and this show
  2468   from h and this show
  2454     "\<exists>stp ln rn.
  2469     "\<exists>stp ln rn.
  2455     steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev lm @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  2470     steps0 (Suc 0, Bk # Bk\<up>(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
  2456     (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin lm + rs * 2 ^ (length lm - 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2471     (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin lm + rs * 2 ^ (length lm - 1)) @ Bk\<up>(rn))"
  2457   proof(case_tac "xs::nat list", simp)
  2472   proof(case_tac "xs::nat list", simp)
  2458     show "\<exists>stp ln rn.
  2473     show "\<exists>stp ln rn.
  2459       steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  2474           steps0 (Suc 0, Bk # Bk \<up> m @ Oc \<up> Suc a @ Bk # Bk # ires, Bk # Oc \<up> Suc rs @ Bk \<up> n) t_wcode_main stp =
  2460       (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<a>) + rs * 2 ^ a\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2475           (0, Bk # ires, Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> (bl_bin (Oc \<up> Suc a) + rs * 2 ^ a) @ Bk \<up> rn)"
  2461     proof(induct "a" arbitrary: m n rs ires, simp)
  2476     proof(induct "a" arbitrary: m n rs ires, simp)
  2462       fix m n rs ires
  2477       fix m n rs ires
  2463       show "\<exists>stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>)
  2478       show "\<exists>stp ln rn.
  2464         t_wcode_main stp  = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin [Oc] + rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2479           steps0 (Suc 0, Bk # Bk \<up> m @ Oc # Bk # Bk # ires, Bk # Oc \<up> Suc rs @ Bk \<up> n) t_wcode_main stp =
  2465         apply(simp add: bl_bin_one)
  2480           (0, Bk # ires, Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> Suc rs @ Bk \<up> rn)"
  2466         apply(rule_tac wcode_halt_case)
  2481           apply(rule_tac wcode_halt_case)
  2467         done
  2482         done
  2468     next
  2483     next
  2469       fix a m n rs ires
  2484       fix a m n rs ires
  2470       assume ind2: 
  2485       assume ind2:
  2471         "\<And>m n rs ires.
  2486         "\<And>m n rs ires.
  2472         \<exists>stp ln rn.
  2487            \<exists>stp ln rn.
  2473         steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  2488               steps0 (Suc 0, Bk # Bk \<up> m @ Oc \<up> Suc a @ Bk # Bk # ires, Bk # Oc \<up> Suc rs @ Bk \<up> n) t_wcode_main stp =
  2474         (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<a>) + rs * 2 ^ a\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2489               (0, Bk # ires, Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> (bl_bin (Oc \<up> Suc a) + rs * 2 ^ a) @ Bk \<up> rn)"
  2475       show "\<exists>stp ln rn.
  2490       show " \<exists>stp ln rn.
  2476         steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev (<Suc a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  2491           steps0 (Suc 0, Bk # Bk \<up> m @ Oc \<up> Suc (Suc a) @ Bk # Bk # ires, Bk # Oc \<up> Suc rs @ Bk \<up> n) t_wcode_main stp =
  2477         (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<Suc a>) + rs * 2 ^ Suc a\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2492           (0, Bk # ires, Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> (bl_bin (Oc \<up> Suc (Suc a)) + rs * 2 ^ Suc a) @ Bk \<up> rn)"
  2478       proof -
  2493       proof -
  2479         have "\<exists>stp ln rn.
  2494         have "\<exists>stp ln rn.
  2480           steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev (<Suc a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  2495           steps0 (Suc 0, Bk # Bk\<up>(m) @ rev (<Suc a>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
  2481           (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2496           (Suc 0, Bk # Bk\<up>(ln) @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc (2 * rs + 2)) @ Bk\<up>(rn))"
  2482           apply(simp add: tape_of_nat)
  2497           apply(simp add: tape_of_nat)
  2483           using wcode_double_case[of m "Oc\<^bsup>a\<^esup> @ Bk # Bk # ires" rs n]
  2498           using wcode_double_case[of m "Oc\<up>(a) @ Bk # Bk # ires" rs n]
  2484           apply(simp add: exp_ind_def)
  2499           apply(simp add: replicate_Suc)
  2485           done
  2500           done
  2486         from this obtain stpa lna rna where stp1:  
  2501         from this obtain stpa lna rna where stp1:  
  2487           "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev (<Suc a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa =
  2502           "steps0 (Suc 0, Bk # Bk\<up>(m) @ rev (<Suc a>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stpa =
  2488           (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast
  2503           (Suc 0, Bk # Bk\<up>(lna) @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc (2 * rs + 2)) @ Bk\<up>(rna))" by blast
  2489         moreover have 
  2504         moreover have 
  2490           "\<exists>stp ln rn.
  2505           "\<exists>stp ln rn.
  2491           steps (Suc 0,  Bk # Bk\<^bsup>lna\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp =
  2506           steps0 (Suc 0,  Bk # Bk\<up>(lna) @ rev (<a::nat>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc (2 * rs + 2)) @ Bk\<up>(rna)) t_wcode_main stp =
  2492           (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<a>) + (2*rs + 2)  * 2 ^ a\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2507           (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin (<a>) + (2*rs + 2)  * 2 ^ a) @ Bk\<up>(rn))"
  2493           using ind2[of lna ires "2*rs + 2" rna] by simp   
  2508           using ind2[of lna ires "2*rs + 2" rna] by(simp add: tape_of_nl_abv tape_of_nat_abv)   
  2494         from this obtain stpb lnb rnb where stp2:  
  2509         from this obtain stpb lnb rnb where stp2:  
  2495           "steps (Suc 0,  Bk # Bk\<^bsup>lna\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb =
  2510           "steps0 (Suc 0,  Bk # Bk\<up>(lna) @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc (2 * rs + 2)) @ Bk\<up>(rna)) t_wcode_main stpb =
  2496           (0, Bk # ires, Bk # Oc # Bk\<^bsup>lnb\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<a>) + (2*rs + 2)  * 2 ^ a\<^esup> @ Bk\<^bsup>rnb\<^esup>)"
  2511           (0, Bk # ires, Bk # Oc # Bk\<up>(lnb) @ Bk # Bk # Oc\<up>(bl_bin (<a>) + (2*rs + 2)  * 2 ^ a) @ Bk\<up>(rnb))"
  2497           by blast
  2512           by blast
  2498         from stp1 and stp2 show "?thesis"
  2513         from stp1 and stp2 show "?thesis"
  2499           apply(rule_tac x = "stpa + stpb" in exI,
  2514           apply(rule_tac x = "stpa + stpb" in exI,
  2500             rule_tac x = lnb in exI, rule_tac x = rnb in exI, simp)
  2515             rule_tac x = lnb in exI, rule_tac x = rnb in exI, simp add: tape_of_nat_abv)
  2501           apply(simp add: steps_add bl_bin_nat_Suc exponent_def)
  2516           apply(simp add:  bl_bin.simps replicate_Suc)
       
  2517           apply(auto)
  2502           done
  2518           done
  2503       qed
  2519       qed
  2504     qed
  2520     qed
  2505   next
  2521   next
  2506     fix aa list
  2522     fix aa list
  2507     assume g: "Suc x = length args" "args \<noteq> []" "lm = <args>" "args = xs @ [a::nat]" "xs = (aa::nat) # list"
  2523     assume g: "Suc x = length args" "args \<noteq> []" "lm = <args>" "args = xs @ [a::nat]" "xs = (aa::nat) # list"
  2508     thus "\<exists>stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev lm @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  2524     thus "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
  2509       (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin lm + rs * 2 ^ (length lm - 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2525       (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin lm + rs * 2 ^ (length lm - 1)) @ Bk\<up>(rn))"
  2510     proof(induct a arbitrary: m n rs args lm, simp_all add: tape_of_nl_rev, 
  2526     proof(induct a arbitrary: m n rs args lm, simp_all add: tape_of_nl_rev, 
  2511         simp only: tape_of_nl_cons_app1, simp)
  2527         simp only: tape_of_nl_cons_app1, simp)
  2512       fix m n rs args lm
  2528       fix m n rs args lm
  2513       have "\<exists>stp ln rn.
  2529       have "\<exists>stp ln rn.
  2514         steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # rev (<(aa::nat) # list>) @ Bk # Bk # ires,
  2530         steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # rev (<(aa::nat) # list>) @ Bk # Bk # ires,
  2515         Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  2531         Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
  2516         (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ rev (<aa # list>) @ Bk # Bk # ires, 
  2532         (Suc 0, Bk # Bk\<up>(ln) @ rev (<aa # list>) @ Bk # Bk # ires, 
  2517         Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2533         Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rn))"
  2518         proof(simp add: tape_of_nl_rev)
  2534         proof(simp add: tape_of_nl_rev)
  2519           have "\<exists> xs. (<rev list @ [aa]>) = Oc # xs" by auto           
  2535           have "\<exists> xs. (<rev list @ [aa]>) = Oc # xs" by auto           
  2520           from this obtain xs where "(<rev list @ [aa]>) = Oc # xs" ..
  2536           from this obtain xs where "(<rev list @ [aa]>) = Oc # xs" ..
  2521           thus "\<exists>stp ln rn.
  2537           thus "\<exists>stp ln rn.
  2522             steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # <rev list @ [aa]> @ Bk # Bk # ires,
  2538             steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # <rev list @ [aa]> @ Bk # Bk # ires,
  2523             Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  2539             Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
  2524             (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ <rev list @ [aa]> @ Bk # Bk # ires, Bk # Oc\<^bsup>5 + 4 * rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2540             (Suc 0, Bk # Bk\<up>(ln) @ <rev list @ [aa]> @ Bk # Bk # ires, Bk # Oc\<up>(5 + 4 * rs) @ Bk\<up>(rn))"
  2525             apply(simp)
  2541             apply(simp)
  2526             using wcode_fourtimes_case[of m "xs @ Bk # Bk # ires" rs n]
  2542             using wcode_fourtimes_case[of m "xs @ Bk # Bk # ires" rs n]
  2527             apply(simp)
  2543             apply(simp)
  2528             done
  2544             done
  2529         qed
  2545         qed
  2530       from this obtain stpa lna rna where stp1:
  2546       from this obtain stpa lna rna where stp1:
  2531         "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # rev (<aa # list>) @ Bk # Bk # ires,
  2547         "steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # rev (<aa # list>) @ Bk # Bk # ires,
  2532         Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa =
  2548         Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stpa =
  2533         (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<aa # list>) @ Bk # Bk # ires, 
  2549         (Suc 0, Bk # Bk\<up>(lna) @ rev (<aa # list>) @ Bk # Bk # ires, 
  2534         Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast
  2550         Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rna))" by blast
  2535       from g have 
  2551       from g have 
  2536         "\<exists> stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, 
  2552         "\<exists> stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(lna) @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, 
  2537         Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp = (0, Bk # ires, 
  2553         Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rna)) t_wcode_main stp = (0, Bk # ires, 
  2538         Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<aa#list>)+ (4*rs + 4) * 2^(length (<aa#list>) - 1) \<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2554         Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin (<aa#list>)+ (4*rs + 4) * 2^(length (<aa#list>) - 1) ) @ Bk\<up>(rn))"
  2539          apply(rule_tac args = "(aa::nat)#list" in ind, simp_all)
  2555          apply(rule_tac args = "(aa::nat)#list" in ind, simp_all)
  2540          done
  2556          done
  2541        from this obtain stpb lnb rnb where stp2:
  2557        from this obtain stpb lnb rnb where stp2:
  2542          "steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, 
  2558          "steps0 (Suc 0, Bk # Bk\<up>(lna) @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, 
  2543          Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb = (0, Bk # ires, 
  2559          Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rna)) t_wcode_main stpb = (0, Bk # ires, 
  2544          Bk # Oc # Bk\<^bsup>lnb\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<aa#list>)+ (4*rs + 4) * 2^(length (<aa#list>) - 1) \<^esup> @ Bk\<^bsup>rnb\<^esup>)"
  2560          Bk # Oc # Bk\<up>(lnb) @ Bk # Bk # Oc\<up>(bl_bin (<aa#list>)+ (4*rs + 4) * 2^(length (<aa#list>) - 1) ) @ Bk\<up>(rnb))"
  2545          by blast
  2561          by blast
  2546        from stp1 and stp2 and h
  2562        from stp1 and stp2 and h
  2547        show "\<exists>stp ln rn.
  2563        show "\<exists>stp ln rn.
  2548          steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # <rev list @ [aa]> @ Bk # Bk # ires,
  2564          steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # <rev list @ [aa]> @ Bk # Bk # ires,
  2549          Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  2565          Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
  2550          (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk #
  2566          (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk #
  2551          Bk # Oc\<^bsup>bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # <list @ [0]>) + rs * (2 * 2 ^ (aa + length (<list @ [0]>)))\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2567          Bk # Oc\<up>(bl_bin (Oc\<up>(Suc aa) @ Bk # <list @ [0]>) + rs * (2 * 2 ^ (aa + length (<list @ [0]>)))) @ Bk\<up>(rn))"
  2552          apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI,
  2568          apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI,
  2553            rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_rev)
  2569            rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_rev)
  2554          done
  2570          done
  2555      next
  2571      next
  2556        fix ab m n rs args lm
  2572        fix ab m n rs args lm
  2557        assume ind2:
  2573        assume ind2:
  2558          "\<And> m n rs args lm.
  2574          "\<And> m n rs args lm.
  2559          \<lbrakk>lm = <aa # list @ [ab]>; args = aa # list @ [ab]\<rbrakk>
  2575          \<lbrakk>lm = <aa # list @ [ab]>; args = aa # list @ [ab]\<rbrakk>
  2560          \<Longrightarrow> \<exists>stp ln rn.
  2576          \<Longrightarrow> \<exists>stp ln rn.
  2561          steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ <ab # rev list @ [aa]> @ Bk # Bk # ires,
  2577          steps0 (Suc 0, Bk # Bk\<up>(m) @ <ab # rev list @ [aa]> @ Bk # Bk # ires,
  2562          Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  2578          Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
  2563          (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk #
  2579          (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk #
  2564          Bk # Oc\<^bsup>bl_bin (<aa # list @ [ab]>) + rs * 2 ^ (length (<aa # list @ [ab]>) - Suc 0)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2580          Bk # Oc\<up>(bl_bin (<aa # list @ [ab]>) + rs * 2 ^ (length (<aa # list @ [ab]>) - Suc 0)) @ Bk\<up>(rn))"
  2565          and k: "args = aa # list @ [Suc ab]" "lm = <aa # list @ [Suc ab]>"
  2581          and k: "args = aa # list @ [Suc ab]" "lm = <aa # list @ [Suc ab]>"
  2566        show "\<exists>stp ln rn.
  2582        show "\<exists>stp ln rn.
  2567          steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ <Suc ab # rev list @ [aa]> @ Bk # Bk # ires,
  2583          steps0 (Suc 0, Bk # Bk\<up>(m) @ <Suc ab # rev list @ [aa]> @ Bk # Bk # ires,
  2568          Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  2584          Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
  2569          (0, Bk # ires,Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # 
  2585          (0, Bk # ires,Bk # Oc # Bk\<up>(ln) @ Bk # 
  2570          Bk # Oc\<^bsup>bl_bin (<aa # list @ [Suc ab]>) + rs * 2 ^ (length (<aa # list @ [Suc ab]>) - Suc 0)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2586          Bk # Oc\<up>(bl_bin (<aa # list @ [Suc ab]>) + rs * 2 ^ (length (<aa # list @ [Suc ab]>) - Suc 0)) @ Bk\<up>(rn))"
  2571        proof(simp add: tape_of_nl_cons_app1)
  2587        proof(simp add: tape_of_nl_cons_app1)
  2572          have "\<exists>stp ln rn.
  2588          have "\<exists>stp ln rn.
  2573            steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc\<^bsup>Suc (Suc ab)\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, 
  2589            steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc\<up>(Suc (Suc ab)) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, 
  2574            Bk # Oc # Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp
  2590            Bk # Oc # Oc\<up>(rs) @ Bk\<up>(n)) t_wcode_main stp
  2575            = (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc\<^bsup>Suc ab\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
  2591            = (Suc 0, Bk # Bk\<up>(ln) @ Oc\<up>(Suc ab) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
  2576            Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2592            Bk # Oc\<up>(Suc (2*rs + 2)) @ Bk\<up>(rn))"
  2577            using wcode_double_case[of m "Oc\<^bsup>ab\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires"
  2593            using wcode_double_case[of m "Oc\<up>(ab) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires"
  2578                                       rs n]
  2594                                       rs n]
  2579            apply(simp add: exp_ind_def)
  2595            apply(simp add: replicate_Suc)
  2580            done
  2596            done
  2581          from this obtain stpa lna rna where stp1:
  2597          from this obtain stpa lna rna where stp1:
  2582            "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc\<^bsup>Suc (Suc ab)\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, 
  2598            "steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc\<up>(Suc (Suc ab)) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, 
  2583            Bk # Oc # Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa
  2599            Bk # Oc # Oc\<up>(rs) @ Bk\<up>(n)) t_wcode_main stpa
  2584            = (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ Oc\<^bsup>Suc ab\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
  2600            = (Suc 0, Bk # Bk\<up>(lna) @ Oc\<up>(Suc ab) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
  2585            Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast
  2601            Bk # Oc\<up>(Suc (2*rs + 2)) @ Bk\<up>(rna))" by blast
  2586          from k have 
  2602          from k have 
  2587            "\<exists> stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ <ab # rev list @ [aa]> @ Bk # Bk # ires,
  2603            "\<exists> stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(lna) @ <ab # rev list @ [aa]> @ Bk # Bk # ires,
  2588            Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp
  2604            Bk # Oc\<up>(Suc (2*rs + 2)) @ Bk\<up>(rna)) t_wcode_main stp
  2589            = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk #
  2605            = (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk #
  2590            Bk # Oc\<^bsup>bl_bin (<aa # list @ [ab]> ) +  (2*rs + 2)* 2^(length (<aa # list @ [ab]>) - Suc 0)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2606            Bk # Oc\<up>(bl_bin (<aa # list @ [ab]> ) +  (2*rs + 2)* 2^(length (<aa # list @ [ab]>) - Suc 0)) @ Bk\<up>(rn))"
  2591            apply(rule_tac ind2, simp_all)
  2607            apply(rule_tac ind2, simp_all)
  2592            done
  2608            done
  2593          from this obtain stpb lnb rnb where stp2: 
  2609          from this obtain stpb lnb rnb where stp2: 
  2594            "steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @  <ab # rev list @ [aa]> @ Bk # Bk # ires,
  2610            "steps0 (Suc 0, Bk # Bk\<up>(lna) @  <ab # rev list @ [aa]> @ Bk # Bk # ires,
  2595            Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb
  2611            Bk # Oc\<up>(Suc (2*rs + 2)) @ Bk\<up>(rna)) t_wcode_main stpb
  2596            = (0, Bk # ires, Bk # Oc # Bk\<^bsup>lnb\<^esup> @ Bk #
  2612            = (0, Bk # ires, Bk # Oc # Bk\<up>(lnb) @ Bk #
  2597            Bk # Oc\<^bsup>bl_bin (<aa # list @ [ab]> ) +  (2*rs + 2)* 2^(length (<aa # list @ [ab]>) - Suc 0)\<^esup> @ Bk\<^bsup>rnb\<^esup>)" 
  2613            Bk # Oc\<up>(bl_bin (<aa # list @ [ab]> ) +  (2*rs + 2)* 2^(length (<aa # list @ [ab]>) - Suc 0)) @ Bk\<up>(rnb))" 
  2598            by blast
  2614            by blast
  2599          from stp1 and stp2 show 
  2615          from stp1 and stp2 show 
  2600            "\<exists>stp ln rn.
  2616            "\<exists>stp ln rn.
  2601            steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc\<^bsup>Suc (Suc ab)\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
  2617            steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc\<up>(Suc (Suc ab)) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
  2602            Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp =
  2618            Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
  2603            (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # 
  2619            (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # 
  2604            Oc\<^bsup>bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # <list @ [Suc ab]>) + rs * (2 * 2 ^ (aa + length (<list @ [Suc ab]>)))\<^esup> 
  2620            Oc\<up>(bl_bin (Oc\<up>(Suc aa) @ Bk # <list @ [Suc ab]>) + rs * (2 * 2 ^ (aa + length (<list @ [Suc ab]>)))) 
  2605            @ Bk\<^bsup>rn\<^esup>)"
  2621            @ Bk\<up>(rn))"
  2606            apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI,
  2622            apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI,
  2607              rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_cons_app1 exp_ind_def)
  2623              rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_cons_app1 replicate_Suc)
  2608            done
  2624            done
  2609        qed
  2625        qed
  2610      qed
  2626      qed
  2611    qed
  2627    qed
  2612  qed
  2628  qed
  2613 
  2629 
  2614 
  2630 
  2615          
  2631 definition t_wcode_prepare :: "instr list"
  2616 (* turing_shift can be used*)
       
  2617 term t_wcode_main
       
  2618 definition t_wcode_prepare :: "tprog"
       
  2619   where
  2632   where
  2620   "t_wcode_prepare \<equiv> 
  2633   "t_wcode_prepare \<equiv> 
  2621          [(W1, 2), (L, 1), (L, 3), (R, 2), (R, 4), (W0, 3),
  2634          [(W1, 2), (L, 1), (L, 3), (R, 2), (R, 4), (W0, 3),
  2622           (R, 4), (R, 5), (R, 6), (R, 5), (R, 7), (R, 5),
  2635           (R, 4), (R, 5), (R, 6), (R, 5), (R, 7), (R, 5),
  2623           (W1, 7), (L, 0)]"
  2636           (W1, 7), (L, 0)]"
  2624 
  2637 
  2625 fun wprepare_add_one :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2638 fun wprepare_add_one :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2626   where
  2639   where
  2627   "wprepare_add_one m lm (l, r) = 
  2640   "wprepare_add_one m lm (l, r) = 
  2628       (\<exists> rn. l = [] \<and>
  2641       (\<exists> rn. l = [] \<and>
  2629                (r = <m # lm> @ Bk\<^bsup>rn\<^esup> \<or> 
  2642                (r = <m # lm> @ Bk\<up>(rn) \<or> 
  2630                 r = Bk # <m # lm> @ Bk\<^bsup>rn\<^esup>))"
  2643                 r = Bk # <m # lm> @ Bk\<up>(rn)))"
  2631 
  2644 
  2632 fun wprepare_goto_first_end :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2645 fun wprepare_goto_first_end :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2633   where
  2646   where
  2634   "wprepare_goto_first_end m lm (l, r) = 
  2647   "wprepare_goto_first_end m lm (l, r) = 
  2635       (\<exists> ml mr rn. l = Oc\<^bsup>ml\<^esup> \<and>
  2648       (\<exists> ml mr rn. l = Oc\<up>(ml) \<and>
  2636                       r = Oc\<^bsup>mr\<^esup> @ Bk # <lm> @ Bk\<^bsup>rn\<^esup> \<and>
  2649                       r = Oc\<up>(mr) @ Bk # <lm> @ Bk\<up>(rn) \<and>
  2637                       ml + mr = Suc (Suc m))"
  2650                       ml + mr = Suc (Suc m))"
  2638 
  2651 
  2639 fun wprepare_erase :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow>  bool"
  2652 fun wprepare_erase :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow>  bool"
  2640   where
  2653   where
  2641   "wprepare_erase m lm (l, r) = 
  2654   "wprepare_erase m lm (l, r) = 
  2642      (\<exists> rn. l = Oc\<^bsup>Suc m\<^esup> \<and> 
  2655      (\<exists> rn. l = Oc\<up>(Suc m) \<and> 
  2643                tl r = Bk # <lm> @ Bk\<^bsup>rn\<^esup>)"
  2656                tl r = Bk # <lm> @ Bk\<up>(rn))"
  2644 
  2657 
  2645 fun wprepare_goto_start_pos_B :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2658 fun wprepare_goto_start_pos_B :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2646   where
  2659   where
  2647   "wprepare_goto_start_pos_B m lm (l, r) = 
  2660   "wprepare_goto_start_pos_B m lm (l, r) = 
  2648      (\<exists> rn. l = Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  2661      (\<exists> rn. l = Bk # Oc\<up>(Suc m) \<and>
  2649                r = Bk # <lm> @ Bk\<^bsup>rn\<^esup>)"
  2662                r = Bk # <lm> @ Bk\<up>(rn))"
  2650 
  2663 
  2651 fun wprepare_goto_start_pos_O :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2664 fun wprepare_goto_start_pos_O :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2652   where
  2665   where
  2653   "wprepare_goto_start_pos_O m lm (l, r) = 
  2666   "wprepare_goto_start_pos_O m lm (l, r) = 
  2654      (\<exists> rn. l = Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  2667      (\<exists> rn. l = Bk # Bk # Oc\<up>(Suc m) \<and>
  2655                r = <lm> @ Bk\<^bsup>rn\<^esup>)"
  2668                r = <lm> @ Bk\<up>(rn))"
  2656 
  2669 
  2657 fun wprepare_goto_start_pos :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2670 fun wprepare_goto_start_pos :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2658   where
  2671   where
  2659   "wprepare_goto_start_pos m lm (l, r) = 
  2672   "wprepare_goto_start_pos m lm (l, r) = 
  2660        (wprepare_goto_start_pos_B m lm (l, r) \<or>
  2673        (wprepare_goto_start_pos_B m lm (l, r) \<or>
  2661         wprepare_goto_start_pos_O m lm (l, r))"
  2674         wprepare_goto_start_pos_O m lm (l, r))"
  2662 
  2675 
  2663 fun wprepare_loop_start_on_rightmost :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2676 fun wprepare_loop_start_on_rightmost :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2664   where
  2677   where
  2665   "wprepare_loop_start_on_rightmost m lm (l, r) = 
  2678   "wprepare_loop_start_on_rightmost m lm (l, r) = 
  2666      (\<exists> rn mr. rev l @ r = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm> @ Bk\<^bsup>rn\<^esup> \<and> l \<noteq> [] \<and>
  2679      (\<exists> rn mr. rev l @ r = Oc\<up>(Suc m) @ Bk # Bk # <lm> @ Bk\<up>(rn) \<and> l \<noteq> [] \<and>
  2667                        r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  2680                        r = Oc\<up>(mr) @ Bk\<up>(rn))"
  2668 
  2681 
  2669 fun wprepare_loop_start_in_middle :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2682 fun wprepare_loop_start_in_middle :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2670   where
  2683   where
  2671   "wprepare_loop_start_in_middle m lm (l, r) =
  2684   "wprepare_loop_start_in_middle m lm (l, r) =
  2672      (\<exists> rn (mr:: nat) (lm1::nat list). 
  2685      (\<exists> rn (mr:: nat) (lm1::nat list). 
  2673   rev l @ r = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm> @ Bk\<^bsup>rn\<^esup> \<and> l \<noteq> [] \<and>
  2686   rev l @ r = Oc\<up>(Suc m) @ Bk # Bk # <lm> @ Bk\<up>(rn) \<and> l \<noteq> [] \<and>
  2674   r = Oc\<^bsup>mr\<^esup> @ Bk # <lm1> @ Bk\<^bsup>rn\<^esup> \<and> lm1 \<noteq> [])"
  2687   r = Oc\<up>(mr) @ Bk # <lm1> @ Bk\<up>(rn) \<and> lm1 \<noteq> [])"
  2675 
  2688 
  2676 fun wprepare_loop_start :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2689 fun wprepare_loop_start :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2677   where
  2690   where
  2678   "wprepare_loop_start m lm (l, r) = (wprepare_loop_start_on_rightmost m lm (l, r) \<or> 
  2691   "wprepare_loop_start m lm (l, r) = (wprepare_loop_start_on_rightmost m lm (l, r) \<or> 
  2679                                       wprepare_loop_start_in_middle m lm (l, r))"
  2692                                       wprepare_loop_start_in_middle m lm (l, r))"
  2680 
  2693 
  2681 fun wprepare_loop_goon_on_rightmost :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2694 fun wprepare_loop_goon_on_rightmost :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2682   where
  2695   where
  2683   "wprepare_loop_goon_on_rightmost m lm (l, r) = 
  2696   "wprepare_loop_goon_on_rightmost m lm (l, r) = 
  2684      (\<exists> rn. l = Bk # <rev lm> @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  2697      (\<exists> rn. l = Bk # <rev lm> @ Bk # Bk # Oc\<up>(Suc m) \<and>
  2685                r = Bk\<^bsup>rn\<^esup>)"
  2698                r = Bk\<up>(rn))"
  2686 
  2699 
  2687 fun wprepare_loop_goon_in_middle :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2700 fun wprepare_loop_goon_in_middle :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2688   where
  2701   where
  2689   "wprepare_loop_goon_in_middle m lm (l, r) = 
  2702   "wprepare_loop_goon_in_middle m lm (l, r) = 
  2690      (\<exists> rn (mr:: nat) (lm1::nat list). 
  2703      (\<exists> rn (mr:: nat) (lm1::nat list). 
  2691   rev l @ r = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm> @ Bk\<^bsup>rn\<^esup> \<and> l \<noteq> [] \<and>
  2704   rev l @ r = Oc\<up>(Suc m) @ Bk # Bk # <lm> @ Bk\<up>(rn) \<and> l \<noteq> [] \<and>
  2692                      (if lm1 = [] then r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> 
  2705                      (if lm1 = [] then r = Oc\<up>(mr) @ Bk\<up>(rn) 
  2693                      else r = Oc\<^bsup>mr\<^esup> @ Bk # <lm1> @ Bk\<^bsup>rn\<^esup>) \<and> mr > 0)"
  2706                      else r = Oc\<up>(mr) @ Bk # <lm1> @ Bk\<up>(rn)) \<and> mr > 0)"
  2694 
  2707 
  2695 fun wprepare_loop_goon :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2708 fun wprepare_loop_goon :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2696   where
  2709   where
  2697   "wprepare_loop_goon m lm (l, r) = 
  2710   "wprepare_loop_goon m lm (l, r) = 
  2698               (wprepare_loop_goon_in_middle m lm (l, r) \<or> 
  2711               (wprepare_loop_goon_in_middle m lm (l, r) \<or> 
  2699                wprepare_loop_goon_on_rightmost m lm (l, r))"
  2712                wprepare_loop_goon_on_rightmost m lm (l, r))"
  2700 
  2713 
  2701 fun wprepare_add_one2 :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2714 fun wprepare_add_one2 :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2702   where
  2715   where
  2703   "wprepare_add_one2 m lm (l, r) =
  2716   "wprepare_add_one2 m lm (l, r) =
  2704           (\<exists> rn. l = Bk # Bk # <rev lm> @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  2717           (\<exists> rn. l = Bk # Bk # <rev lm> @ Bk # Bk # Oc\<up>(Suc m) \<and>
  2705                (r = [] \<or> tl r = Bk\<^bsup>rn\<^esup>))"
  2718                (r = [] \<or> tl r = Bk\<up>(rn)))"
  2706 
  2719 
  2707 fun wprepare_stop :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2720 fun wprepare_stop :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2708   where
  2721   where
  2709   "wprepare_stop m lm (l, r) = 
  2722   "wprepare_stop m lm (l, r) = 
  2710          (\<exists> rn. l = Bk # <rev lm> @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  2723          (\<exists> rn. l = Bk # <rev lm> @ Bk # Bk # Oc\<up>(Suc m) \<and>
  2711                r = Bk # Oc # Bk\<^bsup>rn\<^esup>)"
  2724                r = Bk # Oc # Bk\<up>(rn))"
  2712 
  2725 
  2713 fun wprepare_inv :: "nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2726 fun wprepare_inv :: "nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
  2714   where
  2727   where
  2715   "wprepare_inv st m lm (l, r) = 
  2728   "wprepare_inv st m lm (l, r) = 
  2716         (if st = 0 then wprepare_stop m lm (l, r) 
  2729         (if st = 0 then wprepare_stop m lm (l, r) 
  2721          else if st = 5 then wprepare_loop_start m lm (l, r)
  2734          else if st = 5 then wprepare_loop_start m lm (l, r)
  2722          else if st = 6 then wprepare_loop_goon m lm (l, r)
  2735          else if st = 6 then wprepare_loop_goon m lm (l, r)
  2723          else if st = 7 then wprepare_add_one2 m lm (l, r)
  2736          else if st = 7 then wprepare_add_one2 m lm (l, r)
  2724          else False)"
  2737          else False)"
  2725 
  2738 
  2726 fun wprepare_stage :: "t_conf \<Rightarrow> nat"
  2739 fun wprepare_stage :: "config \<Rightarrow> nat"
  2727   where
  2740   where
  2728   "wprepare_stage (st, l, r) = 
  2741   "wprepare_stage (st, l, r) = 
  2729       (if st \<ge> 1 \<and> st \<le> 4 then 3
  2742       (if st \<ge> 1 \<and> st \<le> 4 then 3
  2730        else if st = 5 \<or> st = 6 then 2
  2743        else if st = 5 \<or> st = 6 then 2
  2731        else 1)"
  2744        else 1)"
  2732 
  2745 
  2733 fun wprepare_state :: "t_conf \<Rightarrow> nat"
  2746 fun wprepare_state :: "config \<Rightarrow> nat"
  2734   where
  2747   where
  2735   "wprepare_state (st, l, r) = 
  2748   "wprepare_state (st, l, r) = 
  2736        (if st = 1 then 4
  2749        (if st = 1 then 4
  2737         else if st = Suc (Suc 0) then 3
  2750         else if st = Suc (Suc 0) then 3
  2738         else if st = Suc (Suc (Suc 0)) then 2
  2751         else if st = Suc (Suc (Suc 0)) then 2
  2739         else if st = 4 then 1
  2752         else if st = 4 then 1
  2740         else if st = 7 then 2
  2753         else if st = 7 then 2
  2741         else 0)"
  2754         else 0)"
  2742 
  2755 
  2743 fun wprepare_step :: "t_conf \<Rightarrow> nat"
  2756 fun wprepare_step :: "config \<Rightarrow> nat"
  2744   where
  2757   where
  2745   "wprepare_step (st, l, r) = 
  2758   "wprepare_step (st, l, r) = 
  2746       (if st = 1 then (if hd r = Oc then Suc (length l)
  2759       (if st = 1 then (if hd r = Oc then Suc (length l)
  2747                        else 0)
  2760                        else 0)
  2748        else if st = Suc (Suc 0) then length r
  2761        else if st = Suc (Suc 0) then length r
  2753        else if st = 6 then (if r = [] then 0 else Suc (length r))
  2766        else if st = 6 then (if r = [] then 0 else Suc (length r))
  2754        else if st = 7 then (if (r \<noteq> [] \<and> hd r = Oc) then 0
  2767        else if st = 7 then (if (r \<noteq> [] \<and> hd r = Oc) then 0
  2755                             else 1)
  2768                             else 1)
  2756        else 0)"
  2769        else 0)"
  2757 
  2770 
  2758 fun wcode_prepare_measure :: "t_conf \<Rightarrow> nat \<times> nat \<times> nat"
  2771 fun wcode_prepare_measure :: "config \<Rightarrow> nat \<times> nat \<times> nat"
  2759   where
  2772   where
  2760   "wcode_prepare_measure (st, l, r) = 
  2773   "wcode_prepare_measure (st, l, r) = 
  2761      (wprepare_stage (st, l, r), 
  2774      (wprepare_stage (st, l, r), 
  2762       wprepare_state (st, l, r), 
  2775       wprepare_state (st, l, r), 
  2763       wprepare_step (st, l, r))"
  2776       wprepare_step (st, l, r))"
  2764 
  2777 
  2765 definition wcode_prepare_le :: "(t_conf \<times> t_conf) set"
  2778 definition wcode_prepare_le :: "(config \<times> config) set"
  2766   where "wcode_prepare_le \<equiv> (inv_image lex_triple wcode_prepare_measure)"
  2779   where "wcode_prepare_le \<equiv> (inv_image lex_triple wcode_prepare_measure)"
  2767 
  2780 
  2768 lemma [intro]: "wf lex_pair"
  2781 lemma [intro]: "wf lex_pair"
  2769 by(auto intro:wf_lex_prod simp:lex_pair_def)
  2782 by(auto intro:wf_lex_prod simp:lex_pair_def)
  2770 
  2783 
  2771 lemma wf_wcode_prepare_le[intro]: "wf wcode_prepare_le"
  2784 lemma wf_wcode_prepare_le[intro]: "wf wcode_prepare_le"
  2772 by(auto intro:wf_inv_image simp: wcode_prepare_le_def 
  2785 by(auto intro:wf_inv_image simp: wcode_prepare_le_def 
  2773            recursive.lex_triple_def)
  2786            lex_triple_def)
  2774 
  2787 
  2775 declare wprepare_add_one.simps[simp del] wprepare_goto_first_end.simps[simp del]
  2788 declare wprepare_add_one.simps[simp del] wprepare_goto_first_end.simps[simp del]
  2776         wprepare_erase.simps[simp del] wprepare_goto_start_pos.simps[simp del]
  2789         wprepare_erase.simps[simp del] wprepare_goto_start_pos.simps[simp del]
  2777         wprepare_loop_start.simps[simp del] wprepare_loop_goon.simps[simp del]
  2790         wprepare_loop_start.simps[simp del] wprepare_loop_goon.simps[simp del]
  2778         wprepare_add_one2.simps[simp del]
  2791         wprepare_add_one2.simps[simp del]
  2806 lemma [simp]: "fetch t_wcode_prepare (Suc (Suc (Suc 0))) Oc = (W0, 3)"
  2819 lemma [simp]: "fetch t_wcode_prepare (Suc (Suc (Suc 0))) Oc = (W0, 3)"
  2807 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
  2820 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
  2808 done
  2821 done
  2809 
  2822 
  2810 lemma [simp]: "fetch t_wcode_prepare 4 Bk = (R, 4)"
  2823 lemma [simp]: "fetch t_wcode_prepare 4 Bk = (R, 4)"
  2811 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
  2824 apply(subgoal_tac "4 = Suc 3")
       
  2825 apply(simp_all only: fetch.simps t_wcode_prepare_def nth_of.simps)
       
  2826 apply(auto)
  2812 done
  2827 done
  2813 
  2828 
  2814 lemma [simp]: "fetch t_wcode_prepare 4 Oc = (R, 5)"
  2829 lemma [simp]: "fetch t_wcode_prepare 4 Oc = (R, 5)"
  2815 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
  2830 apply(subgoal_tac "4 = Suc 3")
  2816 done
  2831 apply(simp_all only: fetch.simps t_wcode_prepare_def nth_of.simps)
       
  2832 apply(auto)
       
  2833 done
       
  2834 
  2817 
  2835 
  2818 lemma [simp]: "fetch t_wcode_prepare 5 Oc = (R, 5)"
  2836 lemma [simp]: "fetch t_wcode_prepare 5 Oc = (R, 5)"
  2819 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
  2837 apply(subgoal_tac "5 = Suc 4")
       
  2838 apply(simp_all only: fetch.simps t_wcode_prepare_def nth_of.simps)
       
  2839 apply(auto)
  2820 done
  2840 done
  2821 
  2841 
  2822 lemma [simp]: "fetch t_wcode_prepare 5 Bk = (R, 6)"
  2842 lemma [simp]: "fetch t_wcode_prepare 5 Bk = (R, 6)"
  2823 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
  2843 apply(subgoal_tac "5 = Suc 4")
       
  2844 apply(simp_all only: fetch.simps t_wcode_prepare_def nth_of.simps)
       
  2845 apply(auto)
  2824 done
  2846 done
  2825 
  2847 
  2826 lemma [simp]: "fetch t_wcode_prepare 6 Oc = (R, 5)"
  2848 lemma [simp]: "fetch t_wcode_prepare 6 Oc = (R, 5)"
  2827 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
  2849 apply(subgoal_tac "6 = Suc 5")
       
  2850 apply(simp_all only: fetch.simps t_wcode_prepare_def nth_of.simps)
       
  2851 apply(auto)
  2828 done
  2852 done
  2829 
  2853 
  2830 lemma [simp]: "fetch t_wcode_prepare 6 Bk = (R, 7)"
  2854 lemma [simp]: "fetch t_wcode_prepare 6 Bk = (R, 7)"
  2831 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
  2855 apply(subgoal_tac "6 = Suc 5")
       
  2856 apply(simp_all only: fetch.simps t_wcode_prepare_def nth_of.simps)
       
  2857 apply(auto)
  2832 done
  2858 done
  2833 
  2859 
  2834 lemma [simp]: "fetch t_wcode_prepare 7 Oc = (L, 0)"
  2860 lemma [simp]: "fetch t_wcode_prepare 7 Oc = (L, 0)"
  2835 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
  2861 apply(subgoal_tac "7 = Suc 6")
       
  2862 apply(simp_all only: fetch.simps t_wcode_prepare_def nth_of.simps)
       
  2863 apply(auto)
  2836 done
  2864 done
  2837 
  2865 
  2838 lemma [simp]: "fetch t_wcode_prepare 7 Bk = (W1, 7)"
  2866 lemma [simp]: "fetch t_wcode_prepare 7 Bk = (W1, 7)"
  2839 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps)
  2867 apply(subgoal_tac "7 = Suc 6")
  2840 done
  2868 apply(simp_all only: fetch.simps t_wcode_prepare_def nth_of.simps)
  2841 
  2869 apply(auto)
  2842 lemma tape_of_nl_not_null: "lm \<noteq> [] \<Longrightarrow> <lm::nat list> \<noteq> []"
       
  2843 apply(case_tac lm, auto)
       
  2844 apply(case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
       
  2845 done
  2870 done
  2846 
  2871 
  2847 lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_add_one m lm (b, []) = False"
  2872 lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_add_one m lm (b, []) = False"
  2848 apply(simp add: wprepare_invs)
  2873 apply(simp add: wprepare_invs)
  2849 apply(simp add: tape_of_nl_not_null)
       
  2850 done
  2874 done
  2851 
  2875 
  2852 lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_goto_first_end m lm (b, []) = False"
  2876 lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_goto_first_end m lm (b, []) = False"
  2853 apply(simp add: wprepare_invs)
  2877 apply(simp add: wprepare_invs)
  2854 done
  2878 done
  2855 
  2879 
  2856 lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_erase m lm (b, []) = False"
  2880 lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_erase m lm (b, []) = False"
  2857 apply(simp add: wprepare_invs)
  2881 apply(simp add: wprepare_invs)
  2858 done
  2882 done
  2859 
  2883 
  2860 
       
  2861 
       
  2862 lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_goto_start_pos m lm (b, []) = False"
  2884 lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_goto_start_pos m lm (b, []) = False"
  2863 apply(simp add: wprepare_invs tape_of_nl_not_null)
  2885 apply(simp add: wprepare_invs)
  2864 done
  2886 done
  2865 
  2887 
  2866 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [])\<rbrakk> \<Longrightarrow> b \<noteq> []"
  2888 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [])\<rbrakk> \<Longrightarrow> b \<noteq> []"
  2867 apply(simp add: wprepare_invs tape_of_nl_not_null, auto)
  2889 apply(simp add: wprepare_invs)
  2868 done
  2890 done
       
  2891 
       
  2892 lemma rev_eq: "rev xs = rev ys \<Longrightarrow> xs = ys"
       
  2893 by auto
  2869 
  2894 
  2870 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [])\<rbrakk> \<Longrightarrow> 
  2895 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [])\<rbrakk> \<Longrightarrow> 
  2871                                   wprepare_loop_goon m lm (Bk # b, [])"
  2896                                   wprepare_loop_goon m lm (Bk # b, [])"
  2872 apply(simp only: wprepare_invs tape_of_nl_not_null)
  2897 apply(simp only: wprepare_invs)
  2873 apply(erule_tac disjE)
  2898 apply(erule_tac disjE)
  2874 apply(rule_tac disjI2)
  2899 apply(rule_tac disjI2)
  2875 apply(simp add: wprepare_loop_start_on_rightmost.simps
  2900 apply(simp add: wprepare_loop_start_on_rightmost.simps
  2876                 wprepare_loop_goon_on_rightmost.simps, auto)
  2901                 wprepare_loop_goon_on_rightmost.simps, auto)
  2877 apply(rule_tac rev_eq, simp add: tape_of_nl_rev)
  2902 apply(rule_tac rev_eq, simp add: tape_of_nl_rev)
  2878 done
  2903 done
  2879 
  2904 
  2880 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, [])\<rbrakk> \<Longrightarrow> b \<noteq> []"
  2905 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, [])\<rbrakk> \<Longrightarrow> b \<noteq> []"
  2881 apply(simp only: wprepare_invs tape_of_nl_not_null, auto)
  2906 apply(simp only: wprepare_invs, auto)
  2882 done
  2907 done
  2883 
  2908 
  2884 lemma [simp]:"\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, [])\<rbrakk> \<Longrightarrow> 
  2909 lemma [simp]:"\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, [])\<rbrakk> \<Longrightarrow> 
  2885   wprepare_add_one2 m lm (Bk # b, [])"
  2910   wprepare_add_one2 m lm (Bk # b, [])"
  2886 apply(simp only: wprepare_invs tape_of_nl_not_null, auto split: if_splits)
  2911 apply(simp only: wprepare_invs, auto split: if_splits)
  2887 apply(case_tac mr, simp, simp add: exp_ind_def)
       
  2888 done
  2912 done
  2889 
  2913 
  2890 lemma [simp]: "wprepare_add_one2 m lm (b, []) \<Longrightarrow> b \<noteq> []"
  2914 lemma [simp]: "wprepare_add_one2 m lm (b, []) \<Longrightarrow> b \<noteq> []"
  2891 apply(simp only: wprepare_invs tape_of_nl_not_null, auto)
  2915 apply(simp only: wprepare_invs, auto)
  2892 done
  2916 done
  2893 
  2917 
  2894 lemma [simp]: "wprepare_add_one2 m lm (b, []) \<Longrightarrow> wprepare_add_one2 m lm (b, [Oc])"
  2918 lemma [simp]: "wprepare_add_one2 m lm (b, []) \<Longrightarrow> wprepare_add_one2 m lm (b, [Oc])"
  2895 apply(simp only: wprepare_invs tape_of_nl_not_null, auto)
  2919 apply(simp only: wprepare_invs, auto)
  2896 done
  2920 done
  2897 
  2921 
  2898 lemma [simp]: "Bk # list = <(m::nat) # lm> @ ys = False"
  2922 lemma [simp]: "Bk # list = <(m::nat) # lm> @ ys = False"
  2899 apply(case_tac lm, auto simp: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
  2923 apply(case_tac lm, auto simp: tape_of_nl_cons replicate_Suc)
  2900 done
  2924 done
  2901 
  2925 
  2902 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_add_one m lm (b, Bk # list)\<rbrakk>
  2926 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_add_one m lm (b, Bk # list)\<rbrakk>
  2903        \<Longrightarrow> (b = [] \<longrightarrow> wprepare_goto_first_end m lm ([], Oc # list)) \<and> 
  2927        \<Longrightarrow> (b = [] \<longrightarrow> wprepare_goto_first_end m lm ([], Oc # list)) \<and> 
  2904            (b \<noteq> [] \<longrightarrow> wprepare_goto_first_end m lm (b, Oc # list))"
  2928            (b \<noteq> [] \<longrightarrow> wprepare_goto_first_end m lm (b, Oc # list))"
  2905 apply(simp only: wprepare_invs, auto)
  2929 apply(simp only: wprepare_invs)
  2906 apply(rule_tac x = 0 in exI, simp add: exp_ind_def)
  2930 apply(auto simp: tape_of_nl_cons split: if_splits)
  2907 apply(case_tac lm, simp, simp add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
  2931 apply(rule_tac x = 0 in exI, simp add: replicate_Suc)
  2908 apply(rule_tac x = rn in exI, simp)
  2932 apply(case_tac lm, simp, simp add: tape_of_nl_abv tape_of_nat_list.simps replicate_Suc)
  2909 done
  2933 done
  2910 
  2934 
  2911 lemma [simp]: "wprepare_goto_first_end m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  2935 lemma [simp]: "wprepare_goto_first_end m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  2912 apply(simp only: wprepare_invs tape_of_nl_not_null, auto)
  2936 apply(simp only: wprepare_invs , auto simp: replicate_Suc)
  2913 apply(case_tac mr, simp_all add: exp_ind_def)
  2937 done
  2914 done
  2938 
       
  2939 declare replicate_Suc[simp]
  2915 
  2940 
  2916 lemma [simp]: "wprepare_goto_first_end m lm (b, Bk # list) \<Longrightarrow>
  2941 lemma [simp]: "wprepare_goto_first_end m lm (b, Bk # list) \<Longrightarrow>
  2917                           wprepare_erase m lm (tl b, hd b # Bk # list)"
  2942                           wprepare_erase m lm (tl b, hd b # Bk # list)"
  2918 apply(simp only: wprepare_invs tape_of_nl_not_null, auto)
  2943 apply(simp only: wprepare_invs, auto)
  2919 apply(case_tac mr, simp_all add: exp_ind_def)
  2944 apply(case_tac mr, simp_all)
  2920 apply(case_tac mr, auto simp: exp_ind_def)
  2945 apply(case_tac mr, auto)
  2921 done
  2946 done
  2922 
  2947 
  2923 lemma [simp]: "wprepare_erase m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  2948 lemma [simp]: "wprepare_erase m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  2924 apply(simp only: wprepare_invs exp_ind_def, auto)
  2949 apply(simp only: wprepare_invs, auto)
  2925 done
  2950 done
  2926 
  2951 
  2927 lemma [simp]: "wprepare_erase m lm (b, Bk # list) \<Longrightarrow> 
  2952 lemma [simp]: "wprepare_erase m lm (b, Bk # list) \<Longrightarrow> 
  2928                            wprepare_goto_start_pos m lm (Bk # b, list)"
  2953                            wprepare_goto_start_pos m lm (Bk # b, list)"
  2929 apply(simp only: wprepare_invs, auto)
  2954 apply(simp only: wprepare_invs, auto)
  2930 done
  2955 done
  2931 
  2956 
  2932 lemma [simp]: "\<lbrakk>wprepare_add_one m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
  2957 lemma [simp]: "\<lbrakk>wprepare_add_one m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
  2933 apply(simp only: wprepare_invs)
  2958 apply(simp only: wprepare_invs)
  2934 apply(case_tac lm, simp_all add: tape_of_nl_abv 
  2959 apply(case_tac lm, simp_all add: tape_of_nl_abv 
  2935                          tape_of_nat_list.simps exp_ind_def, auto)
  2960                          tape_of_nat_list.simps, auto)
  2936 done
  2961 done
  2937     
  2962     
  2938 lemma [simp]: "\<lbrakk>lm \<noteq> [];  wprepare_goto_first_end m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
  2963 lemma [simp]: "\<lbrakk>lm \<noteq> [];  wprepare_goto_first_end m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
  2939 apply(simp only: wprepare_invs, auto)
  2964 apply(simp only: wprepare_invs, auto)
  2940 apply(case_tac mr, simp_all add: exp_ind_def)
  2965 apply(case_tac mr, simp_all)
  2941 apply(simp add: tape_of_nl_not_null)
       
  2942 done
  2966 done
  2943      
  2967      
  2944 lemma [simp]: "\<lbrakk>lm \<noteq> [];  wprepare_goto_first_end m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []"
  2968 lemma [simp]: "\<lbrakk>lm \<noteq> [];  wprepare_goto_first_end m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []"
  2945 apply(simp only: wprepare_invs, auto)
  2969 apply(simp only: wprepare_invs, auto)
  2946 apply(case_tac mr, simp_all add: exp_ind_def tape_of_nl_not_null)
       
  2947 done
  2970 done
  2948 
  2971 
  2949 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_erase m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
  2972 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_erase m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
  2950 apply(simp only: wprepare_invs, auto)
  2973 apply(simp only: wprepare_invs, auto)
  2951 done
  2974 done
  2952 
  2975 
  2953 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_erase m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []"
  2976 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_erase m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []"
  2954 apply(simp only: wprepare_invs, auto simp: exp_ind_def)
  2977 apply(simp only: wprepare_invs, auto)
  2955 done
  2978 done
  2956 
  2979 
  2957 lemma [simp]: "\<lbrakk>lm \<noteq> [];  wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
  2980 lemma [simp]: "\<lbrakk>lm \<noteq> [];  wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
  2958 apply(simp only: wprepare_invs, auto)
  2981 apply(simp only: wprepare_invs, auto)
  2959 apply(simp add: tape_of_nl_not_null)
       
  2960 apply(case_tac lm, simp, case_tac list)
  2982 apply(case_tac lm, simp, case_tac list)
  2961 apply(simp_all add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
  2983 apply(simp_all add: tape_of_nl_abv tape_of_nat_list.simps)
  2962 done
  2984 done
  2963 
  2985 
  2964 lemma [simp]: "\<lbrakk>lm \<noteq> [];  wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []"
  2986 lemma [simp]: "\<lbrakk>lm \<noteq> [];  wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []"
  2965 apply(simp only: wprepare_invs)
  2987 apply(simp only: wprepare_invs)
  2966 apply(auto)
  2988 apply(auto)
  2973 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> 
  2995 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> 
  2974   (list = [] \<longrightarrow> wprepare_add_one2 m lm (Bk # b, [])) \<and> 
  2996   (list = [] \<longrightarrow> wprepare_add_one2 m lm (Bk # b, [])) \<and> 
  2975   (list \<noteq> [] \<longrightarrow> wprepare_add_one2 m lm (Bk # b, list))"
  2997   (list \<noteq> [] \<longrightarrow> wprepare_add_one2 m lm (Bk # b, list))"
  2976 apply(simp only: wprepare_invs, simp)
  2998 apply(simp only: wprepare_invs, simp)
  2977 apply(case_tac list, simp_all split: if_splits, auto)
  2999 apply(case_tac list, simp_all split: if_splits, auto)
  2978 apply(case_tac [1-3] mr, simp_all add: exp_ind_def)
  3000 apply(case_tac [1-3] mr, simp_all add: )
  2979 apply(case_tac mr, simp_all add: exp_ind_def tape_of_nl_not_null)
  3001 apply(case_tac mr, simp_all)
  2980 apply(case_tac [1-2] mr, simp_all add: exp_ind_def)
  3002 apply(case_tac [1-2] mr, simp_all add: )
  2981 apply(case_tac rn, simp, case_tac nat, auto simp: exp_ind_def)
  3003 apply(case_tac rn, simp, case_tac nat, auto simp: )
  2982 done
  3004 done
  2983 
  3005 
  2984 lemma [simp]: "wprepare_add_one2 m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  3006 lemma [simp]: "wprepare_add_one2 m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []"
  2985 apply(simp only: wprepare_invs, simp)
  3007 apply(simp only: wprepare_invs, simp)
  2986 done
  3008 done
  2994 lemma [simp]: "wprepare_goto_first_end m lm (b, Oc # list)
  3016 lemma [simp]: "wprepare_goto_first_end m lm (b, Oc # list)
  2995        \<Longrightarrow> (b = [] \<longrightarrow> wprepare_goto_first_end m lm ([Oc], list)) \<and> 
  3017        \<Longrightarrow> (b = [] \<longrightarrow> wprepare_goto_first_end m lm ([Oc], list)) \<and> 
  2996            (b \<noteq> [] \<longrightarrow> wprepare_goto_first_end m lm (Oc # b, list))"
  3018            (b \<noteq> [] \<longrightarrow> wprepare_goto_first_end m lm (Oc # b, list))"
  2997 apply(simp only:  wprepare_invs, auto)
  3019 apply(simp only:  wprepare_invs, auto)
  2998 apply(rule_tac x = 1 in exI, auto)
  3020 apply(rule_tac x = 1 in exI, auto)
  2999 apply(case_tac mr, simp_all add: exp_ind_def)
  3021 apply(case_tac mr, simp_all add: )
  3000 apply(case_tac ml, simp_all add: exp_ind_def)
  3022 apply(case_tac ml, simp_all add: )
  3001 apply(rule_tac x = rn in exI, simp)
  3023 apply(rule_tac x = "Suc ml" in exI, simp_all add: )
  3002 apply(rule_tac x = "Suc ml" in exI, simp_all add: exp_ind_def)
       
  3003 apply(rule_tac x = "mr - 1" in exI, simp)
  3024 apply(rule_tac x = "mr - 1" in exI, simp)
  3004 apply(case_tac mr, simp_all add: exp_ind_def, auto)
       
  3005 done
  3025 done
  3006 
  3026 
  3007 lemma [simp]: "wprepare_erase m lm (b, Oc # list) \<Longrightarrow> b \<noteq> []"
  3027 lemma [simp]: "wprepare_erase m lm (b, Oc # list) \<Longrightarrow> b \<noteq> []"
  3008 apply(simp only: wprepare_invs, auto simp: exp_ind_def)
  3028 apply(simp only: wprepare_invs, auto simp: )
  3009 done
  3029 done
  3010 
  3030 
  3011 lemma [simp]: "wprepare_erase m lm (b, Oc # list)
  3031 lemma [simp]: "wprepare_erase m lm (b, Oc # list)
  3012   \<Longrightarrow> wprepare_erase m lm (b, Bk # list)"
  3032   \<Longrightarrow> wprepare_erase m lm (b, Bk # list)"
  3013 apply(simp  only:wprepare_invs, auto simp: exp_ind_def)
  3033 apply(simp  only:wprepare_invs, auto simp: )
  3014 done
  3034 done
  3015 
  3035 
  3016 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk>
  3036 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk>
  3017        \<Longrightarrow> wprepare_goto_start_pos m lm (Bk # b, list)"
  3037        \<Longrightarrow> wprepare_goto_start_pos m lm (Bk # b, list)"
  3018 apply(simp only:wprepare_invs, auto)
  3038 apply(simp only:wprepare_invs, auto)
  3020 done
  3040 done
  3021 
  3041 
  3022 lemma [simp]: "wprepare_loop_start m lm (b, aa) \<Longrightarrow> b \<noteq> []"
  3042 lemma [simp]: "wprepare_loop_start m lm (b, aa) \<Longrightarrow> b \<noteq> []"
  3023 apply(simp only:wprepare_invs, auto)
  3043 apply(simp only:wprepare_invs, auto)
  3024 done
  3044 done
  3025 lemma [elim]: "Bk # list = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup>  \<Longrightarrow> \<exists>rn. list = Bk\<^bsup>rn\<^esup>"
  3045 lemma [elim]: "Bk # list = Oc\<up>(mr) @ Bk\<up>(rn)  \<Longrightarrow> \<exists>rn. list = Bk\<up>(rn)"
  3026 apply(case_tac mr, simp_all)
  3046 apply(case_tac mr, simp_all)
  3027 apply(case_tac rn, simp_all add: exp_ind_def, auto)
  3047 apply(case_tac rn, simp_all)
  3028 done
  3048 done
  3029 
  3049 
  3030 lemma rev_equal_iff: "x = y \<Longrightarrow> rev x = rev y"
  3050 lemma rev_equal_iff: "x = y \<Longrightarrow> rev x = rev y"
  3031 by simp
  3051 by simp
  3032 
  3052 
  3033 lemma tape_of_nl_false1:
  3053 lemma tape_of_nl_false1:
  3034   "lm \<noteq> [] \<Longrightarrow> rev b @ [Bk] \<noteq> Bk\<^bsup>ln\<^esup> @ Oc # Oc\<^bsup>m\<^esup> @ Bk # Bk # <lm::nat list>"
  3054   "lm \<noteq> [] \<Longrightarrow> rev b @ [Bk] \<noteq> Bk\<up>(ln) @ Oc # Oc\<up>(m) @ Bk # Bk # <lm::nat list>"
  3035 apply(auto)
  3055 apply(auto)
  3036 apply(drule_tac rev_equal_iff, simp add: tape_of_nl_rev)
  3056 apply(drule_tac rev_equal_iff, simp add: tape_of_nl_rev)
  3037 apply(case_tac "rev lm")
  3057 apply(case_tac "rev lm")
  3038 apply(case_tac [2] list, auto simp: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
  3058 apply(case_tac [2] list, auto simp: tape_of_nl_abv tape_of_nat_list.simps )
  3039 done
  3059 done
  3040 
  3060 
  3041 lemma [simp]: "wprepare_loop_start_in_middle m lm (b, [Bk]) = False"
  3061 lemma [simp]: "wprepare_loop_start_in_middle m lm (b, [Bk]) = False"
  3042 apply(simp add: wprepare_loop_start_in_middle.simps, auto)
  3062 apply(simp add: wprepare_loop_start_in_middle.simps, auto)
  3043 apply(case_tac mr, simp_all add: exp_ind_def)
  3063 apply(case_tac mr, simp_all add: )
  3044 apply(case_tac lm1, simp, simp add: tape_of_nl_not_null)
       
  3045 done
  3064 done
  3046 
  3065 
  3047 declare wprepare_loop_start_in_middle.simps[simp del]
  3066 declare wprepare_loop_start_in_middle.simps[simp del]
  3048 
  3067 
  3049 declare wprepare_loop_start_on_rightmost.simps[simp del] 
  3068 declare wprepare_loop_start_on_rightmost.simps[simp del] 
  3057 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [Bk])\<rbrakk> \<Longrightarrow>
  3076 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [Bk])\<rbrakk> \<Longrightarrow>
  3058   wprepare_loop_goon m lm (Bk # b, [])"
  3077   wprepare_loop_goon m lm (Bk # b, [])"
  3059 apply(simp only: wprepare_invs, simp)
  3078 apply(simp only: wprepare_invs, simp)
  3060 apply(simp add: wprepare_loop_goon_on_rightmost.simps 
  3079 apply(simp add: wprepare_loop_goon_on_rightmost.simps 
  3061   wprepare_loop_start_on_rightmost.simps, auto)
  3080   wprepare_loop_start_on_rightmost.simps, auto)
  3062 apply(case_tac mr, simp_all add: exp_ind_def)
  3081 apply(case_tac mr, simp_all add: )
  3063 apply(rule_tac rev_eq)
  3082 apply(rule_tac rev_eq)
  3064 apply(simp add: tape_of_nl_rev)
  3083 apply(simp add: tape_of_nl_rev)
  3065 apply(simp add: exp_ind_def[THEN sym] exp_ind)
  3084 apply(simp add: exp_ind replicate_Suc[THEN sym] del: replicate_Suc)
  3066 done
  3085 done
  3067 
  3086 
  3068 lemma [simp]: "wprepare_loop_start_on_rightmost m lm (b, Bk # a # lista)
  3087 lemma [simp]: "wprepare_loop_start_on_rightmost m lm (b, Bk # a # lista)
  3069  \<Longrightarrow> wprepare_loop_goon_in_middle m lm (Bk # b, a # lista) = False"
  3088  \<Longrightarrow> wprepare_loop_goon_in_middle m lm (Bk # b, a # lista) = False"
  3070 apply(auto simp: wprepare_loop_start_on_rightmost.simps
  3089 apply(auto simp: wprepare_loop_start_on_rightmost.simps
  3071                  wprepare_loop_goon_in_middle.simps)
  3090                  wprepare_loop_goon_in_middle.simps)
  3072 apply(case_tac [!] mr, simp_all add: exp_ind_def)
  3091 apply(case_tac [!] mr, simp_all)
  3073 done
  3092 done
  3074 
  3093 
  3075 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start_on_rightmost m lm (b, Bk # a # lista)\<rbrakk>
  3094 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start_on_rightmost m lm (b, Bk # a # lista)\<rbrakk>
  3076     \<Longrightarrow> wprepare_loop_goon_on_rightmost m lm (Bk # b, a # lista)"
  3095     \<Longrightarrow> wprepare_loop_goon_on_rightmost m lm (Bk # b, a # lista)"
  3077 apply(simp only: wprepare_loop_start_on_rightmost.simps
  3096 apply(simp only: wprepare_loop_start_on_rightmost.simps
  3078                  wprepare_loop_goon_on_rightmost.simps, auto)
  3097                  wprepare_loop_goon_on_rightmost.simps, auto)
  3079 apply(case_tac mr, simp_all add: exp_ind_def)
  3098 apply(case_tac mr, simp_all add: )
  3080 apply(simp add: tape_of_nl_rev)
  3099 apply(simp add: tape_of_nl_rev)
  3081 apply(simp add: exp_ind_def[THEN sym] exp_ind)
  3100 apply(simp add: replicate_Suc[THEN sym] exp_ind del: replicate_Suc)
  3082 done
  3101 done
  3083 
  3102 
  3084 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start_in_middle m lm (b, Bk # a # lista)\<rbrakk>
  3103 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start_in_middle m lm (b, Bk # a # lista)\<rbrakk>
  3085   \<Longrightarrow> wprepare_loop_goon_on_rightmost m lm (Bk # b, a # lista) = False"
  3104   \<Longrightarrow> wprepare_loop_goon_on_rightmost m lm (Bk # b, a # lista) = False"
  3086 apply(simp add: wprepare_loop_start_in_middle.simps
  3105 apply(simp add: wprepare_loop_start_in_middle.simps
  3087                 wprepare_loop_goon_on_rightmost.simps, auto)
  3106                 wprepare_loop_goon_on_rightmost.simps, auto)
  3088 apply(case_tac mr, simp_all add: exp_ind_def)
  3107 apply(case_tac mr, simp_all add: )
  3089 apply(case_tac  "lm1::nat list", simp_all, case_tac  list, simp)
  3108 apply(case_tac  "lm1::nat list", simp_all, case_tac  list, simp)
  3090 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps tape_of_nat_abv exp_ind_def)
  3109 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps tape_of_nat_abv )
  3091 apply(case_tac [!] rna, simp_all add: exp_ind_def)
  3110 apply(case_tac [!] rna, simp_all add: )
  3092 apply(case_tac mr, simp_all add: exp_ind_def)
  3111 apply(case_tac mr, simp_all add: )
  3093 apply(case_tac lm1, simp, case_tac list, simp)
  3112 apply(case_tac lm1, simp, case_tac list, simp)
  3094 apply(simp_all add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def tape_of_nat_abv)
  3113 apply(simp_all add: tape_of_nl_abv tape_of_nat_list.simps  tape_of_nat_abv)
  3095 done
  3114 done
  3096 
  3115 
  3097 lemma [simp]: 
  3116 lemma [simp]: 
  3098   "\<lbrakk>lm \<noteq> []; wprepare_loop_start_in_middle m lm (b, Bk # a # lista)\<rbrakk> 
  3117   "\<lbrakk>lm \<noteq> []; wprepare_loop_start_in_middle m lm (b, Bk # a # lista)\<rbrakk> 
  3099   \<Longrightarrow> wprepare_loop_goon_in_middle m lm (Bk # b, a # lista)"
  3118   \<Longrightarrow> wprepare_loop_goon_in_middle m lm (Bk # b, a # lista)"
  3100 apply(simp add: wprepare_loop_start_in_middle.simps
  3119 apply(simp add: wprepare_loop_start_in_middle.simps
  3101                wprepare_loop_goon_in_middle.simps, auto)
  3120                wprepare_loop_goon_in_middle.simps, auto)
  3102 apply(rule_tac x = rn in exI, simp)
  3121 apply(rule_tac x = rn in exI, simp)
  3103 apply(case_tac mr, simp_all add: exp_ind_def)
  3122 apply(case_tac mr, simp_all add: )
  3104 apply(case_tac lm1, simp)
  3123 apply(case_tac lm1, simp)
  3105 apply(rule_tac x = "Suc aa" in exI, simp)
  3124 apply(rule_tac x = "Suc aa" in exI, simp)
  3106 apply(rule_tac x = list in exI)
  3125 apply(rule_tac x = list in exI)
  3107 apply(case_tac list, simp_all add: tape_of_nl_abv tape_of_nat_list.simps)
  3126 apply(case_tac list, simp_all add: tape_of_nl_abv tape_of_nat_list.simps)
  3108 done
  3127 done
  3135 
  3154 
  3136 lemma [simp]: "wprepare_loop_start_on_rightmost m lm (b, Oc # list) \<Longrightarrow> 
  3155 lemma [simp]: "wprepare_loop_start_on_rightmost m lm (b, Oc # list) \<Longrightarrow> 
  3137   wprepare_loop_start_on_rightmost m lm (Oc # b, list)"
  3156   wprepare_loop_start_on_rightmost m lm (Oc # b, list)"
  3138 apply(simp add: wprepare_loop_start_on_rightmost.simps, auto)
  3157 apply(simp add: wprepare_loop_start_on_rightmost.simps, auto)
  3139 apply(rule_tac x = rn in exI, auto)
  3158 apply(rule_tac x = rn in exI, auto)
  3140 apply(case_tac mr, simp_all add: exp_ind_def)
  3159 apply(case_tac mr, simp_all add: )
  3141 apply(case_tac rn, auto simp: exp_ind_def)
       
  3142 done
  3160 done
  3143 
  3161 
  3144 lemma [simp]: "wprepare_loop_start_in_middle m lm (b, Oc # list) \<Longrightarrow> 
  3162 lemma [simp]: "wprepare_loop_start_in_middle m lm (b, Oc # list) \<Longrightarrow> 
  3145                 wprepare_loop_start_in_middle m lm (Oc # b, list)"
  3163                 wprepare_loop_start_in_middle m lm (Oc # b, list)"
  3146 apply(simp add: wprepare_loop_start_in_middle.simps, auto)
  3164 apply(simp add: wprepare_loop_start_in_middle.simps, auto)
  3147 apply(rule_tac x = rn in exI, auto)
  3165 apply(rule_tac x = rn in exI, auto)
  3148 apply(case_tac mr, simp, simp add: exp_ind_def)
  3166 apply(case_tac mr, simp, simp add: )
  3149 apply(rule_tac x = nat in exI, simp)
  3167 apply(rule_tac x = nat in exI, simp)
  3150 apply(rule_tac x = lm1 in exI, simp)
  3168 apply(rule_tac x = lm1 in exI, simp)
  3151 done
  3169 done
  3152 
  3170 
  3153 lemma start_2_start: "wprepare_loop_start m lm (b, Oc # list) \<Longrightarrow> 
  3171 lemma start_2_start: "wprepare_loop_start m lm (b, Oc # list) \<Longrightarrow> 
  3168 done
  3186 done
  3169 
  3187 
  3170 lemma [simp]: "wprepare_loop_goon_on_rightmost m lm (b, Oc # list) = False"
  3188 lemma [simp]: "wprepare_loop_goon_on_rightmost m lm (b, Oc # list) = False"
  3171 apply(simp add: wprepare_loop_goon_on_rightmost.simps)
  3189 apply(simp add: wprepare_loop_goon_on_rightmost.simps)
  3172 done
  3190 done
  3173 lemma wprepare_loop1: "\<lbrakk>rev b @ Oc\<^bsup>mr\<^esup> =  Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm>; 
  3191 lemma wprepare_loop1: "\<lbrakk>rev b @ Oc\<up>(mr) =  Oc\<up>(Suc m) @ Bk # Bk # <lm>; 
  3174          b \<noteq> []; 0 < mr; Oc # list = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup>\<rbrakk>
  3192          b \<noteq> []; 0 < mr; Oc # list = Oc\<up>(mr) @ Bk\<up>(rn)\<rbrakk>
  3175        \<Longrightarrow> wprepare_loop_start_on_rightmost m lm (Oc # b, list)"
  3193        \<Longrightarrow> wprepare_loop_start_on_rightmost m lm (Oc # b, list)"
  3176 apply(simp add: wprepare_loop_start_on_rightmost.simps)
  3194 apply(simp add: wprepare_loop_start_on_rightmost.simps)
  3177 apply(rule_tac x = rn in exI, simp)
  3195 apply(rule_tac x = rn in exI, simp)
  3178 apply(case_tac mr, simp, simp add: exp_ind_def, auto)
  3196 apply(case_tac mr, simp, simp)
  3179 done
  3197 done
  3180 
  3198 
  3181 lemma wprepare_loop2: "\<lbrakk>rev b @ Oc\<^bsup>mr\<^esup> @ Bk # <a # lista> = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm>;
  3199 lemma wprepare_loop2: "\<lbrakk>rev b @ Oc\<up>(mr) @ Bk # <a # lista> = Oc\<up>(Suc m) @ Bk # Bk # <lm>;
  3182                 b \<noteq> []; Oc # list = Oc\<^bsup>mr\<^esup> @ Bk # <(a::nat) # lista> @ Bk\<^bsup>rn\<^esup>\<rbrakk>
  3200                 b \<noteq> []; Oc # list = Oc\<up>(mr) @ Bk # <(a::nat) # lista> @ Bk\<up>(rn)\<rbrakk>
  3183        \<Longrightarrow>  wprepare_loop_start_in_middle m lm (Oc # b, list)"
  3201        \<Longrightarrow>  wprepare_loop_start_in_middle m lm (Oc # b, list)"
  3184 apply(simp add: wprepare_loop_start_in_middle.simps)
  3202 apply(simp add: wprepare_loop_start_in_middle.simps)
  3185 apply(rule_tac x = rn in exI, simp)
  3203 apply(rule_tac x = rn in exI, simp)
  3186 apply(case_tac mr, simp_all add: exp_ind_def)
  3204 apply(case_tac mr, simp_all add: )
  3187 apply(rule_tac x = nat in exI, simp)
  3205 apply(rule_tac x = nat in exI, simp)
  3188 apply(rule_tac x = "a#lista" in exI, simp)
  3206 apply(rule_tac x = "a#lista" in exI, simp)
  3189 done
  3207 done
  3190 
  3208 
  3191 lemma [simp]: "wprepare_loop_goon_in_middle m lm (b, Oc # list) \<Longrightarrow>
  3209 lemma [simp]: "wprepare_loop_goon_in_middle m lm (b, Oc # list) \<Longrightarrow>
  3210 lemma [simp]: "wprepare_goto_start_pos m [a] (b, Oc # list)
  3228 lemma [simp]: "wprepare_goto_start_pos m [a] (b, Oc # list)
  3211               \<Longrightarrow> wprepare_loop_start_on_rightmost m [a] (Oc # b, list) "
  3229               \<Longrightarrow> wprepare_loop_start_on_rightmost m [a] (Oc # b, list) "
  3212 apply(auto simp: wprepare_goto_start_pos.simps 
  3230 apply(auto simp: wprepare_goto_start_pos.simps 
  3213                  wprepare_loop_start_on_rightmost.simps)
  3231                  wprepare_loop_start_on_rightmost.simps)
  3214 apply(rule_tac x = rn in exI, simp)
  3232 apply(rule_tac x = rn in exI, simp)
  3215 apply(simp add: tape_of_nat_abv tape_of_nat_list.simps exp_ind_def, auto)
  3233 apply(simp add: replicate_Suc[THEN sym] exp_ind del: replicate_Suc)
  3216 done
  3234 done
  3217 
  3235 
  3218 lemma [simp]:  "wprepare_goto_start_pos m (a # aa # listaa) (b, Oc # list)
  3236 lemma [simp]:  "wprepare_goto_start_pos m (a # aa # listaa) (b, Oc # list)
  3219        \<Longrightarrow>wprepare_loop_start_in_middle m (a # aa # listaa) (Oc # b, list)"
  3237        \<Longrightarrow>wprepare_loop_start_in_middle m (a # aa # listaa) (Oc # b, list)"
  3220 apply(auto simp: wprepare_goto_start_pos.simps
  3238 apply(auto simp: wprepare_goto_start_pos.simps
  3221                  wprepare_loop_start_in_middle.simps)
  3239                  wprepare_loop_start_in_middle.simps)
  3222 apply(rule_tac x = rn in exI, simp)
  3240 apply(rule_tac x = rn in exI, simp)
  3223 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def)
  3241 apply(simp add: exp_ind[THEN sym])
  3224 apply(rule_tac x = a in exI, rule_tac x = "aa#listaa" in exI, simp)
  3242 apply(rule_tac x = a in exI, rule_tac x = "aa#listaa" in exI, simp)
       
  3243 apply(simp add: tape_of_nl_cons)
  3225 done
  3244 done
  3226 
  3245 
  3227 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_start_pos m lm (b, Oc # list)\<rbrakk>
  3246 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_start_pos m lm (b, Oc # list)\<rbrakk>
  3228        \<Longrightarrow> wprepare_loop_start m lm (Oc # b, list)"
  3247        \<Longrightarrow> wprepare_loop_start m lm (Oc # b, list)"
  3229 apply(case_tac lm, simp_all)
  3248 apply(case_tac lm, simp_all)
  3244 
  3263 
  3245 lemma wprepare_correctness:
  3264 lemma wprepare_correctness:
  3246   assumes h: "lm \<noteq> []"
  3265   assumes h: "lm \<noteq> []"
  3247   shows "let P = (\<lambda> (st, l, r). st = 0) in 
  3266   shows "let P = (\<lambda> (st, l, r). st = 0) in 
  3248   let Q = (\<lambda> (st, l, r). wprepare_inv st m lm (l, r)) in 
  3267   let Q = (\<lambda> (st, l, r). wprepare_inv st m lm (l, r)) in 
  3249   let f = (\<lambda> stp. steps (Suc 0, [], (<m # lm>)) t_wcode_prepare stp) in
  3268   let f = (\<lambda> stp. steps0 (Suc 0, [], (<m # lm>)) t_wcode_prepare stp) in
  3250     \<exists> n .P (f n) \<and> Q (f n)"
  3269     \<exists> n .P (f n) \<and> Q (f n)"
  3251 proof -
  3270 proof -
  3252   let ?P = "(\<lambda> (st, l, r). st = 0)"
  3271   let ?P = "(\<lambda> (st, l, r). st = 0)"
  3253   let ?Q = "(\<lambda> (st, l, r). wprepare_inv st m lm (l, r))"
  3272   let ?Q = "(\<lambda> (st, l, r). wprepare_inv st m lm (l, r))"
  3254   let ?f = "(\<lambda> stp. steps (Suc 0, [], (<m # lm>)) t_wcode_prepare stp)"
  3273   let ?f = "(\<lambda> stp. steps0 (Suc 0, [], (<m # lm>)) t_wcode_prepare stp)"
  3255   have "\<exists> n. ?P (?f n) \<and> ?Q (?f n)"
  3274   have "\<exists> n. ?P (?f n) \<and> ?Q (?f n)"
  3256   proof(rule_tac halt_lemma2)
  3275   proof(rule_tac halt_lemma2)
  3257     show "wf wcode_prepare_le" by auto
  3276     show "wf wcode_prepare_le" by auto
  3258   next
  3277   next
  3259     show "\<forall> n. \<not> ?P (?f n) \<and> ?Q (?f n) \<longrightarrow> 
  3278     show "\<forall> n. \<not> ?P (?f n) \<and> ?Q (?f n) \<longrightarrow> 
  3260                  ?Q (?f (Suc n)) \<and> (?f (Suc n), ?f n) \<in> wcode_prepare_le"
  3279                  ?Q (?f (Suc n)) \<and> (?f (Suc n), ?f n) \<in> wcode_prepare_le"
  3261       using h
  3280       using h
  3262       apply(rule_tac allI, rule_tac impI, case_tac "?f n", 
  3281       apply(rule_tac allI, rule_tac impI, case_tac "?f n", 
  3263             simp add: tstep_red tstep.simps)
  3282             simp add: step_red step.simps)
  3264       apply(case_tac c, simp, case_tac [2] aa)
  3283       apply(case_tac c, simp, case_tac [2] aa)
  3265       apply(simp_all add: wprepare_inv.simps wcode_prepare_le_def new_tape.simps
  3284       apply(simp_all add: wprepare_inv.simps wcode_prepare_le_def lex_triple_def lex_pair_def
  3266                           lex_triple_def lex_pair_def
       
  3267 
       
  3268                  split: if_splits)
  3285                  split: if_splits)
  3269       apply(simp_all add: start_2_goon  start_2_start
  3286       apply(simp_all add: start_2_goon  start_2_start
  3270                            add_one_2_add_one add_one_2_stop)
  3287                            add_one_2_add_one add_one_2_stop)
  3271       apply(auto simp: wprepare_add_one2.simps)
  3288       apply(auto simp: wprepare_add_one2.simps)
  3272       done   
  3289       done   
  3282   thus "?thesis"
  3299   thus "?thesis"
  3283     apply(auto)
  3300     apply(auto)
  3284     done
  3301     done
  3285 qed
  3302 qed
  3286 
  3303 
  3287 lemma [intro]: "t_correct t_wcode_prepare"
  3304 lemma [intro]: "tm_wf (t_wcode_prepare, 0)"
  3288 apply(simp add: t_correct.simps t_wcode_prepare_def iseven_def)
  3305 apply(simp add:tm_wf.simps t_wcode_prepare_def)
  3289 apply(rule_tac x = 7 in exI, simp)
  3306 done
  3290 done
  3307    
  3291     
  3308 (* 
  3292 lemma twice_len_even: "length (tm_of abc_twice) mod 2 = 0"
       
  3293 apply(simp add: tm_even)
       
  3294 done
       
  3295 
       
  3296 lemma fourtimes_len_even: "length (tm_of abc_fourtimes) mod 2 = 0"
       
  3297 apply(simp add: tm_even)
       
  3298 done
       
  3299 
       
  3300 lemma t_correct_termi: "t_correct tp \<Longrightarrow> 
  3309 lemma t_correct_termi: "t_correct tp \<Longrightarrow> 
  3301       list_all (\<lambda>(acn, st). (st \<le> Suc (length tp div 2))) (change_termi_state tp)"
  3310       list_all (\<lambda>(acn, st). (st \<le> Suc (length tp div 2))) (change_termi_state tp)"
  3302 apply(auto simp: t_correct.simps List.list_all_length)
  3311 apply(auto simp: t_correct.simps List.list_all_length)
  3303 apply(erule_tac x = n in allE, simp)
  3312 apply(erule_tac x = n in allE, simp)
  3304 apply(case_tac "tp!n", auto simp: change_termi_state.simps split: if_splits)
  3313 apply(case_tac "tp!n", auto simp: change_termi_state.simps split: if_splits)
  3305 done
  3314 done
  3306 
  3315 *)
  3307 
  3316 
  3308 lemma t_correct_shift:
  3317 lemma t_correct_shift:
  3309          "list_all (\<lambda>(acn, st). (st \<le> y)) tp \<Longrightarrow>
  3318          "list_all (\<lambda>(acn, st). (st \<le> y)) tp \<Longrightarrow>
  3310           list_all (\<lambda>(acn, st). (st \<le> y + off)) (tshift tp off) "
  3319           list_all (\<lambda>(acn, st). (st \<le> y + off)) (shift tp off) "
  3311 apply(auto simp: t_correct.simps List.list_all_length)
  3320 apply(auto simp: List.list_all_length)
  3312 apply(erule_tac x = n in allE, simp add: shift_length)
  3321 apply(erule_tac x = n in allE, simp add: length_shift)
  3313 apply(case_tac "tp!n", auto simp: tshift.simps)
  3322 apply(case_tac "tp!n", auto simp: shift.simps)
  3314 done
  3323 done
  3315 
  3324 (*
  3316 lemma [intro]: 
  3325 lemma [intro]: 
  3317   "t_correct (tm_of abc_twice @ tMp (Suc 0) 
  3326   "t_correct (tm_of abc_twice @ tMp (Suc 0) 
  3318         (start_of twice_ly (length abc_twice) - Suc 0))"
  3327         (start_of twice_ly (length abc_twice) - Suc 0))"
  3319 apply(rule_tac t_compiled_correct, simp_all)
  3328 apply(rule_tac t_compiled_correct, simp_all)
  3320 apply(simp add: twice_ly_def)
  3329 apply(simp add: twice_ly_def)
  3323 lemma [intro]: "t_correct (tm_of abc_fourtimes @ tMp (Suc 0) 
  3332 lemma [intro]: "t_correct (tm_of abc_fourtimes @ tMp (Suc 0) 
  3324    (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))"
  3333    (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))"
  3325 apply(rule_tac t_compiled_correct, simp_all)
  3334 apply(rule_tac t_compiled_correct, simp_all)
  3326 apply(simp add: fourtimes_ly_def)
  3335 apply(simp add: fourtimes_ly_def)
  3327 done
  3336 done
  3328 
  3337 *)
  3329 
  3338 
  3330 lemma [intro]: "t_correct t_wcode_main"
  3339 lemma [intro]: "(28 + (length t_twice_compile + length t_fourtimes_compile)) mod 2 = 0"
  3331 apply(auto simp: t_wcode_main_def t_correct.simps shift_length 
  3340 apply(auto simp: t_twice_compile_def t_fourtimes_compile_def)
  3332                  t_twice_def t_fourtimes_def)
  3341 by arith
       
  3342 
       
  3343 lemma [elim]: "(a, b) \<in> set t_wcode_main_first_part \<Longrightarrow>
       
  3344   b \<le> (28 + (length t_twice_compile + length t_fourtimes_compile)) div 2"
       
  3345 apply(auto simp: t_wcode_main_first_part_def t_twice_def)
       
  3346 done
       
  3347 
       
  3348 
       
  3349 
       
  3350 lemma tm_wf_change_termi: "tm_wf (tp, 0) \<Longrightarrow> 
       
  3351       list_all (\<lambda>(acn, st). (st \<le> Suc (length tp div 2))) (adjust tp)"
       
  3352 apply(auto simp: tm_wf.simps List.list_all_length)
       
  3353 apply(case_tac "tp!n", auto simp: adjust.simps split: if_splits)
       
  3354 apply(erule_tac x = "(a, b)" in ballE, auto)
       
  3355 by (metis in_set_conv_nth)
       
  3356 
       
  3357 lemma tm_wf_shift:
       
  3358          "list_all (\<lambda>(acn, st). (st \<le> y)) tp \<Longrightarrow>
       
  3359           list_all (\<lambda>(acn, st). (st \<le> y + off)) (shift tp off) "
       
  3360 apply(auto simp: tm_wf.simps List.list_all_length)
       
  3361 apply(erule_tac x = n in allE, simp add: length_shift)
       
  3362 apply(case_tac "tp!n", auto simp: shift.simps)
       
  3363 done
       
  3364 
       
  3365 declare length_tp'[simp del]
       
  3366 
       
  3367 lemma [simp]: "length (mopup (Suc 0)) = 16"
       
  3368 apply(auto simp: mopup.simps)
       
  3369 done
       
  3370 
       
  3371 lemma [elim]: "(a, b) \<in> set (shift (turing_basic.adjust t_twice_compile) 12) \<Longrightarrow> 
       
  3372   b \<le> (28 + (length t_twice_compile + length t_fourtimes_compile)) div 2"
       
  3373 apply(simp add: t_twice_compile_def t_fourtimes_compile_def)
  3333 proof -
  3374 proof -
  3334   show "iseven (60 + (length (tm_of abc_twice) +
  3375   assume g: "(a, b) \<in> set (shift (turing_basic.adjust (tm_of abc_twice @ shift (mopup (Suc 0)) (length (tm_of abc_twice) div 2))) 12)"
  3335                  length (tm_of abc_fourtimes)))"
  3376   moreover have "length (tm_of abc_twice) mod 2 = 0" by auto
  3336     using twice_len_even fourtimes_len_even
  3377   moreover have "length (tm_of abc_fourtimes) mod 2 = 0" by auto
  3337     apply(auto simp: iseven_def)
  3378   ultimately have "list_all (\<lambda>(acn, st). (st \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2)) 
  3338     apply(rule_tac x = "30 + q + qa" in exI, simp)
  3379     (shift (turing_basic.adjust t_twice_compile) 12)"
       
  3380   proof(auto simp: mod_ex1)
       
  3381     fix q qa
       
  3382     assume h: "length (tm_of abc_twice) = 2 * q" "length (tm_of abc_fourtimes) = 2 * qa"
       
  3383     hence "list_all (\<lambda>(acn, st). st \<le> (18 + (q + qa)) + 12) (shift (turing_basic.adjust t_twice_compile) 12)"
       
  3384     proof(rule_tac tm_wf_shift t_twice_compile_def)
       
  3385       have "list_all (\<lambda>(acn, st). st \<le> Suc (length t_twice_compile div 2)) (adjust t_twice_compile)"
       
  3386         by(rule_tac tm_wf_change_termi, auto)
       
  3387       thus "list_all (\<lambda>(acn, st). st \<le> 18 + (q + qa)) (turing_basic.adjust t_twice_compile)"
       
  3388         using h
       
  3389         apply(simp add: t_twice_compile_def, auto simp: List.list_all_length)
       
  3390         done
       
  3391     qed
       
  3392     thus "list_all (\<lambda>(acn, st). st \<le> 30 + (q + qa)) (shift (turing_basic.adjust t_twice_compile) 12)"
       
  3393       by simp
       
  3394   qed
       
  3395   thus "b \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2"
       
  3396     using g
       
  3397     apply(auto simp:t_twice_compile_def)
       
  3398     apply(simp add: Ball_set[THEN sym])
       
  3399     apply(erule_tac x = "(a, b)" in ballE, simp, simp)
  3339     done
  3400     done
  3340 next
  3401 qed 
  3341   show " list_all (\<lambda>(acn, s). s \<le> (60 + (length (tm_of abc_twice) + 
  3402 
  3342            length (tm_of abc_fourtimes))) div 2) t_wcode_main_first_part"
  3403 lemma [elim]: "(a, b) \<in> set (shift (turing_basic.adjust t_fourtimes_compile) (t_twice_len + 13)) 
  3343     apply(auto simp: t_wcode_main_first_part_def shift_length t_twice_def)
  3404   \<Longrightarrow> b \<le> (28 + (length t_twice_compile + length t_fourtimes_compile)) div 2"
  3344     done
  3405 apply(simp add: t_twice_compile_def t_fourtimes_compile_def t_twice_len_def)
  3345 next
  3406 proof -
  3346   have "list_all (\<lambda>(acn, s). s \<le> Suc (length (tm_of abc_twice @ tMp (Suc 0)
  3407   assume g: "(a, b) \<in> set (shift (turing_basic.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0))
  3347     (start_of twice_ly (length abc_twice) - Suc 0)) div 2))
  3408     (length (tm_of abc_fourtimes) div 2))) (length t_twice div 2 + 13))"
  3348     (change_termi_state (tm_of abc_twice @ tMp (Suc 0) 
  3409   moreover have "length (tm_of abc_twice) mod 2 = 0" by auto
  3349     (start_of twice_ly (length abc_twice) - Suc 0)))"
  3410   moreover have "length (tm_of abc_fourtimes) mod 2 = 0" by auto
  3350     apply(rule_tac t_correct_termi, auto)
  3411   ultimately have "list_all (\<lambda>(acn, st). (st \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2)) 
  3351     done
  3412     (shift (turing_basic.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0))
  3352   hence "list_all (\<lambda>(acn, s). s \<le>  Suc (length (tm_of abc_twice @ tMp (Suc 0)
  3413     (length (tm_of abc_fourtimes) div 2))) (length t_twice div 2 + 13))"
  3353     (start_of twice_ly (length abc_twice) - Suc 0)) div 2) + 12)
  3414   proof(auto simp: mod_ex1 t_twice_def t_twice_compile_def)
  3354      (tshift (change_termi_state (tm_of abc_twice @ tMp (Suc 0) 
  3415     fix q qa
  3355            (start_of twice_ly (length abc_twice) - Suc 0))) 12)"
  3416     assume h: "length (tm_of abc_twice) = 2 * q" "length (tm_of abc_fourtimes) = 2 * qa"
  3356     apply(rule_tac t_correct_shift, simp)
  3417     hence "list_all (\<lambda>(acn, st). st \<le> (9 + qa + (21 + q)))
  3357     done
  3418       (shift (turing_basic.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa)) (21 + q))"
  3358   thus  "list_all (\<lambda>(acn, s). s \<le> 
  3419     proof(rule_tac tm_wf_shift t_twice_compile_def)
  3359            (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2)
  3420       have "list_all (\<lambda>(acn, st). st \<le> Suc (length (tm_of abc_fourtimes @ shift 
  3360      (tshift (change_termi_state (tm_of abc_twice @ tMp (Suc 0)
  3421         (mopup (Suc 0)) qa) div 2)) (adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa))"
  3361                  (start_of twice_ly (length abc_twice) - Suc 0))) 12)"
  3422         apply(rule_tac tm_wf_change_termi)
  3362     apply(simp)
  3423         using wf_fourtimes h
  3363     apply(simp add: list_all_length, auto)
  3424         apply(simp add: t_fourtimes_compile_def)
  3364     done
  3425         done        
  3365 next
  3426       thus "list_all (\<lambda>(acn, st). st \<le> 9 + qa) ((turing_basic.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa)))"
  3366   have "list_all (\<lambda>(acn, s). s \<le> Suc (length (tm_of abc_fourtimes @ tMp (Suc 0) 
  3427         using h
  3367     (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) div 2))
  3428         apply(simp)
  3368       (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) 
  3429         done
  3369     (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) "
  3430     qed
  3370     apply(rule_tac t_correct_termi, auto)
  3431     thus "list_all (\<lambda>(acn, st). st \<le> 30 + (q + qa)) (shift (turing_basic.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa)) (21 + q))"
  3371     done
  3432       apply(subgoal_tac "qa + q = q + qa")
  3372   hence "list_all (\<lambda>(acn, s). s \<le> Suc (length (tm_of abc_fourtimes @ tMp (Suc 0) 
  3433       apply(simp, simp)
  3373     (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) div 2) + (t_twice_len + 13))
  3434       done
  3374     (tshift (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) 
  3435   qed
  3375     (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) (t_twice_len + 13))"
  3436   thus "b \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2"
  3376     apply(rule_tac t_correct_shift, simp)
  3437     using g
  3377     done
  3438     apply(simp add: Ball_set[THEN sym])
  3378   thus "list_all (\<lambda>(acn, s). s \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2)
  3439     apply(erule_tac x = "(a, b)" in ballE, simp, simp)
  3379     (tshift (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0)
       
  3380     (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) (t_twice_len + 13))"
       
  3381     apply(simp add: t_twice_len_def t_twice_def)
       
  3382     using twice_len_even fourtimes_len_even
       
  3383     apply(auto simp: list_all_length)
       
  3384     done
  3440     done
  3385 qed
  3441 qed
  3386 
  3442 
  3387 lemma [intro]: "t_correct (t_wcode_prepare |+| t_wcode_main)"
  3443 lemma [intro]: "tm_wf (t_wcode_main, 0)"
  3388 apply(auto intro: t_correct_add)
  3444 apply(auto simp: t_wcode_main_def tm_wf.simps
       
  3445                  t_twice_def t_fourtimes_def del: List.list_all_iff)
       
  3446 done
       
  3447 
       
  3448 declare tm_comp.simps[simp del]
       
  3449 lemma tm_wf_comp: "\<lbrakk>tm_wf (A, 0); tm_wf (B, 0)\<rbrakk> \<Longrightarrow> tm_wf (A |+| B, 0)"
       
  3450 apply(auto simp: tm_wf.simps shift.simps adjust.simps tm_comp_length
       
  3451                  tm_comp.simps)
       
  3452 done
       
  3453 
       
  3454 lemma [intro]: "tm_wf (t_wcode_prepare |+| t_wcode_main, 0)"
       
  3455 apply(rule_tac tm_wf_comp, auto)
  3389 done
  3456 done
  3390 
  3457 
  3391 lemma prepare_mainpart_lemma:
  3458 lemma prepare_mainpart_lemma:
  3392   "args \<noteq> [] \<Longrightarrow> 
  3459   "args \<noteq> [] \<Longrightarrow> 
  3393   \<exists> stp ln rn. steps (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) stp
  3460   \<exists> stp ln rn. steps0 (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) stp
  3394               = (0,  Bk # Oc\<^bsup>Suc m\<^esup>, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  3461               = (0,  Bk # Oc\<up>(Suc m), Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin (<args>)) @ Bk\<up>(rn))"
  3395 proof -
  3462 proof -
  3396   let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <m # args>"
  3463   let ?P1 = "(\<lambda> (l, r). (l::cell list) = [] \<and> r = <m # args>)"
  3397   let ?Q1 = "\<lambda> (l, r). wprepare_stop m args (l, r)"
  3464   let ?Q1 = "(\<lambda> (l, r). wprepare_stop m args (l, r))"
  3398   let ?P2 = ?Q1
  3465   let ?P2 = ?Q1
  3399   let ?Q2 = "\<lambda> (l, r). (\<exists> ln rn. l = Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  3466   let ?Q2 = "(\<lambda> (l, r). (\<exists> ln rn. l = Bk # Oc\<up>(Suc m) \<and>
  3400                            r =  Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  3467                            r =  Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin (<args>)) @ Bk\<up>(rn)))"
  3401   let ?P3 = "\<lambda> tp. False"
  3468   let ?P3 = "\<lambda> tp. False"
  3402   assume h: "args \<noteq> []"
  3469   assume h: "args \<noteq> []"
  3403   have "?P1 \<turnstile>-> \<lambda> tp. (\<exists> stp tp'. steps (Suc 0, tp)
  3470   have "{?P1} t_wcode_prepare |+| t_wcode_main {?Q2}"
  3404                       (t_wcode_prepare |+| t_wcode_main) stp = (0, tp') \<and> ?Q2 tp')"
  3471   proof(rule_tac Hoare_plus_halt)
  3405   proof(rule_tac turing_merge.t_merge_halt[of t_wcode_prepare t_wcode_main ?P1 ?P2 ?P3 ?P3 ?Q1 ?Q2], 
  3472     show "?Q1 \<mapsto> ?P2"
  3406         auto simp: turing_merge_def)
  3473       by(simp add: assert_imp_def)
  3407     show "\<exists>stp. case steps (Suc 0, [], <m # args>) t_wcode_prepare stp of (st, tp')
       
  3408                   \<Rightarrow> st = 0 \<and> wprepare_stop m args tp'"
       
  3409       using wprepare_correctness[of args m] h
       
  3410       apply(simp, auto)
       
  3411       apply(rule_tac x = n in exI, simp add: wprepare_inv.simps)
       
  3412       done
       
  3413   next
  3474   next
  3414     fix a b
  3475     show "tm_wf (t_wcode_prepare, 0)"
  3415     assume "wprepare_stop m args (a, b)"
  3476       by auto
  3416     thus "\<exists>stp. case steps (Suc 0, a, b) t_wcode_main stp of
  3477   next
  3417       (st, tp') \<Rightarrow> (st = 0) \<and> (case tp' of (l, r) \<Rightarrow> l = Bk # Oc\<^bsup>Suc m\<^esup> \<and> 
  3478     show "{?P1} t_wcode_prepare {?Q1}"
  3418       (\<exists>ln rn. r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>))"
  3479     proof(rule_tac HoareI, auto)
  3419       proof(simp only: wprepare_stop.simps, erule_tac exE)
  3480       show "\<exists>n. is_final (steps0 (Suc 0, [], <m # args>) t_wcode_prepare n) \<and>
       
  3481         wprepare_stop m args holds_for steps0 (Suc 0, [], <m # args>) t_wcode_prepare n"
       
  3482         using wprepare_correctness[of args m] h
       
  3483         apply(auto)
       
  3484         apply(rule_tac x = n in exI, simp add: wprepare_inv.simps)
       
  3485         done
       
  3486     qed
       
  3487   next
       
  3488     show "{?P2} t_wcode_main {?Q2}"
       
  3489     proof(rule_tac HoareI, auto)
       
  3490       fix l r
       
  3491       assume "wprepare_stop m args (l, r)"
       
  3492       thus "\<exists>n. is_final (steps0 (Suc 0, l, r) t_wcode_main n) \<and>
       
  3493               (\<lambda>(l, r). l = Bk # Oc # Oc \<up> m \<and> (\<exists>ln rn. r = Bk # Oc # Bk \<up> ln @ 
       
  3494         Bk # Bk # Oc \<up> bl_bin (<args>) @ Bk \<up> rn)) holds_for steps0 (Suc 0, l, r) t_wcode_main n"
       
  3495       proof(auto simp: wprepare_stop.simps)
  3420         fix rn
  3496         fix rn
  3421         assume "a = Bk # <rev args> @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and> 
  3497         show " \<exists>n. is_final (steps0 (Suc 0, Bk # <rev args> @ Bk # Bk # Oc # Oc \<up> m, Bk # Oc # Bk \<up> rn) t_wcode_main n) \<and>
  3422                    b = Bk # Oc # Bk\<^bsup>rn\<^esup>"
  3498           (\<lambda>(l, r). l = Bk # Oc # Oc \<up> m \<and>
  3423         thus "?thesis"
  3499           (\<exists>ln rn. r = Bk # Oc # Bk \<up> ln @
  3424           using t_wcode_main_lemma_pre[of "args" "<args>" 0 "Oc\<^bsup>Suc m\<^esup>" 0 rn] h
  3500           Bk # Bk # Oc \<up> bl_bin (<args>) @
  3425           apply(simp)
  3501           Bk \<up> rn)) holds_for steps0 (Suc 0, Bk # <rev args> @ Bk # Bk # Oc # Oc \<up> m, Bk # Oc # Bk \<up> rn) t_wcode_main n"
  3426           apply(erule_tac exE)+
  3502           using t_wcode_main_lemma_pre[of "args" "<args>" 0 "Oc\<up>(Suc m)" 0 rn] h
  3427           apply(rule_tac x = stp in exI, simp add: tape_of_nl_rev, auto)
  3503           apply(auto simp: tape_of_nl_rev)
       
  3504           apply(rule_tac x = stp in exI, auto)
  3428           done
  3505           done
  3429       qed
  3506       qed
  3430   next
  3507     qed
  3431     show "wprepare_stop m args \<turnstile>-> wprepare_stop m args"
       
  3432       by(simp add: t_imply_def)
       
  3433   qed
  3508   qed
  3434   thus "\<exists> stp ln rn. steps (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) stp
  3509   thus "?thesis"
  3435               = (0,  Bk # Oc\<^bsup>Suc m\<^esup>, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  3510     apply(auto simp: Hoare_def)
  3436     apply(simp add: t_imply_def)
  3511     apply(rule_tac x = n in exI)
  3437     apply(erule_tac exE)+
  3512     apply(case_tac "(steps0 (Suc 0, [], <m # args>)
  3438     apply(auto)
  3513       (turing_basic.adjust t_wcode_prepare @ shift t_wcode_main (length t_wcode_prepare div 2)) n)")
       
  3514     apply(auto simp: tm_comp.simps)
  3439     done
  3515     done
  3440 qed
  3516 qed
  3441       
  3517    
  3442 
       
  3443 lemma [simp]:  "tinres r r' \<Longrightarrow> 
  3518 lemma [simp]:  "tinres r r' \<Longrightarrow> 
  3444   fetch t ss (case r of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) = 
  3519   fetch t ss (read r) = 
  3445   fetch t ss (case r' of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)"
  3520   fetch t ss (read r')"
  3446 apply(simp add: fetch.simps, auto split: if_splits simp: tinres_def)
  3521 apply(simp add: fetch.simps, auto split: if_splits simp: tinres_def)
  3447 apply(case_tac [!] r', simp_all)
  3522 apply(case_tac [!] n, simp_all)
  3448 apply(case_tac [!] n, simp_all add: exp_ind_def)
  3523 done
  3449 apply(case_tac [!] r, simp_all)
  3524 
  3450 done
  3525 lemma [intro]: "\<exists> n. (a::cell)\<up>(n) = []"
  3451 
       
  3452 lemma [intro]: "\<exists> n. (a::block)\<^bsup>n\<^esup> = []"
       
  3453 by auto
  3526 by auto
  3454 
  3527 
  3455 lemma [simp]: "\<lbrakk>tinres r r'; r \<noteq> []; r' \<noteq> []\<rbrakk> \<Longrightarrow> hd r = hd r'"
  3528 lemma [simp]: "\<lbrakk>tinres r r'; r \<noteq> []; r' \<noteq> []\<rbrakk> \<Longrightarrow> hd r = hd r'"
  3456 apply(auto simp: tinres_def)
  3529 apply(auto simp: tinres_def)
  3457 done
  3530 done
  3458 
  3531 
  3459 lemma [intro]: "hd (Bk\<^bsup>Suc n\<^esup>) = Bk"
  3532 lemma [intro]: "hd (Bk\<up>(Suc n)) = Bk"
  3460 apply(simp add: exp_ind_def)
  3533 apply(simp add: )
  3461 done
  3534 done
  3462 
  3535 
  3463 lemma [simp]: "\<lbrakk>tinres r []; r \<noteq> []\<rbrakk> \<Longrightarrow> hd r = Bk"
  3536 lemma [simp]: "\<lbrakk>tinres r []; r \<noteq> []\<rbrakk> \<Longrightarrow> hd r = Bk"
  3464 apply(auto simp: tinres_def)
  3537 apply(auto simp: tinres_def)
  3465 apply(case_tac n, auto)
       
  3466 done
  3538 done
  3467 
  3539 
  3468 lemma [simp]: "\<lbrakk>tinres [] r'; r' \<noteq> []\<rbrakk> \<Longrightarrow> hd r' = Bk"
  3540 lemma [simp]: "\<lbrakk>tinres [] r'; r' \<noteq> []\<rbrakk> \<Longrightarrow> hd r' = Bk"
  3469 apply(auto simp: tinres_def)
  3541 apply(auto simp: tinres_def)
  3470 done
  3542 done
  3471 
  3543 
  3472 lemma [intro]: "\<exists>na. tl r = tl (r @ Bk\<^bsup>n\<^esup>) @ Bk\<^bsup>na\<^esup> \<or> tl (r @ Bk\<^bsup>n\<^esup>) = tl r @ Bk\<^bsup>na\<^esup>"
  3544 lemma [intro]: "\<exists>na. tl r = tl (r @ Bk\<up>(n)) @ Bk\<up>(na) \<or> tl (r @ Bk\<up>(n)) = tl r @ Bk\<up>(na)"
  3473 apply(case_tac r, simp)
  3545 apply(case_tac r, simp)
  3474 apply(case_tac n, simp)
  3546 apply(case_tac n, simp, simp)
  3475 apply(rule_tac x = 0 in exI, simp)
  3547 apply(rule_tac x = nat in exI, simp)
  3476 apply(rule_tac x = nat in exI, simp add: exp_ind_def)
       
  3477 apply(simp)
       
  3478 apply(rule_tac x = n in exI, simp)
  3548 apply(rule_tac x = n in exI, simp)
  3479 done
  3549 done
  3480 
  3550 
  3481 lemma [simp]: "tinres r r' \<Longrightarrow> tinres (tl r) (tl r')"
  3551 lemma [simp]: "tinres r r' \<Longrightarrow> tinres (tl r) (tl r')"
  3482 apply(auto simp: tinres_def)
  3552 apply(auto simp: tinres_def)
  3483 apply(case_tac r', simp_all)
  3553 apply(case_tac r', simp)
  3484 apply(case_tac n, simp_all add: exp_ind_def)
  3554 apply(case_tac n, simp_all)
  3485 apply(rule_tac x = 0 in exI, simp)
  3555 apply(rule_tac x = nat in exI, simp)
  3486 apply(rule_tac x = nat in exI, simp_all)
       
  3487 apply(rule_tac x = n in exI, simp)
  3556 apply(rule_tac x = n in exI, simp)
  3488 done
  3557 done
  3489 
  3558 
  3490 lemma [simp]: "\<lbrakk>tinres r [];  r \<noteq> []\<rbrakk> \<Longrightarrow> tinres (tl r) []"
  3559 lemma [simp]: "\<lbrakk>tinres r [];  r \<noteq> []\<rbrakk> \<Longrightarrow> tinres (tl r) []"
  3491 apply(case_tac r, auto simp: tinres_def)
  3560 apply(case_tac r, auto simp: tinres_def)
  3492 apply(case_tac n, simp_all add: exp_ind_def)
  3561 apply(case_tac n, simp_all add: )
  3493 apply(rule_tac x = nat in exI, simp)
  3562 apply(rule_tac x = nat in exI, simp)
  3494 done
  3563 done
  3495 
  3564 
  3496 lemma [simp]: "\<lbrakk>tinres [] r'\<rbrakk> \<Longrightarrow> tinres [] (tl r')"
  3565 lemma [simp]: "\<lbrakk>tinres [] r'\<rbrakk> \<Longrightarrow> tinres [] (tl r')"
  3497 apply(case_tac r', auto simp: tinres_def)
  3566 apply(case_tac r', auto simp: tinres_def)
  3498 apply(case_tac n, simp_all add: exp_ind_def)
  3567 apply(case_tac n, simp_all add: )
  3499 apply(rule_tac x = nat in exI, simp)
  3568 apply(rule_tac x = nat in exI, simp)
  3500 done
  3569 done
  3501 
  3570 
  3502 lemma [simp]: "tinres r r' \<Longrightarrow> tinres (b # r) (b # r')"
  3571 lemma [simp]: "tinres r r' \<Longrightarrow> tinres (b # r) (b # r')"
  3503 apply(auto simp: tinres_def)
  3572 apply(auto simp: tinres_def)
  3504 done
  3573 done
  3505 
  3574 
       
  3575 lemma [simp]: "tinres r [] \<Longrightarrow> tinres (Bk # tl r) [Bk]"
       
  3576 apply(auto simp: tinres_def)
       
  3577 done
       
  3578 
       
  3579 lemma [simp]: "tinres r [] \<Longrightarrow> tinres (Oc # tl r) [Oc]"
       
  3580 apply(auto simp: tinres_def)
       
  3581 done
       
  3582 
  3506 lemma tinres_step2: 
  3583 lemma tinres_step2: 
  3507   "\<lbrakk>tinres r r'; tstep (ss, l, r) t = (sa, la, ra); tstep (ss, l, r') t = (sb, lb, rb)\<rbrakk>
  3584   "\<lbrakk>tinres r r'; step0 (ss, l, r) t = (sa, la, ra); step0 (ss, l, r') t = (sb, lb, rb)\<rbrakk>
  3508     \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb"
  3585     \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb"
  3509 apply(case_tac "ss = 0", simp add: tstep_0)
  3586 apply(case_tac "ss = 0", simp add: step_0)
  3510 apply(simp add: tstep.simps [simp del])
  3587 apply(simp add: step.simps [simp del], auto)
  3511 apply(case_tac "fetch t ss (case r of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp)
  3588 apply(case_tac [!] "fetch t ss (read r')", simp)
  3512 apply(auto simp: new_tape.simps)
  3589 apply(auto simp: update.simps)
  3513 apply(simp_all split: taction.splits if_splits)
  3590 apply(case_tac [!] a, auto split: if_splits)
  3514 apply(auto)
  3591 done
  3515 done
       
  3516 
       
  3517 
  3592 
  3518 lemma tinres_steps2: 
  3593 lemma tinres_steps2: 
  3519   "\<lbrakk>tinres r r'; steps (ss, l, r) t stp = (sa, la, ra); steps (ss, l, r') t stp = (sb, lb, rb)\<rbrakk>
  3594   "\<lbrakk>tinres r r'; steps0 (ss, l, r) t stp = (sa, la, ra); steps0 (ss, l, r') t stp = (sb, lb, rb)\<rbrakk>
  3520     \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb"
  3595     \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb"
  3521 apply(induct stp arbitrary: sa la ra sb lb rb, simp add: steps.simps)
  3596 apply(induct stp arbitrary: sa la ra sb lb rb, simp add: steps.simps)
  3522 apply(simp add: tstep_red)
  3597 apply(simp add: step_red)
  3523 apply(case_tac "(steps (ss, l, r) t stp)")
  3598 apply(case_tac "(steps0 (ss, l, r) t stp)")
  3524 apply(case_tac "(steps (ss, l, r') t stp)")
  3599 apply(case_tac "(steps0 (ss, l, r') t stp)")
  3525 proof -
  3600 proof -
  3526   fix stp sa la ra sb lb rb a b c aa ba ca
  3601   fix stp sa la ra sb lb rb a b c aa ba ca
  3527   assume ind: "\<And>sa la ra sb lb rb. \<lbrakk>steps (ss, l, r) t stp = (sa, la, ra); 
  3602   assume ind: "\<And>sa la ra sb lb rb. \<lbrakk>steps0 (ss, l, r) t stp = (sa, la, ra); 
  3528     steps (ss, l, r') t stp = (sb, lb, rb)\<rbrakk> \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb"
  3603     steps0 (ss, l, r') t stp = (sb, lb, rb)\<rbrakk> \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb"
  3529   and h: " tinres r r'" "tstep (steps (ss, l, r) t stp) t = (sa, la, ra)"
  3604   and h: " tinres r r'" "step0 (steps0 (ss, l, r) t stp) t = (sa, la, ra)"
  3530          "tstep (steps (ss, l, r') t stp) t = (sb, lb, rb)" "steps (ss, l, r) t stp = (a, b, c)" 
  3605          "step0 (steps0 (ss, l, r') t stp) t = (sb, lb, rb)" "steps0 (ss, l, r) t stp = (a, b, c)" 
  3531          "steps (ss, l, r') t stp = (aa, ba, ca)"
  3606          "steps0 (ss, l, r') t stp = (aa, ba, ca)"
  3532   have "b = ba \<and> tinres c ca \<and> a = aa"
  3607   have "b = ba \<and> tinres c ca \<and> a = aa"
  3533     apply(rule_tac ind, simp_all add: h)
  3608     apply(rule_tac ind, simp_all add: h)
  3534     done
  3609     done
  3535   thus "la = lb \<and> tinres ra rb \<and> sa = sb"
  3610   thus "la = lb \<and> tinres ra rb \<and> sa = sb"
  3536     apply(rule_tac l = b  and r = c  and ss = a and r' = ca   
  3611     apply(rule_tac l = b  and r = c  and ss = a and r' = ca   
  3537             and t = t in tinres_step2)
  3612             and t = t in tinres_step2)
  3538     using h
  3613     using h
  3539     apply(simp, simp, simp)
  3614     apply(simp, simp, simp)
  3540     done
  3615     done
  3541 qed
  3616 qed
  3542  
  3617 
  3543 definition t_wcode_adjust :: "tprog"
  3618 definition t_wcode_adjust :: "instr list"
  3544   where
  3619   where
  3545   "t_wcode_adjust = [(W1, 1), (R, 2), (Nop, 2), (R, 3), (R, 3), (R, 4), 
  3620   "t_wcode_adjust = [(W1, 1), (R, 2), (Nop, 2), (R, 3), (R, 3), (R, 4), 
  3546                    (L, 8), (L, 5), (L, 6), (W0, 5), (L, 6), (R, 7), 
  3621                    (L, 8), (L, 5), (L, 6), (W0, 5), (L, 6), (R, 7), 
  3547                    (W1, 2), (Nop, 7), (L, 9), (W0, 8), (L, 9), (L, 10), 
  3622                    (W1, 2), (Nop, 7), (L, 9), (W0, 8), (L, 9), (L, 10), 
  3548                     (L, 11), (L, 10), (R, 0), (L, 11)]"
  3623                     (L, 11), (L, 10), (R, 0), (L, 11)]"
  3564 done
  3639 done
  3565 
  3640 
  3566 lemma [simp]: "fetch t_wcode_adjust  (Suc (Suc (Suc 0))) Bk = (R, 3)"
  3641 lemma [simp]: "fetch t_wcode_adjust  (Suc (Suc (Suc 0))) Bk = (R, 3)"
  3567 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3642 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3568 done
  3643 done
  3569    
  3644 
  3570 lemma [simp]: "fetch t_wcode_adjust 4 Bk = (L, 8)"
  3645 lemma [simp]: "fetch t_wcode_adjust (Suc (Suc (Suc (Suc 0)))) Bk = (L, 8)"
  3571 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3646 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_4_eq_4)
  3572 done
  3647 done
  3573 
  3648 
  3574 lemma [simp]: "fetch t_wcode_adjust 4 Oc = (L, 5)"
  3649 lemma [simp]: "fetch t_wcode_adjust (Suc (Suc (Suc (Suc 0)))) Oc = (L, 5)"
  3575 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3650 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_4_eq_4)
  3576 done
  3651 done
       
  3652 
       
  3653 thm numeral_5_eq_5
  3577 
  3654 
  3578 lemma [simp]: "fetch t_wcode_adjust 5 Oc = (W0, 5)"
  3655 lemma [simp]: "fetch t_wcode_adjust 5 Oc = (W0, 5)"
  3579 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3656 apply(simp only: fetch.simps t_wcode_adjust_def nth_of.simps numeral_5_eq_5, simp)
  3580 done
  3657 done
  3581 
  3658 
  3582 lemma [simp]: "fetch t_wcode_adjust 5 Bk = (L, 6)"
  3659 lemma [simp]: "fetch t_wcode_adjust 5 Bk = (L, 6)"
  3583 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3660 apply(simp only: fetch.simps t_wcode_adjust_def nth_of.simps numeral_5_eq_5, auto)
  3584 done
  3661 done
  3585 
  3662 
       
  3663 thm numeral_6_eq_6
  3586 lemma [simp]: "fetch t_wcode_adjust 6 Oc = (R, 7)"
  3664 lemma [simp]: "fetch t_wcode_adjust 6 Oc = (R, 7)"
  3587 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3665 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_6_eq_6)
  3588 done
  3666 done
  3589 
  3667 
  3590 lemma [simp]: "fetch t_wcode_adjust 6 Bk = (L, 6)"
  3668 lemma [simp]: "fetch t_wcode_adjust 6 Bk = (L, 6)"
  3591 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3669 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_6_eq_6)
  3592 done
  3670 done
  3593 
  3671 
  3594 lemma [simp]: "fetch t_wcode_adjust 7 Bk = (W1, 2)"
  3672 lemma [simp]: "fetch t_wcode_adjust 7 Bk = (W1, 2)"
  3595 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3673 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_7_eq_7)
  3596 done
  3674 done
  3597 
  3675 
  3598 lemma [simp]: "fetch t_wcode_adjust 8 Bk = (L, 9)"
  3676 lemma [simp]: "fetch t_wcode_adjust 8 Bk = (L, 9)"
  3599 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3677 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_8_eq_8)
  3600 done
  3678 done
  3601 
  3679 
  3602 lemma [simp]: "fetch t_wcode_adjust 8 Oc = (W0, 8)"
  3680 lemma [simp]: "fetch t_wcode_adjust 8 Oc = (W0, 8)"
  3603 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3681 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_8_eq_8)
  3604 done
  3682 done
  3605 
  3683 
  3606 lemma [simp]: "fetch t_wcode_adjust 9 Oc = (L, 10)"
  3684 lemma [simp]: "fetch t_wcode_adjust 9 Oc = (L, 10)"
  3607 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3685 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_9_eq_9)
  3608 done
  3686 done
  3609 
  3687 
  3610 lemma [simp]: "fetch t_wcode_adjust 9 Bk = (L, 9)"
  3688 lemma [simp]: "fetch t_wcode_adjust 9 Bk = (L, 9)"
  3611 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3689 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_9_eq_9)
  3612 done
  3690 done
  3613 
  3691 
  3614 lemma [simp]: "fetch t_wcode_adjust 10 Bk = (L, 11)"
  3692 lemma [simp]: "fetch t_wcode_adjust 10 Bk = (L, 11)"
  3615 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3693 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps  numeral_10_eq_10)
  3616 done
  3694 done
  3617 
  3695 
  3618 lemma [simp]: "fetch t_wcode_adjust 10 Oc = (L, 10)"
  3696 lemma [simp]: "fetch t_wcode_adjust 10 Oc = (L, 10)"
  3619 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3697 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps eval_nat_numeral)
  3620 done
  3698 done
  3621 
  3699 
  3622 lemma [simp]: "fetch t_wcode_adjust 11 Oc = (L, 11)"
  3700 lemma [simp]: "fetch t_wcode_adjust 11 Oc = (L, 11)"
  3623 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3701 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps eval_nat_numeral)
  3624 done
  3702 done
  3625 
  3703 
  3626 lemma [simp]: "fetch t_wcode_adjust 11 Bk = (R, 0)"
  3704 lemma [simp]: "fetch t_wcode_adjust 11 Bk = (R, 0)"
  3627 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps)
  3705 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps eval_nat_numeral)
  3628 done
  3706 done
  3629 
  3707 
  3630 fun wadjust_start :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3708 fun wadjust_start :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3631   where
  3709   where
  3632   "wadjust_start m rs (l, r) = 
  3710   "wadjust_start m rs (l, r) = 
  3633          (\<exists> ln rn. l = Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  3711          (\<exists> ln rn. l = Bk # Oc\<up>(Suc m) \<and>
  3634                    tl r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  3712                    tl r = Oc # Bk\<up>(ln) @ Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
  3635 
  3713 
  3636 fun wadjust_loop_start :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3714 fun wadjust_loop_start :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3637   where
  3715   where
  3638   "wadjust_loop_start m rs (l, r) = 
  3716   "wadjust_loop_start m rs (l, r) = 
  3639           (\<exists> ln rn ml mr. l = Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup>  \<and>
  3717           (\<exists> ln rn ml mr. l = Oc\<up>(ml) @ Bk # Oc\<up>(Suc m)  \<and>
  3640                           r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
  3718                           r = Oc # Bk\<up>(ln) @ Bk # Oc\<up>(mr) @ Bk\<up>(rn) \<and>
  3641                           ml + mr = Suc (Suc rs) \<and> mr > 0)"
  3719                           ml + mr = Suc (Suc rs) \<and> mr > 0)"
  3642 
  3720 
  3643 fun wadjust_loop_right_move :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3721 fun wadjust_loop_right_move :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3644   where
  3722   where
  3645   "wadjust_loop_right_move m rs (l, r) = 
  3723   "wadjust_loop_right_move m rs (l, r) = 
  3646    (\<exists> ml mr nl nr rn. l = Bk\<^bsup>nl\<^esup> @ Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  3724    (\<exists> ml mr nl nr rn. l = Bk\<up>(nl) @ Oc # Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and>
  3647                       r = Bk\<^bsup>nr\<^esup> @ Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
  3725                       r = Bk\<up>(nr) @ Oc\<up>(mr) @ Bk\<up>(rn) \<and>
  3648                       ml + mr = Suc (Suc rs) \<and> mr > 0 \<and>
  3726                       ml + mr = Suc (Suc rs) \<and> mr > 0 \<and>
  3649                       nl + nr > 0)"
  3727                       nl + nr > 0)"
  3650 
  3728 
  3651 fun wadjust_loop_check :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3729 fun wadjust_loop_check :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3652   where
  3730   where
  3653   "wadjust_loop_check m rs (l, r) = 
  3731   "wadjust_loop_check m rs (l, r) = 
  3654   (\<exists> ml mr ln rn. l = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  3732   (\<exists> ml mr ln rn. l = Oc # Bk\<up>(ln) @ Bk # Oc # Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and>
  3655                   r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> ml + mr = (Suc rs))"
  3733                   r = Oc\<up>(mr) @ Bk\<up>(rn) \<and> ml + mr = (Suc rs))"
  3656 
  3734 
  3657 fun wadjust_loop_erase :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3735 fun wadjust_loop_erase :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3658   where
  3736   where
  3659   "wadjust_loop_erase m rs (l, r) = 
  3737   "wadjust_loop_erase m rs (l, r) = 
  3660     (\<exists> ml mr ln rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  3738     (\<exists> ml mr ln rn. l = Bk\<up>(ln) @ Bk # Oc # Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and>
  3661                     tl r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> ml + mr = (Suc rs) \<and> mr > 0)"
  3739                     tl r = Oc\<up>(mr) @ Bk\<up>(rn) \<and> ml + mr = (Suc rs) \<and> mr > 0)"
  3662 
  3740 
  3663 fun wadjust_loop_on_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3741 fun wadjust_loop_on_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3664   where
  3742   where
  3665   "wadjust_loop_on_left_moving_O m rs (l, r) = 
  3743   "wadjust_loop_on_left_moving_O m rs (l, r) = 
  3666       (\<exists> ml mr ln rn. l = Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m \<^esup>\<and>
  3744       (\<exists> ml mr ln rn. l = Oc\<up>(ml) @ Bk # Oc\<up>(Suc m )\<and>
  3667                       r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
  3745                       r = Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(mr) @ Bk\<up>(rn) \<and>
  3668                       ml + mr = Suc rs \<and> mr > 0)"
  3746                       ml + mr = Suc rs \<and> mr > 0)"
  3669 
  3747 
  3670 fun wadjust_loop_on_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3748 fun wadjust_loop_on_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3671   where
  3749   where
  3672   "wadjust_loop_on_left_moving_B m rs (l, r) = 
  3750   "wadjust_loop_on_left_moving_B m rs (l, r) = 
  3673       (\<exists> ml mr nl nr rn. l = Bk\<^bsup>nl\<^esup> @ Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  3751       (\<exists> ml mr nl nr rn. l = Bk\<up>(nl) @ Oc # Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and>
  3674                          r = Bk\<^bsup>nr\<^esup> @ Bk # Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
  3752                          r = Bk\<up>(nr) @ Bk # Bk # Oc\<up>(mr) @ Bk\<up>(rn) \<and> 
  3675                          ml + mr = Suc rs \<and> mr > 0)"
  3753                          ml + mr = Suc rs \<and> mr > 0)"
  3676 
  3754 
  3677 fun wadjust_loop_on_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3755 fun wadjust_loop_on_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3678   where
  3756   where
  3679   "wadjust_loop_on_left_moving m rs (l, r) = 
  3757   "wadjust_loop_on_left_moving m rs (l, r) = 
  3681        wadjust_loop_on_left_moving_B m rs (l, r))"
  3759        wadjust_loop_on_left_moving_B m rs (l, r))"
  3682 
  3760 
  3683 fun wadjust_loop_right_move2 :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3761 fun wadjust_loop_right_move2 :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3684   where
  3762   where
  3685   "wadjust_loop_right_move2 m rs (l, r) = 
  3763   "wadjust_loop_right_move2 m rs (l, r) = 
  3686         (\<exists> ml mr ln rn. l = Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  3764         (\<exists> ml mr ln rn. l = Oc # Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and>
  3687                         r = Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and>
  3765                         r = Bk\<up>(ln) @ Bk # Bk # Oc\<up>(mr) @ Bk\<up>(rn) \<and>
  3688                         ml + mr = Suc rs \<and> mr > 0)"
  3766                         ml + mr = Suc rs \<and> mr > 0)"
  3689 
  3767 
  3690 fun wadjust_erase2 :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3768 fun wadjust_erase2 :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3691   where
  3769   where
  3692   "wadjust_erase2 m rs (l, r) = 
  3770   "wadjust_erase2 m rs (l, r) = 
  3693      (\<exists> ln rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Oc # Oc\<^bsup>Suc rs\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  3771      (\<exists> ln rn. l = Bk\<up>(ln) @ Bk # Oc # Oc\<up>(Suc rs) @ Bk # Oc\<up>(Suc m) \<and>
  3694                      tl r = Bk\<^bsup>rn\<^esup>)"
  3772                      tl r = Bk\<up>(rn))"
  3695 
  3773 
  3696 fun wadjust_on_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3774 fun wadjust_on_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3697   where
  3775   where
  3698   "wadjust_on_left_moving_O m rs (l, r) = 
  3776   "wadjust_on_left_moving_O m rs (l, r) = 
  3699         (\<exists> rn. l = Oc\<^bsup>Suc rs\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  3777         (\<exists> rn. l = Oc\<up>(Suc rs) @ Bk # Oc\<up>(Suc m) \<and>
  3700                   r = Oc # Bk\<^bsup>rn\<^esup>)"
  3778                   r = Oc # Bk\<up>(rn))"
  3701 
  3779 
  3702 fun wadjust_on_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3780 fun wadjust_on_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3703   where
  3781   where
  3704   "wadjust_on_left_moving_B m rs (l, r) = 
  3782   "wadjust_on_left_moving_B m rs (l, r) = 
  3705          (\<exists> ln rn. l = Bk\<^bsup>ln\<^esup> @ Oc # Oc\<^bsup>Suc rs\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  3783          (\<exists> ln rn. l = Bk\<up>(ln) @ Oc # Oc\<up>(Suc rs) @ Bk # Oc\<up>(Suc m) \<and>
  3706                    r = Bk\<^bsup>rn\<^esup>)"
  3784                    r = Bk\<up>(rn))"
  3707 
  3785 
  3708 fun wadjust_on_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3786 fun wadjust_on_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3709   where
  3787   where
  3710   "wadjust_on_left_moving m rs (l, r) = 
  3788   "wadjust_on_left_moving m rs (l, r) = 
  3711       (wadjust_on_left_moving_O m rs (l, r) \<or>
  3789       (wadjust_on_left_moving_O m rs (l, r) \<or>
  3712        wadjust_on_left_moving_B m rs (l, r))"
  3790        wadjust_on_left_moving_B m rs (l, r))"
  3713 
  3791 
  3714 fun wadjust_goon_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3792 fun wadjust_goon_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3715   where 
  3793   where 
  3716   "wadjust_goon_left_moving_B m rs (l, r) = 
  3794   "wadjust_goon_left_moving_B m rs (l, r) = 
  3717         (\<exists> rn. l = Oc\<^bsup>Suc m\<^esup> \<and> 
  3795         (\<exists> rn. l = Oc\<up>(Suc m) \<and> 
  3718                r = Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  3796                r = Bk # Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn))"
  3719 
  3797 
  3720 fun wadjust_goon_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3798 fun wadjust_goon_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3721   where
  3799   where
  3722   "wadjust_goon_left_moving_O m rs (l, r) = 
  3800   "wadjust_goon_left_moving_O m rs (l, r) = 
  3723       (\<exists> ml mr rn. l = Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  3801       (\<exists> ml mr rn. l = Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and>
  3724                       r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
  3802                       r = Oc\<up>(mr) @ Bk\<up>(rn) \<and> 
  3725                       ml + mr = Suc (Suc rs) \<and> mr > 0)"
  3803                       ml + mr = Suc (Suc rs) \<and> mr > 0)"
  3726 
  3804 
  3727 fun wadjust_goon_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3805 fun wadjust_goon_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3728   where
  3806   where
  3729   "wadjust_goon_left_moving m rs (l, r) = 
  3807   "wadjust_goon_left_moving m rs (l, r) = 
  3732 
  3810 
  3733 fun wadjust_backto_standard_pos_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3811 fun wadjust_backto_standard_pos_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3734   where
  3812   where
  3735   "wadjust_backto_standard_pos_B m rs (l, r) =
  3813   "wadjust_backto_standard_pos_B m rs (l, r) =
  3736         (\<exists> rn. l = [] \<and> 
  3814         (\<exists> rn. l = [] \<and> 
  3737                r = Bk # Oc\<^bsup>Suc m \<^esup>@ Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  3815                r = Bk # Oc\<up>(Suc m )@ Bk # Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn))"
  3738 
  3816 
  3739 fun wadjust_backto_standard_pos_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3817 fun wadjust_backto_standard_pos_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3740   where
  3818   where
  3741   "wadjust_backto_standard_pos_O m rs (l, r) = 
  3819   "wadjust_backto_standard_pos_O m rs (l, r) = 
  3742       (\<exists> ml mr rn. l = Oc\<^bsup>ml\<^esup> \<and>
  3820       (\<exists> ml mr rn. l = Oc\<up>(ml) \<and>
  3743                       r = Oc\<^bsup>mr\<^esup> @ Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> 
  3821                       r = Oc\<up>(mr) @ Bk # Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn) \<and> 
  3744                       ml + mr = Suc m \<and> mr > 0)"
  3822                       ml + mr = Suc m \<and> mr > 0)"
  3745 
  3823 
  3746 fun wadjust_backto_standard_pos :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3824 fun wadjust_backto_standard_pos :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3747   where
  3825   where
  3748   "wadjust_backto_standard_pos m rs (l, r) = 
  3826   "wadjust_backto_standard_pos m rs (l, r) = 
  3751 
  3829 
  3752 fun wadjust_stop :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3830 fun wadjust_stop :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
  3753 where
  3831 where
  3754   "wadjust_stop m rs (l, r) =
  3832   "wadjust_stop m rs (l, r) =
  3755         (\<exists> rn. l = [Bk] \<and> 
  3833         (\<exists> rn. l = [Bk] \<and> 
  3756                r = Oc\<^bsup>Suc m \<^esup>@ Bk # Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  3834                r = Oc\<up>(Suc m )@ Bk # Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn))"
  3757 
  3835 
  3758 declare wadjust_start.simps[simp del]  wadjust_loop_start.simps[simp del]
  3836 declare wadjust_start.simps[simp del]  wadjust_loop_start.simps[simp del]
  3759         wadjust_loop_right_move.simps[simp del]  wadjust_loop_check.simps[simp del]
  3837         wadjust_loop_right_move.simps[simp del]  wadjust_loop_check.simps[simp del]
  3760         wadjust_loop_erase.simps[simp del] wadjust_loop_on_left_moving.simps[simp del]
  3838         wadjust_loop_erase.simps[simp del] wadjust_loop_on_left_moving.simps[simp del]
  3761         wadjust_loop_right_move2.simps[simp del] wadjust_erase2.simps[simp del]
  3839         wadjust_loop_right_move2.simps[simp del] wadjust_erase2.simps[simp del]
  3783         else False
  3861         else False
  3784 )"
  3862 )"
  3785 
  3863 
  3786 declare wadjust_inv.simps[simp del]
  3864 declare wadjust_inv.simps[simp del]
  3787 
  3865 
  3788 fun wadjust_phase :: "nat \<Rightarrow> t_conf \<Rightarrow> nat"
  3866 fun wadjust_phase :: "nat \<Rightarrow> config \<Rightarrow> nat"
  3789   where
  3867   where
  3790   "wadjust_phase rs (st, l, r) = 
  3868   "wadjust_phase rs (st, l, r) = 
  3791          (if st = 1 then 3 
  3869          (if st = 1 then 3 
  3792           else if st \<ge> 2 \<and> st \<le> 7 then 2
  3870           else if st \<ge> 2 \<and> st \<le> 7 then 2
  3793           else if st \<ge> 8 \<and> st \<le> 11 then 1
  3871           else if st \<ge> 8 \<and> st \<le> 11 then 1
  3794           else 0)"
  3872           else 0)"
  3795 
  3873 
  3796 thm dropWhile.simps
  3874 fun wadjust_stage :: "nat \<Rightarrow> config \<Rightarrow> nat"
  3797 
       
  3798 fun wadjust_stage :: "nat \<Rightarrow> t_conf \<Rightarrow> nat"
       
  3799   where
  3875   where
  3800   "wadjust_stage rs (st, l, r) = 
  3876   "wadjust_stage rs (st, l, r) = 
  3801            (if st \<ge> 2 \<and> st \<le> 7 then 
  3877            (if st \<ge> 2 \<and> st \<le> 7 then 
  3802                   rs - length (takeWhile (\<lambda> a. a = Oc) 
  3878                   rs - length (takeWhile (\<lambda> a. a = Oc) 
  3803                           (tl (dropWhile (\<lambda> a. a = Oc) (rev l @ r))))
  3879                           (tl (dropWhile (\<lambda> a. a = Oc) (rev l @ r))))
  3804             else 0)"
  3880             else 0)"
  3805 
  3881 
  3806 fun wadjust_state :: "nat \<Rightarrow> t_conf \<Rightarrow> nat"
  3882 fun wadjust_state :: "nat \<Rightarrow> config \<Rightarrow> nat"
  3807   where
  3883   where
  3808   "wadjust_state rs (st, l, r) = 
  3884   "wadjust_state rs (st, l, r) = 
  3809        (if st \<ge> 2 \<and> st \<le> 7 then 8 - st
  3885        (if st \<ge> 2 \<and> st \<le> 7 then 8 - st
  3810         else if st \<ge> 8 \<and> st \<le> 11 then 12 - st
  3886         else if st \<ge> 8 \<and> st \<le> 11 then 12 - st
  3811         else 0)"
  3887         else 0)"
  3812 
  3888 
  3813 fun wadjust_step :: "nat \<Rightarrow> t_conf \<Rightarrow> nat"
  3889 fun wadjust_step :: "nat \<Rightarrow> config \<Rightarrow> nat"
  3814   where
  3890   where
  3815   "wadjust_step rs (st, l, r) = 
  3891   "wadjust_step rs (st, l, r) = 
  3816        (if st = 1 then (if hd r = Bk then 1
  3892        (if st = 1 then (if hd r = Bk then 1
  3817                         else 0) 
  3893                         else 0) 
  3818         else if st = 3 then length r
  3894         else if st = 3 then length r
  3825         else if st = 10 then length l
  3901         else if st = 10 then length l
  3826         else if st = 11 then (if hd r = Bk then 0
  3902         else if st = 11 then (if hd r = Bk then 0
  3827                               else Suc (length l))
  3903                               else Suc (length l))
  3828         else 0)"
  3904         else 0)"
  3829 
  3905 
  3830 fun wadjust_measure :: "(nat \<times> t_conf) \<Rightarrow> nat \<times> nat \<times> nat \<times> nat"
  3906 fun wadjust_measure :: "(nat \<times> config) \<Rightarrow> nat \<times> nat \<times> nat \<times> nat"
  3831   where
  3907   where
  3832   "wadjust_measure (rs, (st, l, r)) = 
  3908   "wadjust_measure (rs, (st, l, r)) = 
  3833      (wadjust_phase rs (st, l, r), 
  3909      (wadjust_phase rs (st, l, r), 
  3834       wadjust_stage rs (st, l, r),
  3910       wadjust_stage rs (st, l, r),
  3835       wadjust_state rs (st, l, r), 
  3911       wadjust_state rs (st, l, r), 
  3836       wadjust_step rs (st, l, r))"
  3912       wadjust_step rs (st, l, r))"
  3837 
  3913 
  3838 definition wadjust_le :: "((nat \<times> t_conf) \<times> nat \<times> t_conf) set"
  3914 definition wadjust_le :: "((nat \<times> config) \<times> nat \<times> config) set"
  3839   where "wadjust_le \<equiv> (inv_image lex_square wadjust_measure)"
  3915   where "wadjust_le \<equiv> (inv_image lex_square wadjust_measure)"
  3840 
  3916 
  3841 lemma [intro]: "wf lex_square"
  3917 lemma [intro]: "wf lex_square"
  3842 by(auto intro:wf_lex_prod simp: abacus.lex_pair_def lex_square_def 
  3918 by(auto intro:wf_lex_prod simp: abacus.lex_pair_def lex_square_def 
  3843   abacus.lex_triple_def)
  3919   abacus.lex_triple_def)
  3856 
  3932 
  3857 lemma [simp]: "wadjust_loop_right_move m rs (c, [])
  3933 lemma [simp]: "wadjust_loop_right_move m rs (c, [])
  3858         \<Longrightarrow>  wadjust_loop_check m rs (Bk # c, [])"
  3934         \<Longrightarrow>  wadjust_loop_check m rs (Bk # c, [])"
  3859 apply(simp only: wadjust_loop_right_move.simps wadjust_loop_check.simps)
  3935 apply(simp only: wadjust_loop_right_move.simps wadjust_loop_check.simps)
  3860 apply(auto)
  3936 apply(auto)
  3861 apply(case_tac [!] mr, simp_all add: exp_ind_def)
       
  3862 done
  3937 done
  3863 
  3938 
  3864 lemma [simp]: "wadjust_loop_check m rs (c, []) \<Longrightarrow> c \<noteq> []"
  3939 lemma [simp]: "wadjust_loop_check m rs (c, []) \<Longrightarrow> c \<noteq> []"
  3865 apply(simp only: wadjust_loop_check.simps, auto)
  3940 apply(simp only: wadjust_loop_check.simps, auto)
  3866 done
  3941 done
  3872 lemma [simp]: "wadjust_loop_right_move m rs (c, []) \<Longrightarrow> 
  3947 lemma [simp]: "wadjust_loop_right_move m rs (c, []) \<Longrightarrow> 
  3873   wadjust_loop_right_move m rs (Bk # c, [])"
  3948   wadjust_loop_right_move m rs (Bk # c, [])"
  3874 apply(simp only: wadjust_loop_right_move.simps)
  3949 apply(simp only: wadjust_loop_right_move.simps)
  3875 apply(erule_tac exE)+
  3950 apply(erule_tac exE)+
  3876 apply(auto)
  3951 apply(auto)
  3877 apply(case_tac [!] mr, simp_all add: exp_ind_def)
       
  3878 done
  3952 done
  3879 
  3953 
  3880 lemma [simp]: "wadjust_loop_check m rs (c, []) \<Longrightarrow> wadjust_erase2 m rs (tl c, [hd c])"
  3954 lemma [simp]: "wadjust_loop_check m rs (c, []) \<Longrightarrow> wadjust_erase2 m rs (tl c, [hd c])"
  3881 apply(simp only: wadjust_loop_check.simps wadjust_erase2.simps, auto)
  3955 apply(simp only: wadjust_loop_check.simps wadjust_erase2.simps, auto)
  3882 apply(case_tac mr, simp_all add: exp_ind_def, auto)
       
  3883 done
  3956 done
  3884 
  3957 
  3885 lemma [simp]: " wadjust_loop_erase m rs (c, [])
  3958 lemma [simp]: " wadjust_loop_erase m rs (c, [])
  3886     \<Longrightarrow> (c = [] \<longrightarrow> wadjust_loop_on_left_moving m rs ([], [Bk])) \<and>
  3959     \<Longrightarrow> (c = [] \<longrightarrow> wadjust_loop_on_left_moving m rs ([], [Bk])) \<and>
  3887         (c \<noteq> [] \<longrightarrow> wadjust_loop_on_left_moving m rs (tl c, [hd c]))"
  3960         (c \<noteq> [] \<longrightarrow> wadjust_loop_on_left_moving m rs (tl c, [hd c]))"
  3888 apply(simp add: wadjust_loop_erase.simps, auto)
  3961 apply(simp add: wadjust_loop_erase.simps)
  3889 apply(case_tac [!] mr, simp_all add: exp_ind_def)
       
  3890 done
  3962 done
  3891 
  3963 
  3892 lemma [simp]: "wadjust_loop_on_left_moving m rs (c, []) = False"
  3964 lemma [simp]: "wadjust_loop_on_left_moving m rs (c, []) = False"
  3893 apply(auto simp: wadjust_loop_on_left_moving.simps)
  3965 apply(auto simp: wadjust_loop_on_left_moving.simps)
  3894 done
  3966 done
  3901 lemma [simp]: "wadjust_erase2 m rs ([], []) = False"
  3973 lemma [simp]: "wadjust_erase2 m rs ([], []) = False"
  3902 apply(auto simp: wadjust_erase2.simps)
  3974 apply(auto simp: wadjust_erase2.simps)
  3903 done
  3975 done
  3904 
  3976 
  3905 lemma [simp]: "wadjust_on_left_moving_B m rs 
  3977 lemma [simp]: "wadjust_on_left_moving_B m rs 
  3906                  (Oc # Oc # Oc\<^bsup>rs\<^esup> @ Bk # Oc # Oc\<^bsup>m\<^esup>, [Bk])"
  3978                  (Oc # Oc # Oc\<up>(rs) @ Bk # Oc # Oc\<up>(m), [Bk])"
  3907 apply(simp add: wadjust_on_left_moving_B.simps, auto)
  3979 apply(simp add: wadjust_on_left_moving_B.simps, auto)
  3908 apply(rule_tac x = 0 in exI, simp add: exp_ind_def)
       
  3909 done
  3980 done
  3910 
  3981 
  3911 lemma [simp]: "wadjust_on_left_moving_B m rs 
  3982 lemma [simp]: "wadjust_on_left_moving_B m rs 
  3912                  (Bk\<^bsup>n\<^esup> @ Bk # Oc # Oc # Oc\<^bsup>rs\<^esup> @ Bk # Oc # Oc\<^bsup>m\<^esup>, [Bk])"
  3983                  (Bk\<up>(n) @ Bk # Oc # Oc # Oc\<up>(rs) @ Bk # Oc # Oc\<up>(m), [Bk])"
  3913 apply(simp add: wadjust_on_left_moving_B.simps exp_ind_def, auto)
  3984 apply(simp add: wadjust_on_left_moving_B.simps , auto)
  3914 apply(rule_tac x = "Suc n" in exI, simp add: exp_ind)
  3985 apply(rule_tac x = "Suc n" in exI, simp add: exp_ind del: replicate_Suc)
  3915 done
  3986 done
  3916 
  3987 
  3917 lemma [simp]: "\<lbrakk>wadjust_erase2 m rs (c, []); c \<noteq> []\<rbrakk> \<Longrightarrow>
  3988 lemma [simp]: "\<lbrakk>wadjust_erase2 m rs (c, []); c \<noteq> []\<rbrakk> \<Longrightarrow>
  3918             wadjust_on_left_moving m rs (tl c, [hd c])"
  3989             wadjust_on_left_moving m rs (tl c, [hd c])"
  3919 apply(simp only: wadjust_erase2.simps)
  3990 apply(simp only: wadjust_erase2.simps)
  3920 apply(erule_tac exE)+
  3991 apply(erule_tac exE)+
  3921 apply(case_tac ln, simp_all add: exp_ind_def wadjust_on_left_moving.simps)
  3992 apply(case_tac ln, simp_all add:  wadjust_on_left_moving.simps)
  3922 done
  3993 done
  3923 
  3994 
  3924 lemma [simp]: "wadjust_erase2 m rs (c, [])
  3995 lemma [simp]: "wadjust_erase2 m rs (c, [])
  3925     \<Longrightarrow> (c = [] \<longrightarrow> wadjust_on_left_moving m rs ([], [Bk])) \<and> 
  3996     \<Longrightarrow> (c = [] \<longrightarrow> wadjust_on_left_moving m rs ([], [Bk])) \<and> 
  3926        (c \<noteq> [] \<longrightarrow> wadjust_on_left_moving m rs (tl c, [hd c]))"
  3997        (c \<noteq> [] \<longrightarrow> wadjust_on_left_moving m rs (tl c, [hd c]))"
  3937 done
  4008 done
  3938 
  4009 
  3939 lemma [simp]: " \<lbrakk>wadjust_on_left_moving_B m rs (c, []); c \<noteq> []; hd c = Bk\<rbrakk> \<Longrightarrow>
  4010 lemma [simp]: " \<lbrakk>wadjust_on_left_moving_B m rs (c, []); c \<noteq> []; hd c = Bk\<rbrakk> \<Longrightarrow>
  3940                                       wadjust_on_left_moving_B m rs (tl c, [Bk])"
  4011                                       wadjust_on_left_moving_B m rs (tl c, [Bk])"
  3941 apply(simp add: wadjust_on_left_moving_B.simps, auto)
  4012 apply(simp add: wadjust_on_left_moving_B.simps, auto)
  3942 apply(case_tac [!] ln, simp_all add: exp_ind_def, auto)
  4013 apply(case_tac [!] ln, simp_all)
  3943 done
  4014 done
  3944 
  4015 
  3945 lemma [simp]: "\<lbrakk>wadjust_on_left_moving_B m rs (c, []); c \<noteq> []; hd c = Oc\<rbrakk> \<Longrightarrow>
  4016 lemma [simp]: "\<lbrakk>wadjust_on_left_moving_B m rs (c, []); c \<noteq> []; hd c = Oc\<rbrakk> \<Longrightarrow>
  3946                                   wadjust_on_left_moving_O m rs (tl c, [Oc])"
  4017                                   wadjust_on_left_moving_O m rs (tl c, [Oc])"
  3947 apply(simp add: wadjust_on_left_moving_B.simps wadjust_on_left_moving_O.simps, auto)
  4018 apply(simp add: wadjust_on_left_moving_B.simps wadjust_on_left_moving_O.simps, auto)
  3948 apply(case_tac [!] ln, simp_all add: exp_ind_def)
  4019 apply(case_tac [!] ln, simp_all add: )
  3949 done
  4020 done
  3950 
  4021 
  3951 lemma [simp]: "\<lbrakk>wadjust_on_left_moving m rs (c, []); c \<noteq> []\<rbrakk> \<Longrightarrow> 
  4022 lemma [simp]: "\<lbrakk>wadjust_on_left_moving m rs (c, []); c \<noteq> []\<rbrakk> \<Longrightarrow> 
  3952   wadjust_on_left_moving m rs (tl c, [hd c])"
  4023   wadjust_on_left_moving m rs (tl c, [hd c])"
  3953 apply(simp add: wadjust_on_left_moving.simps)
  4024 apply(simp add: wadjust_on_left_moving.simps)
  3989     \<Longrightarrow> wadjust_loop_right_move m rs (Bk # c, list)"
  4060     \<Longrightarrow> wadjust_loop_right_move m rs (Bk # c, list)"
  3990 apply(simp only: wadjust_loop_right_move.simps)
  4061 apply(simp only: wadjust_loop_right_move.simps)
  3991 apply(erule_tac exE)+
  4062 apply(erule_tac exE)+
  3992 apply(rule_tac x = ml in exI, simp)
  4063 apply(rule_tac x = ml in exI, simp)
  3993 apply(rule_tac x = mr in exI, simp)
  4064 apply(rule_tac x = mr in exI, simp)
  3994 apply(rule_tac x = "Suc nl" in exI, simp add: exp_ind_def)
  4065 apply(rule_tac x = "Suc nl" in exI, simp add: )
  3995 apply(case_tac nr, simp, case_tac mr, simp_all add: exp_ind_def)
  4066 apply(case_tac nr, simp, case_tac mr, simp_all add: )
  3996 apply(rule_tac x = nat in exI, auto)
  4067 apply(rule_tac x = nat in exI, auto)
  3997 done
  4068 done
  3998 
  4069 
  3999 lemma [simp]: "wadjust_loop_check m rs (c, b) \<Longrightarrow> c \<noteq> []"
  4070 lemma [simp]: "wadjust_loop_check m rs (c, b) \<Longrightarrow> c \<noteq> []"
  4000 apply(simp only: wadjust_loop_check.simps, auto)
  4071 apply(simp only: wadjust_loop_check.simps, auto)
  4001 done
  4072 done
  4002 
  4073 
  4003 lemma [simp]: "wadjust_loop_check m rs (c, Bk # list)
  4074 lemma [simp]: "wadjust_loop_check m rs (c, Bk # list)
  4004               \<Longrightarrow>  wadjust_erase2 m rs (tl c, hd c # Bk # list)"
  4075               \<Longrightarrow>  wadjust_erase2 m rs (tl c, hd c # Bk # list)"
  4005 apply(auto simp: wadjust_loop_check.simps wadjust_erase2.simps)
  4076 apply(auto simp: wadjust_loop_check.simps wadjust_erase2.simps)
  4006 apply(case_tac [!] mr, simp_all add: exp_ind_def, auto)
  4077 apply(case_tac [!] mr, simp_all)
  4007 done
  4078 done
  4008 
  4079 
  4009 lemma [simp]: "wadjust_loop_erase m rs (c, b) \<Longrightarrow> c \<noteq> []"
  4080 lemma [simp]: "wadjust_loop_erase m rs (c, b) \<Longrightarrow> c \<noteq> []"
  4010 apply(simp only: wadjust_loop_erase.simps, auto)
  4081 apply(simp only: wadjust_loop_erase.simps, auto)
  4011 done
  4082 done
  4018 apply(simp only: wadjust_loop_erase.simps 
  4089 apply(simp only: wadjust_loop_erase.simps 
  4019   wadjust_loop_on_left_moving_B.simps)
  4090   wadjust_loop_on_left_moving_B.simps)
  4020 apply(erule_tac exE)+
  4091 apply(erule_tac exE)+
  4021 apply(rule_tac x = ml in exI, rule_tac x = mr in exI, 
  4092 apply(rule_tac x = ml in exI, rule_tac x = mr in exI, 
  4022       rule_tac x = ln in exI, rule_tac x = 0 in exI, simp)
  4093       rule_tac x = ln in exI, rule_tac x = 0 in exI, simp)
  4023 apply(case_tac ln, simp_all add: exp_ind_def, auto)
  4094 apply(case_tac ln, simp_all add: , auto)
  4024 apply(simp add: exp_ind exp_ind_def[THEN sym])
  4095 apply(simp add: exp_ind [THEN sym])
  4025 done
  4096 done
  4026 
  4097 
  4027 lemma [simp]: "\<lbrakk>wadjust_loop_erase m rs (c, Bk # list); c \<noteq> []; hd c = Oc\<rbrakk> \<Longrightarrow>
  4098 lemma [simp]: "\<lbrakk>wadjust_loop_erase m rs (c, Bk # list); c \<noteq> []; hd c = Oc\<rbrakk> \<Longrightarrow>
  4028              wadjust_loop_on_left_moving_O m rs (tl c, Oc # Bk # list)"
  4099              wadjust_loop_on_left_moving_O m rs (tl c, Oc # Bk # list)"
  4029 apply(simp only: wadjust_loop_erase.simps wadjust_loop_on_left_moving_O.simps,
  4100 apply(simp only: wadjust_loop_erase.simps wadjust_loop_on_left_moving_O.simps,
  4030        auto)
  4101        auto)
  4031 apply(case_tac [!] ln, simp_all add: exp_ind_def)
  4102 apply(case_tac [!] ln, simp_all add: )
  4032 done
  4103 done
  4033 
  4104 
  4034 lemma [simp]: "\<lbrakk>wadjust_loop_erase m rs (c, Bk # list); c \<noteq> []\<rbrakk> \<Longrightarrow> 
  4105 lemma [simp]: "\<lbrakk>wadjust_loop_erase m rs (c, Bk # list); c \<noteq> []\<rbrakk> \<Longrightarrow> 
  4035                 wadjust_loop_on_left_moving m rs (tl c, hd c # Bk # list)"
  4106                 wadjust_loop_on_left_moving m rs (tl c, hd c # Bk # list)"
  4036 apply(case_tac "hd c", simp_all add:wadjust_loop_on_left_moving.simps)
  4107 apply(case_tac "hd c", simp_all add:wadjust_loop_on_left_moving.simps)
  4048 lemma [simp]: "\<lbrakk>wadjust_loop_on_left_moving_B m rs (c, Bk # list); hd c = Bk\<rbrakk>
  4119 lemma [simp]: "\<lbrakk>wadjust_loop_on_left_moving_B m rs (c, Bk # list); hd c = Bk\<rbrakk>
  4049     \<Longrightarrow>  wadjust_loop_on_left_moving_B m rs (tl c, Bk # Bk # list)"
  4120     \<Longrightarrow>  wadjust_loop_on_left_moving_B m rs (tl c, Bk # Bk # list)"
  4050 apply(simp only: wadjust_loop_on_left_moving_B.simps)
  4121 apply(simp only: wadjust_loop_on_left_moving_B.simps)
  4051 apply(erule_tac exE)+
  4122 apply(erule_tac exE)+
  4052 apply(rule_tac x = ml in exI, rule_tac x = mr in exI)
  4123 apply(rule_tac x = ml in exI, rule_tac x = mr in exI)
  4053 apply(case_tac nl, simp_all add: exp_ind_def, auto)
  4124 apply(case_tac nl, simp_all add: , auto)
  4054 apply(rule_tac x = "Suc nr" in exI, auto simp: exp_ind_def)
  4125 apply(rule_tac x = "Suc nr" in exI, auto simp: )
  4055 done
  4126 done
  4056 
  4127 
  4057 lemma [simp]: "\<lbrakk>wadjust_loop_on_left_moving_B m rs (c, Bk # list); hd c = Oc\<rbrakk>
  4128 lemma [simp]: "\<lbrakk>wadjust_loop_on_left_moving_B m rs (c, Bk # list); hd c = Oc\<rbrakk>
  4058     \<Longrightarrow> wadjust_loop_on_left_moving_O m rs (tl c, Oc # Bk # list)"
  4129     \<Longrightarrow> wadjust_loop_on_left_moving_O m rs (tl c, Oc # Bk # list)"
  4059 apply(simp only: wadjust_loop_on_left_moving_O.simps 
  4130 apply(simp only: wadjust_loop_on_left_moving_O.simps 
  4060                  wadjust_loop_on_left_moving_B.simps)
  4131                  wadjust_loop_on_left_moving_B.simps)
  4061 apply(erule_tac exE)+
  4132 apply(erule_tac exE)+
  4062 apply(rule_tac x = ml in exI, rule_tac x = mr in exI)
  4133 apply(rule_tac x = ml in exI, rule_tac x = mr in exI)
  4063 apply(case_tac nl, simp_all add: exp_ind_def, auto)
  4134 apply(case_tac nl, simp_all add: , auto)
  4064 done
  4135 done
  4065 
  4136 
  4066 lemma [simp]: "wadjust_loop_on_left_moving m rs (c, Bk # list)
  4137 lemma [simp]: "wadjust_loop_on_left_moving m rs (c, Bk # list)
  4067             \<Longrightarrow> wadjust_loop_on_left_moving m rs (tl c, hd c # Bk # list)"
  4138             \<Longrightarrow> wadjust_loop_on_left_moving m rs (tl c, hd c # Bk # list)"
  4068 apply(simp add: wadjust_loop_on_left_moving.simps)
  4139 apply(simp add: wadjust_loop_on_left_moving.simps)
  4073 apply(simp only: wadjust_loop_right_move2.simps, auto)
  4144 apply(simp only: wadjust_loop_right_move2.simps, auto)
  4074 done
  4145 done
  4075 
  4146 
  4076 lemma [simp]: "wadjust_loop_right_move2 m rs (c, Bk # list) \<Longrightarrow>  wadjust_loop_start m rs (c, Oc # list)"
  4147 lemma [simp]: "wadjust_loop_right_move2 m rs (c, Bk # list) \<Longrightarrow>  wadjust_loop_start m rs (c, Oc # list)"
  4077 apply(auto simp: wadjust_loop_right_move2.simps wadjust_loop_start.simps)
  4148 apply(auto simp: wadjust_loop_right_move2.simps wadjust_loop_start.simps)
  4078 apply(case_tac ln, simp_all add: exp_ind_def)
  4149 apply(case_tac ln, simp_all add: )
  4079 apply(rule_tac x = 0 in exI, simp)
  4150 apply(rule_tac x = 0 in exI, simp)
  4080 apply(rule_tac x = rn in exI, simp)
  4151 apply(rule_tac x = rn in exI, simp)
  4081 apply(rule_tac x = "Suc ml" in exI, simp add: exp_ind_def, auto)
  4152 apply(rule_tac x = "Suc ml" in exI, simp add: , auto)
  4082 apply(rule_tac x = "Suc nat" in exI, simp add: exp_ind)
  4153 apply(rule_tac x = "Suc nat" in exI, simp add: exp_ind del: replicate_Suc)
  4083 apply(rule_tac x = rn in exI, auto)
  4154 apply(rule_tac x = rn in exI, auto)
  4084 apply(rule_tac x = "Suc ml" in exI, auto simp: exp_ind_def)
  4155 apply(rule_tac x = "Suc ml" in exI, auto )
  4085 done
  4156 done
  4086 
  4157 
  4087 lemma [simp]: "wadjust_erase2 m rs (c, Bk # list) \<Longrightarrow> c \<noteq> []"
  4158 lemma [simp]: "wadjust_erase2 m rs (c, Bk # list) \<Longrightarrow> c \<noteq> []"
  4088 apply(auto simp:wadjust_erase2.simps )
  4159 apply(auto simp:wadjust_erase2.simps )
  4089 done
  4160 done
  4090 
  4161 
  4091 lemma [simp]: "wadjust_erase2 m rs (c, Bk # list) \<Longrightarrow> 
  4162 lemma [simp]: "wadjust_erase2 m rs (c, Bk # list) \<Longrightarrow> 
  4092                  wadjust_on_left_moving m rs (tl c, hd c # Bk # list)"
  4163                  wadjust_on_left_moving m rs (tl c, hd c # Bk # list)"
  4093 apply(auto simp: wadjust_erase2.simps)
  4164 apply(auto simp: wadjust_erase2.simps)
  4094 apply(case_tac ln, simp_all add: exp_ind_def wadjust_on_left_moving.simps 
  4165 apply(case_tac ln, simp_all add:  wadjust_on_left_moving.simps 
  4095         wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps)
  4166         wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps)
  4096 apply(auto)
  4167 apply(auto)
  4097 apply(rule_tac x = "(Suc (Suc rn))" in exI, simp add: exp_ind_def)
  4168 apply(rule_tac x = "(Suc (Suc rn))" in exI, simp add: )
  4098 apply(rule_tac x = "Suc nat" in exI, simp add: exp_ind)
  4169 apply(rule_tac x = "Suc nat" in exI, simp add: exp_ind del: replicate_Suc)
  4099 apply(rule_tac x = "(Suc (Suc rn))" in exI, simp add: exp_ind_def)
  4170 apply(rule_tac x = "(Suc (Suc rn))" in exI, simp add: )
  4100 done
  4171 done
  4101 
  4172 
  4102 lemma [simp]: "wadjust_on_left_moving m rs (c,b) \<Longrightarrow> c \<noteq> []"
  4173 lemma [simp]: "wadjust_on_left_moving m rs (c,b) \<Longrightarrow> c \<noteq> []"
  4103 apply(simp only:wadjust_on_left_moving.simps
  4174 apply(simp only:wadjust_on_left_moving.simps
  4104                 wadjust_on_left_moving_O.simps
  4175                 wadjust_on_left_moving_O.simps
  4111 done
  4182 done
  4112 
  4183 
  4113 lemma [simp]: "\<lbrakk>wadjust_on_left_moving_B m rs (c, Bk # list); hd c = Bk\<rbrakk>
  4184 lemma [simp]: "\<lbrakk>wadjust_on_left_moving_B m rs (c, Bk # list); hd c = Bk\<rbrakk>
  4114     \<Longrightarrow> wadjust_on_left_moving_B m rs (tl c, Bk # Bk # list)"
  4185     \<Longrightarrow> wadjust_on_left_moving_B m rs (tl c, Bk # Bk # list)"
  4115 apply(auto simp: wadjust_on_left_moving_B.simps)
  4186 apply(auto simp: wadjust_on_left_moving_B.simps)
  4116 apply(case_tac ln, simp_all add: exp_ind_def, auto)
  4187 apply(case_tac ln, simp_all)
  4117 done
  4188 done
  4118 
  4189 
  4119 lemma [simp]: "\<lbrakk>wadjust_on_left_moving_B m rs (c, Bk # list); hd c = Oc\<rbrakk>
  4190 lemma [simp]: "\<lbrakk>wadjust_on_left_moving_B m rs (c, Bk # list); hd c = Oc\<rbrakk>
  4120     \<Longrightarrow> wadjust_on_left_moving_O m rs (tl c, Oc # Bk # list)"
  4191     \<Longrightarrow> wadjust_on_left_moving_O m rs (tl c, Oc # Bk # list)"
  4121 apply(auto simp: wadjust_on_left_moving_O.simps
  4192 apply(auto simp: wadjust_on_left_moving_O.simps
  4122                  wadjust_on_left_moving_B.simps)
  4193                  wadjust_on_left_moving_B.simps)
  4123 apply(case_tac ln, simp_all add: exp_ind_def)
  4194 apply(case_tac ln, simp_all add: )
  4124 done
  4195 done
  4125 
  4196 
  4126 lemma [simp]: "wadjust_on_left_moving  m rs (c, Bk # list) \<Longrightarrow>  
  4197 lemma [simp]: "wadjust_on_left_moving  m rs (c, Bk # list) \<Longrightarrow>  
  4127                   wadjust_on_left_moving m rs (tl c, hd c # Bk # list)"
  4198                   wadjust_on_left_moving m rs (tl c, hd c # Bk # list)"
  4128 apply(simp add: wadjust_on_left_moving.simps)
  4199 apply(simp add: wadjust_on_left_moving.simps)
  4130 done
  4201 done
  4131 
  4202 
  4132 lemma [simp]: "wadjust_goon_left_moving m rs (c, b) \<Longrightarrow> c \<noteq> []"
  4203 lemma [simp]: "wadjust_goon_left_moving m rs (c, b) \<Longrightarrow> c \<noteq> []"
  4133 apply(simp add: wadjust_goon_left_moving.simps
  4204 apply(simp add: wadjust_goon_left_moving.simps
  4134                 wadjust_goon_left_moving_B.simps
  4205                 wadjust_goon_left_moving_B.simps
  4135                 wadjust_goon_left_moving_O.simps exp_ind_def, auto)
  4206                 wadjust_goon_left_moving_O.simps , auto)
  4136 done
  4207 done
  4137 
  4208 
  4138 lemma [simp]: "wadjust_goon_left_moving_O m rs (c, Bk # list) = False"
  4209 lemma [simp]: "wadjust_goon_left_moving_O m rs (c, Bk # list) = False"
  4139 apply(simp add: wadjust_goon_left_moving_O.simps, auto)
  4210 apply(simp add: wadjust_goon_left_moving_O.simps, auto)
  4140 apply(case_tac mr, simp_all add: exp_ind_def)
  4211 apply(case_tac mr, simp_all add: )
  4141 done
  4212 done
  4142 
  4213 
  4143 lemma [simp]: "\<lbrakk>wadjust_goon_left_moving_B m rs (c, Bk # list); hd c = Bk\<rbrakk>
  4214 lemma [simp]: "\<lbrakk>wadjust_goon_left_moving_B m rs (c, Bk # list); hd c = Bk\<rbrakk>
  4144     \<Longrightarrow> wadjust_backto_standard_pos_B m rs (tl c, Bk # Bk # list)"
  4215     \<Longrightarrow> wadjust_backto_standard_pos_B m rs (tl c, Bk # Bk # list)"
  4145 apply(auto simp: wadjust_goon_left_moving_B.simps 
  4216 apply(auto simp: wadjust_goon_left_moving_B.simps 
  4146                  wadjust_backto_standard_pos_B.simps exp_ind_def)
  4217                  wadjust_backto_standard_pos_B.simps )
  4147 done
  4218 done
  4148 
  4219 
  4149 lemma [simp]: "\<lbrakk>wadjust_goon_left_moving_B m rs (c, Bk # list); hd c = Oc\<rbrakk>
  4220 lemma [simp]: "\<lbrakk>wadjust_goon_left_moving_B m rs (c, Bk # list); hd c = Oc\<rbrakk>
  4150     \<Longrightarrow> wadjust_backto_standard_pos_O m rs (tl c, Oc # Bk # list)"
  4221     \<Longrightarrow> wadjust_backto_standard_pos_O m rs (tl c, Oc # Bk # list)"
  4151 apply(auto simp: wadjust_goon_left_moving_B.simps 
  4222 apply(auto simp: wadjust_goon_left_moving_B.simps 
  4152                  wadjust_backto_standard_pos_O.simps exp_ind_def)
  4223                  wadjust_backto_standard_pos_O.simps)
  4153 apply(rule_tac x = m in exI, simp, auto)
       
  4154 done
  4224 done
  4155 
  4225 
  4156 lemma [simp]: "wadjust_goon_left_moving m rs (c, Bk # list) \<Longrightarrow>
  4226 lemma [simp]: "wadjust_goon_left_moving m rs (c, Bk # list) \<Longrightarrow>
  4157   wadjust_backto_standard_pos m rs (tl c, hd c # Bk # list)"
  4227   wadjust_backto_standard_pos m rs (tl c, hd c # Bk # list)"
  4158 apply(case_tac "hd c", simp_all add: wadjust_backto_standard_pos.simps 
  4228 apply(case_tac "hd c", simp_all add: wadjust_backto_standard_pos.simps 
  4162 lemma [simp]: "wadjust_backto_standard_pos m rs (c, Bk # list) \<Longrightarrow>
  4232 lemma [simp]: "wadjust_backto_standard_pos m rs (c, Bk # list) \<Longrightarrow>
  4163   (c = [] \<longrightarrow> wadjust_stop m rs ([Bk], list)) \<and> (c \<noteq> [] \<longrightarrow> wadjust_stop m rs (Bk # c, list))"
  4233   (c = [] \<longrightarrow> wadjust_stop m rs ([Bk], list)) \<and> (c \<noteq> [] \<longrightarrow> wadjust_stop m rs (Bk # c, list))"
  4164 apply(auto simp: wadjust_backto_standard_pos.simps 
  4234 apply(auto simp: wadjust_backto_standard_pos.simps 
  4165                  wadjust_backto_standard_pos_B.simps
  4235                  wadjust_backto_standard_pos_B.simps
  4166                  wadjust_backto_standard_pos_O.simps wadjust_stop.simps)
  4236                  wadjust_backto_standard_pos_O.simps wadjust_stop.simps)
  4167 apply(case_tac [!] mr, simp_all add: exp_ind_def)
  4237 apply(case_tac [!] mr, simp_all add: )
  4168 done
  4238 done
  4169 
  4239 
  4170 lemma [simp]: "wadjust_start m rs (c, Oc # list)
  4240 lemma [simp]: "wadjust_start m rs (c, Oc # list)
  4171               \<Longrightarrow> (c = [] \<longrightarrow> wadjust_loop_start m rs ([Oc], list)) \<and>
  4241               \<Longrightarrow> (c = [] \<longrightarrow> wadjust_loop_start m rs ([Oc], list)) \<and>
  4172                 (c \<noteq> [] \<longrightarrow> wadjust_loop_start m rs (Oc # c, list))"
  4242                 (c \<noteq> [] \<longrightarrow> wadjust_loop_start m rs (Oc # c, list))"
  4182 lemma [simp]: "wadjust_loop_start m rs (c, Oc # list)
  4252 lemma [simp]: "wadjust_loop_start m rs (c, Oc # list)
  4183               \<Longrightarrow> wadjust_loop_right_move m rs (Oc # c, list)"
  4253               \<Longrightarrow> wadjust_loop_right_move m rs (Oc # c, list)"
  4184 apply(simp add: wadjust_loop_start.simps wadjust_loop_right_move.simps, auto)
  4254 apply(simp add: wadjust_loop_start.simps wadjust_loop_right_move.simps, auto)
  4185 apply(rule_tac x = ml in exI, rule_tac x = mr in exI, 
  4255 apply(rule_tac x = ml in exI, rule_tac x = mr in exI, 
  4186       rule_tac x = 0 in exI, simp)
  4256       rule_tac x = 0 in exI, simp)
  4187 apply(rule_tac x = "Suc ln" in exI, simp add: exp_ind, auto)
  4257 apply(rule_tac x = "Suc ln" in exI, simp add: exp_ind del: replicate_Suc)
  4188 done
  4258 done
  4189 
  4259 
  4190 lemma [simp]: "wadjust_loop_right_move m rs (c, Oc # list) \<Longrightarrow> 
  4260 lemma [simp]: "wadjust_loop_right_move m rs (c, Oc # list) \<Longrightarrow> 
  4191                        wadjust_loop_check m rs (Oc # c, list)"
  4261                        wadjust_loop_check m rs (Oc # c, list)"
  4192 apply(simp add: wadjust_loop_right_move.simps  
  4262 apply(simp add: wadjust_loop_right_move.simps  
  4193                  wadjust_loop_check.simps, auto)
  4263                  wadjust_loop_check.simps, auto)
  4194 apply(rule_tac [!] x = ml in exI, simp_all, auto)
  4264 apply(rule_tac [!] x = ml in exI, simp_all add: exp_ind del: replicate_Suc, auto)
  4195 apply(case_tac nl, auto simp: exp_ind_def)
  4265 apply(case_tac nl, simp_all add: exp_ind del: replicate_Suc)
  4196 apply(rule_tac x = "mr - 1" in exI, case_tac mr, simp_all add: exp_ind_def)
  4266 apply(rule_tac x = "mr - 1" in exI, case_tac mr, simp_all add: )
  4197 apply(case_tac [!] nr, simp_all add: exp_ind_def, auto)
  4267 apply(case_tac [!] nr, simp_all)
  4198 done
  4268 done
  4199 
  4269 
  4200 lemma [simp]: "wadjust_loop_check m rs (c, Oc # list) \<Longrightarrow> 
  4270 lemma [simp]: "wadjust_loop_check m rs (c, Oc # list) \<Longrightarrow> 
  4201                wadjust_loop_erase m rs (tl c, hd c # Oc # list)"
  4271                wadjust_loop_erase m rs (tl c, hd c # Oc # list)"
  4202 apply(simp only: wadjust_loop_check.simps wadjust_loop_erase.simps)
  4272 apply(simp only: wadjust_loop_check.simps wadjust_loop_erase.simps)
  4203 apply(erule_tac exE)+
  4273 apply(erule_tac exE)+
  4204 apply(rule_tac x = ml in exI, rule_tac x = mr in exI, auto)
  4274 apply(rule_tac x = ml in exI, rule_tac x = mr in exI, auto)
  4205 apply(case_tac mr, simp_all add: exp_ind_def)
  4275 apply(case_tac mr, simp_all add: )
  4206 apply(case_tac rn, simp_all add: exp_ind_def)
       
  4207 done
  4276 done
  4208 
  4277 
  4209 lemma [simp]: "wadjust_loop_erase m rs (c, Oc # list) \<Longrightarrow> 
  4278 lemma [simp]: "wadjust_loop_erase m rs (c, Oc # list) \<Longrightarrow> 
  4210                 wadjust_loop_erase m rs (c, Bk # list)"
  4279                 wadjust_loop_erase m rs (c, Bk # list)"
  4211 apply(auto simp: wadjust_loop_erase.simps)
  4280 apply(auto simp: wadjust_loop_erase.simps)
  4212 done
  4281 done
  4213 
  4282 
  4214 lemma [simp]: "wadjust_loop_on_left_moving_B m rs (c, Oc # list) = False"
  4283 lemma [simp]: "wadjust_loop_on_left_moving_B m rs (c, Oc # list) = False"
  4215 apply(auto simp: wadjust_loop_on_left_moving_B.simps)
  4284 apply(auto simp: wadjust_loop_on_left_moving_B.simps)
  4216 apply(case_tac nr, simp_all add: exp_ind_def)
  4285 apply(case_tac nr, simp_all add: )
  4217 done
  4286 done
  4218 
  4287 
  4219 lemma [simp]: "wadjust_loop_on_left_moving m rs (c, Oc # list)
  4288 lemma [simp]: "wadjust_loop_on_left_moving m rs (c, Oc # list)
  4220            \<Longrightarrow> wadjust_loop_right_move2 m rs (Oc # c, list)"
  4289            \<Longrightarrow> wadjust_loop_right_move2 m rs (Oc # c, list)"
  4221 apply(simp add:wadjust_loop_on_left_moving.simps)
  4290 apply(simp add:wadjust_loop_on_left_moving.simps)
  4223                  wadjust_loop_right_move2.simps)
  4292                  wadjust_loop_right_move2.simps)
  4224 done
  4293 done
  4225 
  4294 
  4226 lemma [simp]: "wadjust_loop_right_move2 m rs (c, Oc # list) = False"
  4295 lemma [simp]: "wadjust_loop_right_move2 m rs (c, Oc # list) = False"
  4227 apply(auto simp: wadjust_loop_right_move2.simps )
  4296 apply(auto simp: wadjust_loop_right_move2.simps )
  4228 apply(case_tac ln, simp_all add: exp_ind_def)
  4297 apply(case_tac ln, simp_all add: )
  4229 done
  4298 done
  4230 
  4299 
  4231 lemma [simp]: "wadjust_erase2 m rs (c, Oc # list)
  4300 lemma [simp]: "wadjust_erase2 m rs (c, Oc # list)
  4232               \<Longrightarrow> (c = [] \<longrightarrow> wadjust_erase2 m rs ([], Bk # list))
  4301               \<Longrightarrow> (c = [] \<longrightarrow> wadjust_erase2 m rs ([], Bk # list))
  4233                \<and> (c \<noteq> [] \<longrightarrow> wadjust_erase2 m rs (c, Bk # list))"
  4302                \<and> (c \<noteq> [] \<longrightarrow> wadjust_erase2 m rs (c, Bk # list))"
  4239 done
  4308 done
  4240 
  4309 
  4241 lemma [simp]: "\<lbrakk>wadjust_on_left_moving_O m rs (c, Oc # list); hd c = Bk\<rbrakk> \<Longrightarrow> 
  4310 lemma [simp]: "\<lbrakk>wadjust_on_left_moving_O m rs (c, Oc # list); hd c = Bk\<rbrakk> \<Longrightarrow> 
  4242          wadjust_goon_left_moving_B m rs (tl c, Bk # Oc # list)"
  4311          wadjust_goon_left_moving_B m rs (tl c, Bk # Oc # list)"
  4243 apply(auto simp: wadjust_on_left_moving_O.simps 
  4312 apply(auto simp: wadjust_on_left_moving_O.simps 
  4244      wadjust_goon_left_moving_B.simps exp_ind_def)
  4313      wadjust_goon_left_moving_B.simps )
  4245 done
  4314 done
  4246 
  4315 
  4247 lemma [simp]: "\<lbrakk>wadjust_on_left_moving_O m rs (c, Oc # list); hd c = Oc\<rbrakk>
  4316 lemma [simp]: "\<lbrakk>wadjust_on_left_moving_O m rs (c, Oc # list); hd c = Oc\<rbrakk>
  4248     \<Longrightarrow> wadjust_goon_left_moving_O m rs (tl c, Oc # Oc # list)"
  4317     \<Longrightarrow> wadjust_goon_left_moving_O m rs (tl c, Oc # Oc # list)"
  4249 apply(auto simp: wadjust_on_left_moving_O.simps 
  4318 apply(auto simp: wadjust_on_left_moving_O.simps 
  4250                  wadjust_goon_left_moving_O.simps exp_ind_def)
  4319                  wadjust_goon_left_moving_O.simps )
  4251 apply(rule_tac x = rs in exI, simp)
  4320 apply(auto simp:  numeral_2_eq_2)
  4252 apply(auto simp: exp_ind_def numeral_2_eq_2)
       
  4253 done
  4321 done
  4254 
  4322 
  4255 
  4323 
  4256 lemma [simp]: "wadjust_on_left_moving m rs (c, Oc # list) \<Longrightarrow> 
  4324 lemma [simp]: "wadjust_on_left_moving m rs (c, Oc # list) \<Longrightarrow> 
  4257               wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)"
  4325               wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)"
  4272 done
  4340 done
  4273 
  4341 
  4274 lemma [simp]: "\<lbrakk>wadjust_goon_left_moving_O m rs (c, Oc # list); hd c = Bk\<rbrakk> 
  4342 lemma [simp]: "\<lbrakk>wadjust_goon_left_moving_O m rs (c, Oc # list); hd c = Bk\<rbrakk> 
  4275                \<Longrightarrow> wadjust_goon_left_moving_B m rs (tl c, Bk # Oc # list)"
  4343                \<Longrightarrow> wadjust_goon_left_moving_B m rs (tl c, Bk # Oc # list)"
  4276 apply(auto simp: wadjust_goon_left_moving_O.simps wadjust_goon_left_moving_B.simps)
  4344 apply(auto simp: wadjust_goon_left_moving_O.simps wadjust_goon_left_moving_B.simps)
  4277 apply(case_tac [!] ml, auto simp: exp_ind_def)
  4345 apply(case_tac [!] ml, auto simp: )
  4278 done
  4346 done
  4279 
  4347 
  4280 lemma  [simp]: "\<lbrakk>wadjust_goon_left_moving_O m rs (c, Oc # list); hd c = Oc\<rbrakk> \<Longrightarrow> 
  4348 lemma  [simp]: "\<lbrakk>wadjust_goon_left_moving_O m rs (c, Oc # list); hd c = Oc\<rbrakk> \<Longrightarrow> 
  4281   wadjust_goon_left_moving_O m rs (tl c, Oc # Oc # list)"
  4349   wadjust_goon_left_moving_O m rs (tl c, Oc # Oc # list)"
  4282 apply(auto simp: wadjust_goon_left_moving_O.simps wadjust_goon_left_moving_B.simps)
  4350 apply(auto simp: wadjust_goon_left_moving_O.simps wadjust_goon_left_moving_B.simps)
  4283 apply(rule_tac x = "ml - 1" in exI, simp)
  4351 apply(rule_tac x = "ml - 1" in exI, simp)
  4284 apply(case_tac ml, simp_all add: exp_ind_def)
  4352 apply(case_tac ml, simp_all add: )
  4285 apply(rule_tac x = "Suc mr" in exI, auto simp: exp_ind_def)
  4353 apply(rule_tac x = "Suc mr" in exI, auto simp: )
  4286 done
  4354 done
  4287 
  4355 
  4288 lemma [simp]: "wadjust_goon_left_moving m rs (c, Oc # list) \<Longrightarrow> 
  4356 lemma [simp]: "wadjust_goon_left_moving m rs (c, Oc # list) \<Longrightarrow> 
  4289   wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)"
  4357   wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)"
  4290 apply(simp add: wadjust_goon_left_moving.simps)
  4358 apply(simp add: wadjust_goon_left_moving.simps)
  4295 apply(simp add: wadjust_backto_standard_pos_B.simps)
  4363 apply(simp add: wadjust_backto_standard_pos_B.simps)
  4296 done
  4364 done
  4297 
  4365 
  4298 lemma [simp]: "wadjust_backto_standard_pos_O m rs (c, Bk # xs) = False"
  4366 lemma [simp]: "wadjust_backto_standard_pos_O m rs (c, Bk # xs) = False"
  4299 apply(simp add: wadjust_backto_standard_pos_O.simps, auto)
  4367 apply(simp add: wadjust_backto_standard_pos_O.simps, auto)
  4300 apply(case_tac mr, simp_all add: exp_ind_def)
  4368 apply(case_tac mr, simp_all add: )
  4301 done
  4369 done
  4302 
       
  4303 
       
  4304 
  4370 
  4305 lemma [simp]: "wadjust_backto_standard_pos_O m rs ([], Oc # list) \<Longrightarrow> 
  4371 lemma [simp]: "wadjust_backto_standard_pos_O m rs ([], Oc # list) \<Longrightarrow> 
  4306   wadjust_backto_standard_pos_B m rs ([], Bk # Oc # list)"
  4372   wadjust_backto_standard_pos_B m rs ([], Bk # Oc # list)"
  4307 apply(auto simp: wadjust_backto_standard_pos_O.simps
  4373 apply(auto simp: wadjust_backto_standard_pos_O.simps
  4308                  wadjust_backto_standard_pos_B.simps)
  4374                  wadjust_backto_standard_pos_B.simps)
  4309 apply(rule_tac x = rn in exI, simp)
  4375 done
  4310 apply(case_tac ml, simp_all add: exp_ind_def)
       
  4311 done
       
  4312 
       
  4313 
  4376 
  4314 lemma [simp]: 
  4377 lemma [simp]: 
  4315   "\<lbrakk>wadjust_backto_standard_pos_O m rs (c, Oc # list); c \<noteq> []; hd c = Bk\<rbrakk>
  4378   "\<lbrakk>wadjust_backto_standard_pos_O m rs (c, Oc # list); c \<noteq> []; hd c = Bk\<rbrakk>
  4316   \<Longrightarrow> wadjust_backto_standard_pos_B m rs (tl c, Bk # Oc # list)"
  4379   \<Longrightarrow> wadjust_backto_standard_pos_B m rs (tl c, Bk # Oc # list)"
  4317 apply(simp add:wadjust_backto_standard_pos_O.simps 
  4380 apply(simp add:wadjust_backto_standard_pos_O.simps 
  4318         wadjust_backto_standard_pos_B.simps, auto)
  4381         wadjust_backto_standard_pos_B.simps, auto)
  4319 apply(case_tac [!] ml, simp_all add: exp_ind_def)
       
  4320 done 
  4382 done 
  4321 
  4383 
  4322 lemma [simp]: "\<lbrakk>wadjust_backto_standard_pos_O m rs (c, Oc # list); c \<noteq> []; hd c = Oc\<rbrakk>
  4384 lemma [simp]: "\<lbrakk>wadjust_backto_standard_pos_O m rs (c, Oc # list); c \<noteq> []; hd c = Oc\<rbrakk>
  4323           \<Longrightarrow>  wadjust_backto_standard_pos_O m rs (tl c, Oc # Oc # list)"
  4385           \<Longrightarrow>  wadjust_backto_standard_pos_O m rs (tl c, Oc # Oc # list)"
  4324 apply(simp add: wadjust_backto_standard_pos_O.simps, auto)
  4386 apply(simp add: wadjust_backto_standard_pos_O.simps, auto)
  4325 apply(case_tac ml, simp_all add: exp_ind_def, auto)
  4387 apply(case_tac ml, simp_all add: , auto)
  4326 apply(rule_tac x = nat in exI, auto simp: exp_ind_def)
       
  4327 done
  4388 done
  4328 
  4389 
  4329 lemma [simp]: "wadjust_backto_standard_pos m rs (c, Oc # list)
  4390 lemma [simp]: "wadjust_backto_standard_pos m rs (c, Oc # list)
  4330   \<Longrightarrow> (c = [] \<longrightarrow> wadjust_backto_standard_pos m rs ([], Bk # Oc # list)) \<and> 
  4391   \<Longrightarrow> (c = [] \<longrightarrow> wadjust_backto_standard_pos m rs ([], Bk # Oc # list)) \<and> 
  4331  (c \<noteq> [] \<longrightarrow> wadjust_backto_standard_pos m rs (tl c, hd c # Oc # list))"
  4392  (c \<noteq> [] \<longrightarrow> wadjust_backto_standard_pos m rs (tl c, hd c # Oc # list))"
  4332 apply(auto simp: wadjust_backto_standard_pos.simps)
  4393 apply(auto simp: wadjust_backto_standard_pos.simps)
  4333 apply(case_tac "hd c", simp_all)
  4394 apply(case_tac "hd c", simp_all)
  4334 done
  4395 done
  4335 thm wadjust_loop_right_move.simps
       
  4336 
  4396 
  4337 lemma [simp]: "wadjust_loop_right_move m rs (c, []) = False"
  4397 lemma [simp]: "wadjust_loop_right_move m rs (c, []) = False"
  4338 apply(simp only: wadjust_loop_right_move.simps)
  4398 apply(simp only: wadjust_loop_right_move.simps)
  4339 apply(rule_tac iffI)
  4399 apply(rule_tac iffI)
  4340 apply(erule_tac exE)+
  4400 apply(erule_tac exE)+
  4341 apply(case_tac nr, simp_all add: exp_ind_def)
  4401 apply(case_tac nr, simp_all add: )
  4342 apply(case_tac mr, simp_all add: exp_ind_def)
  4402 apply(case_tac mr, simp_all add: )
  4343 done
  4403 done
  4344 
  4404 
  4345 lemma [simp]: "wadjust_loop_erase m rs (c, []) = False"
  4405 lemma [simp]: "wadjust_loop_erase m rs (c, []) = False"
  4346 apply(simp only: wadjust_loop_erase.simps, auto)
  4406 apply(simp only: wadjust_loop_erase.simps, auto)
  4347 apply(case_tac mr, simp_all add: exp_ind_def)
       
  4348 done
  4407 done
  4349 
  4408 
  4350 lemma [simp]: "\<lbrakk>Suc (Suc rs) = a;  wadjust_loop_erase m rs (c, Bk # list)\<rbrakk>
  4409 lemma [simp]: "\<lbrakk>Suc (Suc rs) = a;  wadjust_loop_erase m rs (c, Bk # list)\<rbrakk>
  4351   \<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Bk # list))))
  4410   \<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Bk # list))))
  4352   < a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list)))) \<or>
  4411   < a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list)))) \<or>
  4365   a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list))))"
  4424   a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list))))"
  4366 apply(subgoal_tac "c \<noteq> []")
  4425 apply(subgoal_tac "c \<noteq> []")
  4367 apply(case_tac c, simp_all)
  4426 apply(case_tac c, simp_all)
  4368 done
  4427 done
  4369 
  4428 
  4370 lemma dropWhile_exp1: "dropWhile (\<lambda>a. a = Oc) (Oc\<^bsup>n\<^esup> @ xs) = dropWhile (\<lambda>a. a = Oc) xs"
  4429 lemma dropWhile_exp1: "dropWhile (\<lambda>a. a = Oc) (Oc\<up>(n) @ xs) = dropWhile (\<lambda>a. a = Oc) xs"
  4371 apply(induct n, simp_all add: exp_ind_def)
  4430 apply(induct n, simp_all add: )
  4372 done
  4431 done
  4373 lemma takeWhile_exp1: "takeWhile (\<lambda>a. a = Oc) (Oc\<^bsup>n\<^esup> @ xs) = Oc\<^bsup>n\<^esup> @ takeWhile (\<lambda>a. a = Oc) xs"
  4432 lemma takeWhile_exp1: "takeWhile (\<lambda>a. a = Oc) (Oc\<up>(n) @ xs) = Oc\<up>(n) @ takeWhile (\<lambda>a. a = Oc) xs"
  4374 apply(induct n, simp_all add: exp_ind_def)
  4433 apply(induct n, simp_all add: )
  4375 done
  4434 done
  4376 
  4435 
  4377 lemma [simp]: "\<lbrakk>Suc (Suc rs) = a;  wadjust_loop_right_move2 m rs (c, Bk # list)\<rbrakk>
  4436 lemma [simp]: "\<lbrakk>Suc (Suc rs) = a;  wadjust_loop_right_move2 m rs (c, Bk # list)\<rbrakk>
  4378               \<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list))))
  4437               \<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list))))
  4379                  < a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list))))"
  4438                  < a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list))))"
  4380 apply(simp add: wadjust_loop_right_move2.simps, auto)
  4439 apply(simp add: wadjust_loop_right_move2.simps, auto)
  4381 apply(simp add: dropWhile_exp1 takeWhile_exp1)
  4440 apply(simp add: dropWhile_exp1 takeWhile_exp1)
  4382 apply(case_tac ln, simp, simp add: exp_ind_def)
  4441 apply(case_tac ln, simp, simp add: )
  4383 done
  4442 done
  4384 
  4443 
  4385 lemma [simp]: "wadjust_loop_check m rs ([], b) = False"
  4444 lemma [simp]: "wadjust_loop_check m rs ([], b) = False"
  4386 apply(simp add: wadjust_loop_check.simps)
  4445 apply(simp add: wadjust_loop_check.simps)
  4387 done
  4446 done
  4409 declare numeral_2_eq_2[simp del]
  4468 declare numeral_2_eq_2[simp del]
  4410 
  4469 
  4411 lemma wadjust_correctness:
  4470 lemma wadjust_correctness:
  4412   shows "let P = (\<lambda> (len, st, l, r). st = 0) in 
  4471   shows "let P = (\<lambda> (len, st, l, r). st = 0) in 
  4413   let Q = (\<lambda> (len, st, l, r). wadjust_inv st m rs (l, r)) in 
  4472   let Q = (\<lambda> (len, st, l, r). wadjust_inv st m rs (l, r)) in 
  4414   let f = (\<lambda> stp. (Suc (Suc rs),  steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, 
  4473   let f = (\<lambda> stp. (Suc (Suc rs),  steps0 (Suc 0, Bk # Oc\<up>(Suc m), 
  4415                 Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk #  Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust stp)) in
  4474                 Bk # Oc # Bk\<up>(ln) @ Bk #  Oc\<up>(Suc rs) @ Bk\<up>(rn)) t_wcode_adjust stp)) in
  4416     \<exists> n .P (f n) \<and> Q (f n)"
  4475     \<exists> n .P (f n) \<and> Q (f n)"
  4417 proof -
  4476 proof -
  4418   let ?P = "(\<lambda> (len, st, l, r). st = 0)"
  4477   let ?P = "(\<lambda> (len, st, l, r). st = 0)"
  4419   let ?Q = "\<lambda> (len, st, l, r). wadjust_inv st m rs (l, r)"
  4478   let ?Q = "\<lambda> (len, st, l, r). wadjust_inv st m rs (l, r)"
  4420   let ?f = "\<lambda> stp. (Suc (Suc rs),  steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, 
  4479   let ?f = "\<lambda> stp. (Suc (Suc rs),  steps0 (Suc 0, Bk # Oc\<up>(Suc m), 
  4421                 Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust stp)"
  4480                 Bk # Oc # Bk\<up>(ln) @ Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn)) t_wcode_adjust stp)"
  4422   have "\<exists> n. ?P (?f n) \<and> ?Q (?f n)"
  4481   have "\<exists> n. ?P (?f n) \<and> ?Q (?f n)"
  4423   proof(rule_tac halt_lemma2)
  4482   proof(rule_tac halt_lemma2)
  4424     show "wf wadjust_le" by auto
  4483     show "wf wadjust_le" by auto
  4425   next
  4484   next
  4426     show "\<forall> n. \<not> ?P (?f n) \<and> ?Q (?f n) \<longrightarrow> 
  4485     show "\<forall> n. \<not> ?P (?f n) \<and> ?Q (?f n) \<longrightarrow> 
  4427                  ?Q (?f (Suc n)) \<and> (?f (Suc n), ?f n) \<in> wadjust_le"
  4486                  ?Q (?f (Suc n)) \<and> (?f (Suc n), ?f n) \<in> wadjust_le"
  4428     proof(rule_tac allI, rule_tac impI, case_tac "?f n", 
  4487       apply(rule_tac allI, rule_tac impI, case_tac "?f n", simp)
  4429             simp add: tstep_red tstep.simps, rule_tac conjI, erule_tac conjE,
  4488       apply(simp add: step.simps)
  4430           erule_tac conjE)      
  4489       apply(case_tac d, case_tac [2] aa, simp_all)
  4431       fix n a b c d
  4490       apply(simp_all add: wadjust_inv.simps wadjust_le_def
  4432       assume "0 < b" "wadjust_inv b m rs (c, d)" "Suc (Suc rs) = a"
  4491             abacus.lex_triple_def abacus.lex_pair_def lex_square_def numeral_4_eq_4
  4433       thus "case case fetch t_wcode_adjust b (case d of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)
       
  4434         of (ac, ns) \<Rightarrow> (ns, new_tape ac (c, d)) of (st, x) \<Rightarrow> wadjust_inv st m rs x"
       
  4435         apply(case_tac d, simp, case_tac [2] aa)
       
  4436         apply(simp_all add: wadjust_inv.simps wadjust_le_def new_tape.simps
       
  4437           abacus.lex_triple_def abacus.lex_pair_def lex_square_def
       
  4438           split: if_splits)
       
  4439         done
       
  4440     next
       
  4441       fix n a b c d
       
  4442       assume "0 < b \<and> wadjust_inv b m rs (c, d)"
       
  4443         "Suc (Suc rs) = a \<and> steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>,
       
  4444          Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust n = (b, c, d)"
       
  4445       thus "((a, case fetch t_wcode_adjust b (case d of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)
       
  4446         of (ac, ns) \<Rightarrow> (ns, new_tape ac (c, d))), a, b, c, d) \<in> wadjust_le"
       
  4447       proof(erule_tac conjE, erule_tac conjE, erule_tac conjE)
       
  4448         assume "0 < b" "wadjust_inv b m rs (c, d)" "Suc (Suc rs) = a"
       
  4449         thus "?thesis"
       
  4450           apply(case_tac d, case_tac [2] aa)
       
  4451           apply(simp_all add: wadjust_inv.simps wadjust_le_def new_tape.simps
       
  4452             abacus.lex_triple_def abacus.lex_pair_def lex_square_def
       
  4453             split: if_splits)
  4492             split: if_splits)
  4454           done
  4493       done
  4455       qed
       
  4456     qed
       
  4457   next
  4494   next
  4458     show "?Q (?f 0)"
  4495     show "?Q (?f 0)"
  4459       apply(simp add: steps.simps wadjust_inv.simps wadjust_start.simps)
  4496       apply(simp add: steps.simps wadjust_inv.simps wadjust_start.simps, auto)
  4460       apply(rule_tac x = ln in exI,auto)
       
  4461       done
  4497       done
  4462   next
  4498   next
  4463     show "\<not> ?P (?f 0)"
  4499     show "\<not> ?P (?f 0)"
  4464       apply(simp add: steps.simps)
  4500       apply(simp add: steps.simps)
  4465       done
  4501       done
  4466   qed
  4502   qed
  4467   thus "?thesis"
  4503   thus"?thesis"
  4468     apply(auto)
  4504     apply(simp)
  4469     done
  4505     done
  4470 qed
  4506 qed
  4471 
  4507 
  4472 lemma [intro]: "t_correct t_wcode_adjust"
  4508 lemma [intro]: "tm_wf (t_wcode_adjust, 0)"
  4473 apply(auto simp: t_wcode_adjust_def t_correct.simps iseven_def)
  4509 apply(auto simp: t_wcode_adjust_def tm_wf.simps)
  4474 apply(rule_tac x = 11 in exI, simp)
  4510 done
       
  4511 
       
  4512 declare tm_comp.simps[simp del]
       
  4513 
       
  4514 lemma [simp]: "args \<noteq> [] \<Longrightarrow> bl_bin (<args::nat list>) > 0"
       
  4515 apply(case_tac args)
       
  4516 apply(auto simp: tape_of_nl_cons bl_bin.simps split: if_splits)
  4475 done
  4517 done
  4476 
  4518 
  4477 lemma wcode_lemma_pre':
  4519 lemma wcode_lemma_pre':
  4478   "args \<noteq> [] \<Longrightarrow> 
  4520   "args \<noteq> [] \<Longrightarrow> 
  4479   \<exists> stp rn. steps (Suc 0, [], <m # args>) 
  4521   \<exists> stp rn. steps0 (Suc 0, [], <m # args>) 
  4480               ((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) stp
  4522               ((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) stp
  4481   = (0,  [Bk],  Oc\<^bsup>Suc m\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>)" 
  4523   = (0,  [Bk],  Oc\<up>(Suc m) @ Bk # Oc\<up>(Suc (bl_bin (<args>))) @ Bk\<up>(rn))" 
  4482 proof -
  4524 proof -
  4483   let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <m # args>"
  4525   let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <m # args>"
  4484   let ?Q1 = "\<lambda>(l, r). l = Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  4526   let ?Q1 = "\<lambda>(l, r). l = Bk # Oc\<up>(Suc m) \<and>
  4485     (\<exists>ln rn. r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  4527     (\<exists>ln rn. r = Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin (<args>)) @ Bk\<up>(rn))"
  4486   let ?P2 = ?Q1
  4528   let ?P2 = ?Q1
  4487   let ?Q2 = "\<lambda> (l, r). (wadjust_stop m (bl_bin (<args>) - 1) (l, r))"
  4529   let ?Q2 = "\<lambda> (l, r). (wadjust_stop m (bl_bin (<args>) - 1) (l, r))"
  4488   let ?P3 = "\<lambda> tp. False"
  4530   let ?P3 = "\<lambda> tp. False"
  4489   assume h: "args \<noteq> []"
  4531   assume h: "args \<noteq> []"
  4490   have "?P1 \<turnstile>-> \<lambda> tp. (\<exists> stp tp'. steps (Suc 0, tp)
  4532   hence a: "bl_bin (<args>) > 0"
  4491                       ((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) stp = (0, tp') \<and> ?Q2 tp')"
  4533     using h by simp
  4492   proof(rule_tac turing_merge.t_merge_halt[of "t_wcode_prepare |+| t_wcode_main" 
  4534   hence "{?P1} (t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust {?Q2}"
  4493                t_wcode_adjust ?P1 ?P2 ?P3 ?P3 ?Q1 ?Q2], 
  4535   proof(rule_tac Hoare_plus_halt)
  4494         auto simp: turing_merge_def)
  4536     show "?Q1 \<mapsto> ?P2"
  4495 
  4537       by(simp add: assert_imp_def)
  4496     show "\<exists>stp. case steps (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) stp of
  4538   next
  4497           (st, tp') \<Rightarrow> st = 0 \<and> (case tp' of (l, r) \<Rightarrow> l = Bk # Oc\<^bsup>Suc m\<^esup> \<and>
  4539     show "tm_wf (t_wcode_prepare |+| t_wcode_main, 0)"
  4498                 (\<exists>ln rn. r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>))"
  4540       apply(rule_tac tm_wf_comp, auto)
  4499       using h prepare_mainpart_lemma[of args m]
       
  4500       apply(auto)
       
  4501       apply(rule_tac x = stp in exI, simp)
       
  4502       apply(rule_tac x = ln in exI, auto)
       
  4503       done
  4541       done
  4504   next
  4542   next
  4505     fix ln rn
  4543     show "{?P1} t_wcode_prepare |+| t_wcode_main {?Q1}"
  4506     show "\<exists>stp. case steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # 
  4544     proof(rule_tac HoareI, auto)
  4507                                Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust stp of
  4545       show 
  4508       (st, tp') \<Rightarrow> st = 0 \<and> wadjust_stop m (bl_bin (<args>) - Suc 0) tp'"
  4546         "\<exists>n. is_final (steps0 (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) n) \<and>
  4509       using wadjust_correctness[of m "bl_bin (<args>) - 1" "Suc ln" rn]
  4547         (\<lambda>(l, r). l = Bk # Oc # Oc \<up> m \<and>
  4510       apply(subgoal_tac "bl_bin (<args>) > 0", auto simp: wadjust_inv.simps)
  4548         (\<exists>ln rn. r = Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> bl_bin (<args>) @ Bk \<up> rn))
  4511       apply(rule_tac x = n in exI, simp add: exp_ind)
  4549         holds_for steps0 (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) n"
  4512       using h
  4550         using h prepare_mainpart_lemma[of args m]
  4513       apply(case_tac args, simp_all, case_tac list,
  4551         apply(auto)
  4514             simp_all add: tape_of_nl_abv  tape_of_nat_list.simps exp_ind_def
  4552         apply(rule_tac x = stp in exI, simp)
  4515             bl_bin.simps)
  4553         apply(rule_tac x = ln in exI, auto)
  4516       done     
  4554         done
       
  4555     qed
  4517   next
  4556   next
  4518     show "?Q1 \<turnstile>-> ?P2"
  4557     show "{?P2} t_wcode_adjust {?Q2}"
  4519       by(simp add: t_imply_def)
  4558     proof(rule_tac HoareI, auto del: replicate_Suc)
       
  4559       fix ln rn
       
  4560       show "\<exists>n. is_final (steps0 (Suc 0, Bk # Oc # Oc \<up> m, 
       
  4561         Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> bl_bin (<args>) @ Bk \<up> rn) t_wcode_adjust n) \<and>
       
  4562         wadjust_stop m (bl_bin (<args>) - Suc 0) holds_for steps0
       
  4563         (Suc 0, Bk # Oc # Oc \<up> m, Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> bl_bin (<args>) @ Bk \<up> rn) t_wcode_adjust n"
       
  4564         using wadjust_correctness[of m "bl_bin (<args>) - 1" "Suc ln" rn]
       
  4565         apply(simp del: replicate_Suc add: replicate_Suc[THEN sym] exp_ind, auto)
       
  4566         apply(rule_tac x = n in exI)
       
  4567         using a
       
  4568         apply(case_tac "bl_bin (<args>)", simp, simp del: replicate_Suc add: exp_ind wadjust_inv.simps)
       
  4569         done
       
  4570     qed
  4520   qed
  4571   qed
  4521   thus "\<exists>stp rn. steps (Suc 0, [], <m # args>) ((t_wcode_prepare |+| t_wcode_main) |+| 
  4572   thus "?thesis"
  4522         t_wcode_adjust) stp = (0, [Bk], Oc\<^bsup>Suc m\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  4573     apply(simp add: Hoare_def, auto)
  4523     apply(simp add: t_imply_def)
  4574     apply(case_tac "(steps0 (Suc 0, [], <(m::nat) # args>) 
  4524     apply(erule_tac exE)+
  4575       ((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) n)")
  4525     apply(subgoal_tac "bl_bin (<args>) > 0", auto simp: wadjust_stop.simps)
  4576     apply(rule_tac x = n in exI, auto simp: wadjust_stop.simps)
  4526     using h
  4577     using a
  4527     apply(case_tac args, simp_all, case_tac list,  
  4578     apply(case_tac "bl_bin (<args>)", simp_all)
  4528           simp_all add: tape_of_nl_abv  tape_of_nat_list.simps exp_ind_def
       
  4529             bl_bin.simps)
       
  4530     done
  4579     done
  4531 qed
  4580 qed
  4532 
  4581     
  4533 text {*
  4582 text {*
  4534   The initialization TM @{text "t_wcode"}.
  4583   The initialization TM @{text "t_wcode"}.
  4535   *}
  4584   *}
  4536 definition t_wcode :: "tprog"
  4585 definition t_wcode :: "instr list"
  4537   where
  4586   where
  4538   "t_wcode = (t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust"
  4587   "t_wcode = (t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust"
  4539 
  4588 
  4540 
  4589 
  4541 text {*
  4590 text {*
  4542   The correctness of @{text "t_wcode"}.
  4591   The correctness of @{text "t_wcode"}.
  4543   *}
  4592   *}
       
  4593 
  4544 lemma wcode_lemma_1:
  4594 lemma wcode_lemma_1:
  4545   "args \<noteq> [] \<Longrightarrow> 
  4595   "args \<noteq> [] \<Longrightarrow> 
  4546   \<exists> stp ln rn. steps (Suc 0, [], <m # args>)  (t_wcode) stp = 
  4596   \<exists> stp ln rn. steps0 (Suc 0, [], <m # args>)  (t_wcode) stp = 
  4547               (0,  [Bk],  Oc\<^bsup>Suc m\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>)"
  4597               (0,  [Bk],  Oc\<up>(Suc m) @ Bk # Oc\<up>(Suc (bl_bin (<args>))) @ Bk\<up>(rn))"
  4548 apply(simp add: wcode_lemma_pre' t_wcode_def)
  4598 apply(simp add: wcode_lemma_pre' t_wcode_def del: replicate_Suc)
  4549 done
  4599 done
  4550 
  4600 
  4551 lemma wcode_lemma: 
  4601 lemma wcode_lemma: 
  4552   "args \<noteq> [] \<Longrightarrow> 
  4602   "args \<noteq> [] \<Longrightarrow> 
  4553   \<exists> stp ln rn. steps (Suc 0, [], <m # args>)  (t_wcode) stp = 
  4603   \<exists> stp ln rn. steps0 (Suc 0, [], <m # args>)  (t_wcode) stp = 
  4554               (0,  [Bk],  <[m ,bl_bin (<args>)]> @ Bk\<^bsup>rn\<^esup>)"
  4604               (0,  [Bk],  <[m ,bl_bin (<args>)]> @ Bk\<up>(rn))"
  4555 using wcode_lemma_1[of args m]
  4605 using wcode_lemma_1[of args m]
  4556 apply(simp add: t_wcode_def tape_of_nl_abv tape_of_nat_list.simps)
  4606 apply(simp add: t_wcode_def tape_of_nl_abv tape_of_nat_list.simps)
  4557 done
  4607 done
  4558 
  4608 
  4559 section {* The universal TM *}
  4609 section {* The universal TM *}
  4562   This section gives the explicit construction of {\em Universal Turing Machine}, defined as @{text "UTM"} and proves its 
  4612   This section gives the explicit construction of {\em Universal Turing Machine}, defined as @{text "UTM"} and proves its 
  4563   correctness. It is pretty easy by composing the partial results we have got so far.
  4613   correctness. It is pretty easy by composing the partial results we have got so far.
  4564   *}
  4614   *}
  4565 
  4615 
  4566 
  4616 
  4567 definition UTM :: "tprog"
  4617 definition UTM :: "instr list"
  4568   where
  4618   where
  4569   "UTM = (let (aprog, rs_pos, a_md) = rec_ci rec_F in 
  4619   "UTM = (let (aprog, rs_pos, a_md) = rec_ci rec_F in 
  4570           let abc_F = aprog [+] dummy_abc (Suc (Suc 0)) in 
  4620           let abc_F = aprog [+] dummy_abc (Suc (Suc 0)) in 
  4571           (t_wcode |+| (tm_of abc_F @ tMp (Suc (Suc 0)) (start_of (layout_of abc_F) 
  4621           (t_wcode |+| (tm_of abc_F @ shift (mopup (Suc (Suc 0))) (length (tm_of abc_F) div 2))))"
  4572                                                    (length abc_F) - Suc 0))))"
       
  4573 
  4622 
  4574 definition F_aprog :: "abc_prog"
  4623 definition F_aprog :: "abc_prog"
  4575   where
  4624   where
  4576   "F_aprog \<equiv> (let (aprog, rs_pos, a_md) = rec_ci rec_F in 
  4625   "F_aprog \<equiv> (let (aprog, rs_pos, a_md) = rec_ci rec_F in 
  4577                        aprog [+] dummy_abc (Suc (Suc 0)))"
  4626                        aprog [+] dummy_abc (Suc (Suc 0)))"
  4578 
  4627 
  4579 definition F_tprog :: "tprog"
  4628 definition F_tprog :: "instr list"
  4580   where
  4629   where
  4581   "F_tprog = tm_of (F_aprog)"
  4630   "F_tprog = tm_of (F_aprog)"
  4582 
  4631 
  4583 definition t_utm :: "tprog"
  4632 definition t_utm :: "instr list"
  4584   where
  4633   where
  4585   "t_utm \<equiv>
  4634   "t_utm \<equiv>
  4586      (F_tprog) @ tMp (Suc (Suc 0)) (start_of (layout_of (F_aprog)) 
  4635      F_tprog @ shift (mopup (Suc (Suc 0))) (length F_tprog div 2)"
  4587                                   (length (F_aprog)) - Suc 0)"
  4636 
  4588 
  4637 definition UTM_pre :: "instr list"
  4589 definition UTM_pre :: "tprog"
       
  4590   where
  4638   where
  4591   "UTM_pre = t_wcode |+| t_utm"
  4639   "UTM_pre = t_wcode |+| t_utm"
  4592 
  4640 
       
  4641 (*
  4593 lemma F_abc_halt_eq:
  4642 lemma F_abc_halt_eq:
  4594   "\<lbrakk>turing_basic.t_correct tp; 
  4643   "\<lbrakk>turing_basic.t_correct tp; 
  4595     length lm = k;
  4644     length lm = k;
  4596     steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>);
  4645     steps (Suc 0, Bk\<up>(l), <lm>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(n));
  4597     rs > 0\<rbrakk>
  4646     rs > 0\<rbrakk>
  4598     \<Longrightarrow> \<exists> stp m. abc_steps_l (0, [code tp, bl2wc (<lm>)]) (F_aprog) stp =
  4647     \<Longrightarrow> \<exists> stp m. abc_steps_l (0, [code tp, bl2wc (<lm>)]) (F_aprog) stp =
  4599                        (length (F_aprog), code tp # bl2wc (<lm>) # (rs - 1) # 0\<^bsup>m\<^esup>)"
  4648                        (length (F_aprog), code tp # bl2wc (<lm>) # (rs - 1) # 0\<up>(m))"
  4600 apply(drule_tac  F_t_halt_eq, simp, simp, simp)
  4649 apply(drule_tac  F_t_halt_eq, simp, simp, simp)
  4601 apply(case_tac "rec_ci rec_F")
  4650 apply(case_tac "rec_ci rec_F")
  4602 apply(frule_tac abc_append_dummy_complie, simp, simp, erule_tac exE,
  4651 apply(frule_tac abc_append_dummy_complie, simp, simp, erule_tac exE,
  4603       erule_tac exE)
  4652       erule_tac exE)
  4604 apply(rule_tac x = stp in exI, rule_tac x = m in exI)
  4653 apply(rule_tac x = stp in exI, rule_tac x = m in exI)
  4606 done
  4655 done
  4607 
  4656 
  4608 lemma F_abc_utm_halt_eq: 
  4657 lemma F_abc_utm_halt_eq: 
  4609   "\<lbrakk>rs > 0; 
  4658   "\<lbrakk>rs > 0; 
  4610   abc_steps_l (0, [code tp, bl2wc (<lm>)]) F_aprog stp =
  4659   abc_steps_l (0, [code tp, bl2wc (<lm>)]) F_aprog stp =
  4611         (length F_aprog, code tp #  bl2wc (<lm>) # (rs - 1) # 0\<^bsup>m\<^esup>)\<rbrakk>
  4660         (length F_aprog, code tp #  bl2wc (<lm>) # (rs - 1) # 0\<up>(m))\<rbrakk>
  4612   \<Longrightarrow> \<exists>stp m n.(steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp =
  4661   \<Longrightarrow> \<exists>stp m n.(steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp =
  4613                                              (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>))"
  4662                                              (0, Bk\<up>(m), Oc\<up>(rs) @ Bk\<up>(n)))"
  4614   thm abacus_turing_eq_halt
  4663   thm abacus_turing_eq_halt
  4615   using abacus_turing_eq_halt
  4664   using abacus_turing_eq_halt
  4616   [of "layout_of F_aprog" "F_aprog" "F_tprog" "length (F_aprog)" 
  4665   [of "layout_of F_aprog" "F_aprog" "F_tprog" "length (F_aprog)" 
  4617     "[code tp, bl2wc (<lm>)]" stp "code tp # bl2wc (<lm>) # (rs - 1) # 0\<^bsup>m\<^esup>" "Suc (Suc 0)"
  4666     "[code tp, bl2wc (<lm>)]" stp "code tp # bl2wc (<lm>) # (rs - 1) # 0\<up>(m)" "Suc (Suc 0)"
  4618     "start_of (layout_of (F_aprog)) (length (F_aprog))" "[]" 0]
  4667     "start_of (layout_of (F_aprog)) (length (F_aprog))" "[]" 0]
  4619 apply(simp add: F_tprog_def t_utm_def abc_lm_v.simps nth_append)
  4668 apply(simp add: F_tprog_def t_utm_def abc_lm_v.simps nth_append)
  4620 apply(erule_tac exE)+
  4669 apply(erule_tac exE)+
  4621 apply(rule_tac x = stpa in exI, rule_tac x = "Suc (Suc ma)" in exI, 
  4670 apply(rule_tac x = stpa in exI, rule_tac x = "Suc (Suc ma)" in exI, 
  4622        rule_tac x = l in exI, simp add: exp_ind)
  4671        rule_tac x = l in exI, simp add: exp_ind)
  4625 declare tape_of_nl_abv_cons[simp del]
  4674 declare tape_of_nl_abv_cons[simp del]
  4626 
  4675 
  4627 lemma t_utm_halt_eq': 
  4676 lemma t_utm_halt_eq': 
  4628   "\<lbrakk>turing_basic.t_correct tp;
  4677   "\<lbrakk>turing_basic.t_correct tp;
  4629    0 < rs;
  4678    0 < rs;
  4630   steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>)\<rbrakk>
  4679   steps (Suc 0, Bk\<up>(l), <lm::nat list>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(n))\<rbrakk>
  4631   \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp = 
  4680   \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp = 
  4632                                                 (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
  4681                                                 (0, Bk\<up>(m), Oc\<up>(rs) @ Bk\<up>(n))"
  4633 apply(drule_tac  l = l in F_abc_halt_eq, simp, simp, simp)
  4682 apply(drule_tac  l = l in F_abc_halt_eq, simp, simp, simp)
  4634 apply(erule_tac exE, erule_tac exE)
  4683 apply(erule_tac exE, erule_tac exE)
  4635 apply(rule_tac F_abc_utm_halt_eq, simp_all)
  4684 apply(rule_tac F_abc_utm_halt_eq, simp_all)
  4636 done
  4685 done
  4637 
  4686 *)
  4638 lemma [simp]: "tinres xs (xs @ Bk\<^bsup>i\<^esup>)"
  4687 (*
       
  4688 lemma [simp]: "tinres xs (xs @ Bk\<up>(i))"
  4639 apply(auto simp: tinres_def)
  4689 apply(auto simp: tinres_def)
  4640 done
  4690 done
  4641 
  4691 
  4642 lemma [elim]: "\<lbrakk>rs > 0; Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup> = c @ Bk\<^bsup>n\<^esup>\<rbrakk>
  4692 lemma [elim]: "\<lbrakk>rs > 0; Oc\<up>(rs) @ Bk\<up>(na) = c @ Bk\<up>(n)\<rbrakk>
  4643         \<Longrightarrow> \<exists>n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>"
  4693         \<Longrightarrow> \<exists>n. c = Oc\<up>(rs) @ Bk\<up>(n)"
  4644 apply(case_tac "na > n")
  4694 apply(case_tac "na > n")
  4645 apply(subgoal_tac "\<exists> d. na = d + n", auto simp: exp_add)
  4695 apply(subgoal_tac "\<exists> d. na = d + n", auto simp: exp_add)
  4646 apply(rule_tac x = "na - n" in exI, simp)
  4696 apply(rule_tac x = "na - n" in exI, simp)
  4647 apply(subgoal_tac "\<exists> d. n = d + na", auto simp: exp_add)
  4697 apply(subgoal_tac "\<exists> d. n = d + na", auto simp: exp_add)
  4648 apply(case_tac rs, simp_all add: exp_ind, case_tac d, 
  4698 apply(case_tac rs, simp_all add: exp_ind, case_tac d, 
  4649            simp_all add: exp_ind)
  4699            simp_all add: exp_ind)
  4650 apply(rule_tac x = "n - na" in exI, simp)
  4700 apply(rule_tac x = "n - na" in exI, simp)
  4651 done
  4701 done
  4652 
  4702 *)
  4653 
  4703 (*
  4654 lemma t_utm_halt_eq'': 
  4704 lemma t_utm_halt_eq'': 
  4655   "\<lbrakk>turing_basic.t_correct tp;
  4705   "\<lbrakk>turing_basic.t_correct tp;
  4656    0 < rs;
  4706    0 < rs;
  4657    steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>)\<rbrakk>
  4707    steps (Suc 0, Bk\<up>(l), <lm::nat list>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(n))\<rbrakk>
  4658   \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stp = 
  4708   \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>(i)) t_utm stp = 
  4659                                                 (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
  4709                                                 (0, Bk\<up>(m), Oc\<up>(rs) @ Bk\<up>(n))"
  4660 apply(drule_tac t_utm_halt_eq', simp_all)
  4710 apply(drule_tac t_utm_halt_eq', simp_all)
  4661 apply(erule_tac exE)+
  4711 apply(erule_tac exE)+
  4662 proof -
  4712 proof -
  4663   fix stpa ma na
  4713   fix stpa ma na
  4664   assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stpa = (0, Bk\<^bsup>ma\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>)"
  4714   assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stpa = (0, Bk\<up>(ma), Oc\<up>(rs) @ Bk\<up>(na))"
  4665   and gr: "rs > 0"
  4715   and gr: "rs > 0"
  4666   thus "\<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
  4716   thus "\<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>(i)) t_utm stp = (0, Bk\<up>(m), Oc\<up>(rs) @ Bk\<up>(n))"
  4667     apply(rule_tac x = stpa in exI, rule_tac x = ma in exI,  simp)
  4717     apply(rule_tac x = stpa in exI, rule_tac x = ma in exI,  simp)
  4668   proof(case_tac "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa", simp)
  4718   proof(case_tac "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>(i)) t_utm stpa", simp)
  4669     fix a b c
  4719     fix a b c
  4670     assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stpa = (0, Bk\<^bsup>ma\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>)"
  4720     assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stpa = (0, Bk\<up>(ma), Oc\<up>(rs) @ Bk\<up>(na))"
  4671             "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (a, b, c)"
  4721             "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>(i)) t_utm stpa = (a, b, c)"
  4672     thus " a = 0 \<and> b = Bk\<^bsup>ma\<^esup> \<and> (\<exists>n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
  4722     thus " a = 0 \<and> b = Bk\<up>(ma) \<and> (\<exists>n. c = Oc\<up>(rs) @ Bk\<up>(n))"
  4673       using tinres_steps2[of "<[code tp, bl2wc (<lm>)]>" "<[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>" 
  4723       using tinres_steps2[of "<[code tp, bl2wc (<lm>)]>" "<[code tp, bl2wc (<lm>)]> @ Bk\<up>(i)" 
  4674                            "Suc 0" " [Bk, Bk]" t_utm stpa 0 "Bk\<^bsup>ma\<^esup>" "Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>" a b c]
  4724                            "Suc 0" " [Bk, Bk]" t_utm stpa 0 "Bk\<up>(ma)" "Oc\<up>(rs) @ Bk\<up>(na)" a b c]
  4675       apply(simp)
  4725       apply(simp)
  4676       using gr
  4726       using gr
  4677       apply(simp only: tinres_def, auto)
  4727       apply(simp only: tinres_def, auto)
  4678       apply(rule_tac x = "na + n" in exI, simp add: exp_add)
  4728       apply(rule_tac x = "na + n" in exI, simp add: exp_add)
  4679       done
  4729       done
  4682 
  4732 
  4683 lemma [simp]: "tinres [Bk, Bk] [Bk]"
  4733 lemma [simp]: "tinres [Bk, Bk] [Bk]"
  4684 apply(auto simp: tinres_def)
  4734 apply(auto simp: tinres_def)
  4685 done
  4735 done
  4686 
  4736 
  4687 lemma [elim]: "Bk\<^bsup>ma\<^esup> = b @ Bk\<^bsup>n\<^esup>  \<Longrightarrow> \<exists>m. b = Bk\<^bsup>m\<^esup>"
  4737 lemma [elim]: "Bk\<up>(ma) = b @ Bk\<up>(n)  \<Longrightarrow> \<exists>m. b = Bk\<up>(m)"
  4688 apply(subgoal_tac "ma = length b + n")
  4738 apply(subgoal_tac "ma = length b + n")
  4689 apply(rule_tac x = "ma - n" in exI, simp add: exp_add)
  4739 apply(rule_tac x = "ma - n" in exI, simp add: exp_add)
  4690 apply(drule_tac length_equal)
  4740 apply(drule_tac length_equal)
  4691 apply(simp)
  4741 apply(simp)
  4692 done
  4742 done
       
  4743 *)
       
  4744 
       
  4745 
       
  4746 
       
  4747 lemma tinres_step1: 
       
  4748   "\<lbrakk>tinres l l'; step (ss, l, r) (t, 0) = (sa, la, ra); 
       
  4749                  step (ss, l', r) (t, 0) = (sb, lb, rb)\<rbrakk>
       
  4750     \<Longrightarrow> tinres la lb \<and> ra = rb \<and> sa = sb"
       
  4751 apply(case_tac ss, case_tac [!]r, case_tac [!] "a::cell")
       
  4752 apply(auto simp: step.simps fetch.simps nth_of.simps
       
  4753         split: if_splits )
       
  4754 apply(case_tac [!] "t ! (2 * nat)", 
       
  4755      auto simp: tinres_def split: if_splits)
       
  4756 apply(case_tac [1-8] a, auto split: if_splits)
       
  4757 apply(case_tac [!] "t ! (2 * nat)", 
       
  4758      auto simp: tinres_def split: if_splits)
       
  4759 apply(case_tac [1-4] a, auto split: if_splits)
       
  4760 apply(case_tac [!] "t ! Suc (2 * nat)", 
       
  4761      auto simp: if_splits)
       
  4762 apply(case_tac [!] aa, auto split: if_splits)
       
  4763 done
       
  4764 
       
  4765 lemma tinres_steps1: 
       
  4766   "\<lbrakk>tinres l l'; steps (ss, l, r) (t, 0) stp = (sa, la, ra); 
       
  4767                  steps (ss, l', r) (t, 0) stp = (sb, lb, rb)\<rbrakk>
       
  4768     \<Longrightarrow> tinres la lb \<and> ra = rb \<and> sa = sb"
       
  4769 apply(induct stp arbitrary: sa la ra sb lb rb, simp add: steps.simps)
       
  4770 apply(simp add: step_red)
       
  4771 apply(case_tac "(steps (ss, l, r) (t, 0) stp)")
       
  4772 apply(case_tac "(steps (ss, l', r) (t, 0) stp)")
       
  4773 proof -
       
  4774   fix stp sa la ra sb lb rb a b c aa ba ca
       
  4775   assume ind: "\<And>sa la ra sb lb rb. \<lbrakk>steps (ss, l, r) (t, 0) stp = (sa, (la::cell list), ra); 
       
  4776           steps (ss, l', r) (t, 0) stp = (sb, lb, rb)\<rbrakk> \<Longrightarrow> tinres la lb \<and> ra = rb \<and> sa = sb"
       
  4777   and h: " tinres l l'" "step (steps (ss, l, r) (t, 0) stp) (t, 0) = (sa, la, ra)"
       
  4778          "step (steps (ss, l', r) (t, 0) stp) (t, 0) = (sb, lb, rb)" "steps (ss, l, r) (t, 0) stp = (a, b, c)" 
       
  4779          "steps (ss, l', r) (t, 0) stp = (aa, ba, ca)"
       
  4780   have "tinres b ba \<and> c = ca \<and> a = aa"
       
  4781     apply(rule_tac ind, simp_all add: h)
       
  4782     done
       
  4783   thus "tinres la lb \<and> ra = rb \<and> sa = sb"
       
  4784     apply(rule_tac l = b and l' = ba and r = c  and ss = a   
       
  4785             and t = t in tinres_step1)
       
  4786     using h
       
  4787     apply(simp, simp, simp)
       
  4788     done
       
  4789 qed
       
  4790 
       
  4791 lemma [simp]: 
       
  4792   "tinres (Bk \<up> m @ [Bk, Bk]) la \<Longrightarrow> \<exists>m. la = Bk \<up> m"
       
  4793 apply(auto simp: tinres_def)
       
  4794 apply(case_tac n, simp add: exp_ind)
       
  4795 apply(rule_tac  x ="Suc (Suc m)" in exI, simp only: exp_ind, simp)
       
  4796 apply(simp add: exp_ind del: replicate_Suc)
       
  4797 apply(case_tac nat, simp add: exp_ind)
       
  4798 apply(rule_tac x = "Suc m" in exI, simp only: exp_ind)
       
  4799 apply(simp only: exp_ind, simp)
       
  4800 apply(subgoal_tac "m = length la + nata")
       
  4801 apply(rule_tac x = "m - nata" in exI, simp add: exp_add)
       
  4802 apply(drule_tac length_equal, simp)
       
  4803 apply(simp only: exp_ind[THEN sym] replicate_Suc[THEN sym] replicate_add[THEN sym])
       
  4804 apply(rule_tac x = "m + Suc (Suc n)" in exI, simp)
       
  4805 done
  4693 
  4806 
  4694 lemma t_utm_halt_eq: 
  4807 lemma t_utm_halt_eq: 
  4695   "\<lbrakk>turing_basic.t_correct tp;
  4808   assumes tm_wf: "tm_wf (tp, 0)"
  4696    0 < rs;
  4809   and exec: "steps0 (Suc 0, Bk\<up>(l), <lm::nat list>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(n))"
  4697    steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>)\<rbrakk>
  4810   and resutl: "0 < rs"
  4698   \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stp = 
  4811   shows "\<exists>stp m n. steps0 (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>(i)) t_utm stp = 
  4699                                                 (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
  4812                                                 (0, Bk\<up>(m), Oc\<up>(rs) @ Bk\<up>(n))"
  4700 apply(drule_tac i = i in t_utm_halt_eq'', simp_all)
       
  4701 apply(erule_tac exE)+
       
  4702 proof -
  4813 proof -
  4703   fix stpa ma na
  4814   obtain ap arity fp where a: "rec_ci rec_F = (ap, arity, fp)"
  4704   assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (0, Bk\<^bsup>ma\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>)"
  4815     by (metis prod_cases3) 
  4705   and gr: "rs > 0"
  4816   moreover have b: "rec_calc_rel  rec_F [code tp, (bl2wc (<lm>))] (rs - Suc 0)"
  4706   thus "\<exists>stp m n. steps (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
  4817     using assms
  4707     apply(rule_tac x = stpa in exI)
  4818     apply(rule_tac F_correct, simp_all)
  4708   proof(case_tac "steps (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa", simp)
  4819     done 
  4709     fix a b c
  4820   have "\<exists> stp m l. steps0 (Suc 0, Bk # Bk # [], <[code tp, bl2wc (<lm>)]> @ Bk\<up>i)
  4710     assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (0, Bk\<^bsup>ma\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>)"
  4821     (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp
  4711             "steps (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (a, b, c)"
  4822     = (0, Bk\<up>m @ Bk # Bk # [], Oc\<up>Suc (rs - 1) @ Bk\<up>l)"  
  4712     thus "a = 0 \<and> (\<exists>m. b = Bk\<^bsup>m\<^esup>) \<and> (\<exists>n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
  4823   proof(rule_tac recursive_compile_to_tm_correct)
  4713       using tinres_steps[of "[Bk, Bk]" "[Bk]" "Suc 0" "<[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>" t_utm stpa 0
  4824     show "rec_ci rec_F = (ap, arity, fp)" using a by simp
  4714                              "Bk\<^bsup>ma\<^esup>" "Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>" a b c]
  4825   next
  4715       apply(simp)
  4826     show "rec_calc_rel rec_F [code tp, bl2wc (<lm>)] (rs - 1)"
  4716       apply(auto simp: tinres_def)
  4827       using b by simp
  4717       apply(rule_tac x = "ma + n" in exI, simp add: exp_add)
  4828   next
       
  4829     show "length [code tp, bl2wc (<lm>)] = 2" by simp
       
  4830   next
       
  4831     show "layout_of (ap [+] dummy_abc 2) = layout_of (ap [+] dummy_abc 2)"
       
  4832       by simp
       
  4833   next
       
  4834     show "F_tprog = tm_of (ap [+] dummy_abc 2)"
       
  4835       using a
       
  4836       apply(simp add: F_tprog_def F_aprog_def numeral_2_eq_2)
  4718       done
  4837       done
  4719   qed
  4838   qed
  4720 qed
  4839   then obtain stp m l where 
  4721 
  4840     "steps0 (Suc 0, Bk # Bk # [], <[code tp, bl2wc (<lm>)]> @ Bk\<up>i)
  4722 lemma [intro]: "t_correct t_wcode"
  4841     (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp
  4723 apply(simp add: t_wcode_def)
  4842     = (0, Bk\<up>m @ Bk # Bk # [], Oc\<up>Suc (rs - 1) @ Bk\<up>l)" by blast
  4724 apply(auto)
  4843   hence "\<exists> m. steps0 (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>i)
  4725 done
  4844     (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp
  4726       
  4845     = (0, Bk\<up>m, Oc\<up>Suc (rs - 1) @ Bk\<up>l)"
  4727 lemma [intro]: "t_correct t_utm"
  4846   proof -
  4728 apply(simp add: t_utm_def F_tprog_def)
  4847     assume g: "steps0 (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk \<up> i)
  4729 apply(rule_tac t_compiled_correct, auto)
  4848       (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp =
  4730 done   
  4849       (0, Bk \<up> m @ [Bk, Bk], Oc \<up> Suc (rs - 1) @ Bk \<up> l)"
  4731 
  4850    moreover have "tinres [Bk, Bk] [Bk]"
  4732 lemma UTM_halt_lemma_pre: 
  4851      apply(auto simp: tinres_def)
  4733   "\<lbrakk>turing_basic.t_correct tp;
  4852      done
  4734    0 < rs;
  4853     moreover obtain sa la ra where "steps0 (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>i)
  4735    args \<noteq> [];
  4854     (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp = (sa, la, ra)"
  4736    steps (Suc 0, Bk\<^bsup>i\<^esup>, <args::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>k\<^esup>)\<rbrakk>
  4855       apply(case_tac "steps0 (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>i)
  4737   \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [], <code tp # args>) UTM_pre stp = 
  4856     (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp", auto)
  4738                                                 (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
       
  4739 proof -
       
  4740   let ?Q2 = "\<lambda> (l, r). (\<exists> ln rn. l = Bk\<^bsup>ln\<^esup> \<and> r = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>rn\<^esup>)"
       
  4741   term ?Q2
       
  4742   let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <code tp # args>"
       
  4743   let ?Q1 = "\<lambda> (l, r). (l = [Bk] \<and>
       
  4744              (\<exists> rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>))"
       
  4745   let ?P2 = ?Q1
       
  4746   let ?P3 = "\<lambda> (l, r). False"
       
  4747   assume h: "turing_basic.t_correct tp" "0 < rs"
       
  4748             "args \<noteq> []" "steps (Suc 0, Bk\<^bsup>i\<^esup>, <args::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>k\<^esup>)"
       
  4749   have "?P1 \<turnstile>-> \<lambda> tp. (\<exists> stp tp'. steps (Suc 0, tp)
       
  4750                     (t_wcode |+| t_utm) stp = (0, tp') \<and> ?Q2 tp')"
       
  4751   proof(rule_tac turing_merge.t_merge_halt [of "t_wcode" "t_utm"
       
  4752           ?P1 ?P2 ?P3 ?P3 ?Q1 ?Q2], auto simp: turing_merge_def)
       
  4753     show "\<exists>stp. case steps (Suc 0, [], <code tp # args>) t_wcode stp of (st, tp') \<Rightarrow> 
       
  4754        st = 0 \<and> (case tp' of (l, r) \<Rightarrow> l = [Bk] \<and>
       
  4755                    (\<exists>rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>))"
       
  4756       using wcode_lemma_1[of args "code tp"] h
       
  4757       apply(simp, auto)
       
  4758       apply(rule_tac x = stpa in exI, auto)
       
  4759       done      
       
  4760   next
       
  4761     fix rn 
       
  4762     show "\<exists>stp. case steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @
       
  4763       Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp of
       
  4764       (st, tp') \<Rightarrow> st = 0 \<and> (case tp' of (l, r) \<Rightarrow>
       
  4765       (\<exists>ln. l = Bk\<^bsup>ln\<^esup>) \<and> (\<exists>rn. r = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>rn\<^esup>))"
       
  4766       using t_utm_halt_eq[of tp rs i args stp m k rn] h
       
  4767       apply(auto)
       
  4768       apply(rule_tac x = stpa in exI, simp add: bin_wc_eq 
       
  4769         tape_of_nat_list.simps tape_of_nl_abv)
       
  4770       apply(auto)
       
  4771       done
  4857       done
  4772   next
  4858     ultimately show "?thesis"
  4773     show "?Q1 \<turnstile>-> ?P2"
  4859       apply(drule_tac tinres_steps1, auto)
  4774       apply(simp add: t_imply_def)
       
  4775       done
  4860       done
  4776   qed
  4861   qed
  4777   thus "?thesis"
  4862   thus "?thesis"
  4778     apply(simp add: t_imply_def)
  4863     apply(auto)
  4779     apply(auto simp: UTM_pre_def)
  4864     apply(rule_tac x = stp in exI, simp add: t_utm_def)
       
  4865     using assms
       
  4866     apply(case_tac rs, simp_all add: numeral_2_eq_2)
       
  4867     done
       
  4868 qed
       
  4869 
       
  4870 lemma [intro]: "tm_wf (t_wcode, 0)"
       
  4871 apply(simp add: t_wcode_def)
       
  4872 apply(rule_tac tm_wf_comp)
       
  4873 apply(rule_tac tm_wf_comp, auto)
       
  4874 done
       
  4875       
       
  4876 lemma [intro]: "tm_wf (t_utm, 0)"
       
  4877 apply(simp only: t_utm_def F_tprog_def)
       
  4878 apply(rule_tac t_compiled_correct, auto)
       
  4879 done   
       
  4880 
       
  4881 lemma UTM_halt_lemma_pre: 
       
  4882   assumes wf_tm: "tm_wf (tp, 0)"
       
  4883   and result: "0 < rs"
       
  4884   and args: "args \<noteq> []"
       
  4885   and exec: "steps0 (Suc 0, Bk\<up>(i), <args::nat list>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(k))"
       
  4886   shows "\<exists>stp m n. steps0 (Suc 0, [], <code tp # args>) UTM_pre stp = 
       
  4887                                                 (0, Bk\<up>(m), Oc\<up>(rs) @ Bk\<up>(n))"
       
  4888 proof -
       
  4889   let ?Q2 = "\<lambda> (l, r). (\<exists> ln rn. l = Bk\<up>(ln) \<and> r = Oc\<up>(rs) @ Bk\<up>(rn))"
       
  4890   let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <code tp # args>"
       
  4891   let ?Q1 = "\<lambda> (l, r). (l = [Bk] \<and>
       
  4892     (\<exists> rn. r = Oc\<up>(Suc (code tp)) @ Bk # Oc\<up>(Suc (bl_bin (<args>))) @ Bk\<up>(rn)))"
       
  4893   let ?P2 = ?Q1
       
  4894   let ?P3 = "\<lambda> (l, r). False"
       
  4895   have "{?P1} (t_wcode |+| t_utm) {?Q2}"
       
  4896   proof(rule_tac Hoare_plus_halt)
       
  4897     show "?Q1 \<mapsto> ?P2"
       
  4898       by(simp add: assert_imp_def)
       
  4899   next
       
  4900     show "tm_wf (t_wcode, 0)" by auto
       
  4901   next
       
  4902     show "{?P1} t_wcode {?Q1}"
       
  4903       apply(rule_tac HoareI, auto)
       
  4904       using wcode_lemma_1[of args "code tp"] args
       
  4905       apply(auto)
       
  4906       apply(rule_tac x = stp in exI, simp)
       
  4907       done
       
  4908   next
       
  4909     show "{?P2} t_utm {?Q2}"
       
  4910     proof(rule_tac HoareI, auto)
       
  4911       fix rn
       
  4912       show "\<exists>n. is_final (steps0 (Suc 0, [Bk], Oc # Oc \<up> code tp @ Bk # Oc # Oc \<up> bl_bin (<args>) @ Bk \<up> rn) t_utm n) \<and>
       
  4913         (\<lambda>(l, r). (\<exists>ln. l = Bk \<up> ln) \<and>
       
  4914         (\<exists>rn. r = Oc \<up> rs @ Bk \<up> rn)) holds_for steps0 (Suc 0, [Bk],
       
  4915         Oc # Oc \<up> code tp @ Bk # Oc # Oc \<up> bl_bin (<args>) @ Bk \<up> rn) t_utm n"
       
  4916         using t_utm_halt_eq[of tp i "args" stp m rs k rn] assms
       
  4917       apply(auto simp: bin_wc_eq)
       
  4918       apply(rule_tac x = stpa in exI, simp add: tape_of_nl_abv)
       
  4919       done
       
  4920     qed
       
  4921   qed
       
  4922   thus "?thesis"
       
  4923     apply(auto simp: Hoare_def UTM_pre_def)
       
  4924     apply(case_tac "steps0 (Suc 0, [], <code tp # args>) (t_wcode |+| t_utm) n")
       
  4925     apply(rule_tac x = n in exI, simp)
  4780     done
  4926     done
  4781 qed
  4927 qed
  4782 
  4928 
  4783 text {*
  4929 text {*
  4784   The correctness of @{text "UTM"}, the halt case.
  4930   The correctness of @{text "UTM"}, the halt case.
  4785 *}
  4931 *}
  4786 lemma UTM_halt_lemma: 
  4932 lemma UTM_halt_lemma: 
  4787   "\<lbrakk>turing_basic.t_correct tp;
  4933   assumes tm_wf: "tm_wf (tp, 0)"
  4788    0 < rs;
  4934   and result: "0 < rs"
  4789    args \<noteq> [];
  4935   and args: "args \<noteq> []"
  4790    steps (Suc 0, Bk\<^bsup>i\<^esup>, <args::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>k\<^esup>)\<rbrakk>
  4936   and exec: "steps0 (Suc 0, Bk\<up>(i), <args::nat list>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(k))"
  4791   \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [], <code tp # args>) UTM stp = 
  4937   shows "\<exists>stp m n. steps0 (Suc 0, [], <code tp # args>) UTM stp = 
  4792                                                 (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)"
  4938                                                 (0, Bk\<up>(m), Oc\<up>(rs) @ Bk\<up>(n))"
  4793 using UTM_halt_lemma_pre[of tp rs args i stp m k]
  4939 using UTM_halt_lemma_pre[of tp rs args i stp m k] assms
  4794 apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def)
  4940 apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def)
  4795 apply(case_tac "rec_ci rec_F", simp)
  4941 apply(case_tac "rec_ci rec_F", simp)
  4796 done
  4942 done
  4797 
  4943 
  4798 definition TSTD:: "t_conf \<Rightarrow> bool"
  4944 definition TSTD:: "config \<Rightarrow> bool"
  4799   where
  4945   where
  4800   "TSTD c = (let (st, l, r) = c in 
  4946   "TSTD c = (let (st, l, r) = c in 
  4801              st = 0 \<and> (\<exists> m. l = Bk\<^bsup>m\<^esup>) \<and> (\<exists> rs n. r = Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>))"
  4947              st = 0 \<and> (\<exists> m. l = Bk\<up>(m)) \<and> (\<exists> rs n. r = Oc\<up>(Suc rs) @ Bk\<up>(n)))"
  4802 
       
  4803 thm abacus_turing_eq_uhalt
       
  4804 
  4948 
  4805 lemma nstd_case1: "0 < a \<Longrightarrow> NSTD (trpl_code (a, b, c))"
  4949 lemma nstd_case1: "0 < a \<Longrightarrow> NSTD (trpl_code (a, b, c))"
  4806 apply(simp add: NSTD.simps trpl_code.simps)
  4950 apply(simp add: NSTD.simps trpl_code.simps)
  4807 done
  4951 done
  4808 
  4952 
  4809 lemma [simp]: "\<forall>m. b \<noteq> Bk\<^bsup>m\<^esup> \<Longrightarrow> 0 < bl2wc b"
  4953 lemma [simp]: "\<forall>m. b \<noteq> Bk\<up>(m) \<Longrightarrow> 0 < bl2wc b"
  4810 apply(rule classical, simp)
  4954 apply(rule classical, simp)
  4811 apply(induct b, erule_tac x = 0 in allE, simp)
  4955 apply(induct b, erule_tac x = 0 in allE, simp)
  4812 apply(simp add: bl2wc.simps, case_tac a, simp_all 
  4956 apply(simp add: bl2wc.simps, case_tac a, simp_all 
  4813   add: bl2nat.simps bl2nat_double)
  4957   add: bl2nat.simps bl2nat_double)
  4814 apply(case_tac "\<exists> m. b = Bk\<^bsup>m\<^esup>",  erule exE)
  4958 apply(case_tac "\<exists> m. b = Bk\<up>(m)",  erule exE)
  4815 apply(erule_tac x = "Suc m" in allE, simp add: exp_ind_def, simp)
  4959 apply(erule_tac x = "Suc m" in allE, simp add: , simp)
  4816 done
  4960 done
  4817 lemma nstd_case2: "\<forall>m. b \<noteq> Bk\<^bsup>m\<^esup> \<Longrightarrow> NSTD (trpl_code (a, b, c))"
  4961 
       
  4962 lemma nstd_case2: "\<forall>m. b \<noteq> Bk\<up>(m) \<Longrightarrow> NSTD (trpl_code (a, b, c))"
  4818 apply(simp add: NSTD.simps trpl_code.simps)
  4963 apply(simp add: NSTD.simps trpl_code.simps)
  4819 done
  4964 done
  4820 
       
  4821 thm lg.simps
       
  4822 thm lgR.simps
       
  4823 
  4965 
  4824 lemma [elim]: "Suc (2 * x) = 2 * y \<Longrightarrow> RR"
  4966 lemma [elim]: "Suc (2 * x) = 2 * y \<Longrightarrow> RR"
  4825 apply(induct x arbitrary: y, simp, simp)
  4967 apply(induct x arbitrary: y, simp, simp)
  4826 apply(case_tac y, simp, simp)
  4968 apply(case_tac y, simp, simp)
  4827 done
  4969 done
  4828 
  4970 
  4829 lemma bl2nat_zero_eq[simp]: "(bl2nat c 0 = 0) = (\<exists>n. c = Bk\<^bsup>n\<^esup>)"
  4971 declare replicate_Suc[simp del]
       
  4972 
       
  4973 lemma bl2nat_zero_eq[simp]: "(bl2nat c 0 = 0) = (\<exists>n. c = Bk\<up>(n))"
  4830 apply(auto)
  4974 apply(auto)
  4831 apply(induct c, simp add: bl2nat.simps)
  4975 apply(induct c, simp_all add: bl2nat.simps)
  4832 apply(rule_tac x = 0 in exI, simp)
       
  4833 apply(case_tac a, auto simp: bl2nat.simps bl2nat_double)
  4976 apply(case_tac a, auto simp: bl2nat.simps bl2nat_double)
  4834 done
  4977 done
  4835 
  4978 
  4836 lemma bl2wc_exp_ex: 
  4979 lemma bl2wc_exp_ex: 
  4837   "\<lbrakk>Suc (bl2wc c) = 2 ^  m\<rbrakk> \<Longrightarrow> \<exists> rs n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>"
  4980   "\<lbrakk>Suc (bl2wc c) = 2 ^  m\<rbrakk> \<Longrightarrow> \<exists> rs n. c = Oc\<up>(rs) @ Bk\<up>(n)"
  4838 apply(induct c arbitrary: m, simp add: bl2wc.simps bl2nat.simps)
  4981 apply(induct c arbitrary: m, simp add: bl2wc.simps bl2nat.simps)
  4839 apply(case_tac a, auto)
  4982 apply(case_tac a, auto)
  4840 apply(case_tac m, simp_all add: bl2wc.simps, auto)
  4983 apply(case_tac m, simp_all add: bl2wc.simps, auto)
  4841 apply(rule_tac x = 0 in exI, rule_tac x = "Suc n" in exI, 
  4984 apply(rule_tac x = 0 in exI, rule_tac x = "Suc n" in exI, 
  4842   simp add: exp_ind_def)
  4985   simp add: replicate_Suc)
  4843 apply(simp add: bl2wc.simps bl2nat.simps bl2nat_double)
  4986 apply(simp add: bl2wc.simps bl2nat.simps bl2nat_double)
  4844 apply(case_tac m, simp, simp)
  4987 apply(case_tac m, simp, simp)
  4845 proof -
  4988 proof -
  4846   fix c m nat
  4989   fix c m nat
  4847   assume ind: 
  4990   assume ind: 
  4848     "\<And>m. Suc (bl2nat c 0) = 2 ^ m \<Longrightarrow> \<exists>rs n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>"
  4991     "\<And>m. Suc (bl2nat c 0) = 2 ^ m \<Longrightarrow> \<exists>rs n. c = Oc\<up>(rs) @ Bk\<up>(n)"
  4849   and h: 
  4992   and h: 
  4850     "Suc (Suc (2 * bl2nat c 0)) = 2 * 2 ^ nat"
  4993     "Suc (Suc (2 * bl2nat c 0)) = 2 * 2 ^ nat"
  4851   have "\<exists>rs n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>"
  4994   have "\<exists>rs n. c = Oc\<up>(rs) @ Bk\<up>(n)"
  4852     apply(rule_tac m = nat in ind)
  4995     apply(rule_tac m = nat in ind)
  4853     using h
  4996     using h
  4854     apply(simp)
  4997     apply(simp)
  4855     done
  4998     done
  4856   from this obtain rs n where " c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>" by blast 
  4999   from this obtain rs n where " c = Oc\<up>(rs) @ Bk\<up>(n)" by blast 
  4857   thus "\<exists>rs n. Oc # c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>"
  5000   thus "\<exists>rs n. Oc # c = Oc\<up>(rs) @ Bk\<up>(n)"
  4858     apply(rule_tac x = "Suc rs" in exI, simp add: exp_ind_def)
  5001     apply(rule_tac x = "Suc rs" in exI, simp add: )
  4859     apply(rule_tac x = n in exI, simp)
  5002     apply(rule_tac x = n in exI, simp add: replicate_Suc)
  4860     done
  5003     done
  4861 qed
  5004 qed
  4862 
  5005 
  4863 lemma [elim]: 
  5006 lemma lg_bin: 
  4864   "\<lbrakk>\<forall>rs n. c \<noteq> Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>; 
  5007   "\<lbrakk>\<forall>rs n. c \<noteq> Oc\<up>(Suc rs) @ Bk\<up>(n); 
  4865   bl2wc c = 2 ^ lg (Suc (bl2wc c)) 2 - Suc 0\<rbrakk> \<Longrightarrow> bl2wc c = 0"
  5008   bl2wc c = 2 ^ lg (Suc (bl2wc c)) 2 - Suc 0\<rbrakk> \<Longrightarrow> bl2wc c = 0"
  4866 apply(subgoal_tac "\<exists> m. Suc (bl2wc c) = 2^m", erule_tac exE)
  5009 apply(subgoal_tac "\<exists> m. Suc (bl2wc c) = 2^m", erule_tac exE)
  4867 apply(drule_tac bl2wc_exp_ex, simp, erule_tac exE, erule_tac exE)
  5010 apply(drule_tac bl2wc_exp_ex, simp, erule_tac exE, erule_tac exE)
  4868 apply(case_tac rs, simp, simp, erule_tac x = nat in allE,
  5011 apply(case_tac rs, simp, simp, erule_tac x = nat in allE,
  4869   erule_tac x = n in allE, simp)
  5012   erule_tac x = n in allE, simp)
  4874 apply(rule_tac x = rs in exI)
  5017 apply(rule_tac x = rs in exI)
  4875 apply(case_tac "(2::nat)^rs", simp, simp)
  5018 apply(case_tac "(2::nat)^rs", simp, simp)
  4876 done
  5019 done
  4877 
  5020 
  4878 lemma nstd_case3: 
  5021 lemma nstd_case3: 
  4879   "\<forall>rs n. c \<noteq> Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup> \<Longrightarrow>  NSTD (trpl_code (a, b, c))"
  5022   "\<forall>rs n. c \<noteq> Oc\<up>(Suc rs) @ Bk\<up>(n) \<Longrightarrow>  NSTD (trpl_code (a, b, c))"
  4880 apply(simp add: NSTD.simps trpl_code.simps)
  5023 apply(simp add: NSTD.simps trpl_code.simps)
  4881 apply(rule_tac impI)
  5024 apply(auto)
  4882 apply(rule_tac disjI2, rule_tac disjI2, auto)
  5025 apply(drule_tac lg_bin, simp_all)
  4883 done
  5026 done
  4884 
  5027 
  4885 lemma NSTD_1: "\<not> TSTD (a, b, c)
  5028 lemma NSTD_1: "\<not> TSTD (a, b, c)
  4886     \<Longrightarrow> rec_exec rec_NSTD [trpl_code (a, b, c)] = Suc 0"
  5029     \<Longrightarrow> rec_exec rec_NSTD [trpl_code (a, b, c)] = Suc 0"
  4887   using NSTD_lemma1[of "trpl_code (a, b, c)"]
  5030   using NSTD_lemma1[of "trpl_code (a, b, c)"]
  4891   apply(erule_tac disjE, erule_tac nstd_case2)
  5034   apply(erule_tac disjE, erule_tac nstd_case2)
  4892   apply(erule_tac nstd_case3)
  5035   apply(erule_tac nstd_case3)
  4893   done
  5036   done
  4894  
  5037  
  4895 lemma nonstop_t_uhalt_eq:
  5038 lemma nonstop_t_uhalt_eq:
  4896       "\<lbrakk>turing_basic.t_correct tp;
  5039   "\<lbrakk>tm_wf (tp, 0);
  4897         steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp = (a, b, c);
  5040   steps0 (Suc 0, Bk\<up>(l), <lm>) tp stp = (a, b, c);
  4898        \<not> TSTD (a, b, c)\<rbrakk>
  5041   \<not> TSTD (a, b, c)\<rbrakk>
  4899        \<Longrightarrow> rec_exec rec_nonstop [code tp, bl2wc (<lm>), stp] = Suc 0"
  5042   \<Longrightarrow> rec_exec rec_nonstop [code tp, bl2wc (<lm>), stp] = Suc 0"
  4900 apply(simp add: rec_nonstop_def rec_exec.simps)
  5043 apply(simp add: rec_nonstop_def rec_exec.simps)
  4901 apply(subgoal_tac 
  5044 apply(subgoal_tac 
  4902   "rec_exec rec_conf [code tp, bl2wc (<lm>), stp] =
  5045   "rec_exec rec_conf [code tp, bl2wc (<lm>), stp] =
  4903   trpl_code (a, b, c)", simp)
  5046   trpl_code (a, b, c)", simp)
  4904 apply(erule_tac NSTD_1)
  5047 apply(erule_tac NSTD_1)
  4905 using rec_t_eq_steps[of tp l lm stp]
  5048 using rec_t_eq_steps[of tp l lm stp]
  4906 apply(simp)
  5049 apply(simp)
  4907 done
  5050 done
  4908 
  5051 
  4909 lemma nonstop_true:
  5052 lemma nonstop_true:
  4910   "\<lbrakk>turing_basic.t_correct tp;
  5053   "\<lbrakk>tm_wf (tp, 0);
  4911   \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp))\<rbrakk>
  5054   \<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <lm>) tp stp))\<rbrakk>
  4912      \<Longrightarrow> \<forall>y. rec_calc_rel rec_nonstop 
  5055   \<Longrightarrow> \<forall>y. rec_calc_rel rec_nonstop 
  4913                         ([code tp, bl2wc (<lm>), y]) (Suc 0)"
  5056   ([code tp, bl2wc (<lm>), y]) (Suc 0)"
  4914 apply(rule_tac allI, erule_tac x = y in allE)
  5057 apply(rule_tac allI, erule_tac x = y in allE)
  4915 apply(case_tac "steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp y", simp)
  5058 apply(case_tac "steps0 (Suc 0, Bk\<up>(l), <lm>) tp y", simp)
  4916 apply(rule_tac nonstop_t_uhalt_eq, simp_all)
  5059 apply(rule_tac nonstop_t_uhalt_eq, simp_all)
  4917 done
  5060 done
  4918 
  5061 
  4919 (*
  5062 (*
  4920 lemma [simp]: 
  5063 lemma [simp]: 
  4926 done
  5069 done
  4927 *)
  5070 *)
  4928 declare ci_cn_para_eq[simp]
  5071 declare ci_cn_para_eq[simp]
  4929 
  5072 
  4930 lemma F_aprog_uhalt: 
  5073 lemma F_aprog_uhalt: 
  4931   "\<lbrakk>turing_basic.t_correct tp; 
  5074   "\<lbrakk>tm_wf (tp,0); 
  4932     \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp)); 
  5075     \<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <lm>) tp stp)); 
  4933     rec_ci rec_F = (F_ap, rs_pos, a_md)\<rbrakk>
  5076     rec_ci rec_F = (F_ap, rs_pos, a_md)\<rbrakk>
  4934   \<Longrightarrow> \<forall> stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)] @ 0\<^bsup>a_md - rs_pos \<^esup>
  5077   \<Longrightarrow> \<forall> stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)] @ 0\<up>(a_md - rs_pos )
  4935                @ suflm) (F_ap) stp of (ss, e) \<Rightarrow> ss < length (F_ap)"
  5078                @ suflm) (F_ap) stp of (ss, e) \<Rightarrow> ss < length (F_ap)"
  4936 apply(case_tac "rec_ci (Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) rec_conf 
  5079 apply(case_tac "rec_ci (Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0)) rec_conf 
  4937                ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])])")
  5080                ([id (Suc (Suc 0)) 0, id (Suc (Suc 0)) (Suc 0), rec_halt])])")
  4938 apply(simp only: rec_F_def, rule_tac i = 0  and ga = a and gb = b and 
  5081 apply(simp only: rec_F_def, rule_tac i = 0  and ga = a and gb = b and 
  4939   gc = c in cn_gi_uhalt, simp, simp, simp, simp, simp, simp, simp)
  5082   gc = c in cn_gi_uhalt, simp, simp, simp, simp, simp, simp, simp)
  4972 apply(rule_tac allI)
  5115 apply(rule_tac allI)
  4973 apply(erule_tac x = y in allE)+
  5116 apply(erule_tac x = y in allE)+
  4974 apply(simp)
  5117 apply(simp)
  4975 done
  5118 done
  4976 
  5119 
  4977 thm abc_list_crsp_steps
       
  4978 
       
  4979 lemma uabc_uhalt': 
  5120 lemma uabc_uhalt': 
  4980   "\<lbrakk>turing_basic.t_correct tp;
  5121   "\<lbrakk>tm_wf (tp, 0);
  4981   \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp));
  5122   \<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <lm>) tp stp));
  4982   rec_ci rec_F = (ap, pos, md)\<rbrakk>
  5123   rec_ci rec_F = (ap, pos, md)\<rbrakk>
  4983   \<Longrightarrow> \<forall> stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)]) ap stp of (ss, e)
  5124   \<Longrightarrow> \<forall> stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)]) ap stp of (ss, e)
  4984            \<Rightarrow>  ss < length ap"
  5125            \<Rightarrow>  ss < length ap"
  4985 proof(frule_tac F_ap = ap and rs_pos = pos and a_md = md
  5126 proof(frule_tac F_ap = ap and rs_pos = pos and a_md = md
  4986     and suflm = "[]" in F_aprog_uhalt, auto)
  5127     and suflm = "[]" in F_aprog_uhalt, auto)
  4987   fix stp a b
  5128   fix stp a b
  4988   assume h: 
  5129   assume h: 
  4989     "\<forall>stp. case abc_steps_l (0, code tp # bl2wc (<lm>) # 0\<^bsup>md - pos\<^esup>) ap stp of 
  5130     "\<forall>stp. case abc_steps_l (0, code tp # bl2wc (<lm>) # 0\<up>(md - pos)) ap stp of 
  4990     (ss, e) \<Rightarrow> ss < length ap"
  5131     (ss, e) \<Rightarrow> ss < length ap"
  4991     "abc_steps_l (0, [code tp, bl2wc (<lm>)]) ap stp = (a, b)" 
  5132     "abc_steps_l (0, [code tp, bl2wc (<lm>)]) ap stp = (a, b)" 
  4992     "turing_basic.t_correct tp" 
  5133     "tm_wf (tp, 0)" 
  4993     "rec_ci rec_F = (ap, pos, md)"
  5134     "rec_ci rec_F = (ap, pos, md)"
  4994   moreover have "ap \<noteq> []"
  5135   moreover have "ap \<noteq> []"
  4995     using h apply(rule_tac rec_ci_not_null, simp)
  5136     using h apply(rule_tac rec_ci_not_null, simp)
  4996     done
  5137     done
  4997   ultimately show "a < length ap"
  5138   ultimately show "a < length ap"
  4998   proof(erule_tac x = stp in allE,
  5139   proof(erule_tac x = stp in allE,
  4999   case_tac "abc_steps_l (0, code tp # bl2wc (<lm>) # 0\<^bsup>md - pos\<^esup>) ap stp", simp)
  5140   case_tac "abc_steps_l (0, code tp # bl2wc (<lm>) # 0\<up>(md - pos)) ap stp", simp)
  5000     fix aa ba
  5141     fix aa ba
  5001     assume g: "aa < length ap" 
  5142     assume g: "aa < length ap" 
  5002       "abc_steps_l (0, code tp # bl2wc (<lm>) # 0\<^bsup>md - pos\<^esup>) ap stp = (aa, ba)" 
  5143       "abc_steps_l (0, code tp # bl2wc (<lm>) # 0\<up>(md - pos)) ap stp = (aa, ba)" 
  5003       "ap \<noteq> []"
  5144       "ap \<noteq> []"
  5004     thus "?thesis"
  5145     thus "?thesis"
  5005       using abc_list_crsp_steps[of "[code tp, bl2wc (<lm>)]"
  5146       using abc_list_crsp_steps[of "[code tp, bl2wc (<lm>)]"
  5006                                    "md - pos" ap stp aa ba] h
  5147                                    "md - pos" ap stp aa ba] h
  5007       apply(simp)
  5148       apply(simp)
  5008       done
  5149       done
  5009   qed
  5150   qed
  5010 qed
  5151 qed
  5011 
  5152 
  5012 lemma uabc_uhalt: 
  5153 lemma uabc_uhalt: 
  5013   "\<lbrakk>turing_basic.t_correct tp; 
  5154   "\<lbrakk>tm_wf (tp, 0); 
  5014   \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp))\<rbrakk>
  5155   \<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <lm>) tp stp))\<rbrakk>
  5015   \<Longrightarrow> \<forall> stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)]) F_aprog 
  5156   \<Longrightarrow> \<forall> stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)]) F_aprog 
  5016        stp of (ss, e) \<Rightarrow> ss < length F_aprog"
  5157        stp of (ss, e) \<Rightarrow> ss < length F_aprog"
  5017 apply(case_tac "rec_ci rec_F", simp add: F_aprog_def)
  5158 apply(case_tac "rec_ci rec_F", simp add: F_aprog_def)
  5018 thm uabc_uhalt'
  5159 thm uabc_uhalt'
  5019 apply(drule_tac ap = a and pos = b and md = c in uabc_uhalt', simp_all)
  5160 apply(drule_tac ap = a and pos = b and md = c in uabc_uhalt', simp_all)
  5032     apply(simp)
  5173     apply(simp)
  5033     done
  5174     done
  5034 qed
  5175 qed
  5035 
  5176 
  5036 lemma tutm_uhalt': 
  5177 lemma tutm_uhalt': 
  5037   "\<lbrakk>turing_basic.t_correct tp;
  5178 assumes tm_wf:  "tm_wf (tp,0)"
  5038     \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp))\<rbrakk>
  5179   and unhalt: "\<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <lm>) tp stp))"
  5039   \<Longrightarrow> \<forall> stp. \<not> isS0 (steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp)"
  5180   shows "\<forall> stp. \<not> is_final (steps0 (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp)"
  5040   using abacus_turing_eq_uhalt[of "layout_of (F_aprog)" 
  5181 apply(simp add: t_utm_def)
  5041                "F_aprog" "F_tprog" "[code tp, bl2wc (<lm>)]" 
  5182 proof(rule_tac compile_correct_unhalt)
  5042                "start_of (layout_of (F_aprog )) (length (F_aprog))" 
  5183   show "layout_of F_aprog = layout_of F_aprog" by simp
  5043                "Suc (Suc 0)"]
  5184 next
  5044 apply(simp add: F_tprog_def)
  5185   show "F_tprog = tm_of F_aprog"
  5045 apply(subgoal_tac "\<forall>stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)])
  5186     by(simp add:  F_tprog_def)
  5046   (F_aprog) stp of (as, am) \<Rightarrow> as < length (F_aprog)", simp)
  5187 next
  5047 thm abacus_turing_eq_uhalt
  5188   show "crsp (layout_of F_aprog) (0, [code tp, bl2wc (<lm>)]) (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>)  []"
  5048 apply(simp add: t_utm_def F_tprog_def)
  5189     by(auto simp: crsp.simps start_of.simps)
  5049 apply(rule_tac uabc_uhalt, simp_all)
  5190 next
  5050 done
  5191   show "length F_tprog div 2 = length F_tprog div 2" by simp
  5051 
  5192 next
       
  5193   show "\<forall>stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)]) F_aprog stp of (as, am) \<Rightarrow> as < length F_aprog"
       
  5194     using assms
       
  5195     apply(erule_tac uabc_uhalt, simp)
       
  5196     done
       
  5197 qed
       
  5198 
       
  5199  
  5052 lemma tinres_commute: "tinres r r' \<Longrightarrow> tinres r' r"
  5200 lemma tinres_commute: "tinres r r' \<Longrightarrow> tinres r' r"
  5053 apply(auto simp: tinres_def)
  5201 apply(auto simp: tinres_def)
  5054 done
  5202 done
  5055 
  5203 
  5056 lemma inres_tape:
  5204 lemma inres_tape:
  5057   "\<lbrakk>steps (st, l, r) tp stp = (a, b, c); steps (st, l', r') tp stp = (a', b', c'); 
  5205   "\<lbrakk>steps0 (st, l, r) tp stp = (a, b, c); steps0 (st, l', r') tp stp = (a', b', c'); 
  5058   tinres l l'; tinres r r'\<rbrakk>
  5206   tinres l l'; tinres r r'\<rbrakk>
  5059   \<Longrightarrow> a = a' \<and> tinres b b' \<and> tinres c c'"
  5207   \<Longrightarrow> a = a' \<and> tinres b b' \<and> tinres c c'"
  5060 proof(case_tac "steps (st, l', r) tp stp")
  5208 proof(case_tac "steps0 (st, l', r) tp stp")
  5061   fix aa ba ca
  5209   fix aa ba ca
  5062   assume h: "steps (st, l, r) tp stp = (a, b, c)" 
  5210   assume h: "steps0 (st, l, r) tp stp = (a, b, c)" 
  5063             "steps (st, l', r') tp stp = (a', b', c')"
  5211             "steps0 (st, l', r') tp stp = (a', b', c')"
  5064             "tinres l l'" "tinres r r'"
  5212             "tinres l l'" "tinres r r'"
  5065             "steps (st, l', r) tp stp = (aa, ba, ca)"
  5213             "steps0 (st, l', r) tp stp = (aa, ba, ca)"
  5066   have "tinres b ba \<and> c = ca \<and> a = aa"
  5214   have "tinres b ba \<and> c = ca \<and> a = aa"
  5067     using h
  5215     using h
  5068     apply(rule_tac tinres_steps, auto)
  5216     apply(rule_tac tinres_steps1, auto)
  5069     done
  5217     done
  5070 
       
  5071   thm tinres_steps2
       
  5072   moreover have "b' = ba \<and> tinres c' ca \<and> a' =  aa"
  5218   moreover have "b' = ba \<and> tinres c' ca \<and> a' =  aa"
  5073     using h
  5219     using h
  5074     apply(rule_tac tinres_steps2, auto intro: tinres_commute)
  5220     apply(rule_tac tinres_steps2, auto intro: tinres_commute)
  5075     done
  5221     done
  5076   ultimately show "?thesis"
  5222   ultimately show "?thesis"
  5077     apply(auto intro: tinres_commute)
  5223     apply(auto intro: tinres_commute)
  5078     done
  5224     done
  5079 qed
  5225 qed
  5080 
  5226 
  5081 lemma tape_normalize: "\<forall> stp. \<not> isS0 (steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp)
  5227 lemma tape_normalize: "\<forall> stp. \<not> is_final(steps0 (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp)
  5082       \<Longrightarrow> \<forall> stp. \<not> isS0 (steps (Suc 0, Bk\<^bsup>m\<^esup>, <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>n\<^esup>) t_utm stp)"
  5228       \<Longrightarrow> \<forall> stp. \<not> is_final (steps0 (Suc 0, Bk\<up>(m), <[code tp, bl2wc (<lm>)]> @ Bk\<up>(n)) t_utm stp)"
  5083 apply(rule_tac allI, case_tac "(steps (Suc 0, Bk\<^bsup>m\<^esup>, 
  5229 apply(rule_tac allI, case_tac "(steps0 (Suc 0, Bk\<up>(m), 
  5084                <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>n\<^esup>) t_utm stp)", simp add: isS0_def)
  5230                <[code tp, bl2wc (<lm>)]> @ Bk\<up>(n)) t_utm stp)", simp)
  5085 apply(erule_tac x = stp in allE)
  5231 apply(erule_tac x = stp in allE)
  5086 apply(case_tac "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp", simp)
  5232 apply(case_tac "steps0 (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp", simp)
  5087 apply(drule_tac inres_tape, auto)
  5233 apply(drule_tac inres_tape, auto)
  5088 apply(auto simp: tinres_def)
  5234 apply(auto simp: tinres_def)
  5089 apply(case_tac "m > Suc (Suc 0)")
  5235 apply(case_tac "m > Suc (Suc 0)")
  5090 apply(rule_tac x = "m - Suc (Suc 0)" in exI) 
  5236 apply(rule_tac x = "m - Suc (Suc 0)" in exI) 
  5091 apply(case_tac m, simp_all add: exp_ind_def, case_tac nat, simp_all add: exp_ind_def)
  5237 apply(case_tac m, simp_all add: , case_tac nat, simp_all add: replicate_Suc)
  5092 apply(rule_tac x = "2 - m" in exI, simp add: exp_ind_def[THEN sym] exp_add[THEN sym])
  5238 apply(rule_tac x = "2 - m" in exI, simp add: exp_add[THEN sym])
  5093 apply(simp only: numeral_2_eq_2, simp add: exp_ind_def)
  5239 apply(simp only: numeral_2_eq_2, simp add: replicate_Suc)
  5094 done
  5240 done
  5095 
  5241 
  5096 lemma tutm_uhalt: 
  5242 lemma tutm_uhalt: 
  5097   "\<lbrakk>turing_basic.t_correct tp;
  5243   "\<lbrakk>tm_wf (tp,0);
  5098     \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <args>) tp stp))\<rbrakk>
  5244     \<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <args>) tp stp))\<rbrakk>
  5099   \<Longrightarrow> \<forall> stp. \<not> isS0 (steps (Suc 0, Bk\<^bsup>m\<^esup>, <[code tp, bl2wc (<args>)]> @ Bk\<^bsup>n\<^esup>) t_utm stp)"
  5245   \<Longrightarrow> \<forall> stp. \<not> is_final (steps0 (Suc 0, Bk\<up>(m), <[code tp, bl2wc (<args>)]> @ Bk\<up>(n)) t_utm stp)"
  5100 apply(rule_tac tape_normalize)
  5246 apply(rule_tac tape_normalize)
  5101 apply(rule_tac tutm_uhalt', simp_all)
  5247 apply(rule_tac tutm_uhalt', simp_all)
  5102 done
  5248 done
  5103 
  5249 
  5104 lemma UTM_uhalt_lemma_pre:
  5250 lemma UTM_uhalt_lemma_pre:
  5105   "\<lbrakk>turing_basic.t_correct tp;
  5251   assumes tm_wf: "tm_wf (tp, 0)"
  5106    \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <args>) tp stp));
  5252   and exec: "\<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <args>) tp stp))"
  5107    args \<noteq> []\<rbrakk>
  5253   and args: "args \<noteq> []"
  5108   \<Longrightarrow>  \<forall> stp. \<not> isS0 (steps (Suc 0, [], <code tp # args>)  UTM_pre stp)"
  5254   shows "\<forall> stp. \<not> is_final (steps0 (Suc 0, [], <code tp # args>)  UTM_pre stp)"
  5109 proof -
  5255 proof -
  5110   let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <code tp # args>"
  5256   let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <code tp # args>"
  5111   let ?Q1 = "\<lambda> (l, r). (l = [Bk] \<and>
  5257   let ?Q1 = "\<lambda> (l, r). (l = [Bk] \<and>
  5112              (\<exists> rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>))"
  5258              (\<exists> rn. r = Oc\<up>(Suc (code tp)) @ Bk # Oc\<up>(Suc (bl_bin (<args>))) @ Bk\<up>(rn)))"
  5113   let ?P4 = ?Q1
  5259   let ?P2 = ?Q1
  5114   let ?P3 = "\<lambda> (l, r). False"
  5260   have "{?P1} (t_wcode |+| t_utm) \<up>"
  5115   assume h: "turing_basic.t_correct tp" "\<forall>stp. \<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <args>) tp stp)"
  5261   proof(rule_tac Hoare_plus_unhalt)
  5116             "args \<noteq> []"
  5262     show "?Q1 \<mapsto> ?P2"
  5117   have "?P1 \<turnstile>-> \<lambda> tp. \<not> (\<exists> stp. isS0 (steps (Suc 0, tp) (t_wcode |+| t_utm) stp))"
  5263       by(simp add: assert_imp_def)
  5118   proof(rule_tac turing_merge.t_merge_uhalt [of "t_wcode" "t_utm"
       
  5119           ?P1 ?P3 ?P3 ?P4 ?Q1 ?P3], auto simp: turing_merge_def)
       
  5120     show "\<exists>stp. case steps (Suc 0, [], <code tp # args>) t_wcode stp of (st, tp') \<Rightarrow> 
       
  5121        st = 0 \<and> (case tp' of (l, r) \<Rightarrow> l = [Bk] \<and>
       
  5122                    (\<exists>rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>))"
       
  5123       using wcode_lemma_1[of args "code tp"] h
       
  5124       apply(simp, auto)
       
  5125       apply(rule_tac x = stp in exI, auto)
       
  5126       done      
       
  5127   next
  5264   next
  5128     fix rn  stp
  5265     show "tm_wf (t_wcode, 0)" by auto
  5129     show " isS0 (steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp)
  5266   next
  5130           \<Longrightarrow> False"
  5267     show "{?P1} t_wcode {?Q1}"
  5131       using tutm_uhalt[of tp l args "Suc 0" rn] h
  5268       apply(rule_tac HoareI, auto)
  5132       apply(simp)
  5269       using wcode_lemma_1[of args "code tp"] args
  5133       apply(erule_tac x = stp in allE)
  5270       apply(auto)
  5134       apply(simp add: tape_of_nl_abv tape_of_nat_list.simps bin_wc_eq)
  5271       apply(rule_tac x = stp in exI, simp)
  5135       done
  5272       done
  5136   next
  5273   next
  5137     fix rn stp
  5274     show "{?P2} t_utm \<up>"
  5138     show "isS0 (steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp) \<Longrightarrow>
  5275     proof(rule_tac Hoare_unhalt_I, auto)
  5139       isS0 (steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp)"
  5276       fix n rn
  5140       by simp
  5277       assume h: "is_final (steps0 (Suc 0, [Bk], Oc \<up> Suc (code tp) @ Bk # Oc \<up> Suc (bl_bin (<args>)) @ Bk \<up> rn) t_utm n)"
  5141   next
  5278       have "\<forall> stp. \<not> is_final (steps0 (Suc 0, Bk\<up>(Suc 0), <[code tp, bl2wc (<args>)]> @ Bk\<up>(rn)) t_utm stp)"
  5142     show "?Q1 \<turnstile>-> ?P4"
  5279         using assms
  5143       apply(simp add: t_imply_def)
  5280         apply(rule_tac tutm_uhalt, simp_all)
  5144       done
  5281         done
       
  5282       thus "False"
       
  5283         using h
       
  5284         apply(erule_tac x = n in allE)
       
  5285         apply(simp add: tape_of_nl_abv bin_wc_eq)
       
  5286         done
       
  5287     qed
  5145   qed
  5288   qed
  5146   thus "?thesis"
  5289   thus "?thesis"
  5147     apply(simp add: t_imply_def UTM_pre_def)
  5290     apply(simp add: Hoare_unhalt_def UTM_pre_def)
  5148     done
  5291     done
  5149 qed
  5292 qed
  5150 
  5293 
  5151 text {*
  5294 text {*
  5152   The correctness of @{text "UTM"}, the unhalt case.
  5295   The correctness of @{text "UTM"}, the unhalt case.
  5153   *}
  5296   *}
  5154 
  5297 
  5155 lemma UTM_uhalt_lemma:
  5298 lemma UTM_uhalt_lemma:
  5156   "\<lbrakk>turing_basic.t_correct tp;
  5299   assumes tm_wf: "tm_wf (tp, 0)"
  5157    \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <args>) tp stp));
  5300   and unhalt: "\<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <args>) tp stp))"
  5158    args \<noteq> []\<rbrakk>
  5301   and args: "args \<noteq> []"
  5159   \<Longrightarrow>  \<forall> stp. \<not> isS0 (steps (Suc 0, [], <code tp # args>)  UTM stp)"
  5302   shows " \<forall> stp. \<not> is_final (steps0 (Suc 0, [], <code tp # args>)  UTM stp)"
  5160 using UTM_uhalt_lemma_pre[of tp l args]
  5303   using UTM_uhalt_lemma_pre[of tp l args] assms
  5161 apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def)
  5304 apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def)
  5162 apply(case_tac "rec_ci rec_F", simp)
  5305 apply(case_tac "rec_ci rec_F", simp)
  5163 done
  5306 done
  5164 
  5307 
  5165 end                               
  5308 end