changed theory names to uppercase
authorChristian Urban <christian dot urban at kcl dot ac dot uk>
Sun, 10 Feb 2013 19:49:07 +0000
changeset 163 67063c5365e1
parent 162 a63c3f8d7234
child 164 8a3e63163910
changed theory names to uppercase
IsaMakefile
README
ROOT.ML
paper.pdf
thys/Abacus.thy
thys/Rec_Def.thy
thys/Recursive.thy
thys/Turing.thy
thys/Turing_Hoare.thy
thys/UF.thy
thys/UTM.thy
thys/Uncomputable.thy
thys/abacus.thy
thys/rec_def.thy
thys/recursive.thy
thys/turing_basic.thy
thys/turing_hoare.thy
thys/uncomputable.thy
--- a/IsaMakefile	Thu Feb 07 06:39:06 2013 +0000
+++ b/IsaMakefile	Sun Feb 10 19:49:07 2013 +0000
@@ -14,7 +14,7 @@
 OUT = $(ISABELLE_OUTPUT)
 LOG = $(OUT)/log
 
-USEDIR = $(ISABELLE_TOOL) usedir -v true -i false -D generated
+USEDIR = $(ISABELLE_TOOL) usedir -v true -i false 
 
 
 ## utm
@@ -24,13 +24,13 @@
 $(OUT)/utm: thys/*.thy ROOT.ML
 	@$(USEDIR) -f ROOT.ML -b HOL UTM
 
-paper:  ROOT.ML \
-	document/root* \
-	*.thy
-	rm -rf generated # otherwise latex will fall over 
-	@$(USEDIR) -f ROOT1.ML UTM .
-	$(ISABELLE_TOOL) document -o pdf  generated
-	cp generated/root.pdf paper.pdf  
+#paper:  ROOT.ML \
+#	document/root* \
+#	*.thy
+#	rm -rf generated # otherwise latex will fall over 
+#	@$(USEDIR) -f ROOT1.ML UTM .
+#	$(ISABELLE_TOOL) document -o pdf  generated
+#	cp generated/root.pdf paper.pdf  
 
 ## ITP itp
 
--- a/README	Thu Feb 07 06:39:06 2013 +0000
+++ b/README	Sun Feb 10 19:49:07 2013 +0000
@@ -1,9 +1,24 @@
-Formalisation of Turing Machines
-================================
+Formalisation of Turing Machines and Computability Theory
+=========================================================
 
 thys    - contains the formalisation
+
+   Turing.thy:       Basic definitions of Turing machines.
+   Turing_Hoare.thy: Contains the Hoare rules	
+   Uncomputable.thy: The existence of Turing uncomputable functions
+   Abacus.thy:       Basic definitions of abacus machines (an intermediate
+                     "language" for compiling recursive functions into 
+                     Turing machines)
+   Rec_Def.thy:      Basic definitions of recursive functions.
+   Recursive.thy:    The compilation of recursive functions into
+		     abacus machines.
+   UF.thy:           The construction of the Universal Function, named "rec_F" and
+                     the proof of its correctness.
+   UTM.thy:          Obtaining a Universal Turing machine by translating the
+                     Turing machine compiled from "rec_F" with some 
+	             
+
 Paper   - contains the files for the paper
 
-
 Attic      - old files
 Literature - related work
--- a/ROOT.ML	Thu Feb 07 06:39:06 2013 +0000
+++ b/ROOT.ML	Sun Feb 10 19:49:07 2013 +0000
@@ -1,22 +1,10 @@
-(*
-	turing_basic.thy : The basic definitions of Turing Machine.
-	uncomputable.thy : The existence of Turing uncomputable functions.
-	abacus.thy : The basic definitions of Abacus machine (An intermediate langauge underneath recursive functions) and
-	             the compilation of Abacus machines into Turing Machines.
-	rec_def.thy: The basic definitions of Recursive Functions.
-	recursive.thy : The compilation of Recursive Functions into
-		     Abacus machines.
-	UF.thy : The construction of Universal Function, named "rec_F" and the proof of its correctness.
-	UTM.thy: Obtaining Uinversal Turing Machine by scarfolding the Turing Machine compiled from "rec_F" with some 
-	         initialization and termination processing Turing Machines.
-*)
-  
+ 
 no_document 
-use_thys ["thys/turing_basic",
-          "thys/turing_hoare", 
-	  "thys/uncomputable", 
-	  "thys/abacus", 
-	  "thys/rec_def", 
-	  "thys/recursive",
+use_thys ["thys/Turing",
+          "thys/Turing_Hoare", 
+	  "thys/Uncomputable", 
+	  "thys/Abacus", 
+	  "thys/Rec_Def", 
+	  "thys/Recursive",
           "thys/UF",
           "thys/UTM"]
Binary file paper.pdf has changed
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/Abacus.thy	Sun Feb 10 19:49:07 2013 +0000
@@ -0,0 +1,4974 @@
+header {* 
+ {\em abacus} a kind of register machine
+*}
+
+theory Abacus
+imports Uncomputable
+begin
+
+(*
+declare tm_comp.simps [simp add] 
+declare adjust.simps[simp add] 
+declare shift.simps[simp add]
+declare tm_wf.simps[simp add]
+declare step.simps[simp add]
+declare steps.simps[simp add]
+*)
+declare replicate_Suc[simp add]
+
+text {*
+  {\em Abacus} instructions:
+*}
+
+datatype abc_inst =
+  -- {* @{text "Inc n"} increments the memory cell (or register) with address @{text "n"} by one.
+     *}
+     Inc nat
+  -- {*
+     @{text "Dec n label"} decrements the memory cell with address @{text "n"} by one. 
+      If cell @{text "n"} is already zero, no decrements happens and the executio jumps to
+      the instruction labeled by @{text "label"}.
+     *}
+   | Dec nat nat
+  -- {*
+  @{text "Goto label"} unconditionally jumps to the instruction labeled by @{text "label"}.
+  *}
+   | Goto nat
+  
+
+text {*
+  Abacus programs are defined as lists of Abacus instructions.
+*}
+type_synonym abc_prog = "abc_inst list"
+
+section {*
+  Sample Abacus programs
+  *}
+
+text {*
+  Abacus for addition and clearance.
+*}
+fun plus_clear :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"  
+  where
+  "plus_clear m n e = [Dec m e, Inc n, Goto 0]"
+
+text {*
+  Abacus for clearing memory untis.
+*}
+fun clear :: "nat \<Rightarrow> nat \<Rightarrow> abc_prog"
+  where
+  "clear n e = [Dec n e, Goto 0]"
+
+fun plus:: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
+  where
+  "plus m n p e = [Dec m 4, Inc n, Inc p,
+                   Goto 0, Dec p e, Inc m, Goto 4]"
+
+fun mult :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
+  where
+  "mult m1 m2 n p e = [Dec m1 e]@ plus m1 m2 p 1"
+
+fun expo :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
+  where
+  "expo n m1 m2 p e = [Inc n, Dec m1 e] @ mult m2 n n p 2"
+
+
+text {*
+  The state of Abacus machine.
+  *}
+type_synonym abc_state = nat
+
+(* text {*
+  The memory of Abacus machine is defined as a function from address to contents.
+*}
+type_synonym abc_mem = "nat \<Rightarrow> nat" *)
+
+text {*
+  The memory of Abacus machine is defined as a list of contents, with 
+  every units addressed by index into the list.
+  *}
+type_synonym abc_lm = "nat list"
+
+text {*
+  Fetching contents out of memory. Units not represented by list elements are considered
+  as having content @{text "0"}.
+*}
+fun abc_lm_v :: "abc_lm \<Rightarrow> nat \<Rightarrow> nat"
+  where 
+    "abc_lm_v lm n = (if (n < length lm) then (lm!n) else 0)"         
+
+
+text {*
+  Set the content of memory unit @{text "n"} to value @{text "v"}.
+  @{text "am"} is the Abacus memory before setting.
+  If address @{text "n"} is outside to scope of @{text "am"}, @{text "am"} 
+  is extended so that @{text "n"} becomes in scope.
+*}
+
+fun abc_lm_s :: "abc_lm \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_lm"
+  where
+    "abc_lm_s am n v = (if (n < length am) then (am[n:=v]) else 
+                           am@ (replicate (n - length am) 0) @ [v])"
+
+
+text {*
+  The configuration of Abaucs machines consists of its current state and its
+  current memory:
+*}
+type_synonym abc_conf = "abc_state \<times> abc_lm"
+
+text {*
+  Fetch instruction out of Abacus program:
+*}
+
+fun abc_fetch :: "nat \<Rightarrow> abc_prog \<Rightarrow> abc_inst option" 
+  where
+  "abc_fetch s p = (if (s < length p) then Some (p ! s)
+                    else None)"
+
+text {*
+  Single step execution of Abacus machine. If no instruction is feteched, 
+  configuration does not change.
+*}
+fun abc_step_l :: "abc_conf \<Rightarrow> abc_inst option \<Rightarrow> abc_conf"
+  where
+  "abc_step_l (s, lm) a = (case a of 
+               None \<Rightarrow> (s, lm) |
+               Some (Inc n)  \<Rightarrow> (let nv = abc_lm_v lm n in
+                       (s + 1, abc_lm_s lm n (nv + 1))) |
+               Some (Dec n e) \<Rightarrow> (let nv = abc_lm_v lm n in
+                       if (nv = 0) then (e, abc_lm_s lm n 0) 
+                       else (s + 1,  abc_lm_s lm n (nv - 1))) |
+               Some (Goto n) \<Rightarrow> (n, lm) 
+               )"
+
+text {*
+  Multi-step execution of Abacus machine.
+*}
+fun abc_steps_l :: "abc_conf \<Rightarrow> abc_prog \<Rightarrow> nat \<Rightarrow> abc_conf"
+  where
+  "abc_steps_l (s, lm) p 0 = (s, lm)" |
+  "abc_steps_l (s, lm) p (Suc n) = 
+      abc_steps_l (abc_step_l (s, lm) (abc_fetch s p)) p n"
+
+section {*
+  Compiling Abacus machines into Truing machines
+*}
+
+subsection {*
+  Compiling functions
+*}
+
+text {*
+  @{text "findnth n"} returns the TM which locates the represention of
+  memory cell @{text "n"} on the tape and changes representation of zero
+  on the way.
+*}
+
+fun findnth :: "nat \<Rightarrow> instr list"
+  where
+  "findnth 0 = []" |
+  "findnth (Suc n) = (findnth n @ [(W1, 2 * n + 1), 
+           (R, 2 * n + 2), (R, 2 * n + 3), (R, 2 * n + 2)])"
+
+text {*
+  @{text "tinc_b"} returns the TM which increments the representation 
+  of the memory cell under rw-head by one and move the representation 
+  of cells afterwards to the right accordingly.
+  *}
+
+definition tinc_b :: "instr list"
+  where
+  "tinc_b \<equiv> [(W1, 1), (R, 2), (W1, 3), (R, 2), (W1, 3), (R, 4), 
+             (L, 7), (W0, 5), (R, 6), (W0, 5), (W1, 3), (R, 6),
+             (L, 8), (L, 7), (R, 9), (L, 7), (R, 10), (W0, 9)]" 
+
+text {*
+  @{text "tinc ss n"} returns the TM which simulates the execution of 
+  Abacus instruction @{text "Inc n"}, assuming that TM is located at
+  location @{text "ss"} in the final TM complied from the whole
+  Abacus program.
+*}
+
+fun tinc :: "nat \<Rightarrow> nat \<Rightarrow> instr list"
+  where
+  "tinc ss n = shift (findnth n @ shift tinc_b (2 * n)) (ss - 1)"
+
+text {*
+  @{text "tinc_b"} returns the TM which decrements the representation 
+  of the memory cell under rw-head by one and move the representation 
+  of cells afterwards to the left accordingly.
+  *}
+
+definition tdec_b :: "instr list"
+  where
+  "tdec_b \<equiv>  [(W1, 1), (R, 2), (L, 14), (R, 3), (L, 4), (R, 3),
+              (R, 5), (W0, 4), (R, 6), (W0, 5), (L, 7), (L, 8),
+              (L, 11), (W0, 7), (W1, 8), (R, 9), (L, 10), (R, 9),
+              (R, 5), (W0, 10), (L, 12), (L, 11), (R, 13), (L, 11),
+              (R, 17), (W0, 13), (L, 15), (L, 14), (R, 16), (L, 14),
+              (R, 0), (W0, 16)]"
+
+text {*
+  @{text "sete tp e"} attaches the termination edges (edges leading to state @{text "0"}) 
+  of TM @{text "tp"} to the intruction labelled by @{text "e"}.
+  *}
+
+fun sete :: "instr list \<Rightarrow> nat \<Rightarrow> instr list"
+  where
+  "sete tp e = map (\<lambda> (action, state). (action, if state = 0 then e else state)) tp"
+
+text {*
+  @{text "tdec ss n label"} returns the TM which simulates the execution of 
+  Abacus instruction @{text "Dec n label"}, assuming that TM is located at
+  location @{text "ss"} in the final TM complied from the whole
+  Abacus program.
+*}
+
+fun tdec :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> instr list"
+  where
+  "tdec ss n e = shift (findnth n) (ss - 1) @ sete (shift (shift tdec_b (2 * n)) (ss - 1)) e"
+ 
+text {*
+  @{text "tgoto f(label)"} returns the TM simulating the execution of Abacus instruction
+  @{text "Goto label"}, where @{text "f(label)"} is the corresponding location of
+  @{text "label"} in the final TM compiled from the overall Abacus program.
+*}
+
+fun tgoto :: "nat \<Rightarrow> instr list"
+  where
+  "tgoto n = [(Nop, n), (Nop, n)]"
+
+text {*
+  The layout of the final TM compiled from an Abacus program is represented
+  as a list of natural numbers, where the list element at index @{text "n"} represents the 
+  starting state of the TM simulating the execution of @{text "n"}-th instruction
+  in the Abacus program.
+*}
+
+type_synonym layout = "nat list"
+
+text {*
+  @{text "length_of i"} is the length of the 
+  TM simulating the Abacus instruction @{text "i"}.
+*}
+fun length_of :: "abc_inst \<Rightarrow> nat"
+  where
+  "length_of i = (case i of 
+                    Inc n   \<Rightarrow> 2 * n + 9 |
+                    Dec n e \<Rightarrow> 2 * n + 16 |
+                    Goto n  \<Rightarrow> 1)"
+
+text {*
+  @{text "layout_of ap"} returns the layout of Abacus program @{text "ap"}.
+*}
+fun layout_of :: "abc_prog \<Rightarrow> layout"
+  where "layout_of ap = map length_of ap"
+
+
+text {*
+  @{text "start_of layout n"} looks out the starting state of @{text "n"}-th
+  TM in the finall TM.
+*}
+thm listsum_def
+
+fun start_of :: "nat list \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "start_of ly x = (Suc (listsum (take x ly))) "
+
+text {*
+  @{text "ci lo ss i"} complies Abacus instruction @{text "i"}
+  assuming the TM of @{text "i"} starts from state @{text "ss"} 
+  within the overal layout @{text "lo"}.
+*}
+
+fun ci :: "layout \<Rightarrow> nat \<Rightarrow> abc_inst \<Rightarrow> instr list"
+  where
+  "ci ly ss (Inc n) = tinc ss n"
+| "ci ly ss (Dec n e) = tdec ss n (start_of ly e)"
+| "ci ly ss (Goto n) = tgoto (start_of ly n)"
+
+text {*
+  @{text "tpairs_of ap"} transfroms Abacus program @{text "ap"} pairing
+  every instruction with its starting state.
+*}
+
+fun tpairs_of :: "abc_prog \<Rightarrow> (nat \<times> abc_inst) list"
+  where "tpairs_of ap = (zip (map (start_of (layout_of ap)) 
+                         [0..<(length ap)]) ap)"
+
+text {*
+  @{text "tms_of ap"} returns the list of TMs, where every one of them simulates
+  the corresponding Abacus intruction in @{text "ap"}.
+*}
+
+fun tms_of :: "abc_prog \<Rightarrow> (instr list) list"
+  where "tms_of ap = map (\<lambda> (n, tm). ci (layout_of ap) n tm) 
+                         (tpairs_of ap)"
+
+text {*
+  @{text "tm_of ap"} returns the final TM machine compiled from Abacus program @{text "ap"}.
+*}
+fun tm_of :: "abc_prog \<Rightarrow> instr list"
+  where "tm_of ap = concat (tms_of ap)"
+
+text {*
+  The following two functions specify the well-formedness of complied TM.
+*}
+(*
+fun t_ncorrect :: "tprog \<Rightarrow> bool"
+  where
+  "t_ncorrect tp = (length tp mod 2 = 0)"
+
+fun abc2t_correct :: "abc_prog \<Rightarrow> bool"
+  where 
+  "abc2t_correct ap = list_all (\<lambda> (n, tm). 
+             t_ncorrect (ci (layout_of ap) n tm)) (tpairs_of ap)"
+*)
+
+lemma length_findnth: 
+  "length (findnth n) = 4 * n"
+apply(induct n, auto)
+done
+
+lemma ci_length : "length (ci ns n ai) div 2 = length_of ai"
+apply(auto simp: ci.simps tinc_b_def tdec_b_def length_findnth
+                 split: abc_inst.splits)
+done
+
+subsection {*
+  Representation of Abacus memory by TM tape
+*}
+
+text {*
+  @{text "crsp acf tcf"} meams the abacus configuration @{text "acf"}
+  is corretly represented by the TM configuration @{text "tcf"}.
+*}
+
+fun crsp :: "layout \<Rightarrow> abc_conf \<Rightarrow> config \<Rightarrow> cell list \<Rightarrow> bool"
+  where 
+  "crsp ly (as, lm) (s, l, r) inres = 
+           (s = start_of ly as \<and> (\<exists> x. r = <lm> @ Bk\<up>x) \<and> 
+            l = Bk # Bk # inres)"
+
+declare crsp.simps[simp del]
+
+subsection {*
+  A more general definition of TM execution. 
+*}
+
+(*
+fun nnth_of :: "(taction \<times> nat) list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> (taction \<times> nat)"
+  where
+  "nnth_of p s b = (if 2*s < length p 
+                    then (p ! (2*s + b)) else (Nop, 0))"
+
+thm nth_of.simps
+
+fun nfetch :: "tprog \<Rightarrow> nat \<Rightarrow> block \<Rightarrow> taction \<times> nat"
+  where
+  "nfetch p 0 b = (Nop, 0)" |
+  "nfetch p (Suc s) b = 
+             (case b of 
+                Bk \<Rightarrow> nnth_of p s 0 |
+                Oc \<Rightarrow> nnth_of p s 1)"
+*)
+
+                    
+text {*
+  The type of invarints expressing correspondence between 
+  Abacus configuration and TM configuration.
+*}
+
+type_synonym inc_inv_t = "abc_conf \<Rightarrow> config \<Rightarrow> cell list \<Rightarrow> bool"
+
+declare tms_of.simps[simp del] tm_of.simps[simp del]
+        layout_of.simps[simp del] abc_fetch.simps [simp del]  
+        tpairs_of.simps[simp del] start_of.simps[simp del]
+        ci.simps [simp del] length_of.simps[simp del] 
+        layout_of.simps[simp del]
+
+(*
+subsection {* The compilation of @{text "Inc n"} *}
+*)
+
+text {*
+  The lemmas in this section lead to the correctness of 
+  the compilation of @{text "Inc n"} instruction.
+*}
+
+declare abc_step_l.simps[simp del] abc_steps_l.simps[simp del]
+lemma [simp]: "start_of ly as > 0"
+apply(simp add: start_of.simps)
+done
+
+lemma abc_steps_l_0: "abc_steps_l ac ap 0 = ac"
+by(case_tac ac, simp add: abc_steps_l.simps)
+
+lemma abc_step_red: 
+ "abc_steps_l (as, am) ap stp = (bs, bm) \<Longrightarrow> 
+  abc_steps_l (as, am) ap (Suc stp) = abc_step_l (bs, bm) (abc_fetch bs ap) "
+proof(induct stp arbitrary: as am bs bm)
+  case 0
+  thus "?case"
+    by(simp add: abc_steps_l.simps abc_steps_l_0)
+next
+  case (Suc stp as am bs bm)
+  have ind: "\<And>as am bs bm. abc_steps_l (as, am) ap stp = (bs, bm) \<Longrightarrow> 
+    abc_steps_l (as, am) ap (Suc stp) = abc_step_l (bs, bm) (abc_fetch bs ap)"
+    by fact
+  have h:" abc_steps_l (as, am) ap (Suc stp) = (bs, bm)" by fact
+  obtain as' am' where g: "abc_step_l (as, am) (abc_fetch as ap) = (as', am')"
+    by(case_tac "abc_step_l (as, am) (abc_fetch as ap)", auto)
+  then have "abc_steps_l (as', am') ap (Suc stp) = abc_step_l (bs, bm) (abc_fetch bs ap)"
+    using h
+    by(rule_tac ind, simp add: abc_steps_l.simps)
+  thus "?case"
+    using g
+    by(simp add: abc_steps_l.simps)
+qed
+
+lemma tm_shift_fetch: 
+  "\<lbrakk>fetch A s b = (ac, ns); ns \<noteq> 0 \<rbrakk>
+  \<Longrightarrow> fetch (shift A off) s b = (ac, ns + off)"
+apply(case_tac b)
+apply(case_tac [!] s, auto simp: fetch.simps shift.simps)
+done
+
+lemma tm_shift_eq_step:
+  assumes exec: "step (s, l, r) (A, 0) = (s', l', r')"
+  and notfinal: "s' \<noteq> 0"
+  shows "step (s + off, l, r) (shift A off, off) = (s' + off, l', r')"
+using assms
+apply(simp add: step.simps)
+apply(case_tac "fetch A s (read r)", auto)
+apply(drule_tac [!] off = off in tm_shift_fetch, simp_all)
+done
+
+declare step.simps[simp del] steps.simps[simp del] shift.simps[simp del]
+
+lemma tm_shift_eq_steps: 
+  assumes exec: "steps (s, l, r) (A, 0) stp = (s', l', r')"
+  and notfinal: "s' \<noteq> 0"
+  shows "steps (s + off, l, r) (shift A off, off) stp = (s' + off, l', r')"
+  using exec notfinal
+proof(induct stp arbitrary: s' l' r', simp add: steps.simps)
+  fix stp s' l' r'
+  assume ind: "\<And>s' l' r'. \<lbrakk>steps (s, l, r) (A, 0) stp = (s', l', r'); s' \<noteq> 0\<rbrakk> 
+     \<Longrightarrow> steps (s + off, l, r) (shift A off, off) stp = (s' + off, l', r')"
+  and h: " steps (s, l, r) (A, 0) (Suc stp) = (s', l', r')" "s' \<noteq> 0"
+  obtain s1 l1 r1 where g: "steps (s, l, r) (A, 0) stp = (s1, l1, r1)" 
+    apply(case_tac "steps (s, l, r) (A, 0) stp") by blast
+  moreover then have "s1 \<noteq> 0"
+    using h
+    apply(simp add: step_red)
+    apply(case_tac "0 < s1", auto)
+    done
+  ultimately have "steps (s + off, l, r) (shift A off, off) stp =
+                   (s1 + off, l1, r1)"
+    apply(rule_tac ind, simp_all)
+    done
+  thus "steps (s + off, l, r) (shift A off, off) (Suc stp) = (s' + off, l', r')"
+    using h g assms
+    apply(simp add: step_red)
+    apply(rule_tac tm_shift_eq_step, auto)
+    done
+qed
+
+lemma startof_not0[simp]: "0 < start_of ly as"
+apply(simp add: start_of.simps)
+done
+
+lemma startof_ge1[simp]: "Suc 0 \<le> start_of ly as"
+apply(simp add: start_of.simps)
+done
+
+lemma start_of_Suc1: "\<lbrakk>ly = layout_of ap; 
+       abc_fetch as ap = Some (Inc n)\<rbrakk>
+       \<Longrightarrow> start_of ly (Suc as) = start_of ly as + 2 * n + 9"
+apply(auto simp: start_of.simps layout_of.simps  
+                 length_of.simps abc_fetch.simps 
+                 take_Suc_conv_app_nth split: if_splits)
+done
+
+lemma start_of_Suc2:
+  "\<lbrakk>ly = layout_of ap;
+  abc_fetch as ap = Some (Dec n e)\<rbrakk> \<Longrightarrow> 
+        start_of ly (Suc as) = 
+            start_of ly as + 2 * n + 16"
+apply(auto simp: start_of.simps layout_of.simps  
+                 length_of.simps abc_fetch.simps 
+                 take_Suc_conv_app_nth split: if_splits)
+done
+
+lemma start_of_Suc3:
+  "\<lbrakk>ly = layout_of ap;
+  abc_fetch as ap = Some (Goto n)\<rbrakk> \<Longrightarrow> 
+  start_of ly (Suc as) = start_of ly as + 1"
+apply(auto simp: start_of.simps layout_of.simps  
+                 length_of.simps abc_fetch.simps 
+                 take_Suc_conv_app_nth split: if_splits)
+done
+
+lemma length_ci_inc: 
+  "length (ci ly ss (Inc n)) = 4*n + 18"
+apply(auto simp: ci.simps length_findnth tinc_b_def)
+done
+
+lemma length_ci_dec: 
+  "length (ci ly ss (Dec n e)) = 4*n + 32"
+apply(auto simp: ci.simps length_findnth tdec_b_def)
+done
+
+lemma length_ci_goto: 
+  "length (ci ly ss (Goto n )) = 2"
+apply(auto simp: ci.simps length_findnth tdec_b_def)
+done
+
+lemma take_Suc_last[elim]: "Suc as \<le> length xs \<Longrightarrow> 
+            take (Suc as) xs = take as xs @ [xs ! as]"
+apply(induct xs arbitrary: as, simp, simp)
+apply(case_tac as, simp, simp)
+done
+
+lemma concat_suc: "Suc as \<le> length xs \<Longrightarrow> 
+       concat (take (Suc as) xs) = concat (take as xs) @ xs! as"
+apply(subgoal_tac "take (Suc as) xs = take as xs @ [xs ! as]", simp)
+by auto
+
+lemma concat_take_suc_iff: "Suc n \<le> length tps \<Longrightarrow> 
+       concat (take n tps) @ (tps ! n) = concat (take (Suc n) tps)"
+apply(drule_tac concat_suc, simp)
+done
+
+lemma concat_drop_suc_iff: 
+   "Suc n < length tps \<Longrightarrow> concat (drop (Suc n) tps) = 
+           tps ! Suc n @ concat (drop (Suc (Suc n)) tps)"
+apply(induct tps arbitrary: n, simp, simp)
+apply(case_tac tps, simp, simp)
+apply(case_tac n, simp, simp)
+done
+
+declare append_assoc[simp del]
+
+lemma  tm_append:
+  "\<lbrakk>n < length tps; tp = tps ! n\<rbrakk> \<Longrightarrow> 
+  \<exists> tp1 tp2. concat tps = tp1 @ tp @ tp2 \<and> tp1 = 
+  concat (take n tps) \<and> tp2 = concat (drop (Suc n) tps)"
+apply(rule_tac x = "concat (take n tps)" in exI)
+apply(rule_tac x = "concat (drop (Suc n) tps)" in exI)
+apply(auto)
+apply(induct n, simp)
+apply(case_tac tps, simp, simp, simp)
+apply(subgoal_tac "concat (take n tps) @ (tps ! n) = 
+               concat (take (Suc n) tps)")
+apply(simp only: append_assoc[THEN sym], simp only: append_assoc)
+apply(subgoal_tac " concat (drop (Suc n) tps) = tps ! Suc n @ 
+                  concat (drop (Suc (Suc n)) tps)", simp)
+apply(rule_tac concat_drop_suc_iff, simp)
+apply(rule_tac concat_take_suc_iff, simp)
+done
+
+declare append_assoc[simp]
+
+lemma map_of:  "n < length xs \<Longrightarrow> (map f xs) ! n = f (xs ! n)"
+by(auto)
+
+lemma [simp]: "length (tms_of aprog) = length aprog"
+apply(auto simp: tms_of.simps tpairs_of.simps)
+done
+
+lemma ci_nth: 
+  "\<lbrakk>ly = layout_of aprog; 
+  abc_fetch as aprog = Some ins\<rbrakk>
+  \<Longrightarrow> ci ly (start_of ly as) ins = tms_of aprog ! as"
+apply(simp add: tms_of.simps tpairs_of.simps 
+      abc_fetch.simps  map_of del: map_append split: if_splits)
+done
+
+lemma t_split:"\<lbrakk>
+        ly = layout_of aprog;
+        abc_fetch as aprog = Some ins\<rbrakk>
+      \<Longrightarrow> \<exists> tp1 tp2. concat (tms_of aprog) = 
+            tp1 @ (ci ly (start_of ly as) ins) @ tp2
+            \<and> tp1 = concat (take as (tms_of aprog)) \<and> 
+              tp2 = concat (drop (Suc as) (tms_of aprog))"
+apply(insert tm_append[of "as" "tms_of aprog" 
+                             "ci ly (start_of ly as) ins"], simp)
+apply(subgoal_tac "ci ly (start_of ly as) ins = (tms_of aprog) ! as")
+apply(subgoal_tac "length (tms_of aprog) = length aprog")
+apply(simp add: abc_fetch.simps split: if_splits, simp)
+apply(rule_tac ci_nth, auto)
+done
+
+lemma math_sub: "\<lbrakk>x >= Suc 0; x - 1 = z\<rbrakk> \<Longrightarrow> x + y - Suc 0 = z + y"
+by auto
+
+lemma start_more_one: "as \<noteq> 0 \<Longrightarrow> start_of ly as >= Suc 0"
+apply(induct as, simp add: start_of.simps)
+apply(case_tac as, auto simp: start_of.simps)
+done
+
+lemma div_apart: "\<lbrakk>x mod (2::nat) = 0; y mod 2 = 0\<rbrakk> 
+          \<Longrightarrow> (x + y) div 2 = x div 2 + y div 2"
+apply(drule mod_eqD)+
+apply(auto)
+done
+
+lemma div_apart_iff: "\<lbrakk>x mod (2::nat) = 0; y mod 2 = 0\<rbrakk> \<Longrightarrow> 
+           (x + y) mod 2 = 0"
+apply(auto)
+done
+
+lemma [simp]: "length (layout_of aprog) = length aprog"
+apply(auto simp: layout_of.simps)
+done
+
+lemma start_of_ind: "\<lbrakk>as < length aprog; ly = layout_of aprog\<rbrakk> \<Longrightarrow> 
+       start_of ly (Suc as) = start_of ly as + 
+                          length ((tms_of aprog) ! as) div 2"
+apply(simp only: start_of.simps, simp)
+apply(auto simp: start_of.simps tms_of.simps layout_of.simps 
+                 tpairs_of.simps)
+apply(simp add: ci_length take_Suc take_Suc_conv_app_nth)
+done
+
+lemma concat_take_suc: "Suc n \<le> length xs \<Longrightarrow>
+  concat (take (Suc n) xs) = concat (take n xs) @ (xs ! n)"
+apply(subgoal_tac "take (Suc n) xs =
+                   take n xs @ [xs ! n]")
+apply(auto)
+done
+
+lemma [simp]: 
+  "\<lbrakk>as < length aprog; (abc_fetch as aprog) = Some ins\<rbrakk>
+  \<Longrightarrow> ci (layout_of aprog) 
+  (start_of (layout_of aprog) as) (ins) \<in> set (tms_of aprog)"
+apply(insert ci_nth[of "layout_of aprog" aprog as], simp)
+done
+
+lemma [simp]: "length (tms_of ap) = length ap"
+by(auto simp: tms_of.simps tpairs_of.simps)
+
+lemma [intro]:  "n < length ap \<Longrightarrow> length (tms_of ap ! n) mod 2 = 0"
+apply(auto simp: tms_of.simps tpairs_of.simps)
+apply(case_tac "ap ! n", auto simp: ci.simps length_findnth tinc_b_def tdec_b_def)
+apply arith
+by arith
+
+lemma compile_mod2: "length (concat (take n (tms_of ap))) mod 2 = 0"
+apply(induct n, auto)
+apply(case_tac "n < length (tms_of ap)", simp add: take_Suc_conv_app_nth, auto)
+apply(subgoal_tac "length (tms_of ap ! n) mod 2 = 0")
+apply arith
+by auto
+
+lemma tpa_states:
+  "\<lbrakk>tp = concat (take as (tms_of ap));
+  as \<le> length ap\<rbrakk> \<Longrightarrow> 
+  start_of (layout_of ap) as = Suc (length tp div 2)"
+proof(induct as arbitrary: tp)
+  case 0
+  thus "?case"
+    by(simp add: start_of.simps)
+next
+  case (Suc as tp)
+  have ind: "\<And>tp. \<lbrakk>tp = concat (take as (tms_of ap)); as \<le> length ap\<rbrakk> \<Longrightarrow>
+    start_of (layout_of ap) as = Suc (length tp div 2)" by fact
+  have tp: "tp = concat (take (Suc as) (tms_of ap))" by fact
+  have le: "Suc as \<le> length ap" by fact
+  have a: "start_of (layout_of ap) as = Suc (length (concat (take as (tms_of ap))) div 2)"
+    using le
+    by(rule_tac ind, simp_all)
+  from a tp le show "?case"
+    apply(simp add: start_of.simps take_Suc_conv_app_nth)
+    apply(subgoal_tac "length (concat (take as (tms_of ap))) mod 2= 0")
+    apply(subgoal_tac " length (tms_of ap ! as) mod 2 = 0")
+    apply(simp add: Abacus.div_apart) 
+    apply(simp add: layout_of.simps ci_length  tms_of.simps tpairs_of.simps)
+    apply(auto  intro: compile_mod2)
+    done
+qed
+
+lemma append_append_fetch: 
+    "\<lbrakk>length tp1 mod 2 = 0; length tp mod 2 = 0;
+      length tp1 div 2 < a \<and> a \<le> length tp1 div 2 + length tp div 2\<rbrakk>
+    \<Longrightarrow>fetch (tp1 @ tp @ tp2) a b = fetch tp (a - length tp1 div 2) b "
+apply(subgoal_tac "\<exists> x. a = length tp1 div 2 + x", erule exE, simp)
+apply(case_tac x, simp)
+apply(subgoal_tac "length tp1 div 2 + Suc nat = 
+             Suc (length tp1 div 2 + nat)")
+apply(simp only: fetch.simps nth_of.simps, auto)
+apply(case_tac b, simp)
+apply(subgoal_tac "2 * (length tp1 div 2) = length tp1", simp)
+apply(subgoal_tac "2 * nat < length tp", simp add: nth_append, simp)
+apply(subgoal_tac "2 * (length tp1 div 2) = length tp1", simp)
+apply(subgoal_tac "2 * nat < length tp", simp add: nth_append, auto)
+apply(auto simp: nth_append)
+apply(rule_tac x = "a - length tp1 div 2" in exI, simp)
+done
+
+lemma step_eq_fetch':
+  assumes layout: "ly = layout_of ap"
+  and compile: "tp = tm_of ap"
+  and fetch: "abc_fetch as ap = Some ins"
+  and range1: "s \<ge> start_of ly as"
+  and range2: "s < start_of ly (Suc as)"
+  shows "fetch tp s b = fetch (ci ly (start_of ly as) ins)
+       (Suc s - start_of ly as) b "
+proof -
+  have "\<exists>tp1 tp2. concat (tms_of ap) = tp1 @ ci ly (start_of ly as) ins @ tp2 \<and>
+    tp1 = concat (take as (tms_of ap)) \<and> tp2 = concat (drop (Suc as) (tms_of ap))"
+    using assms
+    by(rule_tac t_split, simp_all)
+  then obtain tp1 tp2 where a: "concat (tms_of ap) = tp1 @ ci ly (start_of ly as) ins @ tp2 \<and>
+    tp1 = concat (take as (tms_of ap)) \<and> tp2 = concat (drop (Suc as) (tms_of ap))" by blast
+  then have b: "start_of (layout_of ap) as = Suc (length tp1 div 2)"
+    using fetch
+    apply(rule_tac tpa_states, simp, simp add: abc_fetch.simps split: if_splits)
+    done
+  have "fetch (tp1 @ (ci ly (start_of ly as) ins) @ tp2)  s b = 
+        fetch (ci ly (start_of ly as) ins) (s - length tp1 div 2) b"
+  proof(rule_tac append_append_fetch)
+    show "length tp1 mod 2 = 0"
+      using a
+      by(auto, rule_tac compile_mod2)
+  next
+    show "length (ci ly (start_of ly as) ins) mod 2 = 0"
+      apply(case_tac ins, auto simp: ci.simps length_findnth tinc_b_def tdec_b_def)
+      by(arith, arith)
+  next
+    show "length tp1 div 2 < s \<and> s \<le> 
+      length tp1 div 2 + length (ci ly (start_of ly as) ins) div 2"
+    proof -
+      have "length (ci ly (start_of ly as) ins) div 2 = length_of ins"
+        using ci_length by simp
+      moreover have "start_of ly (Suc as) = start_of ly as + length_of ins"
+        using fetch layout
+        apply(simp add: start_of.simps abc_fetch.simps List.take_Suc_conv_app_nth 
+          split: if_splits)
+        apply(simp add: layout_of.simps)
+        done
+      ultimately show "?thesis"
+        using b layout range1 range2
+        apply(simp)
+        done
+    qed
+  qed
+  thus "?thesis"
+    using b layout a compile  
+    apply(simp add: tm_of.simps)
+    done
+qed
+
+lemma step_eq_fetch: 
+  assumes layout: "ly = layout_of ap"
+  and compile: "tp = tm_of ap"
+  and abc_fetch: "abc_fetch as ap = Some ins" 
+  and fetch: "fetch (ci ly (start_of ly as) ins)
+       (Suc s - start_of ly as) b = (ac, ns)"
+  and notfinal: "ns \<noteq> 0"
+  shows "fetch tp s b = (ac, ns)"
+proof -
+  have "s \<ge> start_of ly as"
+  proof(cases "s \<ge> start_of ly as")
+    case True thus "?thesis" by simp
+  next
+    case False 
+    have "\<not> start_of ly as \<le> s" by fact
+    then have "Suc s - start_of ly as = 0"
+      by arith
+    then have "fetch (ci ly (start_of ly as) ins)
+       (Suc s - start_of ly as) b = (Nop, 0)"
+      by(simp add: fetch.simps)
+    with notfinal fetch show "?thesis"
+      by(simp)
+  qed
+  moreover have "s < start_of ly (Suc as)"
+  proof(cases "s < start_of ly (Suc as)")
+    case True thus "?thesis" by simp
+  next
+    case False
+    have h: "\<not> s < start_of ly (Suc as)"
+      by fact
+    then have "s > start_of ly as"
+      using abc_fetch layout
+      apply(simp add: start_of.simps abc_fetch.simps split: if_splits)
+      apply(simp add: List.take_Suc_conv_app_nth, auto)
+      apply(subgoal_tac "layout_of ap ! as > 0") 
+      apply arith
+      apply(simp add: layout_of.simps)
+      apply(case_tac "ap!as", auto simp: length_of.simps)
+      done
+    from this and h have "fetch (ci ly (start_of ly as) ins) (Suc s - start_of ly as) b = (Nop, 0)"
+      using abc_fetch layout
+      apply(case_tac b, simp_all add: Suc_diff_le)
+      apply(case_tac [!] ins, simp_all add: start_of_Suc2 start_of_Suc1 start_of_Suc3)
+      apply(simp_all only: length_ci_inc length_ci_dec length_ci_goto, auto)
+      using layout
+      apply(subgoal_tac [!] "start_of ly (Suc as) = start_of ly as + 2*nat1 + 16", simp_all)
+      apply(rule_tac [!] start_of_Suc2, auto)
+      done
+    from fetch and notfinal this show "?thesis"by simp
+  qed
+  ultimately show "?thesis"
+    using assms
+    apply(drule_tac b= b and ins = ins in step_eq_fetch', auto)
+    done
+qed
+
+
+lemma step_eq_in:
+  assumes layout: "ly = layout_of ap"
+  and compile: "tp = tm_of ap"
+  and fetch: "abc_fetch as ap = Some ins"    
+  and exec: "step (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) 
+  = (s', l', r')"
+  and notfinal: "s' \<noteq> 0"
+  shows "step (s, l, r) (tp, 0) = (s', l', r')"
+  using assms
+  apply(simp add: step.simps)
+  apply(case_tac "fetch (ci (layout_of ap) (start_of (layout_of ap) as) ins)
+    (Suc s - start_of (layout_of ap) as) (read r)", simp)
+  using layout
+  apply(drule_tac s = s and b = "read r" and ac = a in step_eq_fetch, auto)
+  done
+
+lemma steps_eq_in:
+  assumes layout: "ly = layout_of ap"
+  and compile: "tp = tm_of ap"
+  and crsp: "crsp ly (as, lm) (s, l, r) ires"
+  and fetch: "abc_fetch as ap = Some ins"    
+  and exec: "steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp 
+  = (s', l', r')"
+  and notfinal: "s' \<noteq> 0"
+  shows "steps (s, l, r) (tp, 0) stp = (s', l', r')"
+  using exec notfinal
+proof(induct stp arbitrary: s' l' r', simp add: steps.simps)
+  fix stp s' l' r'
+  assume ind: 
+    "\<And>s' l' r'. \<lbrakk>steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp = (s', l', r'); s' \<noteq> 0\<rbrakk>
+              \<Longrightarrow> steps (s, l, r) (tp, 0) stp = (s', l', r')"
+  and h: "steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) (Suc stp) = (s', l', r')" "s' \<noteq> 0"
+  obtain s1 l1 r1 where g: "steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp = 
+                        (s1, l1, r1)"
+    apply(case_tac "steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp") by blast
+  moreover hence "s1 \<noteq> 0"
+    using h
+    apply(simp add: step_red)
+    apply(case_tac "0 < s1", simp_all)
+    done
+  ultimately have "steps (s, l, r) (tp, 0) stp = (s1, l1, r1)"
+    apply(rule_tac ind, auto)
+    done
+  thus "steps (s, l, r) (tp, 0) (Suc stp) = (s', l', r')"
+    using h g assms
+    apply(simp add: step_red)
+    apply(rule_tac step_eq_in, auto)
+    done
+qed
+
+lemma tm_append_fetch_first: 
+  "\<lbrakk>fetch A s b = (ac, ns); ns \<noteq> 0\<rbrakk> \<Longrightarrow> 
+    fetch (A @ B) s b = (ac, ns)"
+apply(case_tac b)
+apply(case_tac [!] s, auto simp: fetch.simps nth_append split: if_splits)
+done
+
+lemma tm_append_first_step_eq: 
+  assumes "step (s, l, r) (A, off) = (s', l', r')"
+  and "s' \<noteq> 0"
+  shows "step (s, l, r) (A @ B, off) = (s', l', r')"
+using assms
+apply(simp add: step.simps)
+apply(case_tac "fetch A (s - off) (read r)")
+apply(frule_tac  B = B and b = "read r" in tm_append_fetch_first, auto)
+done
+
+lemma tm_append_first_steps_eq: 
+  assumes "steps (s, l, r) (A, off) stp = (s', l', r')"
+  and "s' \<noteq> 0"
+  shows "steps (s, l, r) (A @ B, off) stp = (s', l', r')"
+using assms
+proof(induct stp arbitrary: s' l' r', simp add: steps.simps)
+  fix stp s' l' r'
+  assume ind: "\<And>s' l' r'. \<lbrakk>steps (s, l, r) (A, off) stp = (s', l', r'); s' \<noteq> 0\<rbrakk>
+    \<Longrightarrow> steps (s, l, r) (A @ B, off) stp = (s', l', r')"
+    and h: "steps (s, l, r) (A, off) (Suc stp) = (s', l', r')" "s' \<noteq> 0"
+  obtain sa la ra where a: "steps (s, l, r) (A, off) stp = (sa, la, ra)"
+    apply(case_tac "steps (s, l, r) (A, off) stp") by blast
+  hence "steps (s, l, r) (A @ B, off) stp = (sa, la, ra) \<and> sa \<noteq> 0"
+    using h ind[of sa la ra]
+    apply(case_tac sa, simp_all)
+    done
+  thus "steps (s, l, r) (A @ B, off) (Suc stp) = (s', l', r')"
+    using h a
+    apply(simp add: step_red)
+    apply(rule_tac tm_append_first_step_eq, simp_all)
+    done
+qed
+
+lemma tm_append_second_fetch_eq:
+  assumes
+  even: "length A mod 2 = 0"
+  and off: "off = length A div 2"
+  and fetch: "fetch B s b = (ac, ns)"
+  and notfinal: "ns \<noteq> 0"
+  shows "fetch (A @ shift B off) (s + off) b = (ac, ns + off)"
+using assms
+apply(case_tac b)
+apply(case_tac [!] s, auto simp: fetch.simps nth_append shift.simps 
+  split: if_splits)
+done
+
+
+lemma tm_append_second_step_eq: 
+  assumes 
+  exec: "step0 (s, l, r) B = (s', l', r')"
+  and notfinal: "s' \<noteq> 0"
+  and off: "off = length A div 2"
+  and even: "length A mod 2 = 0"
+  shows "step0 (s + off, l, r) (A @ shift B off) = (s' + off, l', r')"
+using assms
+apply(simp add: step.simps)
+apply(case_tac "fetch B s (read r)")
+apply(frule_tac tm_append_second_fetch_eq, simp_all, auto)
+done
+
+
+lemma tm_append_second_steps_eq: 
+  assumes 
+  exec: "steps (s, l, r) (B, 0) stp = (s', l', r')"
+  and notfinal: "s' \<noteq> 0"
+  and off: "off = length A div 2"
+  and even: "length A mod 2 = 0"
+  shows "steps (s + off, l, r) (A @ shift B off, 0) stp = (s' + off, l', r')"
+using exec notfinal
+proof(induct stp arbitrary: s' l' r')
+  case 0
+  thus "steps0 (s + off, l, r) (A @ shift B off) 0 = (s' + off, l', r')"
+    by(simp add: steps.simps)
+next
+  case (Suc stp s' l' r')
+  have ind: "\<And>s' l' r'. \<lbrakk>steps0 (s, l, r) B stp = (s', l', r'); s' \<noteq> 0\<rbrakk> \<Longrightarrow> 
+    steps0 (s + off, l, r) (A @ shift B off) stp = (s' + off, l', r')"
+    by fact
+  have h: "steps0 (s, l, r) B (Suc stp) = (s', l', r')" by fact
+  have k: "s' \<noteq> 0" by fact
+  obtain s'' l'' r'' where a: "steps0 (s, l, r) B stp = (s'', l'', r'')"
+    by (metis prod_cases3)
+  then have b: "s'' \<noteq> 0"
+    using h k
+    by(rule_tac notI, auto simp: step_red)
+  from a b have c: "steps0 (s + off, l, r) (A @ shift B off) stp = (s'' + off, l'', r'')"
+    by(erule_tac ind, simp)
+  from c b h a k assms show "?case"
+    thm tm_append_second_step_eq
+    apply(simp add: step_red) by(rule tm_append_second_step_eq, simp_all)
+qed
+
+lemma tm_append_second_fetch0_eq:
+  assumes
+  even: "length A mod 2 = 0"
+  and off: "off = length A div 2"
+  and fetch: "fetch B s b = (ac, 0)"
+  and notfinal: "s \<noteq> 0"
+  shows "fetch (A @ shift B off) (s + off) b = (ac, 0)"
+using assms
+apply(case_tac b)
+apply(case_tac [!] s, auto simp: fetch.simps nth_append shift.simps 
+  split: if_splits)
+done
+
+lemma tm_append_second_halt_eq:
+  assumes 
+  exec: "steps (Suc 0, l, r) (B, 0) stp = (0, l', r')"
+  and wf_B: "tm_wf (B, 0)"
+  and off: "off = length A div 2"
+  and even: "length A mod 2 = 0"
+  shows "steps (Suc off, l, r) (A @ shift B off, 0) stp = (0, l', r')"
+proof -
+  thm before_final
+  have "\<exists>n. \<not> is_final (steps0 (1, l, r) B n) \<and> steps0 (1, l, r) B (Suc n) = (0, l', r')"
+    using exec by(rule_tac before_final, simp)
+ then obtain n where a: 
+   "\<not> is_final (steps0 (1, l, r) B n) \<and> steps0 (1, l, r) B (Suc n) = (0, l', r')" ..
+ obtain s'' l'' r'' where b: "steps0 (1, l, r) B n = (s'', l'', r'') \<and> s'' >0"
+   using a
+   by(case_tac "steps0 (1, l, r) B n", auto)
+ have c: "steps (Suc 0 + off, l, r) (A @ shift B off, 0) n = (s'' + off, l'', r'')"
+   using a b assms
+   by(rule_tac tm_append_second_steps_eq, simp_all)
+ obtain ac where d: "fetch B s'' (read r'') = (ac, 0)"
+   using  b a
+   by(case_tac "fetch B s'' (read r'')", auto simp: step_red step.simps)
+ then have "fetch (A @ shift B off) (s'' + off) (read r'') = (ac, 0)"
+   using assms b
+   by(rule_tac tm_append_second_fetch0_eq, simp_all)
+ then have e: "steps (Suc 0 + off, l, r) (A @ shift B off, 0) (Suc n) = (0, l', r')"
+   using a b assms c d
+   by(simp add: step_red step.simps)
+ from a have "n < stp"
+   using exec
+ proof(cases "n < stp")
+   case  True thus "?thesis" by simp
+ next
+   case False
+   have "\<not> n < stp" by fact
+   then obtain d where  "n = stp + d"
+     by (metis add.comm_neutral less_imp_add_positive nat_neq_iff)
+   thus "?thesis"
+     using a e exec
+     by(simp add: steps_add)
+ qed
+ then obtain d where "stp = Suc n + d"
+   by(metis add_Suc less_iff_Suc_add)
+ thus "?thesis"
+   using e
+   by(simp only: steps_add, simp)
+qed
+
+lemma tm_append_steps: 
+  assumes 
+  aexec: "steps (s, l, r) (A, 0) stpa = (Suc (length A div 2), la, ra)"
+  and bexec: "steps (Suc 0, la, ra) (B, 0) stpb =  (sb, lb, rb)"
+  and notfinal: "sb \<noteq> 0"
+  and off: "off = length A div 2"
+  and even: "length A mod 2 = 0"
+  shows "steps (s, l, r) (A @ shift B off, 0) (stpa + stpb) = (sb + off, lb, rb)"
+proof -
+  have "steps (s, l, r) (A@shift B off, 0) stpa = (Suc (length A div 2), la, ra)"
+    apply(rule_tac tm_append_first_steps_eq)
+    apply(auto simp: assms)
+    done
+  moreover have "steps (1 + off, la, ra) (A @ shift B off, 0) stpb = (sb + off, lb, rb)"
+    apply(rule_tac tm_append_second_steps_eq)
+    apply(auto simp: assms bexec)
+    done
+  ultimately show "steps (s, l, r) (A @ shift B off, 0) (stpa + stpb) = (sb + off, lb, rb)"
+    apply(simp add: steps_add off)
+    done
+qed
+       
+subsection {* Crsp of Inc*}
+
+fun at_begin_fst_bwtn :: "inc_inv_t"
+  where
+  "at_begin_fst_bwtn (as, lm) (s, l, r) ires = 
+      (\<exists> lm1 tn rn. lm1 = (lm @ 0\<up>tn) \<and> length lm1 = s \<and> 
+          (if lm1 = [] then l = Bk # Bk # ires
+           else l = [Bk]@<rev lm1>@Bk#Bk#ires) \<and> r = Bk\<up>rn)" 
+
+
+fun at_begin_fst_awtn :: "inc_inv_t"
+  where
+  "at_begin_fst_awtn (as, lm) (s, l, r) ires = 
+      (\<exists> lm1 tn rn. lm1 = (lm @ 0\<up>tn) \<and> length lm1 = s \<and>
+         (if lm1 = []  then l = Bk # Bk # ires
+          else l = [Bk]@<rev lm1>@Bk#Bk#ires) \<and> r = [Oc]@Bk\<up>rn)"
+
+fun at_begin_norm :: "inc_inv_t"
+  where
+  "at_begin_norm (as, lm) (s, l, r) ires= 
+      (\<exists> lm1 lm2 rn. lm = lm1 @ lm2 \<and> length lm1 = s \<and> 
+        (if lm1 = [] then l = Bk # Bk # ires
+         else l = Bk # <rev lm1> @ Bk # Bk # ires ) \<and> r = <lm2>@Bk\<up>rn)"
+
+fun in_middle :: "inc_inv_t"
+  where
+  "in_middle (as, lm) (s, l, r) ires = 
+      (\<exists> lm1 lm2 tn m ml mr rn. lm @ 0\<up>tn = lm1 @ [m] @ lm2
+       \<and> length lm1 = s \<and> m + 1 = ml + mr \<and>  
+         ml \<noteq> 0 \<and> tn = s + 1 - length lm \<and> 
+       (if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires 
+        else l = Oc\<up>ml@[Bk]@<rev lm1>@
+                 Bk # Bk # ires) \<and> (r = Oc\<up>mr @ [Bk] @ <lm2>@ Bk\<up>rn \<or> 
+      (lm2 = [] \<and> r = Oc\<up>mr))
+      )"
+
+fun inv_locate_a :: "inc_inv_t"
+  where "inv_locate_a (as, lm) (s, l, r) ires = 
+     (at_begin_norm (as, lm) (s, l, r) ires \<or>
+      at_begin_fst_bwtn (as, lm) (s, l, r) ires \<or>
+      at_begin_fst_awtn (as, lm) (s, l, r) ires
+      )"
+
+fun inv_locate_b :: "inc_inv_t"
+  where "inv_locate_b (as, lm) (s, l, r) ires = 
+        (in_middle (as, lm) (s, l, r)) ires "
+
+fun inv_after_write :: "inc_inv_t"
+  where "inv_after_write (as, lm) (s, l, r) ires = 
+           (\<exists> rn m lm1 lm2. lm = lm1 @ m # lm2 \<and>
+             (if lm1 = [] then l = Oc\<up>m @ Bk # Bk # ires
+              else Oc # l = Oc\<up>Suc m@ Bk # <rev lm1> @ 
+                      Bk # Bk # ires) \<and> r = [Oc] @ <lm2> @ Bk\<up>rn)"
+
+fun inv_after_move :: "inc_inv_t"
+  where "inv_after_move (as, lm) (s, l, r) ires = 
+      (\<exists> rn m lm1 lm2. lm = lm1 @ m # lm2 \<and>
+        (if lm1 = [] then l = Oc\<up>Suc m @ Bk # Bk # ires
+         else l = Oc\<up>Suc m@ Bk # <rev lm1> @ Bk # Bk # ires) \<and> 
+        r = <lm2> @ Bk\<up>rn)"
+
+fun inv_after_clear :: "inc_inv_t"
+  where "inv_after_clear (as, lm) (s, l, r) ires =
+       (\<exists> rn m lm1 lm2 r'. lm = lm1 @ m # lm2 \<and> 
+        (if lm1 = [] then l = Oc\<up>Suc m @ Bk # Bk # ires
+         else l = Oc\<up>Suc m @ Bk # <rev lm1> @ Bk # Bk # ires) \<and> 
+          r = Bk # r' \<and> Oc # r' = <lm2> @ Bk\<up>rn)"
+
+fun inv_on_right_moving :: "inc_inv_t"
+  where "inv_on_right_moving (as, lm) (s, l, r) ires = 
+       (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
+            ml + mr = m \<and> 
+          (if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires
+          else l = Oc\<up>ml  @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
+         ((r = Oc\<up>mr @ [Bk] @ <lm2> @ Bk\<up>rn) \<or> 
+          (r = Oc\<up>mr \<and> lm2 = [])))"
+
+fun inv_on_left_moving_norm :: "inc_inv_t"
+  where "inv_on_left_moving_norm (as, lm) (s, l, r) ires =
+      (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and>  
+             ml + mr = Suc m \<and> mr > 0 \<and> (if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires
+                                         else l =  Oc\<up>ml @ Bk # <rev lm1> @ Bk # Bk # ires)
+        \<and> (r = Oc\<up>mr @ Bk # <lm2> @ Bk\<up>rn \<or> 
+           (lm2 = [] \<and> r = Oc\<up>mr)))"
+
+fun inv_on_left_moving_in_middle_B:: "inc_inv_t"
+  where "inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires =
+                (\<exists> lm1 lm2 rn. lm = lm1 @ lm2 \<and>  
+                     (if lm1 = [] then l = Bk # ires
+                      else l = <rev lm1> @ Bk # Bk # ires) \<and> 
+                      r = Bk # <lm2> @ Bk\<up>rn)"
+
+fun inv_on_left_moving :: "inc_inv_t"
+  where "inv_on_left_moving (as, lm) (s, l, r) ires = 
+       (inv_on_left_moving_norm  (as, lm) (s, l, r) ires \<or>
+        inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires)"
+
+
+fun inv_check_left_moving_on_leftmost :: "inc_inv_t"
+  where "inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires = 
+                (\<exists> rn. l = ires \<and> r = [Bk, Bk] @ <lm> @  Bk\<up>rn)"
+
+fun inv_check_left_moving_in_middle :: "inc_inv_t"
+  where "inv_check_left_moving_in_middle (as, lm) (s, l, r) ires = 
+              (\<exists> lm1 lm2 r' rn. lm = lm1 @ lm2 \<and>
+                 (Oc # l = <rev lm1> @ Bk # Bk # ires) \<and> r = Oc # Bk # r' \<and> 
+                           r' = <lm2> @  Bk\<up>rn)"
+
+fun inv_check_left_moving :: "inc_inv_t"
+  where "inv_check_left_moving (as, lm) (s, l, r) ires = 
+             (inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires \<or>
+             inv_check_left_moving_in_middle (as, lm) (s, l, r) ires)"
+
+fun inv_after_left_moving :: "inc_inv_t"
+  where "inv_after_left_moving (as, lm) (s, l, r) ires= 
+              (\<exists> rn. l = Bk # ires \<and> r = Bk # <lm> @  Bk\<up>rn)"
+
+fun inv_stop :: "inc_inv_t"
+  where "inv_stop (as, lm) (s, l, r) ires= 
+              (\<exists> rn. l = Bk # Bk # ires \<and> r = <lm> @  Bk\<up>rn)"
+
+
+lemma halt_lemma2': 
+  "\<lbrakk>wf LE;  \<forall> n. ((\<not> P (f n) \<and> Q (f n)) \<longrightarrow> 
+    (Q (f (Suc n)) \<and> (f (Suc n), (f n)) \<in> LE)); Q (f 0)\<rbrakk> 
+      \<Longrightarrow> \<exists> n. P (f n)"
+apply(intro exCI, simp)
+apply(subgoal_tac "\<forall> n. Q (f n)", simp)
+apply(drule_tac f = f in wf_inv_image)
+apply(simp add: inv_image_def)
+apply(erule wf_induct, simp)
+apply(erule_tac x = x in allE)
+apply(erule_tac x = n in allE, erule_tac x = n in allE)
+apply(erule_tac x = "Suc x" in allE, simp)
+apply(rule_tac allI)
+apply(induct_tac n, simp)
+apply(erule_tac x = na in allE, simp)
+done
+
+lemma halt_lemma2'': 
+  "\<lbrakk>P (f n); \<not> P (f (0::nat))\<rbrakk> \<Longrightarrow> 
+         \<exists> n. (P (f n) \<and> (\<forall> i < n. \<not> P (f i)))"
+apply(induct n rule: nat_less_induct, auto)
+done
+
+lemma halt_lemma2''':
+ "\<lbrakk>\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> LE;
+                 Q (f 0);  \<forall>i<na. \<not> P (f i)\<rbrakk> \<Longrightarrow> Q (f na)"
+apply(induct na, simp, simp)
+done
+
+lemma halt_lemma2: 
+  "\<lbrakk>wf LE;  
+    Q (f 0); \<not> P (f 0);
+    \<forall> n. ((\<not> P (f n) \<and> Q (f n)) \<longrightarrow> (Q (f (Suc n)) \<and> (f (Suc n), (f n)) \<in> LE))\<rbrakk> 
+  \<Longrightarrow> \<exists> n. P (f n) \<and> Q (f n)"
+apply(insert halt_lemma2' [of LE P f Q], simp, erule_tac exE)
+apply(subgoal_tac "\<exists> n. (P (f n) \<and> (\<forall> i < n. \<not> P (f i)))")
+apply(erule_tac exE)+
+apply(rule_tac x = na in exI, auto)
+apply(rule halt_lemma2''', simp, simp, simp)
+apply(erule_tac halt_lemma2'', simp)
+done
+
+
+fun findnth_inv :: "layout \<Rightarrow> nat \<Rightarrow> inc_inv_t"
+  where
+  "findnth_inv ly n (as, lm) (s, l, r) ires =
+              (if s = 0 then False
+               else if s \<le> Suc (2*n) then 
+                  if s mod 2 = 1 then inv_locate_a (as, lm) ((s - 1) div 2, l, r) ires
+                  else inv_locate_b (as, lm) ((s - 1) div 2, l, r) ires
+               else False)"
+
+
+fun findnth_state :: "config \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "findnth_state (s, l, r) n = (Suc (2*n) - s)"
+  
+fun findnth_step :: "config \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "findnth_step (s, l, r) n = 
+           (if s mod 2 = 1 then
+                   (if (r \<noteq> [] \<and> hd r = Oc) then 0
+                    else 1)
+            else length r)"
+
+fun findnth_measure :: "config \<times> nat \<Rightarrow> nat \<times> nat"
+  where
+  "findnth_measure (c, n) = 
+     (findnth_state c n, findnth_step c n)"
+
+definition lex_pair :: "((nat \<times> nat) \<times> nat \<times> nat) set"
+  where
+  "lex_pair \<equiv> less_than <*lex*> less_than"
+
+definition findnth_LE :: "((config \<times> nat) \<times> (config \<times> nat)) set"
+  where
+   "findnth_LE \<equiv> (inv_image lex_pair findnth_measure)"
+
+lemma wf_findnth_LE: "wf findnth_LE"
+by(auto intro:wf_inv_image simp: findnth_LE_def lex_pair_def)
+
+declare findnth_inv.simps[simp del]
+
+lemma [simp]: 
+  "\<lbrakk>x < Suc (Suc (2 * n)); Suc x mod 2 = Suc 0; \<not> x < 2 * n\<rbrakk>
+ \<Longrightarrow> x = 2*n"
+by arith
+
+lemma [simp]: 
+  "\<lbrakk>0 < a; a < Suc (2 * n); a mod 2 = Suc 0\<rbrakk> 
+      \<Longrightarrow> fetch (findnth n) a Bk = (W1, a)"
+apply(case_tac a, simp_all)
+apply(induct n, auto simp: findnth.simps length_findnth nth_append)
+apply arith
+done
+
+lemma [simp]: 
+  "\<lbrakk>0 < a; a < Suc (2 * n); a mod 2 = Suc 0\<rbrakk> 
+      \<Longrightarrow> fetch (findnth n) a Oc = (R, Suc a)"
+apply(case_tac a, simp_all)
+apply(induct n, auto simp: findnth.simps length_findnth nth_append)
+apply(subgoal_tac "nat = 2 * n", simp)
+by arith
+
+lemma [simp]: 
+  "\<lbrakk>0 < a; a < Suc (2*n); a mod 2 \<noteq> Suc 0\<rbrakk>
+     \<Longrightarrow> fetch (findnth n) a Oc = (R, a)"
+apply(case_tac a, simp_all)
+apply(induct n, auto simp: findnth.simps length_findnth nth_append)
+apply(subgoal_tac "nat = Suc (2 * n)", simp)
+apply arith
+done
+
+lemma [simp]: 
+  "\<lbrakk>0 < a; a < Suc (2*n); a mod 2 \<noteq> Suc 0\<rbrakk>
+     \<Longrightarrow> fetch (findnth n) a Bk = (R, Suc a)"
+apply(case_tac a, simp_all)
+apply(induct n, auto simp: findnth.simps length_findnth nth_append)
+apply(subgoal_tac "nat = Suc (2 * n)", simp)
+by arith
+
+declare at_begin_norm.simps[simp del] at_begin_fst_bwtn.simps[simp del] 
+   at_begin_fst_awtn.simps[simp del] in_middle.simps[simp del] 
+   abc_lm_s.simps[simp del] abc_lm_v.simps[simp del]  
+   ci.simps[simp del] inv_after_move.simps[simp del] 
+   inv_on_left_moving_norm.simps[simp del] 
+   inv_on_left_moving_in_middle_B.simps[simp del]
+   inv_after_clear.simps[simp del] 
+   inv_after_write.simps[simp del] inv_on_left_moving.simps[simp del]
+   inv_on_right_moving.simps[simp del] 
+   inv_check_left_moving.simps[simp del] 
+   inv_check_left_moving_in_middle.simps[simp del]
+   inv_check_left_moving_on_leftmost.simps[simp del] 
+   inv_after_left_moving.simps[simp del]
+   inv_stop.simps[simp del] inv_locate_a.simps[simp del] 
+   inv_locate_b.simps[simp del]
+
+lemma [intro]: "\<exists>rn. [Bk] = Bk \<up> rn"
+by (metis replicate_0 replicate_Suc)
+
+lemma [intro]:  "at_begin_norm (as, am) (q, aaa, []) ires
+             \<Longrightarrow> at_begin_norm (as, am) (q, aaa, [Bk]) ires"
+apply(simp add: at_begin_norm.simps, erule_tac exE, erule_tac exE)
+apply(rule_tac x = lm1 in exI, simp, auto)
+done
+
+lemma [intro]: "at_begin_fst_bwtn (as, am) (q, aaa, []) ires 
+            \<Longrightarrow> at_begin_fst_bwtn (as, am) (q, aaa, [Bk]) ires"
+apply(simp only: at_begin_fst_bwtn.simps, erule_tac exE, erule_tac exE, erule_tac exE)
+apply(rule_tac x = "am @ 0\<up>tn" in exI, auto)
+done
+
+lemma [intro]: "at_begin_fst_awtn (as, am) (q, aaa, []) ires
+           \<Longrightarrow> at_begin_fst_awtn (as, am) (q, aaa, [Bk]) ires"
+apply(auto simp: at_begin_fst_awtn.simps)
+done 
+
+lemma [intro]: "inv_locate_a (as, am) (q, aaa, []) ires
+            \<Longrightarrow> inv_locate_a (as, am) (q, aaa, [Bk]) ires"
+apply(simp only: inv_locate_a.simps)
+apply(erule disj_forward)
+defer
+apply(erule disj_forward, auto)
+done
+
+lemma tape_of_nl_cons: "<m # lm> = (if lm = [] then Oc\<up>(Suc m)
+                    else Oc\<up>(Suc m) @ Bk # <lm>)"
+apply(case_tac lm, simp_all add: tape_of_nl_abv  tape_of_nat_abv split: if_splits)
+done
+
+
+lemma locate_a_2_locate_a[simp]: "inv_locate_a (as, am) (q, aaa, Bk # xs) ires
+       \<Longrightarrow> inv_locate_a (as, am) (q, aaa, Oc # xs) ires"
+apply(simp only: inv_locate_a.simps at_begin_norm.simps 
+                 at_begin_fst_bwtn.simps at_begin_fst_awtn.simps)
+apply(erule_tac disjE, erule exE, erule exE, erule exE, 
+      rule disjI2, rule disjI2)
+defer
+apply(erule_tac disjE, erule exE, erule exE, 
+      erule exE, rule disjI2, rule disjI2)
+prefer 2
+apply(simp)
+proof-
+  fix lm1 tn rn
+  assume k: "lm1 = am @ 0\<up>tn \<and> length lm1 = q \<and> (if lm1 = [] then aaa = Bk # Bk # 
+    ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> Bk # xs = Bk\<up>rn"
+  thus "\<exists>lm1 tn rn. lm1 = am @ 0 \<up> tn \<and> length lm1 = q \<and> 
+    (if lm1 = [] then aaa = Bk # Bk # ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> Oc # xs = [Oc] @ Bk \<up> rn"
+    (is "\<exists>lm1 tn rn. ?P lm1 tn rn")
+  proof -
+    from k have "?P lm1 tn (rn - 1)"
+      apply(auto simp: Oc_def)
+      by(case_tac [!] "rn::nat", auto)
+    thus ?thesis by blast
+  qed
+next
+  fix lm1 lm2 rn
+  assume h1: "am = lm1 @ lm2 \<and> length lm1 = q \<and> (if lm1 = [] 
+    then aaa = Bk # Bk # ires else aaa = Bk # <rev lm1> @ Bk # Bk # ires) \<and>
+    Bk # xs = <lm2> @ Bk\<up>rn"
+  from h1 have h2: "lm2 = []"
+    apply(auto split: if_splits)
+    apply(case_tac [!] lm2, simp_all add: tape_of_nl_cons split: if_splits)
+    done
+  from h1 and h2 show "\<exists>lm1 tn rn. lm1 = am @ 0\<up>tn \<and> length lm1 = q \<and> 
+    (if lm1 = [] then aaa = Bk # Bk # ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and>
+    Oc # xs = [Oc] @ Bk\<up>rn" 
+    (is "\<exists>lm1 tn rn. ?P lm1 tn rn")
+  proof -
+    from h1 and h2  have "?P lm1 0 (rn - 1)"
+      apply(auto simp: Oc_def
+                      tape_of_nl_abv tape_of_nat_list.simps)
+      by(case_tac "rn::nat", simp, simp)
+    thus ?thesis by blast
+  qed
+qed
+
+lemma [simp]: "inv_locate_a (as, am) (q, aaa, []) ires \<Longrightarrow> 
+               inv_locate_a (as, am) (q, aaa, [Oc]) ires"
+apply(insert locate_a_2_locate_a [of as am q aaa "[]"])
+apply(subgoal_tac "inv_locate_a (as, am) (q, aaa, [Bk]) ires", auto)
+done
+
+(*inv: from locate_b to locate_b*)
+lemma [simp]: "inv_locate_b (as, am) (q, aaa, Oc # xs) ires
+         \<Longrightarrow> inv_locate_b (as, am) (q, Oc # aaa, xs) ires"
+apply(simp only: inv_locate_b.simps in_middle.simps)
+apply(erule exE)+
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
+      rule_tac x = tn in exI, rule_tac x = m in exI)
+apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - 1" in exI,
+      rule_tac x = rn in exI)
+apply(case_tac mr, simp_all, auto)
+done
+
+(*
+lemma zero_and_nil[intro]: "(Bk # Bk\<^bsup>n\<^esup> = Oc\<^bsup>mr\<^esup> @ Bk # <lm::nat list> @ 
+                             Bk\<^bsup>rn \<^esup>) \<or> (lm2 = [] \<and> Bk # Bk\<^bsup>n\<^esup> = Oc\<^bsup>mr\<^esup>)
+       \<Longrightarrow> mr = 0 \<and> lm = []"
+apply(rule context_conjI)
+apply(case_tac mr, auto simp:exponent_def)
+apply(insert BkCons_nil[of "replicate (n - 1) Bk" lm rn])
+apply(case_tac n, auto simp: exponent_def Bk_def  tape_of_nl_nil_eq)
+done
+
+lemma tape_of_nat_def: "<[m::nat]> =  Oc # Oc\<^bsup>m\<^esup>"
+apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
+done
+*)
+lemma [simp]:  "<[x::nat]> = Oc\<up>(Suc x)"
+apply(simp add: tape_of_nat_abv tape_of_nl_abv)
+done
+
+lemma [simp]: " <([]::nat list)> = []"
+apply(simp add: tape_of_nl_abv)
+done
+
+lemma [simp]: "\<lbrakk>inv_locate_b (as, am) (q, aaa, Bk # xs) ires; \<exists>n. xs = Bk\<up>n\<rbrakk>
+            \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires"
+apply(simp add: inv_locate_b.simps inv_locate_a.simps)
+apply(rule_tac disjI2, rule_tac disjI1)
+apply(simp only: in_middle.simps at_begin_fst_bwtn.simps)
+apply(erule_tac exE)+
+apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = tn in exI, simp split: if_splits)
+apply(case_tac mr, simp_all)
+apply(case_tac "length am", simp_all, case_tac tn, simp_all)
+apply(case_tac lm2, simp_all add: tape_of_nl_cons split: if_splits)
+apply(case_tac am, simp_all)
+apply(case_tac n, simp_all)
+apply(case_tac n, simp_all)
+apply(case_tac mr, simp_all)
+apply(case_tac lm2, simp_all add: tape_of_nl_cons split: if_splits, auto)
+apply(case_tac [!] n, simp_all)
+done
+
+lemma [simp]: "(Oc # r = Bk \<up> rn) = False"
+apply(case_tac rn, simp_all)
+done
+
+lemma [simp]: "(\<exists>rna. Bk \<up> rn = Bk # Bk \<up> rna) \<or> rn = 0"
+apply(case_tac rn, auto)
+done
+
+lemma [simp]: "(\<forall> x. a \<noteq> x) = False"
+by auto
+
+lemma exp_ind: "a\<up>(Suc x) = a\<up>x @ [a]"
+apply(induct x, auto)
+done
+
+lemma [simp]: 
+      "inv_locate_a (as, lm) (q, l, Oc # r) ires
+       \<Longrightarrow> inv_locate_b (as, lm) (q, Oc # l, r) ires"
+apply(simp only: inv_locate_a.simps inv_locate_b.simps in_middle.simps
+          at_begin_norm.simps at_begin_fst_bwtn.simps
+          at_begin_fst_awtn.simps)
+apply(erule disjE, erule exE, erule exE, erule exE)
+apply(rule_tac x = lm1 in exI, rule_tac x = "tl lm2" in exI, simp)
+apply(rule_tac x = 0 in exI, rule_tac x = "hd lm2" in exI)
+apply(case_tac lm2, auto simp: tape_of_nl_cons )
+apply(rule_tac x = 1 in exI, rule_tac x = a in exI, auto)
+apply(case_tac list, simp_all)
+apply(case_tac rn, simp_all)
+apply(rule_tac x = "lm @ replicate tn 0" in exI, 
+      rule_tac x = "[]" in exI, 
+      rule_tac x = "Suc tn" in exI, 
+      rule_tac x = 0 in exI, auto)
+apply(simp only: replicate_Suc[THEN sym] exp_ind)
+apply(rule_tac x = "Suc 0" in exI, auto)
+done
+
+lemma length_equal: "xs = ys \<Longrightarrow> length xs = length ys"
+by auto
+
+lemma [simp]: "\<lbrakk>inv_locate_b (as, am) (q, aaa, Bk # xs) ires; 
+                \<not> (\<exists>n. xs = Bk\<up>n)\<rbrakk> 
+       \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires"
+apply(simp add: inv_locate_b.simps inv_locate_a.simps)
+apply(rule_tac disjI1)
+apply(simp only: in_middle.simps at_begin_norm.simps)
+apply(erule_tac exE)+
+apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = lm2 in exI, simp)
+apply(subgoal_tac "tn = 0", simp , auto split: if_splits)
+apply(case_tac [!] mr, simp_all, auto)
+apply(simp add: tape_of_nl_cons)
+apply(drule_tac length_equal, simp)
+apply(case_tac "length am", simp_all, erule_tac x = rn in allE, simp)
+apply(drule_tac length_equal, simp)
+apply(case_tac "(Suc (length lm1) - length am)", simp_all)
+apply(case_tac lm2, simp, simp)
+done
+
+lemma locate_b_2_a[intro]: 
+       "inv_locate_b (as, am) (q, aaa, Bk # xs) ires
+    \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires"
+apply(case_tac "\<exists> n. xs = Bk\<up>n", simp, simp)
+done
+
+
+lemma [simp]:  "inv_locate_b (as, am) (q, l, []) ires 
+           \<Longrightarrow>  inv_locate_b (as, am) (q, l, [Bk]) ires"
+apply(simp only: inv_locate_b.simps in_middle.simps)
+apply(erule exE)+
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
+      rule_tac x = tn in exI, rule_tac x = m in exI, 
+      rule_tac x = ml in exI, rule_tac x = mr in exI)
+apply(auto)
+done
+
+(*inv: from locate_b to after_write*)
+
+lemma [simp]: "(a mod 2 \<noteq> Suc 0) = (a mod 2 = 0)  "
+by arith
+
+lemma [simp]: "(a mod 2 \<noteq> 0) = (a mod 2 = Suc 0)  "
+by arith
+
+lemma mod_ex1: "(a mod 2 = Suc 0) = (\<exists> q. a = Suc (2 * q))"
+by arith
+
+lemma mod_ex2: "(a mod (2::nat) = 0) = (\<exists> q. a = 2 * q)"
+by arith
+
+lemma [simp]: "(2*q - Suc 0) div 2 = (q - 1)"
+by arith
+
+lemma [simp]: "(Suc (2*q)) div 2 = q"
+by arith
+
+lemma mod_2: "x mod 2  = 0 \<or>  x mod 2 = Suc 0"
+by arith
+
+lemma [simp]: "x mod 2 = 0 \<Longrightarrow> Suc x mod 2 = Suc 0"
+by arith
+
+lemma [simp]: "x mod 2 = Suc 0 \<Longrightarrow> Suc x mod 2 = 0"
+by arith
+
+lemma [simp]:  "inv_locate_b (as, am) (q, l, []) ires 
+           \<Longrightarrow>  inv_locate_b (as, am) (q, l, [Bk]) ires"
+apply(simp only: inv_locate_b.simps in_middle.simps)
+apply(erule exE)+
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
+      rule_tac x = tn in exI, rule_tac x = m in exI, 
+      rule_tac x = ml in exI, rule_tac x = mr in exI)
+apply(auto)
+done
+
+lemma locate_b_2_locate_a[simp]: 
+    "\<lbrakk>q > 0;  inv_locate_b (as, am) (q - Suc 0, aaa, Bk # xs) ires\<rbrakk>
+   \<Longrightarrow>  inv_locate_a (as, am) (q, Bk # aaa, xs) ires"
+apply(insert locate_b_2_a [of as am "q - 1" aaa xs ires], simp)
+done
+
+
+lemma [simp]:  "inv_locate_b (as, am) (q, l, []) ires 
+           \<Longrightarrow>  inv_locate_b (as, am) (q, l, [Bk]) ires"
+apply(simp only: inv_locate_b.simps in_middle.simps)
+apply(erule exE)+
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
+      rule_tac x = tn in exI, rule_tac x = m in exI, 
+      rule_tac x = ml in exI, rule_tac x = mr in exI)
+apply(auto)
+done
+
+(*inv: from locate_b to after_write*)
+
+lemma [simp]:
+  "crsp (layout_of ap) (as, lm) (s, l, r) ires
+  \<Longrightarrow> findnth_inv (layout_of ap) n (as, lm) (Suc 0, l, r) ires"
+apply(auto simp: crsp.simps findnth_inv.simps inv_locate_a.simps
+               at_begin_norm.simps at_begin_fst_awtn.simps at_begin_fst_bwtn.simps)
+done
+
+lemma findnth_correct_pre: 
+  assumes layout: "ly = layout_of ap"
+  and crsp: "crsp ly (as, lm) (s, l, r) ires"
+  and not0: "n > 0"
+  and f: "f = (\<lambda> stp. (steps (Suc 0, l, r) (findnth n, 0) stp, n))"
+  and P: "P = (\<lambda> ((s, l, r), n). s = Suc (2 * n))"
+  and Q: "Q = (\<lambda> ((s, l, r), n). findnth_inv ly n (as, lm) (s, l, r) ires)"
+  shows "\<exists> stp. P (f stp) \<and> Q (f stp)"
+thm halt_lemma2
+proof(rule_tac LE = findnth_LE in halt_lemma2)
+  show "wf findnth_LE"  by(intro wf_findnth_LE)
+next
+  show "Q (f 0)"
+    using crsp layout
+    apply(simp add: f P Q steps.simps)
+    done
+next
+  show "\<not> P (f 0)"
+    using not0
+    apply(simp add: f P steps.simps)
+    done
+next
+  show "\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) 
+        \<in> findnth_LE"
+  proof(rule_tac allI, rule_tac impI, simp add: f, 
+      case_tac "steps (Suc 0, l, r) (findnth n, 0) na", simp add: P)
+    fix na a b c
+    assume "a \<noteq> Suc (2 * n) \<and> Q ((a, b, c), n)"
+    thus  "Q (step (a, b, c) (findnth n, 0), n) \<and> 
+        ((step (a, b, c) (findnth n, 0), n), (a, b, c), n) \<in> findnth_LE"
+      apply(case_tac c, case_tac [2] aa)
+      apply(simp_all add: step.simps findnth_LE_def Q findnth_inv.simps mod_2  lex_pair_def split: if_splits)
+      apply(auto simp: mod_ex1 mod_ex2)
+      done
+  qed
+qed
+            
+lemma [intro]: "inv_locate_a (as, lm) (0, Bk # Bk # ires, <lm> @ Bk \<up> x) ires"
+apply(auto simp: crsp.simps inv_locate_a.simps at_begin_norm.simps)
+done
+lemma [simp]: "crsp ly (as, lm) (s, l, r) ires \<Longrightarrow> inv_locate_a (as, lm) (0, l, r) ires"
+apply(auto simp: crsp.simps inv_locate_a.simps at_begin_norm.simps)
+done
+
+lemma findnth_correct: 
+  assumes layout: "ly = layout_of ap"
+  and crsp: "crsp ly (as, lm) (s, l, r) ires"
+  shows "\<exists> stp l' r'. steps (Suc 0, l, r) (findnth n, 0) stp = (Suc (2 * n), l', r')
+              \<and> inv_locate_a (as, lm) (n, l', r') ires"
+  using crsp
+  apply(case_tac "n = 0")
+  apply(rule_tac x = 0 in exI, auto simp: steps.simps)
+  using assms
+  apply(drule_tac findnth_correct_pre, auto)
+  apply(rule_tac x = stp in exI, simp add: findnth_inv.simps)
+  done
+
+
+fun inc_inv :: "nat \<Rightarrow> inc_inv_t"
+  where
+  "inc_inv n (as, lm) (s, l, r) ires =
+              (let lm' = abc_lm_s lm n (Suc (abc_lm_v lm n)) in
+                if s = 0 then False
+                else if s = 1 then 
+                   inv_locate_a (as, lm) (n, l, r) ires
+                else if s = 2 then 
+                   inv_locate_b (as, lm) (n, l, r) ires
+                else if s = 3 then 
+                   inv_after_write (as, lm') (s, l, r) ires
+                else if s = Suc 3 then 
+                   inv_after_move (as, lm') (s, l, r) ires
+                else if s = Suc 4 then 
+                   inv_after_clear (as, lm') (s, l, r) ires
+                else if s = Suc (Suc 4) then 
+                   inv_on_right_moving (as, lm') (s, l, r) ires
+                else if s = Suc (Suc 5) then 
+                   inv_on_left_moving (as, lm') (s, l, r) ires
+                else if s = Suc (Suc (Suc 5)) then 
+                   inv_check_left_moving (as, lm') (s, l, r) ires
+                else if s = Suc (Suc (Suc (Suc 5))) then 
+                   inv_after_left_moving (as, lm') (s, l, r) ires
+                else if s = Suc (Suc (Suc (Suc (Suc 5)))) then 
+                   inv_stop (as, lm') (s, l, r) ires
+                else False)"
+
+
+fun abc_inc_stage1 :: "config \<Rightarrow> nat"
+  where
+  "abc_inc_stage1 (s, l, r) = 
+            (if s = 0 then 0
+             else if s \<le> 2 then 5
+             else if s \<le> 6 then 4
+             else if s \<le> 8 then 3
+             else if s = 9 then 2
+             else 1)"
+
+fun abc_inc_stage2 :: "config \<Rightarrow> nat"
+  where
+  "abc_inc_stage2 (s, l, r) =
+                (if s = 1 then 2
+                 else if s = 2 then 1
+                 else if s = 3 then length r
+                 else if s = 4 then length r
+                 else if s = 5 then length r
+                 else if s = 6 then 
+                                  if r \<noteq> [] then length r
+                                  else 1
+                 else if s = 7 then length l
+                 else if s = 8 then length l
+                 else 0)"
+
+fun abc_inc_stage3 :: "config \<Rightarrow>  nat"
+  where
+  "abc_inc_stage3 (s, l, r) = (
+              if s = 4 then 4
+              else if s = 5 then 3
+              else if s = 6 then 
+                   if r \<noteq> [] \<and> hd r = Oc then 2
+                   else 1
+              else if s = 3 then 0
+              else if s = 2 then length r
+              else if s = 1 then 
+                      if (r \<noteq> [] \<and> hd r = Oc) then 0
+                      else 1
+              else 10 - s)"
+
+
+definition inc_measure :: "config \<Rightarrow> nat \<times> nat \<times> nat"
+  where
+  "inc_measure c = 
+    (abc_inc_stage1 c, abc_inc_stage2 c, abc_inc_stage3 c)"
+
+definition lex_triple :: 
+   "((nat \<times> (nat \<times> nat)) \<times> (nat \<times> (nat \<times> nat))) set"
+  where "lex_triple \<equiv> less_than <*lex*> lex_pair"
+
+definition inc_LE :: "(config \<times> config) set"
+  where
+  "inc_LE \<equiv> (inv_image lex_triple inc_measure)"
+
+declare inc_inv.simps[simp del]
+
+lemma wf_inc_le[intro]: "wf inc_LE"
+by(auto intro:wf_inv_image simp: inc_LE_def lex_triple_def lex_pair_def)
+
+lemma numeral_5_eq_5: "5 = Suc (Suc (Suc (Suc (Suc 0))))"
+by arith
+
+lemma numeral_6_eq_6: "6 = Suc (Suc (Suc (Suc (Suc 1))))"
+by arith
+
+lemma numeral_7_eq_7: "7 = Suc (Suc (Suc (Suc (Suc 2))))"
+by arith
+
+lemma numeral_8_eq_8: "8 = Suc (Suc (Suc (Suc (Suc 3))))"
+by arith
+
+lemma numeral_9_eq_9: "9 = Suc (Suc (Suc (Suc (Suc (Suc 3)))))"
+by arith
+
+lemma numeral_10_eq_10: "10 = Suc (Suc (Suc (Suc (Suc (Suc (Suc 3))))))"
+by arith
+
+lemma inv_locate_b_2_after_write[simp]: 
+      "inv_locate_b (as, am) (n, aaa, Bk # xs) ires
+      \<Longrightarrow> inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n)))
+          (s, aaa, Oc # xs) ires"
+apply(auto simp: in_middle.simps inv_after_write.simps 
+                 abc_lm_v.simps abc_lm_s.simps  inv_locate_b.simps)
+apply(case_tac [!] mr, auto split: if_splits)
+apply(rule_tac x = rn in exI, rule_tac x = "Suc m" in exI,
+      rule_tac x = "lm1" in exI, simp)
+apply(rule_tac x = "lm2" in exI, simp)
+apply(simp only: Suc_diff_le exp_ind)
+apply(subgoal_tac "lm2 = []", simp)
+apply(drule_tac length_equal, simp)
+done
+
+lemma [simp]: "inv_locate_b (as, am) (n, aaa, []) ires \<Longrightarrow> 
+     inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) 
+                     (s, aaa, [Oc]) ires"
+apply(insert inv_locate_b_2_after_write [of as am n aaa "[]"])
+by(simp)
+
+
+
+(*inv: from after_write to after_move*)
+lemma [simp]: "inv_after_write (as, lm) (x, l, Oc # r) ires
+                \<Longrightarrow> inv_after_move (as, lm) (y, Oc # l, r) ires"
+apply(auto simp:inv_after_move.simps inv_after_write.simps split: if_splits)
+done
+
+lemma [simp]: "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n)
+                )) (x, aaa, Bk # xs) ires = False"
+apply(simp add: inv_after_write.simps )
+done
+
+lemma [simp]: 
+ "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) 
+                        (x, aaa, []) ires = False"
+apply(simp add: inv_after_write.simps )
+done
+
+(*inv: from after_move to after_clear*)
+lemma [simp]: "inv_after_move (as, lm) (s, l, Oc # r) ires
+                \<Longrightarrow> inv_after_clear (as, lm) (s', l, Bk # r) ires"
+apply(auto simp: inv_after_move.simps inv_after_clear.simps split: if_splits)
+done
+
+(*inv: from after_move to on_leftmoving*)
+lemma [intro]: "Bk \<up> rn = Bk # Bk \<up> (rn - Suc 0) \<or> rn = 0"
+apply(case_tac rn, auto)
+done
+
+lemma inv_after_move_2_inv_on_left_moving[simp]:  
+   "inv_after_move (as, lm) (s, l, Bk # r) ires
+   \<Longrightarrow> (l = [] \<longrightarrow> 
+         inv_on_left_moving (as, lm) (s', [], Bk # Bk # r) ires) \<and>
+      (l \<noteq> [] \<longrightarrow> 
+         inv_on_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)"
+apply(simp only: inv_after_move.simps inv_on_left_moving.simps)
+apply(subgoal_tac "l \<noteq> []", rule conjI, simp, rule impI, 
+                rule disjI1, simp only: inv_on_left_moving_norm.simps)
+apply(erule exE)+
+apply(subgoal_tac "lm2 = []")
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,  
+    rule_tac x = m in exI, rule_tac x = m in exI, 
+    rule_tac x = 1 in exI,  
+    rule_tac x = "rn - 1" in exI, auto)
+apply(auto split: if_splits)
+apply(case_tac [1-2] rn, simp_all)
+apply(case_tac [!] lm2, simp_all add: tape_of_nl_cons split: if_splits)
+done
+
+
+lemma inv_after_move_2_inv_on_left_moving_B[simp]: 
+    "inv_after_move (as, lm) (s, l, []) ires
+      \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s', [], [Bk]) ires) \<and>
+          (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s', tl l, [hd l]) ires)"
+apply(simp only: inv_after_move.simps inv_on_left_moving.simps)
+apply(subgoal_tac "l \<noteq> []", rule conjI, simp, rule impI, rule disjI1,
+        simp only: inv_on_left_moving_norm.simps)
+apply(erule exE)+
+apply(subgoal_tac "lm2 = []")
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,  
+      rule_tac x = m in exI, rule_tac x = m in exI, 
+      rule_tac x = 1 in exI, rule_tac x = "rn - 1" in exI, simp, case_tac rn)
+apply(auto split: if_splits)
+apply(case_tac [!] lm2, auto simp: tape_of_nl_cons split: if_splits)
+done
+
+(*inv: from after_clear to on_right_moving*)
+lemma [simp]: "Oc # r = replicate rn Bk = False"
+apply(case_tac rn, simp, simp)
+done
+
+lemma inv_after_clear_2_inv_on_right_moving[simp]: 
+     "inv_after_clear (as, lm) (x, l, Bk # r) ires
+      \<Longrightarrow> inv_on_right_moving (as, lm) (y, Bk # l, r) ires"
+apply(auto simp: inv_after_clear.simps inv_on_right_moving.simps )
+apply(subgoal_tac "lm2 \<noteq> []")
+apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI, 
+      rule_tac x = "hd lm2" in exI, simp)
+apply(rule_tac x = 0 in exI, rule_tac x = "hd lm2" in exI)
+apply(simp, rule conjI)
+apply(case_tac [!] "lm2::nat list", auto)
+apply(case_tac rn, auto split: if_splits simp: tape_of_nl_cons)
+apply(case_tac [!] rn, simp_all)
+done
+
+lemma [simp]: "inv_after_clear (as, lm) (x, l, []) ires\<Longrightarrow> 
+               inv_after_clear (as, lm) (y, l, [Bk]) ires" 
+by(auto simp: inv_after_clear.simps)
+
+lemma [simp]: "inv_after_clear (as, lm) (x, l, []) ires
+             \<Longrightarrow> inv_on_right_moving (as, lm) (y, Bk # l, []) ires"
+by(insert 
+    inv_after_clear_2_inv_on_right_moving[of as lm n l "[]"], simp)
+
+(*inv: from on_right_moving to on_right_movign*)
+lemma [simp]: "inv_on_right_moving (as, lm) (x, l, Oc # r) ires
+      \<Longrightarrow> inv_on_right_moving (as, lm) (y, Oc # l, r) ires"
+apply(auto simp: inv_on_right_moving.simps)
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
+           rule_tac x = "ml + mr" in exI, simp)
+apply(rule_tac x = "Suc ml" in exI, 
+           rule_tac x = "mr - 1" in exI, simp)
+apply(case_tac mr, auto)
+apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, 
+      rule_tac x = "ml + mr" in exI, simp)
+apply(rule_tac x = "Suc ml" in exI, 
+      rule_tac x = "mr - 1" in exI, simp)
+apply(case_tac mr, auto split: if_splits)
+done
+
+lemma inv_on_right_moving_2_inv_on_right_moving[simp]: 
+     "inv_on_right_moving (as, lm) (x, l, Bk # r) ires
+     \<Longrightarrow> inv_after_write (as, lm) (y, l, Oc # r) ires"
+apply(auto simp: inv_on_right_moving.simps inv_after_write.simps )
+apply(case_tac mr, auto simp: split: if_splits)
+done
+      
+lemma [simp]: "inv_on_right_moving (as, lm) (x, l, []) ires\<Longrightarrow> 
+             inv_on_right_moving (as, lm) (y, l, [Bk]) ires"
+apply(auto simp: inv_on_right_moving.simps)
+apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, simp)
+done
+
+(*inv: from on_right_moving to after_write*)
+lemma [simp]: "inv_on_right_moving (as, lm) (x, l, []) ires
+       \<Longrightarrow> inv_after_write (as, lm) (y, l, [Oc]) ires"
+apply(rule_tac inv_on_right_moving_2_inv_on_right_moving, simp)
+done
+
+(*inv: from on_left_moving to on_left_moving*)
+lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) 
+               (s, l, Oc # r) ires = False"
+apply(auto simp: inv_on_left_moving_in_middle_B.simps )
+done
+
+lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, l, Bk # r) ires 
+             = False"
+apply(auto simp: inv_on_left_moving_norm.simps)
+apply(case_tac [!] mr, auto simp: )
+done
+
+lemma [simp]: 
+  "\<lbrakk>inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires;
+    hd l = Bk; l \<noteq> []\<rbrakk> \<Longrightarrow> 
+     inv_on_left_moving_in_middle_B (as, lm) (s, tl l, Bk # Oc # r) ires"
+apply(case_tac l, simp, simp)
+apply(simp only: inv_on_left_moving_norm.simps 
+                 inv_on_left_moving_in_middle_B.simps)
+apply(erule_tac exE)+
+apply(rule_tac x = lm1 in exI, rule_tac x = "m # lm2" in exI, auto)
+apply(case_tac [!] ml, auto)
+apply(auto simp: tape_of_nl_cons split: if_splits)
+apply(rule_tac [!] x = "Suc rn" in exI, simp_all)
+done
+
+lemma [simp]: "\<lbrakk>inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires; 
+                hd l = Oc; l \<noteq> []\<rbrakk>
+            \<Longrightarrow> inv_on_left_moving_norm (as, lm) 
+                                        (s, tl l, Oc # Oc # r) ires"
+apply(simp only: inv_on_left_moving_norm.simps)
+apply(erule exE)+
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
+      rule_tac x = m in exI, rule_tac x = "ml - 1" in exI,
+      rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI, simp)
+apply(case_tac ml, auto simp: split: if_splits)
+done
+
+lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, [], Oc # r) ires
+     \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # Oc # r) ires"
+apply(auto simp: inv_on_left_moving_norm.simps 
+                 inv_on_left_moving_in_middle_B.simps split: if_splits)
+done
+
+lemma [simp]:"inv_on_left_moving (as, lm) (s, l, Oc # r) ires
+    \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s, [], Bk # Oc # r) ires)
+ \<and>  (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s, tl l, hd l # Oc # r) ires)"
+apply(simp add: inv_on_left_moving.simps)
+apply(case_tac "l \<noteq> []", rule conjI, simp, simp)
+apply(case_tac "hd l", simp, simp, simp)
+done
+
+(*inv: from on_left_moving to check_left_moving*)
+lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) 
+                                      (s, Bk # list, Bk # r) ires
+          \<Longrightarrow> inv_check_left_moving_on_leftmost (as, lm) 
+                                      (s', list, Bk # Bk # r) ires"
+apply(auto simp: inv_on_left_moving_in_middle_B.simps 
+                 inv_check_left_moving_on_leftmost.simps split: if_splits)
+apply(case_tac [!] "rev lm1", simp_all)
+apply(case_tac [!] lista, simp_all add: tape_of_nl_abv tape_of_nat_abv tape_of_nat_list.simps)
+done
+
+lemma [simp]:
+    "inv_check_left_moving_in_middle (as, lm) (s, l, Bk # r) ires= False"
+by(auto simp: inv_check_left_moving_in_middle.simps )
+
+lemma [simp]: 
+ "inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # r) ires\<Longrightarrow> 
+  inv_check_left_moving_on_leftmost (as, lm) (s', [], Bk # Bk # r) ires"
+apply(auto simp: inv_on_left_moving_in_middle_B.simps 
+                 inv_check_left_moving_on_leftmost.simps split: if_splits)
+done
+
+lemma [simp]: "inv_check_left_moving_on_leftmost (as, lm) 
+                                       (s, list, Oc # r) ires= False"
+by(auto simp: inv_check_left_moving_on_leftmost.simps split: if_splits)
+
+lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) 
+                                         (s, Oc # list, Bk # r) ires
+ \<Longrightarrow> inv_check_left_moving_in_middle (as, lm) (s', list, Oc # Bk # r) ires"
+apply(auto simp: inv_on_left_moving_in_middle_B.simps 
+                 inv_check_left_moving_in_middle.simps  split: if_splits)
+done
+
+lemma inv_on_left_moving_2_check_left_moving[simp]:
+ "inv_on_left_moving (as, lm) (s, l, Bk # r) ires
+ \<Longrightarrow> (l = [] \<longrightarrow> inv_check_left_moving (as, lm) (s', [], Bk # Bk # r) ires)
+ \<and> (l \<noteq> [] \<longrightarrow> 
+      inv_check_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)"
+apply(simp add: inv_on_left_moving.simps inv_check_left_moving.simps)
+apply(case_tac l, simp, simp)
+apply(case_tac a, simp, simp)
+done
+
+lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, l, []) ires = False"
+apply(auto simp: inv_on_left_moving_norm.simps)
+done
+
+lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires\<Longrightarrow> 
+     inv_on_left_moving (as, lm) (6 + 2 * n, l, [Bk]) ires"
+apply(simp add: inv_on_left_moving.simps)
+apply(auto simp: inv_on_left_moving_in_middle_B.simps)
+done
+
+lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires = False"
+apply(simp add: inv_on_left_moving.simps)
+apply(simp add: inv_on_left_moving_in_middle_B.simps)
+done
+
+lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires
+ \<Longrightarrow> (l = [] \<longrightarrow> inv_check_left_moving (as, lm) (s', [], [Bk]) ires) \<and>
+    (l \<noteq> [] \<longrightarrow> inv_check_left_moving (as, lm) (s', tl l, [hd l]) ires)"
+by simp
+
+lemma [intro]: "\<exists>rna. Bk # Bk \<up> rn = Bk \<up> rna"
+apply(rule_tac x = "Suc rn" in exI, simp)
+done
+
+lemma 
+inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[simp]:
+"inv_check_left_moving_in_middle (as, lm) (s, Bk # list, Oc # r) ires
+  \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s', list, Bk # Oc # r) ires"
+apply(simp only: inv_check_left_moving_in_middle.simps 
+                 inv_on_left_moving_in_middle_B.simps)
+apply(erule_tac exE)+
+apply(rule_tac x = "rev (tl (rev lm1))" in exI, 
+      rule_tac x = "[hd (rev lm1)] @ lm2" in exI, auto)
+apply(case_tac [!] "rev lm1",simp_all add: tape_of_nat_abv tape_of_nl_abv tape_of_nat_list.simps)
+apply(case_tac [!] a, simp_all)
+apply(case_tac [1] lm2, simp_all add: tape_of_nat_list.simps tape_of_nat_abv, auto)
+apply(case_tac [3] lm2, simp_all add: tape_of_nat_list.simps tape_of_nat_abv, auto)
+apply(case_tac [!] lista, simp_all add: tape_of_nat_abv tape_of_nat_list.simps)
+done
+
+lemma [simp]: 
+ "inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires\<Longrightarrow>
+     inv_check_left_moving_in_middle (as, lm) (s', [Bk], Oc # r) ires"
+apply(auto simp: inv_check_left_moving_in_middle.simps )
+done
+
+lemma [simp]: 
+ "inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires
+   \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s', [], Bk # Oc # r) ires"
+apply(insert 
+inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[of 
+                  as lm n "[]" r], simp)
+done 
+
+lemma [simp]: "inv_check_left_moving_in_middle (as, lm) 
+                       (s, Oc # list, Oc # r) ires
+   \<Longrightarrow> inv_on_left_moving_norm (as, lm) (s', list, Oc # Oc # r) ires"
+apply(auto simp: inv_check_left_moving_in_middle.simps 
+                 inv_on_left_moving_norm.simps)
+apply(rule_tac x = "rev (tl (rev lm1))" in exI, 
+      rule_tac x = lm2 in exI, rule_tac x = "hd (rev lm1)" in exI)
+apply(rule_tac conjI)
+apply(case_tac "rev lm1", simp, simp)
+apply(rule_tac x = "hd (rev lm1) - 1" in exI, auto)
+apply(rule_tac [!] x = "Suc (Suc 0)" in exI, simp)
+apply(case_tac [!] "rev lm1", simp_all)
+apply(case_tac [!] a, simp_all add: tape_of_nl_cons split: if_splits)
+done
+
+lemma [simp]: "inv_check_left_moving (as, lm) (s, l, Oc # r) ires
+\<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s', [], Bk # Oc # r) ires) \<and>
+   (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s', tl l, hd l # Oc # r) ires)"
+apply(case_tac l, 
+      auto simp: inv_check_left_moving.simps inv_on_left_moving.simps)
+apply(case_tac a, simp, simp)
+done
+
+(*inv: check_left_moving to after_left_moving*)
+lemma [simp]: "inv_check_left_moving (as, lm) (s, l, Bk # r) ires
+                \<Longrightarrow> inv_after_left_moving (as, lm) (s', Bk # l, r) ires"
+apply(auto simp: inv_check_left_moving.simps 
+ inv_check_left_moving_on_leftmost.simps inv_after_left_moving.simps)
+done
+
+
+lemma [simp]:"inv_check_left_moving (as, lm) (s, l, []) ires
+      \<Longrightarrow> inv_after_left_moving (as, lm) (s', Bk # l, []) ires"
+by(simp add: inv_check_left_moving.simps  
+inv_check_left_moving_in_middle.simps 
+inv_check_left_moving_on_leftmost.simps)
+
+(*inv: after_left_moving to inv_stop*)
+lemma [simp]: "inv_after_left_moving (as, lm) (s, l, Bk # r) ires
+       \<Longrightarrow> inv_stop (as, lm) (s', Bk # l, r) ires"
+apply(auto simp: inv_after_left_moving.simps inv_stop.simps)
+done
+
+lemma [simp]: "inv_after_left_moving (as, lm) (s, l, []) ires
+             \<Longrightarrow> inv_stop (as, lm) (s', Bk # l, []) ires"
+by(auto simp: inv_after_left_moving.simps)
+
+(*inv: stop to stop*)
+lemma [simp]: "inv_stop (as, lm) (x, l, r) ires \<Longrightarrow> 
+               inv_stop (as, lm) (y, l, r) ires"
+apply(simp add: inv_stop.simps)
+done
+
+lemma [simp]: "inv_after_clear (as, lm) (s, aaa, Oc # xs) ires= False"
+apply(auto simp: inv_after_clear.simps )
+done
+
+lemma [simp]: 
+  "inv_after_left_moving (as, lm) (s, aaa, Oc # xs) ires = False"
+by(auto simp: inv_after_left_moving.simps  )
+
+lemma [simp]: "inv_after_clear (as, abc_lm_s lm n (Suc (abc_lm_v lm n))) (s, b, []) ires = False"
+apply(auto simp: inv_after_clear.simps)
+done
+
+lemma [simp]: "inv_on_left_moving (as, abc_lm_s lm n (Suc (abc_lm_v lm n))) 
+           (s, b, Oc # list) ires \<Longrightarrow> b \<noteq> []"
+apply(auto simp: inv_on_left_moving.simps inv_on_left_moving_norm.simps split: if_splits)
+done
+
+lemma [simp]: "inv_check_left_moving (as, abc_lm_s lm n (Suc (abc_lm_v lm n))) (s, b, Oc # list) ires \<Longrightarrow> b \<noteq> []"
+apply(auto simp: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps split: if_splits)
+done
+
+lemma tinc_correct_pre:
+  assumes layout: "ly = layout_of ap"
+  and inv_start: "inv_locate_a (as, lm) (n, l, r) ires"
+  and lm': "lm' = abc_lm_s lm n (Suc (abc_lm_v lm n))"
+  and f: "f = steps (Suc 0, l, r) (tinc_b, 0)"
+  and P: "P = (\<lambda> (s, l, r). s = 10)"
+  and Q: "Q = (\<lambda> (s, l, r). inc_inv n (as, lm) (s, l, r) ires)" 
+  shows "\<exists> stp. P (f stp) \<and> Q (f stp)"
+proof(rule_tac LE = inc_LE in halt_lemma2)
+  show "wf inc_LE" by(auto)
+next
+  show "Q (f 0)"
+    using inv_start
+    apply(simp add: f P Q steps.simps inc_inv.simps)
+    done
+next
+  show "\<not> P (f 0)"
+    apply(simp add: f P steps.simps)
+    done
+next
+  show "\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) 
+        \<in> inc_LE"
+  proof(rule_tac allI, rule_tac impI, simp add: f, 
+      case_tac "steps (Suc 0, l, r) (tinc_b, 0) n", simp add: P)
+    fix n a b c
+    assume "a \<noteq> 10 \<and> Q (a, b, c)"
+    thus  "Q (step (a, b, c) (tinc_b, 0)) \<and> (step (a, b, c) (tinc_b, 0), a, b, c) \<in> inc_LE"
+      apply(simp add:Q)
+      apply(simp add: inc_inv.simps)
+      apply(case_tac c, case_tac [2] aa)
+      apply(auto simp: Let_def step.simps tinc_b_def numeral_2_eq_2 numeral_3_eq_3  split: if_splits)
+      apply(simp_all add: inc_inv.simps inc_LE_def lex_triple_def lex_pair_def inc_measure_def numeral_5_eq_5
+                          numeral_6_eq_6 numeral_7_eq_7 numeral_8_eq_8 numeral_9_eq_9)         
+      done
+  qed
+qed
+         
+
+lemma tinc_correct: 
+  assumes layout: "ly = layout_of ap"
+  and inv_start: "inv_locate_a (as, lm) (n, l, r) ires"
+  and lm': "lm' = abc_lm_s lm n (Suc (abc_lm_v lm n))"
+  shows "\<exists> stp l' r'. steps (Suc 0, l, r) (tinc_b, 0) stp = (10, l', r')
+              \<and> inv_stop (as, lm') (10, l', r') ires"
+  using assms
+  apply(drule_tac tinc_correct_pre, auto)
+  apply(rule_tac x = stp in exI, simp)
+  apply(simp add: inc_inv.simps)
+  done
+
+declare inv_locate_a.simps[simp del] abc_lm_s.simps[simp del]
+        abc_lm_v.simps[simp del]
+
+lemma [simp]: "(4::nat) * n mod 2 = 0"
+apply(arith)
+done
+
+lemma crsp_step_inc_pre:
+  assumes layout: "ly = layout_of ap"
+  and crsp: "crsp ly (as, lm) (s, l, r) ires"
+  and aexec: "abc_step_l (as, lm) (Some (Inc n)) = (asa, lma)"
+  shows "\<exists> stp k. steps (Suc 0, l, r) (findnth n @ shift tinc_b (2 * n), 0) stp 
+        = (2*n + 10, Bk # Bk # ires, <lma> @ Bk\<up>k) \<and> stp > 0"
+proof -
+  thm tm_append_steps
+  have "\<exists> stp l' r'. steps (Suc 0, l, r) (findnth n, 0) stp = (Suc (2 * n), l', r')
+    \<and> inv_locate_a (as, lm) (n, l', r') ires"
+    using assms
+    apply(rule_tac findnth_correct, simp_all add: crsp layout)
+    done
+  from this obtain stp l' r' where a:
+    "steps (Suc 0, l, r) (findnth n, 0) stp = (Suc (2 * n), l', r')
+    \<and> inv_locate_a (as, lm) (n, l', r') ires" by blast
+  moreover have
+    "\<exists> stp la ra. steps (Suc 0, l', r') (tinc_b, 0) stp = (10, la, ra)
+                        \<and> inv_stop (as, lma) (10, la, ra) ires"
+    using assms a
+  proof(rule_tac lm' = lma and n = n and lm = lm and ly = ly and ap = ap in tinc_correct,
+      simp, simp)
+    show "lma = abc_lm_s lm n (Suc (abc_lm_v lm n))"
+      using aexec
+      apply(simp add: abc_step_l.simps)
+      done
+  qed
+  from this obtain stpa la ra where b:
+    "steps (Suc 0, l', r') (tinc_b, 0) stpa = (10, la, ra)
+    \<and> inv_stop (as, lma) (10, la, ra) ires" by blast
+  from a b show "\<exists>stp k. steps (Suc 0, l, r) (findnth n @ shift tinc_b (2 * n), 0) stp
+    = (2 * n + 10, Bk # Bk # ires, <lma> @ Bk \<up> k) \<and> stp > 0"
+    apply(rule_tac x = "stp + stpa" in exI)
+    using tm_append_steps[of "Suc 0" l r "findnth n" stp l' r' tinc_b stpa 10 la ra "length (findnth n) div 2"]
+    apply(simp add: length_findnth inv_stop.simps)
+    apply(case_tac stpa, simp_all add: steps.simps)
+    done
+qed 
+     
+lemma crsp_step_inc:
+  assumes layout: "ly = layout_of ap"
+  and crsp: "crsp ly (as, lm) (s, l, r) ires"
+  and fetch: "abc_fetch as ap = Some (Inc n)"
+  shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Inc n)))
+  (steps (s, l, r) (ci ly (start_of ly as) (Inc n), start_of ly as - Suc 0) stp) ires"
+proof(case_tac "(abc_step_l (as, lm) (Some (Inc n)))")
+  fix a b
+  assume aexec: "abc_step_l (as, lm) (Some (Inc n)) = (a, b)"
+  then have "\<exists> stp k. steps (Suc 0, l, r) (findnth n @ shift tinc_b (2 * n), 0) stp 
+        = (2*n + 10, Bk # Bk # ires, <b> @ Bk\<up>k) \<and> stp > 0"
+    using assms
+    apply(rule_tac crsp_step_inc_pre, simp_all)
+    done
+  thus "?thesis"
+    using assms aexec
+    apply(erule_tac exE)
+    apply(erule_tac exE)
+    apply(erule_tac conjE)
+    apply(rule_tac x = stp in exI, simp add: ci.simps tm_shift_eq_steps)
+    apply(drule_tac off = "(start_of (layout_of ap) as - Suc 0)" in tm_shift_eq_steps)
+    apply(auto simp: crsp.simps abc_step_l.simps fetch start_of_Suc1)
+    done
+qed
+    
+subsection{* Crsp of Dec n e*}
+declare sete.simps[simp del]
+
+type_synonym dec_inv_t = "(nat * nat list) \<Rightarrow> config \<Rightarrow> cell list \<Rightarrow>  bool"
+
+fun dec_first_on_right_moving :: "nat \<Rightarrow> dec_inv_t"
+  where
+  "dec_first_on_right_moving n (as, lm) (s, l, r) ires = 
+               (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
+         ml + mr = Suc m \<and> length lm1 = n \<and> ml > 0 \<and> m > 0 \<and>
+             (if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires
+                          else  l = Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
+    ((r = Oc\<up>mr @ [Bk] @ <lm2> @ Bk\<up>rn) \<or> (r = Oc\<up>mr \<and> lm2 = [])))"
+
+fun dec_on_right_moving :: "dec_inv_t"
+  where
+  "dec_on_right_moving (as, lm) (s, l, r) ires =  
+   (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
+                             ml + mr = Suc (Suc m) \<and>
+   (if lm1 = [] then l = Oc\<up>ml@ Bk # Bk # ires
+                else  l = Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
+   ((r = Oc\<up>mr @ [Bk] @ <lm2> @ Bk\<up>rn) \<or> (r = Oc\<up>mr \<and> lm2 = [])))"
+
+fun dec_after_clear :: "dec_inv_t"
+  where
+  "dec_after_clear (as, lm) (s, l, r) ires = 
+              (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
+                ml + mr = Suc m \<and> ml = Suc m \<and> r \<noteq> [] \<and> r \<noteq> [] \<and>
+               (if lm1 = [] then l = Oc\<up>ml@ Bk # Bk # ires
+                            else l = Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
+               (tl r = Bk # <lm2> @ Bk\<up>rn \<or> tl r = [] \<and> lm2 = []))"
+
+fun dec_after_write :: "dec_inv_t"
+  where
+  "dec_after_write (as, lm) (s, l, r) ires = 
+         (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
+       ml + mr = Suc m \<and> ml = Suc m \<and> lm2 \<noteq> [] \<and>
+       (if lm1 = [] then l = Bk # Oc\<up>ml @ Bk # Bk # ires
+                    else l = Bk # Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
+       tl r = <lm2> @ Bk\<up>rn)"
+
+fun dec_right_move :: "dec_inv_t"
+  where
+  "dec_right_move (as, lm) (s, l, r) ires = 
+        (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 
+            \<and> ml = Suc m \<and> mr = (0::nat) \<and> 
+              (if lm1 = [] then l = Bk # Oc\<up>ml @ Bk # Bk # ires
+                          else l = Bk # Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) 
+           \<and> (r = Bk # <lm2> @ Bk\<up>rn \<or> r = [] \<and> lm2 = []))"
+
+fun dec_check_right_move :: "dec_inv_t"
+  where
+  "dec_check_right_move (as, lm) (s, l, r) ires = 
+        (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
+           ml = Suc m \<and> mr = (0::nat) \<and> 
+           (if lm1 = [] then l = Bk # Bk # Oc\<up>ml @ Bk # Bk # ires
+                       else l = Bk # Bk # Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
+           r = <lm2> @ Bk\<up>rn)"
+
+fun dec_left_move :: "dec_inv_t"
+  where
+  "dec_left_move (as, lm) (s, l, r) ires = 
+    (\<exists> lm1 m rn. (lm::nat list) = lm1 @ [m::nat] \<and>   
+    rn > 0 \<and> 
+   (if lm1 = [] then l = Bk # Oc\<up>Suc m @ Bk # Bk # ires
+    else l = Bk # Oc\<up>Suc m @ Bk # <rev lm1> @ Bk # Bk # ires) \<and> r = Bk\<up>rn)"
+
+declare
+  dec_on_right_moving.simps[simp del] dec_after_clear.simps[simp del] 
+  dec_after_write.simps[simp del] dec_left_move.simps[simp del] 
+  dec_check_right_move.simps[simp del] dec_right_move.simps[simp del] 
+  dec_first_on_right_moving.simps[simp del]
+
+fun inv_locate_n_b :: "inc_inv_t"
+  where
+  "inv_locate_n_b (as, lm) (s, l, r) ires= 
+    (\<exists> lm1 lm2 tn m ml mr rn. lm @ 0\<up>tn = lm1 @ [m] @ lm2 \<and> 
+     length lm1 = s \<and> m + 1 = ml + mr \<and> 
+     ml = 1 \<and> tn = s + 1 - length lm \<and>
+     (if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires
+      else l = Oc\<up>ml @ Bk # <rev lm1> @ Bk # Bk # ires) \<and> 
+     (r = Oc\<up>mr @ [Bk] @ <lm2>@ Bk\<up>rn \<or> (lm2 = [] \<and> r = Oc\<up>mr))
+  )"
+(*
+fun dec_inv_1 :: "layout \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> dec_inv_t"
+  where
+  "dec_inv_1 ly n e (as, am) (s, l, r) ires = 
+           (let ss = start_of ly as in
+            let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in
+            let am'' = abc_lm_s am n (abc_lm_v am n) in
+              if s = start_of ly e then inv_stop (as, am'') (s, l, r) ires
+              else if s = ss then False
+              else if s = ss + 2 * n then 
+                  inv_locate_a (as, am) (n, l, r) ires
+                \<or> inv_locate_a (as, am'') (n, l, r) ires
+              else if s = ss + 2 * n + 1 then 
+                  inv_locate_b (as, am) (n, l, r) ires
+              else if s = ss + 2 * n + 13 then 
+                  inv_on_left_moving (as, am'') (s, l, r) ires
+              else if s = ss + 2 * n + 14 then 
+                  inv_check_left_moving (as, am'') (s, l, r) ires
+              else if s = ss + 2 * n + 15 then 
+                  inv_after_left_moving (as, am'') (s, l, r) ires
+              else False)"
+*)
+
+
+fun dec_inv_1 :: "layout \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> dec_inv_t"
+  where
+  "dec_inv_1 ly n e (as, am) (s, l, r) ires = 
+           (let ss = start_of ly as in
+            let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in
+            let am'' = abc_lm_s am n (abc_lm_v am n) in
+              if s = start_of ly e then inv_stop (as, am'') (s, l, r) ires
+              else if s = ss then False
+              else if s = ss + 2 * n + 1 then 
+                  inv_locate_b (as, am) (n, l, r) ires
+              else if s = ss + 2 * n + 13 then 
+                  inv_on_left_moving (as, am'') (s, l, r) ires
+              else if s = ss + 2 * n + 14 then 
+                  inv_check_left_moving (as, am'') (s, l, r) ires
+              else if s = ss + 2 * n + 15 then 
+                  inv_after_left_moving (as, am'') (s, l, r) ires
+              else False)"
+
+declare fetch.simps[simp del]
+lemma [simp]:
+  "fetch (ci ly (start_of ly as) (Dec n e)) (Suc (2 * n)) Bk = (W1,  start_of ly as + 2 *n)"
+apply(auto simp: fetch.simps length_ci_dec)
+apply(auto simp: ci.simps nth_append length_findnth sete.simps shift.simps tdec_b_def)
+using startof_not0[of ly as] by simp
+
+lemma [simp]:
+  "fetch (ci ly (start_of ly as) (Dec n e)) (Suc (2 * n)) Oc = (R,  Suc (start_of ly as) + 2 *n)"
+apply(auto simp: fetch.simps length_ci_dec)
+apply(auto simp: ci.simps nth_append length_findnth sete.simps shift.simps tdec_b_def)
+done
+
+lemma [simp]: 
+  "\<lbrakk>r = [] \<or> hd r = Bk; inv_locate_a (as, lm) (n, l, r) ires\<rbrakk>
+    \<Longrightarrow> \<exists>stp la ra.
+  steps (start_of ly as + 2 * n, l, r) (ci ly (start_of ly as) (Dec n e), 
+  start_of ly as - Suc 0) stp = (Suc (start_of ly as + 2 * n), la, ra) \<and>
+  inv_locate_b (as, lm) (n, la, ra) ires"
+apply(rule_tac x = "Suc (Suc 0)" in exI)
+apply(auto simp: steps.simps step.simps length_ci_dec)
+apply(case_tac r, simp_all)
+done
+
+lemma [simp]: 
+  "\<lbrakk>inv_locate_a (as, lm) (n, l, r) ires; r \<noteq> [] \<and> hd r \<noteq> Bk\<rbrakk>
+    \<Longrightarrow> \<exists>stp la ra.
+  steps (start_of ly as + 2 * n, l, r) (ci ly (start_of ly as) (Dec n e), 
+  start_of ly as - Suc 0) stp = (Suc (start_of ly as + 2 * n), la, ra) \<and>
+  inv_locate_b (as, lm) (n, la, ra) ires"
+apply(rule_tac x = "(Suc 0)" in exI, case_tac "hd r", simp_all)
+apply(auto simp: steps.simps step.simps length_ci_dec)
+apply(case_tac r, simp_all)
+done
+
+fun abc_dec_1_stage1:: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "abc_dec_1_stage1 (s, l, r) ss n = 
+       (if s > ss \<and> s \<le> ss + 2*n + 1 then 4
+        else if s = ss + 2 * n + 13 \<or> s = ss + 2*n + 14 then 3
+        else if s = ss + 2*n + 15 then 2
+        else 0)"
+
+fun abc_dec_1_stage2:: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "abc_dec_1_stage2 (s, l, r) ss n = 
+       (if s \<le> ss + 2 * n + 1 then (ss + 2 * n + 16 - s)
+        else if s = ss + 2*n + 13 then length l
+        else if s = ss + 2*n + 14 then length l
+        else 0)"
+
+fun abc_dec_1_stage3 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "abc_dec_1_stage3 (s, l, r) ss n  = 
+        (if s \<le> ss + 2*n + 1 then 
+             if (s - ss) mod 2 = 0 then 
+                         if r \<noteq> [] \<and> hd r = Oc then 0 else 1  
+                         else length r
+         else if s = ss + 2 * n + 13 then 
+             if r \<noteq> [] \<and> hd r = Oc then 2 
+             else 1
+         else if s = ss + 2 * n + 14 then 
+             if r \<noteq> [] \<and> hd r = Oc then 3 else 0 
+         else 0)"
+
+fun abc_dec_1_measure :: "(config \<times> nat \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)"
+  where
+  "abc_dec_1_measure (c, ss, n) = (abc_dec_1_stage1 c ss n, 
+                   abc_dec_1_stage2 c ss n, abc_dec_1_stage3 c ss n)"
+
+definition abc_dec_1_LE ::
+  "((config \<times> nat \<times>
+  nat) \<times> (config \<times> nat \<times> nat)) set"
+  where "abc_dec_1_LE \<equiv> (inv_image lex_triple abc_dec_1_measure)"
+
+lemma wf_dec_le: "wf abc_dec_1_LE"
+by(auto intro:wf_inv_image simp:abc_dec_1_LE_def lex_triple_def lex_pair_def)
+
+lemma startof_Suc2:
+  "abc_fetch as ap = Some (Dec n e) \<Longrightarrow> 
+        start_of (layout_of ap) (Suc as) = 
+            start_of (layout_of ap) as + 2 * n + 16"
+apply(auto simp: start_of.simps layout_of.simps  
+                 length_of.simps abc_fetch.simps 
+                 take_Suc_conv_app_nth split: if_splits)
+done
+
+lemma start_of_less_2: 
+  "start_of ly e \<le> start_of ly (Suc e)"
+thm take_Suc
+apply(case_tac "e < length ly")
+apply(auto simp: start_of.simps take_Suc take_Suc_conv_app_nth)
+done
+
+lemma start_of_less_1: "start_of ly e \<le> start_of ly (e + d)"
+proof(induct d)
+  case 0 thus "?case" by simp
+next
+  case (Suc d)
+  have "start_of ly e \<le> start_of ly (e + d)"  by fact
+  moreover have "start_of ly (e + d) \<le> start_of ly (Suc (e + d))"
+    by(rule_tac start_of_less_2)
+  ultimately show"?case"
+    by(simp)
+qed
+
+lemma start_of_less: 
+  assumes "e < as"
+  shows "start_of ly e \<le> start_of ly as"
+proof -
+  obtain d where " as = e + d"
+    using assms by (metis less_imp_add_positive)
+  thus "?thesis"
+    by(simp add: start_of_less_1)
+qed
+
+lemma start_of_ge: 
+  assumes fetch: "abc_fetch as ap = Some (Dec n e)"
+  and layout: "ly = layout_of ap"
+  and great: "e > as"
+  shows "start_of ly e \<ge> start_of ly as + 2*n + 16"
+proof(cases "e = Suc as")
+  case True
+  have "e = Suc as" by fact
+  moreover hence "start_of ly (Suc as) = start_of ly as + 2*n + 16"
+    using layout fetch
+    by(simp add: startof_Suc2)
+  ultimately show "?thesis" by (simp)
+next
+  case False
+  have "e \<noteq> Suc as" by fact
+  then have "e > Suc as" using great by arith
+  then have "start_of ly (Suc as) \<le> start_of ly e"
+    by(simp add: start_of_less)
+  moreover have "start_of ly (Suc as) = start_of ly as + 2*n + 16"
+    using layout fetch
+    by(simp add: startof_Suc2)
+  ultimately show "?thesis"
+    by arith
+qed
+    
+lemma [elim]: "\<lbrakk>abc_fetch as ap = Some (Dec n e); as < e; 
+  Suc (start_of (layout_of ap) as + 2 * n) = start_of (layout_of ap) e\<rbrakk> \<Longrightarrow> RR"
+apply(drule_tac start_of_ge, simp_all)
+apply(auto)
+done
+
+lemma [elim]: "\<lbrakk>abc_fetch as ap = Some (Dec n e); as > e;
+  Suc (start_of (layout_of ap) as + 2 * n) = start_of (layout_of ap) e\<rbrakk> \<Longrightarrow> RR"
+apply(drule_tac ly = "layout_of ap" in start_of_less[of])
+apply(arith)
+done
+
+lemma [elim]: "\<lbrakk>abc_fetch as ap = Some (Dec n e);
+  Suc (start_of (layout_of ap) as + 2 * n) = start_of (layout_of ap) e\<rbrakk> \<Longrightarrow> RR"
+apply(subgoal_tac "as = e \<or> as < e \<or> as > e", auto)
+done
+
+lemma [simp]:"fetch (ci (ly) (start_of ly as) (Dec n e)) (Suc (2 * n))  Oc
+  = (R, start_of ly as + 2*n + 1)"
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]: "(start_of ly as = 0) = False"
+apply(simp add: start_of.simps)
+done
+
+lemma [simp]: "fetch (ci (ly) 
+  (start_of ly as) (Dec n e)) (Suc (2 * n))  Bk
+  = (W1, start_of ly as + 2*n)"
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]: 
+  "fetch (ci (ly) 
+                (start_of ly as) (Dec n e)) (Suc (Suc (2 * n)))  Oc
+      = (R, start_of ly as + 2*n + 2)"
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+
+lemma [simp]: "fetch (ci (ly) 
+                  (start_of ly as) (Dec n e)) (Suc (Suc (2 * n))) Bk
+      = (L, start_of ly as + 2*n + 13)"
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+
+lemma [simp]: "fetch (ci (ly) 
+             (start_of ly as) (Dec n e)) (Suc (Suc (Suc (2 * n))))  Oc
+     = (R, start_of ly as + 2*n + 2)"
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+
+lemma [simp]: "fetch (ci (ly) (start_of ly as) (Dec n e)) 
+                             (Suc (Suc (Suc (2 * n))))  Bk
+     = (L, start_of ly as + 2*n + 3)"
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]:
+     "fetch (ci (ly) 
+                      (start_of ly as) (Dec n e)) (2 * n + 4) Oc
+    = (W0, start_of ly as + 2*n + 3)"
+apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]: "fetch (ci (ly) 
+                   (start_of ly as) (Dec n e)) (2 * n + 4) Bk
+    = (R, start_of ly as + 2*n + 4)"
+apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]:"fetch (ci (ly) 
+                          (start_of ly as) (Dec n e)) (2 * n + 5) Bk
+    = (R, start_of ly as + 2*n + 5)"
+apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+
+lemma [simp]: 
+  "fetch (ci (ly)
+                (start_of ly as) (Dec n e)) (2 * n + 6) Bk
+    = (L, start_of ly as + 2*n + 6)"
+apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]: 
+  "fetch (ci (ly) (start_of ly as) 
+                      (Dec n e)) (2 * n + 6) Oc
+    = (L, start_of ly as + 2*n + 7)"
+apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]:"fetch (ci (ly) 
+             (start_of ly as) (Dec n e)) (2 * n + 7) Bk
+    = (L, start_of ly as + 2*n + 10)"
+apply(subgoal_tac "2*n + 7 = Suc (2*n + 6)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]:
+  "fetch (ci (ly) 
+                   (start_of ly as) (Dec n e)) (2 * n + 8) Bk
+    = (W1, start_of ly as + 2*n + 7)"
+apply(subgoal_tac "2*n + 8 = Suc (2*n + 7)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+
+lemma [simp]: 
+  "fetch (ci (ly) 
+                   (start_of ly as) (Dec n e)) (2 * n + 8) Oc
+    = (R, start_of ly as + 2*n + 8)"
+apply(subgoal_tac "2*n + 8 = Suc (2*n + 7)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]: 
+  "fetch (ci (ly) 
+  (start_of ly as) (Dec n e)) (2 * n + 9) Bk
+  = (L, start_of ly as + 2*n + 9)"
+apply(subgoal_tac "2*n + 9 = Suc (2*n + 8)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]: 
+  "fetch (ci (ly) 
+  (start_of ly as) (Dec n e)) (2 * n + 9) Oc
+  = (R, start_of ly as + 2*n + 8)"
+apply(subgoal_tac "2*n + 9 = Suc (2*n + 8)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+
+lemma [simp]: 
+  "fetch (ci (ly) 
+  (start_of ly as) (Dec n e)) (2 * n + 10) Bk
+  = (R, start_of ly as + 2*n + 4)" 
+apply(subgoal_tac "2*n + 10 = Suc (2*n + 9)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]: "fetch (ci (ly) 
+             (start_of ly as) (Dec n e)) (2 * n + 10) Oc
+    = (W0, start_of ly as + 2*n + 9)"
+apply(subgoal_tac "2*n + 10 = Suc (2*n + 9)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+
+lemma [simp]: 
+  "fetch (ci (ly) 
+  (start_of ly as) (Dec n e)) (2 * n + 11) Oc
+  = (L, start_of ly as + 2*n + 10)"
+apply(subgoal_tac "2*n + 11 = Suc (2*n + 10)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+
+lemma [simp]: 
+  "fetch (ci (ly) 
+  (start_of ly as) (Dec n e)) (2 * n + 11) Bk
+  = (L, start_of ly as + 2*n + 11)"
+apply(subgoal_tac "2*n + 11 = Suc (2*n + 10)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]:
+  "fetch (ci (ly) 
+  (start_of ly as) (Dec n e)) (2 * n + 12) Oc
+  = (L, start_of ly as + 2*n + 10)"
+apply(subgoal_tac "2*n + 12 = Suc (2*n + 11)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+
+lemma [simp]: 
+  "fetch (ci (ly) 
+  (start_of ly as) (Dec n e)) (2 * n + 12) Bk
+  = (R, start_of ly as + 2*n + 12)"
+apply(subgoal_tac "2*n + 12 = Suc (2*n + 11)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]: 
+  "fetch (ci (ly) 
+  (start_of ly as) (Dec n e)) (2 * n + 13) Bk
+  = (R, start_of ly as + 2*n + 16)"
+apply(subgoal_tac "2*n + 13 = Suc (2*n + 12)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+
+lemma [simp]: 
+  "fetch (ci (ly) 
+  (start_of ly as) (Dec n e)) (14 + 2 * n) Oc
+  = (L, start_of ly as + 2*n + 13)"
+apply(subgoal_tac "14 + 2*n = Suc (2*n + 13)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]: 
+  "fetch (ci (ly) 
+  (start_of ly as) (Dec n e)) (14 + 2 * n) Bk
+  = (L, start_of ly as + 2*n + 14)"
+apply(subgoal_tac "14 + 2*n = Suc (2*n + 13)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]: 
+  "fetch (ci (ly) 
+  (start_of ly as) (Dec n e)) (15 + 2 * n)  Oc
+  = (L, start_of ly as + 2*n + 13)"
+apply(subgoal_tac "15 + 2*n = Suc (2*n + 14)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]: 
+  "fetch (ci (ly) 
+  (start_of ly as) (Dec n e)) (15 + 2 * n)  Bk
+ = (R, start_of ly as + 2*n + 15)"
+apply(subgoal_tac "15 + 2*n = Suc (2*n + 14)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+lemma [simp]: 
+  "abc_fetch as aprog = Some (Dec n e) \<Longrightarrow> 
+     fetch (ci (ly) (start_of (ly) as) 
+              (Dec n e)) (16 + 2 * n)  Bk
+ = (R, start_of (ly) e)"
+apply(subgoal_tac "16 + 2*n = Suc (2*n + 15)", simp only: fetch.simps)
+apply(auto simp: ci.simps findnth.simps fetch.simps
+                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
+done
+
+declare dec_inv_1.simps[simp del]
+
+
+lemma [simp]: 
+ "\<lbrakk>abc_fetch as aprog = Some (Dec n e); ly = layout_of aprog\<rbrakk>
+   \<Longrightarrow> (start_of ly e \<noteq> Suc (start_of ly as + 2 * n) \<and>
+        start_of ly e \<noteq> Suc (Suc (start_of ly as + 2 * n)) \<and>  
+        start_of ly e \<noteq> start_of ly as + 2 * n + 3 \<and> 
+        start_of ly e \<noteq> start_of ly as + 2 * n + 4 \<and>
+        start_of ly e \<noteq> start_of ly as + 2 * n + 5 \<and> 
+        start_of ly e \<noteq> start_of ly as + 2 * n + 6 \<and> 
+        start_of ly e \<noteq> start_of ly as + 2 * n + 7 \<and>
+        start_of ly e \<noteq> start_of ly as + 2 * n + 8 \<and> 
+        start_of ly e \<noteq> start_of ly as + 2 * n + 9 \<and> 
+        start_of ly e \<noteq> start_of ly as + 2 * n + 10 \<and>
+        start_of ly e \<noteq> start_of ly as + 2 * n + 11 \<and> 
+        start_of ly e \<noteq> start_of ly as + 2 * n + 12 \<and> 
+        start_of ly e \<noteq> start_of ly as + 2 * n + 13 \<and>
+        start_of ly e \<noteq> start_of ly as + 2 * n + 14 \<and> 
+        start_of ly e \<noteq> start_of ly as + 2 * n + 15)"
+using start_of_ge[of as aprog n e ly] start_of_less[of e as ly]
+apply(case_tac "e < as", simp)
+apply(case_tac "e = as", simp, simp)
+done
+
+lemma [simp]: "\<lbrakk>abc_fetch as aprog = Some (Dec n e); ly = layout_of aprog\<rbrakk>
+      \<Longrightarrow> (Suc (start_of ly as + 2 * n) \<noteq> start_of ly e \<and>
+          Suc (Suc (start_of ly as + 2 * n)) \<noteq> start_of ly e \<and> 
+          start_of ly as + 2 * n + 3 \<noteq> start_of ly e \<and> 
+          start_of ly as + 2 * n + 4 \<noteq> start_of ly e \<and>
+          start_of ly as + 2 * n + 5 \<noteq>start_of ly e \<and> 
+          start_of ly as + 2 * n + 6 \<noteq> start_of ly e \<and>
+          start_of ly as + 2 * n + 7 \<noteq> start_of ly e \<and> 
+          start_of ly as + 2 * n + 8 \<noteq> start_of ly e \<and> 
+          start_of ly as + 2 * n + 9 \<noteq> start_of ly e \<and> 
+          start_of ly as + 2 * n + 10 \<noteq> start_of ly e \<and>
+          start_of ly as + 2 * n + 11 \<noteq> start_of ly e \<and> 
+          start_of ly as + 2 * n + 12 \<noteq> start_of ly e \<and> 
+          start_of ly as + 2 * n + 13 \<noteq> start_of ly e \<and> 
+          start_of ly as + 2 * n + 14 \<noteq> start_of ly e \<and> 
+          start_of ly as + 2 * n + 15 \<noteq> start_of ly e)"
+using start_of_ge[of as aprog n e ly] start_of_less[of e as ly]
+apply(case_tac "e < as", simp, simp)
+apply(case_tac "e = as", simp, simp)
+done
+
+lemma [simp]: "inv_locate_b (as, lm) (n, [], []) ires = False"
+apply(auto simp: inv_locate_b.simps in_middle.simps split: if_splits)
+done
+
+lemma [simp]: "inv_locate_b (as, lm) (n, [], Bk # list) ires = False"
+apply(auto simp: inv_locate_b.simps in_middle.simps split: if_splits)
+done
+
+(*
+
+lemma  inv_locate_b_2_on_left_moving_b[simp]: 
+   "inv_locate_b (as, am) (n, l, []) ires
+     \<Longrightarrow> inv_on_left_moving (as, 
+                  abc_lm_s am n (abc_lm_v am n)) (s, [], [Bk]) ires"
+apply(auto simp: inv_locate_b.simps inv_on_left_moving.simps inv_on_left_moving_in_middle_B.simps
+                 in_middle.simps split: if_splits)
+apply(drule_tac length_equal, simp)
+
+apply(insert inv_locate_b_2_on_left_moving[of as am n l "[]" ires s])
+apply(simp only: inv_on_left_moving.simps, simp)
+apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B 
+         (as, abc_lm_s am n (abc_lm_v am n)) (s, tl l, [hd l]) ires", simp)
+*)
+
+(*
+lemma [simp]:
+  "inv_locate_b (as, am) (n, l, []) ires; l \<noteq> []\<rbrakk>
+ \<Longrightarrow> inv_on_left_moving (as, abc_lm_s am n 
+                  (abc_lm_v am n)) (s, tl l, [hd l]) ires"
+apply(auto simp: inv_locate_b.simps inv_on_left_moving.simps inv_on_left_moving_in_middle_B.simps
+                 in_middle.simps split: if_splits)
+apply(drule_tac length_equal, simp)
+
+apply(insert inv_locate_b_2_on_left_moving[of as am n l "[]" ires s])
+apply(simp only: inv_on_left_moving.simps, simp)
+apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B 
+         (as, abc_lm_s am n (abc_lm_v am n)) (s, tl l, [hd l]) ires", simp)
+
+apply(insert inv_locate_b_2_on_left_moving[of as am n l "[]" ires s])
+apply(simp only: inv_on_left_moving.simps, simp)
+apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B 
+         (as, abc_lm_s am n (abc_lm_v am n)) (s, tl l, [hd l]) ires", simp)
+apply(simp only: inv_on_left_moving_norm.simps)
+apply(erule_tac exE)+
+apply(erule_tac conjE)+
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
+      rule_tac x = m in exI, rule_tac x = ml in exI, 
+      rule_tac x = mr in exI, simp)
+apply(case_tac mr, simp, simp, case_tac nat, auto intro: nil_2_nil)
+done
+*)
+
+lemma [simp]: 
+ "\<lbrakk>dec_first_on_right_moving n (as, am) (s, aaa, Oc # xs) ires\<rbrakk>
+   \<Longrightarrow> dec_first_on_right_moving n (as, am) (s', Oc # aaa, xs) ires"
+apply(simp only: dec_first_on_right_moving.simps)
+apply(erule exE)+
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
+      rule_tac x = m in exI, simp)
+apply(rule_tac x = "Suc ml" in exI, 
+      rule_tac x = "mr - 1" in exI, auto)
+apply(case_tac [!] mr, auto)
+done
+
+lemma [simp]: 
+  "dec_first_on_right_moving n (as, am) (s, l, Bk # xs) ires \<Longrightarrow> l \<noteq> []"
+apply(auto simp: dec_first_on_right_moving.simps split: if_splits)
+done
+
+lemma [elim]: 
+  "\<lbrakk>\<not> length lm1 < length am; 
+    am @ replicate (length lm1 - length am) 0 @ [0::nat] = 
+                                                lm1 @ m # lm2;
+    0 < m\<rbrakk>
+   \<Longrightarrow> RR"
+apply(subgoal_tac "lm2 = []", simp)
+apply(drule_tac length_equal, simp)
+done
+
+lemma [simp]: 
+ "\<lbrakk>dec_first_on_right_moving n (as, 
+                   abc_lm_s am n (abc_lm_v am n)) (s, l, Bk # xs) ires\<rbrakk>
+\<Longrightarrow> dec_after_clear (as, abc_lm_s am n 
+                 (abc_lm_v am n - Suc 0)) (s', tl l, hd l # Bk # xs) ires"
+apply(simp only: dec_first_on_right_moving.simps 
+                 dec_after_clear.simps abc_lm_s.simps abc_lm_v.simps)
+apply(erule_tac exE)+
+apply(case_tac "n < length am")
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
+      rule_tac x = "m - 1" in exI, auto simp: )
+apply(case_tac [!] mr, auto)
+done
+
+lemma [simp]: 
+ "\<lbrakk>dec_first_on_right_moving n (as, 
+                   abc_lm_s am n (abc_lm_v am n)) (s, l, []) ires\<rbrakk>
+\<Longrightarrow> (l = [] \<longrightarrow> dec_after_clear (as, 
+             abc_lm_s am n (abc_lm_v am n - Suc 0)) (s', [], [Bk]) ires) \<and>
+    (l \<noteq> [] \<longrightarrow> dec_after_clear (as, abc_lm_s am n 
+                      (abc_lm_v am n - Suc 0)) (s', tl l, [hd l]) ires)"
+apply(subgoal_tac "l \<noteq> []", 
+      simp only: dec_first_on_right_moving.simps 
+                 dec_after_clear.simps abc_lm_s.simps abc_lm_v.simps)
+apply(erule_tac exE)+
+apply(case_tac "n < length am", simp)
+apply(rule_tac x = lm1 in exI, rule_tac x = "m - 1" in exI, auto)
+apply(case_tac [1-2] m, auto)
+apply(auto simp: dec_first_on_right_moving.simps split: if_splits)
+done
+
+lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, Oc # r) ires\<rbrakk>
+                \<Longrightarrow> dec_after_clear (as, am) (s', l, Bk # r) ires"
+apply(auto simp: dec_after_clear.simps)
+done
+
+lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, Bk # r) ires\<rbrakk>
+                \<Longrightarrow> dec_right_move (as, am) (s', Bk # l, r) ires"
+apply(auto simp: dec_after_clear.simps dec_right_move.simps split: if_splits)
+done
+
+lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, []) ires\<rbrakk>
+             \<Longrightarrow> dec_right_move (as, am) (s', Bk # l, []) ires"
+apply(auto simp: dec_after_clear.simps dec_right_move.simps )
+done
+
+lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, []) ires\<rbrakk>
+             \<Longrightarrow> dec_right_move (as, am) (s', Bk # l, [Bk]) ires"
+apply(auto simp: dec_after_clear.simps dec_right_move.simps split: if_splits)
+done
+
+lemma [simp]:"dec_right_move (as, am) (s, l, Oc # r) ires = False"
+apply(auto simp: dec_right_move.simps)
+done
+              
+lemma dec_right_move_2_check_right_move[simp]:
+     "\<lbrakk>dec_right_move (as, am) (s, l, Bk # r) ires\<rbrakk>
+      \<Longrightarrow> dec_check_right_move (as, am) (s', Bk # l, r) ires"
+apply(auto simp: dec_right_move.simps dec_check_right_move.simps split: if_splits)
+done
+
+lemma [simp]: "(<lm::nat list> = []) = (lm = [])"
+apply(case_tac lm, simp_all add: tape_of_nl_cons)
+done
+
+lemma [simp]: 
+ "dec_right_move (as, am) (s, l, []) ires= 
+  dec_right_move (as, am) (s, l, [Bk]) ires"
+apply(simp add: dec_right_move.simps)
+done
+
+lemma [simp]: "\<lbrakk>dec_right_move (as, am) (s, l, []) ires\<rbrakk>
+             \<Longrightarrow> dec_check_right_move (as, am) (s, Bk # l, []) ires"
+apply(insert dec_right_move_2_check_right_move[of as am s l "[]" s'], 
+      simp)
+done
+
+lemma [simp]: "dec_check_right_move (as, am) (s, l, r) ires\<Longrightarrow> l \<noteq> []"
+apply(auto simp: dec_check_right_move.simps split: if_splits)
+done
+ 
+lemma [simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, Oc # r) ires\<rbrakk>
+             \<Longrightarrow> dec_after_write (as, am) (s', tl l, hd l # Oc # r) ires"
+apply(auto simp: dec_check_right_move.simps dec_after_write.simps)
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
+      rule_tac x = m in exI, auto)
+done
+
+
+
+lemma [simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, Bk # r) ires\<rbrakk>
+                \<Longrightarrow> dec_left_move (as, am) (s', tl l, hd l # Bk # r) ires"
+apply(auto simp: dec_check_right_move.simps 
+                 dec_left_move.simps inv_after_move.simps)
+apply(rule_tac x = lm1 in exI, rule_tac x = m in exI, auto split: if_splits)
+apply(case_tac [!] lm2, simp_all add: tape_of_nl_cons split: if_splits)
+apply(rule_tac [!] x = "(Suc rn)" in exI, simp_all)
+done
+
+lemma [simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, []) ires\<rbrakk>
+             \<Longrightarrow> dec_left_move (as, am) (s', tl l, [hd l]) ires"
+apply(auto simp: dec_check_right_move.simps 
+                 dec_left_move.simps inv_after_move.simps)
+apply(rule_tac x = lm1 in exI, rule_tac x = m in exI, auto)
+done
+
+lemma [simp]: "dec_left_move (as, am) (s, aaa, Oc # xs) ires = False"
+apply(auto simp: dec_left_move.simps inv_after_move.simps)
+done
+
+lemma [simp]: "dec_left_move (as, am) (s, l, r) ires
+             \<Longrightarrow> l \<noteq> []"
+apply(auto simp: dec_left_move.simps split: if_splits)
+done
+
+lemma [simp]: "inv_on_left_moving_in_middle_B (as, [m])
+  (s', Oc # Oc\<up>m @ Bk # Bk # ires, Bk # Bk\<up>rn) ires"
+apply(simp add: inv_on_left_moving_in_middle_B.simps)
+apply(rule_tac x = "[m]" in exI, auto)
+done
+
+lemma [simp]: "inv_on_left_moving_in_middle_B (as, [m])
+  (s', Oc # Oc\<up>m @ Bk # Bk # ires, [Bk]) ires"
+apply(simp add: inv_on_left_moving_in_middle_B.simps)
+done
+
+
+lemma [simp]: "lm1 \<noteq> [] \<Longrightarrow> 
+  inv_on_left_moving_in_middle_B (as, lm1 @ [m]) (s', 
+  Oc # Oc\<up>m @ Bk # <rev lm1> @ Bk # Bk # ires, Bk # Bk\<up>rn) ires"
+apply(simp only: inv_on_left_moving_in_middle_B.simps)
+apply(rule_tac x = "lm1 @ [m ]" in exI, rule_tac x = "[]" in exI, simp)
+apply(simp add: tape_of_nl_cons split: if_splits)
+done
+
+lemma [simp]: "lm1 \<noteq> [] \<Longrightarrow> 
+  inv_on_left_moving_in_middle_B (as, lm1 @ [m]) (s', 
+  Oc # Oc\<up> m @ Bk # <rev lm1> @ Bk # Bk # ires, [Bk]) ires"
+apply(simp only: inv_on_left_moving_in_middle_B.simps)
+apply(rule_tac x = "lm1 @ [m ]" in exI, rule_tac x = "[]" in exI, simp)
+apply(simp add: tape_of_nl_cons split: if_splits)
+done
+
+lemma [simp]: "dec_left_move (as, am) (s, l, Bk # r) ires
+       \<Longrightarrow> inv_on_left_moving (as, am) (s', tl l, hd l # Bk # r) ires"
+apply(auto simp: dec_left_move.simps inv_on_left_moving.simps split: if_splits)
+done
+
+(*
+lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm1 @ [m]) 
+                        (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # <rev lm1> @ Bk\<^bsup>ln\<^esup>, [Bk])  ires"
+apply(auto simp: inv_on_left_moving_in_middle_B.simps)
+apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "[]" in exI, auto)
+done
+*)
+
+lemma [simp]: "dec_left_move (as, am) (s, l, []) ires
+             \<Longrightarrow> inv_on_left_moving (as, am) (s', tl l, [hd l]) ires"
+apply(auto simp: dec_left_move.simps inv_on_left_moving.simps split: if_splits)
+done
+
+lemma [simp]: "dec_after_write (as, am) (s, l, Oc # r) ires
+       \<Longrightarrow> dec_on_right_moving (as, am) (s', Oc # l, r) ires"
+apply(auto simp: dec_after_write.simps dec_on_right_moving.simps)
+apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI, 
+      rule_tac x = "hd lm2" in exI, simp)
+apply(rule_tac x = "Suc 0" in exI,rule_tac x =  "Suc (hd lm2)" in exI)
+apply(case_tac lm2, auto split: if_splits simp: tape_of_nl_cons)
+done
+
+lemma [simp]: "dec_after_write (as, am) (s, l, Bk # r) ires
+       \<Longrightarrow> dec_after_write (as, am) (s', l, Oc # r) ires"
+apply(auto simp: dec_after_write.simps)
+done
+
+lemma [simp]: "dec_after_write (as, am) (s, aaa, []) ires
+             \<Longrightarrow> dec_after_write (as, am) (s', aaa, [Oc]) ires"
+apply(auto simp: dec_after_write.simps)
+done
+
+lemma [simp]: "dec_on_right_moving (as, am) (s, l, Oc # r) ires
+       \<Longrightarrow> dec_on_right_moving (as, am) (s', Oc # l, r) ires"
+apply(simp only: dec_on_right_moving.simps)
+apply(erule_tac exE)+
+apply(erule conjE)+
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,
+      rule_tac x = "m" in exI, rule_tac x = "Suc ml" in exI, 
+      rule_tac x = "mr - 1" in exI, simp)
+apply(case_tac mr, auto)
+done
+
+lemma [simp]: "dec_on_right_moving (as, am) (s, l, r) ires\<Longrightarrow>  l \<noteq> []"
+apply(auto simp: dec_on_right_moving.simps split: if_splits)
+done
+
+lemma [simp]: "dec_on_right_moving (as, am) (s, l, Bk # r) ires
+      \<Longrightarrow>  dec_after_clear (as, am) (s', tl l, hd l # Bk # r) ires"
+apply(auto simp: dec_on_right_moving.simps dec_after_clear.simps)
+apply(case_tac [!] mr, auto split: if_splits)
+done
+
+lemma [simp]: "dec_on_right_moving (as, am) (s, l, []) ires
+             \<Longrightarrow> dec_after_clear (as, am) (s', tl l, [hd l]) ires"
+apply(auto simp: dec_on_right_moving.simps dec_after_clear.simps)
+apply(simp_all split: if_splits)
+apply(rule_tac x = lm1 in exI, simp)
+done
+
+lemma [simp]: 
+ "inv_stop (as, abc_lm_s am n (abc_lm_v am n)) (s, l, r) ires \<Longrightarrow> l \<noteq> []"
+apply(auto simp: inv_stop.simps)
+done
+
+lemma dec_false_1[simp]:
+ "\<lbrakk>abc_lm_v am n = 0; inv_locate_b (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
+  \<Longrightarrow> False"
+apply(auto simp: inv_locate_b.simps in_middle.simps)
+apply(case_tac "length lm1 \<ge> length am", auto)
+apply(subgoal_tac "lm2 = []", simp, subgoal_tac "m = 0", simp)
+apply(case_tac mr, auto simp: )
+apply(subgoal_tac "Suc (length lm1) - length am = 
+                   Suc (length lm1 - length am)", 
+      simp add: exp_ind del: replicate.simps, simp)
+apply(drule_tac xs = "am @ replicate (Suc (length lm1) - length am) 0"
+                and ys = "lm1 @ m # lm2" in length_equal, simp)
+apply(case_tac mr, auto simp: abc_lm_v.simps)
+apply(case_tac "mr = 0", simp_all split: if_splits)
+apply(subgoal_tac "Suc (length lm1) - length am = 
+                       Suc (length lm1 - length am)", 
+      simp add: exp_ind del: replicate.simps, simp)
+done
+
+lemma [simp]: 
+ "\<lbrakk>inv_locate_b (as, am) (n, aaa, Bk # xs) ires; 
+   abc_lm_v am n = 0\<rbrakk>
+   \<Longrightarrow> inv_on_left_moving (as, abc_lm_s am n 0) 
+                         (s, tl aaa, hd aaa # Bk # xs) ires"
+apply(simp add: inv_on_left_moving.simps)
+apply(simp only: inv_locate_b.simps in_middle.simps) 
+apply(erule_tac exE)+
+apply(simp add: inv_on_left_moving.simps)
+apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B 
+         (as, abc_lm_s am n 0) (s, tl aaa, hd aaa # Bk # xs) ires", simp)
+apply(simp only: inv_on_left_moving_norm.simps)
+apply(erule_tac conjE)+
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
+      rule_tac x =  m in exI, rule_tac x = m in exI, 
+      rule_tac x = "Suc 0" in exI, simp add: abc_lm_s.simps)
+apply(case_tac mr, simp_all, auto simp: abc_lm_v.simps)
+apply(simp only: exp_ind[THEN sym] replicate_Suc Nat.Suc_diff_le)
+apply(auto simp: inv_on_left_moving_in_middle_B.simps split: if_splits)
+done
+
+
+lemma [simp]:
+ "\<lbrakk>abc_lm_v am n = 0; inv_locate_b (as, am) (n, aaa, []) ires\<rbrakk>
+   \<Longrightarrow> inv_on_left_moving (as, abc_lm_s am n 0) (s, tl aaa, [hd aaa]) ires"
+apply(simp add: inv_on_left_moving.simps)
+apply(simp only: inv_locate_b.simps in_middle.simps) 
+apply(erule_tac exE)+
+apply(simp add: inv_on_left_moving.simps)
+apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B 
+         (as, abc_lm_s am n 0) (s, tl aaa, [hd aaa]) ires", simp)
+apply(simp only: inv_on_left_moving_norm.simps)
+apply(erule_tac conjE)+
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
+      rule_tac x =  m in exI, rule_tac x = m in exI, 
+      rule_tac x = "Suc 0" in exI, simp add: abc_lm_s.simps)
+apply(case_tac mr, simp_all, auto simp: abc_lm_v.simps)
+apply(simp_all only: exp_ind Nat.Suc_diff_le del: replicate_Suc, simp_all)
+apply(auto simp: inv_on_left_moving_in_middle_B.simps split: if_splits)
+apply(case_tac [!] m, simp_all)
+done
+
+lemma [simp]: "\<lbrakk>am ! n = (0::nat); n < length am\<rbrakk> \<Longrightarrow> am[n := 0] = am"
+apply(simp add: list_update_same_conv)
+done
+
+lemma  [intro]: "\<lbrakk>abc_lm_v (a # list) 0 = 0\<rbrakk> \<Longrightarrow> a = 0"
+apply(simp add: abc_lm_v.simps split: if_splits)
+done
+
+lemma [simp]: 
+ "inv_stop (as, abc_lm_s am n 0) 
+          (start_of (layout_of aprog) e, aaa, Oc # xs) ires
+  \<Longrightarrow> inv_locate_a (as, abc_lm_s am n 0) (0, aaa, Oc # xs) ires"
+apply(simp add: inv_locate_a.simps)
+apply(rule disjI1)
+apply(auto simp: inv_stop.simps at_begin_norm.simps)
+done
+
+lemma [simp]: 
+ "\<lbrakk>inv_stop (as, abc_lm_s am n 0) 
+          (start_of (layout_of aprog) e, aaa, Oc # xs) ires\<rbrakk>
+  \<Longrightarrow> inv_locate_b (as, am) (0, Oc # aaa, xs) ires \<or> 
+      inv_locate_b (as, abc_lm_s am n 0) (0, Oc # aaa, xs) ires"
+apply(simp)
+done
+
+lemma dec_false2: 
+ "inv_stop (as, abc_lm_s am n 0) 
+  (start_of (layout_of aprog) e, aaa, Bk # xs) ires = False"
+apply(auto simp: inv_stop.simps abc_lm_s.simps)
+apply(case_tac [!] am, auto)
+apply(case_tac [!] n, auto simp: tape_of_nl_cons split: if_splits)
+done
+
+lemma dec_false3:
+   "inv_stop (as, abc_lm_s am n 0) 
+              (start_of (layout_of aprog) e, aaa, []) ires = False"
+apply(auto simp: inv_stop.simps abc_lm_s.simps)
+done
+
+lemma [simp]:
+  "fetch (ci (layout_of aprog) 
+       (start_of (layout_of aprog) as) (Dec n e)) 0 b = (Nop, 0)"
+by(simp add: fetch.simps)
+
+declare dec_inv_1.simps[simp del]
+
+declare inv_locate_n_b.simps [simp del]
+
+lemma [simp]:
+  "\<lbrakk>0 < abc_lm_v am n; 0 < n; 
+    at_begin_fst_bwtn (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
+ \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires"
+apply(simp add: at_begin_fst_bwtn.simps inv_locate_n_b.simps )
+done
+ 
+lemma Suc_minus:"length am + tn = n
+       \<Longrightarrow> Suc tn = Suc n - length am "
+apply(arith)
+done
+
+lemma [simp]: 
+ "\<lbrakk>0 < abc_lm_v am n; 0 < n; 
+   at_begin_fst_awtn (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
+ \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires"
+apply(simp only: at_begin_fst_awtn.simps inv_locate_n_b.simps )
+apply(erule exE)+
+apply(erule conjE)+
+apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, 
+      rule_tac x = "Suc tn" in exI, rule_tac x = 0 in exI)
+apply(simp add: exp_ind del: replicate.simps)
+apply(rule conjI)+
+apply(auto)
+done
+
+lemma [simp]:
+ "\<lbrakk>inv_locate_n_b (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
+ \<Longrightarrow> dec_first_on_right_moving n (as, abc_lm_s am n (abc_lm_v am n))  
+                                      (s, Oc # aaa, xs) ires"
+apply(auto simp: inv_locate_n_b.simps dec_first_on_right_moving.simps 
+                 abc_lm_s.simps abc_lm_v.simps)
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
+      rule_tac x = m in exI, simp)
+apply(rule_tac x = "Suc (Suc 0)" in exI, 
+      rule_tac x = "m - 1" in exI, simp)
+apply(case_tac m, auto)
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
+      rule_tac x = m in exI, 
+      simp add: Suc_diff_le exp_ind del: replicate.simps)
+apply(rule_tac x = "Suc (Suc 0)" in exI, 
+      rule_tac x = "m - 1" in exI, simp)
+apply(case_tac m, auto)
+apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, 
+      rule_tac x = m in exI, simp)
+apply(rule_tac x = "Suc (Suc 0)" in exI, 
+      rule_tac x = "m - 1" in exI, simp)
+apply(case_tac m, auto)
+apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
+      rule_tac x = m in exI, 
+      simp add: Suc_diff_le exp_ind del: replicate.simps, simp)
+done
+
+lemma [simp]: "inv_on_left_moving (as, am) (s, [], r) ires 
+  = False"
+apply(simp add: inv_on_left_moving.simps inv_on_left_moving_norm.simps
+                inv_on_left_moving_in_middle_B.simps)
+done
+
+lemma [simp]: 
+  "inv_check_left_moving (as, abc_lm_s am n 0)
+  (start_of (layout_of aprog) as + 2 * n + 14, [], Oc # xs) ires
+ = False"
+apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps)
+done
+
+lemma [simp]: "inv_check_left_moving (as, abc_lm_s lm n (abc_lm_v lm n)) (s, [], Oc # list) ires = False"
+apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps)
+done
+
+lemma [elim]: "\<lbrakk>abc_fetch as ap = Some (Dec n e);
+                start_of (layout_of ap) as < start_of (layout_of ap) e; 
+                start_of (layout_of ap) e \<le> Suc (start_of (layout_of ap) as + 2 * n)\<rbrakk>
+       \<Longrightarrow> RR"
+  using start_of_less[of e as "layout_of ap"] start_of_ge[of as ap n e "layout_of ap"]
+apply(case_tac "as < e", simp)
+apply(case_tac "as = e", simp, simp)
+done
+
+lemma crsp_step_dec_b_e_pre':
+  assumes layout: "ly = layout_of ap"
+  and inv_start: "inv_locate_b (as, lm) (n, la, ra) ires"
+  and fetch: "abc_fetch as ap = Some (Dec n e)"
+  and dec_0: "abc_lm_v lm n = 0"
+  and f: "f = (\<lambda> stp. (steps (Suc (start_of ly as) + 2 * n, la, ra) (ci ly (start_of ly as) (Dec n e), 
+            start_of ly as - Suc 0) stp, start_of ly as, n))"
+  and P: "P = (\<lambda> ((s, l, r), ss, x). s = start_of ly e)"
+  and Q: "Q = (\<lambda> ((s, l, r), ss, x). dec_inv_1 ly x e (as, lm) (s, l, r) ires)"
+  shows "\<exists> stp. P (f stp) \<and> Q (f stp)"
+proof(rule_tac LE = abc_dec_1_LE in halt_lemma2)
+  show "wf abc_dec_1_LE" by(intro wf_dec_le)
+next
+  show "Q (f 0)"
+    using layout fetch
+    apply(simp add: f steps.simps Q dec_inv_1.simps)
+    apply(subgoal_tac "e > as \<or> e = as \<or> e < as")
+    apply(auto simp: Let_def start_of_ge start_of_less inv_start)
+    done
+next
+  show "\<not> P (f 0)"
+    using layout fetch
+    apply(simp add: f steps.simps P)
+    done
+next
+  show "\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> abc_dec_1_LE"
+    using fetch
+  proof(rule_tac allI, rule_tac impI)
+    fix na
+    assume "\<not> P (f na) \<and> Q (f na)"
+    thus "Q (f (Suc na)) \<and> (f (Suc na), f na) \<in> abc_dec_1_LE"
+      apply(simp add: f)
+      apply(case_tac "steps (Suc (start_of ly as + 2 * n), la, ra)
+        (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) na", simp)
+    proof -
+      fix a b c 
+      assume "\<not> P ((a, b, c), start_of ly as, n) \<and> Q ((a, b, c), start_of ly as, n)"
+      thus "Q (step (a, b, c) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0), start_of ly as, n) \<and>
+               ((step (a, b, c) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0), start_of ly as, n), 
+                   (a, b, c), start_of ly as, n) \<in> abc_dec_1_LE"
+        apply(simp add: Q)
+        apply(case_tac c, case_tac [2] aa)
+        apply(simp_all add: dec_inv_1.simps Let_def split: if_splits)
+        using fetch layout dec_0
+        apply(auto simp: step.simps P dec_inv_1.simps Let_def abc_dec_1_LE_def lex_triple_def lex_pair_def)
+        using dec_0
+        apply(drule_tac dec_false_1, simp_all)
+        done
+    qed
+  qed
+qed
+
+lemma crsp_step_dec_b_e_pre:
+  assumes "ly = layout_of ap"
+  and inv_start: "inv_locate_b (as, lm) (n, la, ra) ires"
+  and dec_0: "abc_lm_v lm n  = 0"
+  and fetch: "abc_fetch as ap = Some (Dec n e)"
+  shows "\<exists>stp lb rb.
+       steps (Suc (start_of ly as) + 2 * n, la, ra) (ci ly (start_of ly as) (Dec n e), 
+       start_of ly as - Suc 0) stp = (start_of ly e, lb, rb) \<and>
+       dec_inv_1 ly n e (as, lm) (start_of ly e, lb, rb) ires"
+  using assms
+  apply(drule_tac crsp_step_dec_b_e_pre', auto)
+  apply(rule_tac x = stp in exI, simp)
+  done
+
+lemma [simp]:
+  "\<lbrakk>abc_lm_v lm n = 0;
+  inv_stop (as, abc_lm_s lm n (abc_lm_v lm n)) (start_of ly e, lb, rb) ires\<rbrakk>
+  \<Longrightarrow> crsp ly (abc_step_l (as, lm) (Some (Dec n e))) (start_of ly e, lb, rb) ires"
+apply(auto simp: crsp.simps abc_step_l.simps inv_stop.simps)
+done
+
+lemma crsp_step_dec_b_e:
+  assumes layout: "ly = layout_of ap"
+  and inv_start: "inv_locate_a (as, lm) (n, l, r) ires"
+  and dec_0: "abc_lm_v lm n = 0"
+  and fetch: "abc_fetch as ap = Some (Dec n e)"
+  shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
+  (steps (start_of ly as + 2 * n, l, r) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires"
+proof -
+  let ?P = "ci ly (start_of ly as) (Dec n e)"
+  let ?off = "start_of ly as - Suc 0"
+  have "\<exists> stp la ra. steps (start_of ly as + 2 * n, l, r) (?P, ?off) stp = (Suc (start_of ly as) + 2*n, la, ra)
+             \<and>  inv_locate_b (as, lm) (n, la, ra) ires"
+    using inv_start
+    apply(case_tac "r = [] \<or> hd r = Bk", simp_all)
+    done
+  from this obtain stpa la ra where a:
+    "steps (start_of ly as + 2 * n, l, r) (?P, ?off) stpa = (Suc (start_of ly as) + 2*n, la, ra)
+             \<and>  inv_locate_b (as, lm) (n, la, ra) ires" by blast
+  term dec_inv_1
+  have "\<exists> stp lb rb. steps (Suc (start_of ly as) + 2 * n, la, ra) (?P, ?off) stp = (start_of ly e, lb, rb)
+             \<and>  dec_inv_1 ly n e (as, lm) (start_of ly e, lb, rb) ires"
+    using assms a
+    apply(rule_tac crsp_step_dec_b_e_pre, auto)
+    done
+  from this obtain stpb lb rb where b:
+    "steps (Suc (start_of ly as) + 2 * n, la, ra) (?P, ?off) stpb = (start_of ly e, lb, rb)
+             \<and>  dec_inv_1 ly n e (as, lm) (start_of ly e, lb, rb) ires"  by blast
+  from a b show "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e))) 
+    (steps (start_of ly as + 2 * n, l, r) (?P, ?off) stp) ires"
+    apply(rule_tac x = "stpa + stpb" in exI)
+    apply(simp add: steps_add)
+    using dec_0
+    apply(simp add: dec_inv_1.simps)
+    apply(case_tac stpa, simp_all add: steps.simps)
+    done
+qed    
+  
+fun dec_inv_2 :: "layout \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> dec_inv_t"
+  where
+  "dec_inv_2 ly n e (as, am) (s, l, r) ires =
+           (let ss = start_of ly as in
+            let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in
+            let am'' = abc_lm_s am n (abc_lm_v am n) in
+              if s = 0 then False
+              else if s = ss + 2 * n then 
+                      inv_locate_a (as, am) (n, l, r) ires
+              else if s = ss + 2 * n + 1 then 
+                      inv_locate_n_b (as, am) (n, l, r) ires
+              else if s = ss + 2 * n + 2 then 
+                      dec_first_on_right_moving n (as, am'') (s, l, r) ires
+              else if s = ss + 2 * n + 3 then 
+                      dec_after_clear (as, am') (s, l, r) ires
+              else if s = ss + 2 * n + 4 then 
+                      dec_right_move (as, am') (s, l, r) ires
+              else if s = ss + 2 * n + 5 then 
+                      dec_check_right_move (as, am') (s, l, r) ires
+              else if s = ss + 2 * n + 6 then 
+                      dec_left_move (as, am') (s, l, r) ires
+              else if s = ss + 2 * n + 7 then 
+                      dec_after_write (as, am') (s, l, r) ires
+              else if s = ss + 2 * n + 8 then 
+                      dec_on_right_moving (as, am') (s, l, r) ires
+              else if s = ss + 2 * n + 9 then 
+                      dec_after_clear (as, am') (s, l, r) ires
+              else if s = ss + 2 * n + 10 then 
+                      inv_on_left_moving (as, am') (s, l, r) ires
+              else if s = ss + 2 * n + 11 then 
+                      inv_check_left_moving (as, am') (s, l, r) ires
+              else if s = ss + 2 * n + 12 then 
+                      inv_after_left_moving (as, am') (s, l, r) ires
+              else if s = ss + 2 * n + 16 then 
+                      inv_stop (as, am') (s, l, r) ires
+              else False)"
+
+declare dec_inv_2.simps[simp del]
+fun abc_dec_2_stage1 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "abc_dec_2_stage1 (s, l, r) ss n = 
+              (if s \<le> ss + 2*n + 1 then 7
+               else if s = ss + 2*n + 2 then 6 
+               else if s = ss + 2*n + 3 then 5
+               else if s \<ge> ss + 2*n + 4 \<and> s \<le> ss + 2*n + 9 then 4
+               else if s = ss + 2*n + 6 then 3
+               else if s = ss + 2*n + 10 \<or> s = ss + 2*n + 11 then 2
+               else if s = ss + 2*n + 12 then 1
+               else 0)"
+
+fun abc_dec_2_stage2 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "abc_dec_2_stage2 (s, l, r) ss n = 
+       (if s \<le> ss + 2 * n + 1 then (ss + 2 * n + 16 - s)
+        else if s = ss + 2*n + 10 then length l
+        else if s = ss + 2*n + 11 then length l
+        else if s = ss + 2*n + 4 then length r - 1
+        else if s = ss + 2*n + 5 then length r 
+        else if s = ss + 2*n + 7 then length r - 1
+        else if s = ss + 2*n + 8 then  
+              length r + length (takeWhile (\<lambda> a. a = Oc) l) - 1
+        else if s = ss + 2*n + 9 then 
+              length r + length (takeWhile (\<lambda> a. a = Oc) l) - 1
+        else 0)"
+
+fun abc_dec_2_stage3 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "abc_dec_2_stage3 (s, l, r) ss n  =
+        (if s \<le> ss + 2*n + 1 then 
+            if (s - ss) mod 2 = 0 then if r \<noteq> [] \<and> 
+                                          hd r = Oc then 0 else 1  
+            else length r
+         else if s = ss + 2 * n + 10 then 
+             if  r \<noteq> [] \<and> hd r = Oc then 2
+             else 1
+         else if s = ss + 2 * n + 11 then 
+             if r \<noteq> [] \<and> hd r = Oc then 3 
+             else 0 
+         else (ss + 2 * n + 16 - s))"
+
+fun abc_dec_2_stage4 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "abc_dec_2_stage4 (s, l, r) ss n = 
+          (if s = ss + 2*n + 2 then length r
+           else if s = ss + 2*n + 8 then length r
+           else if s = ss + 2*n + 3 then 
+               if r \<noteq> [] \<and> hd r = Oc then 1
+               else 0
+           else if s = ss + 2*n + 7 then 
+               if r \<noteq> [] \<and> hd r = Oc then 0 
+               else 1
+           else if s = ss + 2*n + 9 then 
+               if r \<noteq> [] \<and> hd r = Oc then 1
+               else 0 
+           else 0)"
+
+fun abc_dec_2_measure :: "(config \<times> nat \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat \<times> nat)"
+  where
+  "abc_dec_2_measure (c, ss, n) = 
+  (abc_dec_2_stage1 c ss n, 
+  abc_dec_2_stage2 c ss n, abc_dec_2_stage3 c ss n,  abc_dec_2_stage4 c ss n)"
+
+definition lex_square:: 
+   "((nat \<times> nat \<times> nat \<times> nat) \<times> (nat \<times> nat \<times> nat \<times> nat)) set"
+  where "lex_square \<equiv> less_than <*lex*> lex_triple"
+
+definition abc_dec_2_LE ::
+  "((config \<times> nat \<times>
+  nat) \<times> (config \<times> nat \<times> nat)) set"
+  where "abc_dec_2_LE \<equiv> (inv_image lex_square abc_dec_2_measure)"
+
+lemma wf_dec2_le: "wf abc_dec_2_LE"
+by(auto intro:wf_inv_image simp:abc_dec_2_LE_def lex_square_def lex_triple_def lex_pair_def)
+
+lemma fix_add: "fetch ap ((x::nat) + 2*n) b = fetch ap (2*n + x) b"
+by (metis Suc_1 mult_2 nat_add_commute)
+
+lemma [elim]: 
+ "\<lbrakk>0 < abc_lm_v am n; inv_locate_n_b (as, am) (n, aaa, Bk # xs) ires\<rbrakk> 
+ \<Longrightarrow> RR"
+apply(auto simp: inv_locate_n_b.simps abc_lm_v.simps split: if_splits)
+apply(case_tac [!] m, auto)
+done
+ 
+lemma [elim]:
+ "\<lbrakk>0 < abc_lm_v am n; inv_locate_n_b (as, am) 
+                                (n, aaa, []) ires\<rbrakk> \<Longrightarrow> RR"
+apply(auto simp: inv_locate_n_b.simps abc_lm_v.simps split: if_splits)
+done
+
+lemma [simp]: "dec_after_write (as, am) (s, aa, r) ires
+           \<Longrightarrow> takeWhile (\<lambda>a. a = Oc) aa = []"
+apply(simp only : dec_after_write.simps)
+apply(erule exE)+
+apply(erule_tac conjE)+
+apply(case_tac aa, simp)
+apply(case_tac a, simp only: takeWhile.simps , simp_all split: if_splits)
+done
+
+lemma [simp]: 
+     "\<lbrakk>dec_on_right_moving (as, lm) (s, aa, []) ires; 
+       length (takeWhile (\<lambda>a. a = Oc) (tl aa)) 
+           \<noteq> length (takeWhile (\<lambda>a. a = Oc) aa) - Suc 0\<rbrakk>
+    \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (tl aa)) < 
+                       length (takeWhile (\<lambda>a. a = Oc) aa) - Suc 0"
+apply(simp only: dec_on_right_moving.simps)
+apply(erule_tac exE)+
+apply(erule_tac conjE)+
+apply(case_tac mr, auto split: if_splits)
+done
+
+lemma [simp]: 
+  "dec_after_clear (as, abc_lm_s am n (abc_lm_v am n - Suc 0)) 
+             (start_of (layout_of aprog) as + 2 * n + 9, aa, Bk # xs) ires
+ \<Longrightarrow> length xs - Suc 0 < length xs + 
+                             length (takeWhile (\<lambda>a. a = Oc) aa)"
+apply(simp only: dec_after_clear.simps)
+apply(erule_tac exE)+
+apply(erule conjE)+
+apply(simp split: if_splits )
+done
+
+lemma [simp]: 
+ "\<lbrakk>dec_after_clear (as, abc_lm_s am n (abc_lm_v am n - Suc 0))
+       (start_of (layout_of aprog) as + 2 * n + 9, aa, []) ires\<rbrakk>
+    \<Longrightarrow> Suc 0 < length (takeWhile (\<lambda>a. a = Oc) aa)"
+apply(simp add: dec_after_clear.simps split: if_splits)
+done
+
+lemma [elim]: 
+  "inv_check_left_moving (as, lm)
+  (s, [], Oc # xs) ires
+ \<Longrightarrow> RR"
+apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps)
+done
+
+lemma [simp]:
+"\<lbrakk>0 < abc_lm_v am n; 
+  at_begin_norm (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
+  \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires"
+apply(simp only: at_begin_norm.simps inv_locate_n_b.simps)
+apply(erule_tac exE)+
+apply(rule_tac x = lm1 in exI, simp)
+apply(case_tac "length lm2", simp)
+apply(case_tac "lm2", simp, simp)
+apply(case_tac "lm2", auto simp: tape_of_nl_cons split: if_splits)
+done
+
+lemma [simp]: 
+ "\<lbrakk>0 < abc_lm_v am n; 
+   at_begin_fst_awtn (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
+ \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires"
+apply(simp only: at_begin_fst_awtn.simps inv_locate_n_b.simps )
+apply(erule exE)+
+apply(erule conjE)+
+apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, 
+      rule_tac x = "Suc tn" in exI, rule_tac x = 0 in exI)
+apply(simp add: exp_ind del: replicate.simps)
+apply(rule conjI)+
+apply(auto)
+done
+
+lemma [simp]: 
+ "\<lbrakk>0 < abc_lm_v am n; inv_locate_a (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
+ \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc#aaa, xs) ires"
+apply(auto simp: inv_locate_a.simps at_begin_fst_bwtn.simps)
+done
+
+lemma [simp]: 
+ "\<lbrakk>dec_on_right_moving (as, am) (s, aa, Bk # xs) ires; 
+   Suc (length (takeWhile (\<lambda>a. a = Oc) (tl aa)))
+   \<noteq> length (takeWhile (\<lambda>a. a = Oc) aa)\<rbrakk>
+  \<Longrightarrow> Suc (length (takeWhile (\<lambda>a. a = Oc) (tl aa))) 
+    < length (takeWhile (\<lambda>a. a = Oc) aa)"
+apply(simp only: dec_on_right_moving.simps)
+apply(erule exE)+
+apply(erule conjE)+
+apply(case_tac ml, auto split: if_splits )
+done
+
+lemma crsp_step_dec_b_suc_pre:
+  assumes layout: "ly = layout_of ap"
+  and crsp: "crsp ly (as, lm) (s, l, r) ires"
+  and inv_start: "inv_locate_a (as, lm) (n, la, ra) ires"
+  and fetch: "abc_fetch as ap = Some (Dec n e)"
+  and dec_suc: "0 < abc_lm_v lm n"
+  and f: "f = (\<lambda> stp. (steps (start_of ly as + 2 * n, la, ra) (ci ly (start_of ly as) (Dec n e), 
+            start_of ly as - Suc 0) stp, start_of ly as, n))"
+  and P: "P = (\<lambda> ((s, l, r), ss, x). s = start_of ly as + 2*n + 16)"
+  and Q: "Q = (\<lambda> ((s, l, r), ss, x). dec_inv_2 ly x e (as, lm) (s, l, r) ires)"
+  shows "\<exists> stp. P (f stp) \<and> Q(f stp)"
+  proof(rule_tac LE = abc_dec_2_LE in halt_lemma2)
+  show "wf abc_dec_2_LE" by(intro wf_dec2_le)
+next
+  show "Q (f 0)"
+    using layout fetch inv_start
+    apply(simp add: f steps.simps Q)
+    apply(simp only: dec_inv_2.simps)
+    apply(auto simp: Let_def start_of_ge start_of_less inv_start dec_inv_2.simps)
+    done
+next
+  show "\<not> P (f 0)"
+    using layout fetch
+    apply(simp add: f steps.simps P)
+    done
+next
+  show "\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> abc_dec_2_LE"
+    using fetch
+  proof(rule_tac allI, rule_tac impI)
+    fix na
+    assume "\<not> P (f na) \<and> Q (f na)"
+    thus "Q (f (Suc na)) \<and> (f (Suc na), f na) \<in> abc_dec_2_LE"
+      apply(simp add: f)
+      apply(case_tac "steps ((start_of ly as + 2 * n), la, ra)
+        (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) na", simp)
+    proof -
+      fix a b c 
+      assume "\<not> P ((a, b, c), start_of ly as, n) \<and> Q ((a, b, c), start_of ly as, n)"
+      thus "Q (step (a, b, c) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0), start_of ly as, n) \<and>
+               ((step (a, b, c) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0), start_of ly as, n), 
+                   (a, b, c), start_of ly as, n) \<in> abc_dec_2_LE"
+        apply(simp add: Q)
+        apply(erule_tac conjE)
+        apply(case_tac c, case_tac [2] aa)
+        apply(simp_all add: dec_inv_2.simps Let_def)
+        apply(simp_all split: if_splits)
+        using fetch layout dec_suc
+        apply(auto simp: step.simps P dec_inv_2.simps Let_def abc_dec_2_LE_def lex_triple_def lex_pair_def lex_square_def
+                         fix_add numeral_3_eq_3) 
+        done
+    qed
+  qed
+qed
+
+lemma [simp]: 
+  "\<lbrakk>inv_stop (as, abc_lm_s lm n (abc_lm_v lm n - Suc 0)) 
+  (start_of (layout_of ap) as + 2 * n + 16, a, b) ires;
+   abc_lm_v lm n > 0;
+   abc_fetch as ap = Some (Dec n e)\<rbrakk>
+  \<Longrightarrow> crsp (layout_of ap) (abc_step_l (as, lm) (Some (Dec n e))) 
+  (start_of (layout_of ap) as + 2 * n + 16, a, b) ires"
+apply(auto simp: inv_stop.simps crsp.simps  abc_step_l.simps startof_Suc2)
+apply(drule_tac startof_Suc2, simp)
+done
+  
+lemma crsp_step_dec_b_suc:
+  assumes layout: "ly = layout_of ap"
+  and crsp: "crsp ly (as, lm) (s, l, r) ires"
+  and inv_start: "inv_locate_a (as, lm) (n, la, ra) ires"
+  and fetch: "abc_fetch as ap = Some (Dec n e)"
+  and dec_suc: "0 < abc_lm_v lm n"
+  shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
+              (steps (start_of ly as + 2 * n, la, ra) (ci (layout_of ap) 
+                  (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires"
+  using assms
+  apply(drule_tac crsp_step_dec_b_suc_pre, auto)
+  apply(rule_tac x = stp in exI, simp)
+  apply(simp add: dec_inv_2.simps)
+  apply(case_tac stp, simp_all add: steps.simps)
+  done
+
+lemma crsp_step_dec_b:
+  assumes layout: "ly = layout_of ap"
+  and crsp: "crsp ly (as, lm) (s, l, r) ires"
+  and inv_start: "inv_locate_a (as, lm) (n, la, ra) ires"
+  and fetch: "abc_fetch as ap = Some (Dec n e)"
+  shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
+  (steps (start_of ly as + 2 * n, la, ra) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires"
+using assms
+apply(case_tac "abc_lm_v lm n = 0")
+apply(rule_tac crsp_step_dec_b_e, simp_all)
+apply(rule_tac crsp_step_dec_b_suc, simp_all)
+done
+
+lemma crsp_step_dec: 
+  assumes layout: "ly = layout_of ap"
+  and crsp: "crsp ly (as, lm) (s, l, r) ires"
+  and fetch: "abc_fetch as ap = Some (Dec n e)"
+  shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
+  (steps (s, l, r) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires"
+proof(simp add: ci.simps)
+  let ?off = "start_of ly as - Suc 0"
+  let ?A = "findnth n"
+  let ?B = "sete (shift (shift tdec_b (2 * n)) ?off) (start_of ly e)"
+  have "\<exists> stp la ra. steps (s, l, r) (shift ?A ?off @ ?B, ?off) stp = (start_of ly as + 2*n, la, ra)
+                    \<and> inv_locate_a (as, lm) (n, la, ra) ires"
+  proof -
+    have "\<exists>stp l' r'. steps (Suc 0, l, r) (?A, 0) stp = (Suc (2 * n), l', r') \<and> 
+                     inv_locate_a (as, lm) (n, l', r') ires"
+      using assms
+      apply(rule_tac findnth_correct, simp_all)
+      done
+    then obtain stp l' r' where a: 
+      "steps (Suc 0, l, r) (?A, 0) stp = (Suc (2 * n), l', r') \<and> 
+      inv_locate_a (as, lm) (n, l', r') ires" by blast
+    then have "steps (Suc 0 + ?off, l, r) (shift ?A ?off, ?off) stp = (Suc (2 * n) + ?off, l', r')"
+      apply(rule_tac tm_shift_eq_steps, simp_all)
+      done
+    moreover have "s = start_of ly as"
+      using crsp
+      apply(auto simp: crsp.simps)
+      done
+    ultimately show "\<exists> stp la ra. steps (s, l, r) (shift ?A ?off @ ?B, ?off) stp = (start_of ly as + 2*n, la, ra)
+                    \<and> inv_locate_a (as, lm) (n, la, ra) ires"
+      using a
+      apply(drule_tac B = ?B in tm_append_first_steps_eq, auto)
+      apply(rule_tac x = stp in exI, simp)
+      done
+  qed
+  from this obtain stpa la ra where a: 
+    "steps (s, l, r) (shift ?A ?off @ ?B, ?off) stpa = (start_of ly as + 2*n, la, ra)
+                    \<and> inv_locate_a (as, lm) (n, la, ra) ires" by blast
+  have "\<exists>stp. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
+           (steps (start_of ly as + 2*n, la, ra) (shift ?A ?off @ ?B, ?off) stp) ires \<and> stp > 0"
+    using assms a
+    apply(drule_tac crsp_step_dec_b, auto)
+    apply(rule_tac x = stp in exI, simp add: ci.simps)
+    done
+  then obtain stpb where b: 
+    "crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
+    (steps (start_of ly as + 2*n, la, ra) (shift ?A ?off @ ?B, ?off) stpb) ires \<and> stpb > 0" ..
+  from a b show "\<exists> stp>0. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
+    (steps (s, l, r) (shift ?A ?off @ ?B, ?off) stp) ires"
+    apply(rule_tac x = "stpa + stpb" in exI)
+    apply(simp add: steps_add)
+    done
+qed    
+  
+subsection{*Crsp of Goto*}
+
+lemma crsp_step_goto:
+  assumes layout: "ly = layout_of ap"
+  and crsp: "crsp ly (as, lm) (s, l, r) ires"
+  shows "\<exists>stp>0. crsp ly (abc_step_l (as, lm) (Some (Goto n)))
+  (steps (s, l, r) (ci ly (start_of ly as) (Goto n), 
+            start_of ly as - Suc 0) stp) ires"
+using crsp
+apply(rule_tac x = "Suc 0" in exI)
+apply(case_tac r, case_tac [2] a)
+apply(simp_all add: ci.simps steps.simps step.simps crsp.simps fetch.simps
+  crsp.simps abc_step_l.simps)
+done
+
+lemma crsp_step_in:
+  assumes layout: "ly = layout_of ap"
+  and compile: "tp = tm_of ap"
+  and crsp: "crsp ly (as, lm) (s, l, r) ires"
+  and fetch: "abc_fetch as ap = Some ins"
+  shows "\<exists> stp>0. crsp ly (abc_step_l (as, lm) (Some ins))
+                      (steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp) ires"
+  using assms
+  apply(case_tac ins, simp_all)
+  apply(rule crsp_step_inc, simp_all)
+  apply(rule crsp_step_dec, simp_all)
+  apply(rule_tac crsp_step_goto, simp_all)
+  done
+
+lemma crsp_step:
+  assumes layout: "ly = layout_of ap"
+  and compile: "tp = tm_of ap"
+  and crsp: "crsp ly (as, lm) (s, l, r) ires"
+  and fetch: "abc_fetch as ap = Some ins"
+  shows "\<exists> stp>0. crsp ly (abc_step_l (as, lm) (Some ins))
+                      (steps (s, l, r) (tp, 0) stp) ires"
+proof -
+  have "\<exists> stp>0. crsp ly (abc_step_l (as, lm) (Some ins))
+                      (steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp) ires"
+    using assms
+    apply(rule_tac crsp_step_in, simp_all)
+    done
+  from this obtain stp where d: "stp > 0 \<and> crsp ly (abc_step_l (as, lm) (Some ins))
+                      (steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp) ires" ..
+  obtain s' l' r' where e:
+    "(steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp) = (s', l', r')"
+    apply(case_tac "(steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp)")
+    by blast
+  then have "steps (s, l, r) (tp, 0) stp = (s', l', r')"
+    using assms d
+    apply(rule_tac steps_eq_in)
+    apply(simp_all)
+    apply(case_tac "(abc_step_l (as, lm) (Some ins))", simp add: crsp.simps)
+    done    
+  thus " \<exists>stp>0. crsp ly (abc_step_l (as, lm) (Some ins)) (steps (s, l, r) (tp, 0) stp) ires"
+    using d e
+    apply(rule_tac x = stp in exI, simp)
+    done
+qed
+
+lemma crsp_steps:
+  assumes layout: "ly = layout_of ap"
+  and compile: "tp = tm_of ap"
+  and crsp: "crsp ly (as, lm) (s, l, r) ires"
+  shows "\<exists> stp. crsp ly (abc_steps_l (as, lm) ap n)
+                      (steps (s, l, r) (tp, 0) stp) ires"
+(*
+proof(induct n)
+  case 0
+  have "crsp ly (abc_steps_l (as, lm) ap 0) (steps (s, l, r) (tp, 0) 0) ires"
+    using crsp by(simp add: steps.simps abc_steps_l.simps)
+  thus "?case"
+    by(rule_tac x = 0 in exI, simp)
+next
+  case (Suc n)
+  obtain as' lm' where a: "abc_steps_l (as, lm) ap n = (as', lm')"
+    by(case_tac "abc_steps_l (as, lm) ap n", auto) 
+  have "\<exists>stp\<ge>n. crsp ly (abc_steps_l (as, lm) ap n) (steps (s, l, r) (tp, 0) stp) ires"
+    by fact
+  from this a obtain stpa where b:
+    "stpa\<ge>n \<and> crsp ly (as', lm') (steps (s, l, r) (tp, 0) stpa) ires" by auto
+  obtain s' l' r' where "steps (s, l, r) (tp, 0) stpa = (s', l', r')"
+    by(case_tac "steps (s, l, r) (tp, 0) stpa")
+  then have "stpa\<ge>n \<and> crsp ly (as', lm') (s', l', r') ires" using b by simp
+  from a and this show "?case"
+  proof(cases "abc_fetch as' ap")
+    case None
+    
+  
+
+    have "crsp ly (abc_steps_l (as, lm) ap 0) (steps (s, l, r) (tp, 0) stp) ires"
+    apply(simp add: steps.simps abc_steps_l.simps)
+*)
+  using crsp
+  apply(induct n)
+  apply(rule_tac x = 0 in exI) 
+  apply(simp add: steps.simps abc_steps_l.simps, simp)
+  apply(case_tac "(abc_steps_l (as, lm) ap n)", auto)
+  apply(frule_tac abc_step_red, simp)
+  apply(case_tac "abc_fetch a ap", simp add: abc_step_l.simps, auto)
+  apply(case_tac "steps (s, l, r) (tp, 0) stp", simp)
+  using assms
+  apply(drule_tac s = ab and l = ba and r = c in crsp_step, auto)
+  apply(rule_tac x = "stp + stpa" in exI, simp add: steps_add)
+  done
+
+lemma tp_correct': 
+  assumes layout: "ly = layout_of ap"
+  and compile: "tp = tm_of ap"
+  and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires"
+  and abc_halt: "abc_steps_l (0, lm) ap stp = (length ap, am)"
+  shows "\<exists> stp k. steps (Suc 0, l, r) (tp, 0) stp = (start_of ly (length ap), Bk # Bk # ires, <am> @ Bk\<up>k)"
+  using assms
+  apply(drule_tac n = stp in crsp_steps, auto)
+  apply(rule_tac x = stpa in exI)
+  apply(case_tac "steps (Suc 0, l, r) (tm_of ap, 0) stpa", simp add: crsp.simps)
+  done
+
+text{*The tp @ [(Nop, 0), (Nop, 0)] is nomoral turing machines, so we can use Hoare_plus when composing with Mop machine*}
+
+thm layout_of.simps
+lemma layout_id_cons: "layout_of (ap @ [p]) = layout_of ap @ [length_of p]"
+apply(simp add: layout_of.simps)
+done
+
+lemma [simp]: "length (layout_of xs) = length xs"
+by(simp add: layout_of.simps)
+
+thm tms_of.simps
+term ci
+thm tms_of.simps
+thm tpairs_of.simps
+
+lemma [simp]:  
+  "map (start_of (layout_of xs @ [length_of x])) [0..<length xs] =  (map (start_of (layout_of xs)) [0..<length xs])"
+apply(auto)
+apply(simp add: layout_of.simps start_of.simps)
+done
+
+lemma tpairs_id_cons: 
+  "tpairs_of (xs @ [x]) = tpairs_of xs @ [(start_of (layout_of (xs @ [x])) (length xs), x)]"
+apply(auto simp: tpairs_of.simps layout_id_cons )
+done
+
+lemma map_length_ci:
+  "(map (length \<circ> (\<lambda>(xa, y). ci (layout_of xs @ [length_of x]) xa y)) (tpairs_of xs)) = 
+  (map (length \<circ> (\<lambda>(x, y). ci (layout_of xs) x y)) (tpairs_of xs)) "
+apply(auto)
+apply(case_tac b, auto simp: ci.simps sete.simps)
+done
+
+lemma length_tp'[simp]: 
+  "\<lbrakk>ly = layout_of ap; tp = tm_of ap\<rbrakk> \<Longrightarrow>
+       length tp = 2 * listsum (take (length ap) (layout_of ap))"
+proof(induct ap arbitrary: ly tp rule: rev_induct)
+  case Nil
+  thus "?case"
+    by(simp add: tms_of.simps tm_of.simps tpairs_of.simps)
+next
+  fix x xs ly tp
+  assume ind: "\<And>ly tp. \<lbrakk>ly = layout_of xs; tp = tm_of xs\<rbrakk> \<Longrightarrow> 
+    length tp = 2 * listsum (take (length xs) (layout_of xs))"
+  and layout: "ly = layout_of (xs @ [x])"
+  and tp: "tp = tm_of (xs @ [x])"
+  obtain ly' where a: "ly' = layout_of xs"
+    by metis
+  obtain tp' where b: "tp' = tm_of xs"
+    by metis
+  have c: "length tp' = 2 * listsum (take (length xs) (layout_of xs))"
+    using a b
+    by(erule_tac ind, simp)
+  thus "length tp = 2 * 
+    listsum (take (length (xs @ [x])) (layout_of (xs @ [x])))"
+    using tp b
+    apply(auto simp: layout_id_cons tm_of.simps tms_of.simps length_concat tpairs_id_cons map_length_ci)
+    apply(case_tac x)
+    apply(auto simp: ci.simps tinc_b_def tdec_b_def length_findnth sete.simps length_of.simps
+                 split: abc_inst.splits)
+    done
+qed
+
+lemma [simp]: 
+  "\<lbrakk>ly = layout_of ap; tp = tm_of ap\<rbrakk> \<Longrightarrow>
+        fetch (tp @ [(Nop, 0), (Nop, 0)]) (start_of ly (length ap)) b = 
+       (Nop, 0)"
+apply(case_tac b)
+apply(simp_all add: start_of.simps fetch.simps nth_append)
+done
+(*
+lemma tp_correct: 
+  assumes layout: "ly = layout_of ap"
+  and compile: "tp = tm_of ap"
+  and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires"
+  and abc_halt: "abc_steps_l (0, lm) ap stp = (length ap, am)"
+  shows "\<exists> stp k. steps (Suc 0, l, r) (tp @ [(Nop, 0), (Nop, 0)], 0) stp = (0, Bk # Bk # ires, <am> @ Bk\<up>k)"
+  using assms
+proof -
+  have "\<exists> stp k. steps (Suc 0, l, r) (tp @ [(Nop, 0), (Nop, 0)], 0) stp =
+    (start_of ly (length ap), Bk # Bk # ires, <am> @ Bk\<up>k)"
+  proof -
+    have "\<exists> stp k. steps (Suc 0, l, r) (tp, 0) stp =
+    (start_of ly (length ap), Bk # Bk # ires, <am> @ Bk\<up>k)"
+      using assms
+      apply(rule_tac tp_correct', simp_all)
+      done
+    from this obtain stp k where "steps (Suc 0, l, r) (tp, 0) stp =
+    (start_of ly (length ap), Bk # Bk # ires, <am> @ Bk\<up>k)" by blast
+    thus "?thesis"
+      apply(rule_tac x = stp in exI, rule_tac x = k in exI)
+      apply(drule_tac tm_append_first_steps_eq, simp_all)
+      done
+  qed
+  from this obtain stp k where 
+    "steps (Suc 0, l, r) (tp @ [(Nop, 0), (Nop, 0)], 0) stp =
+    (start_of ly (length ap), Bk # Bk # ires, <am> @ Bk\<up>k)"
+    by blast
+  thus "\<exists>stp k. steps (Suc 0, l, r) (tp @ [(Nop, 0), (Nop, 0)], 0) stp 
+    = (0, Bk # Bk # ires, <am> @ Bk \<up> k)"
+    using assms
+    apply(rule_tac x = "stp + Suc 0" in exI)
+    apply(simp add: steps_add)
+    apply(auto simp: step.simps)
+    done
+qed
+ *)   
+(********for mopup***********)
+fun mopup_a :: "nat \<Rightarrow> instr list"
+  where
+  "mopup_a 0 = []" |
+  "mopup_a (Suc n) = mopup_a n @ 
+       [(R, 2*n + 3), (W0, 2*n + 2), (R, 2*n + 1), (W1, 2*n + 2)]"
+
+definition mopup_b :: "instr list"
+  where
+  "mopup_b \<equiv> [(R, 2), (R, 1), (L, 5), (W0, 3), (R, 4), (W0, 3),
+            (R, 2), (W0, 3), (L, 5), (L, 6), (R, 0), (L, 6)]"
+
+fun mopup :: "nat \<Rightarrow> instr list"
+  where 
+  "mopup n = mopup_a n @ shift mopup_b (2*n)"
+(****)
+
+type_synonym mopup_type = "config \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> cell list \<Rightarrow> bool"
+
+fun mopup_stop :: "mopup_type"
+  where
+  "mopup_stop (s, l, r) lm n ires= 
+        (\<exists> ln rn. l = Bk\<up>ln @ Bk # Bk # ires \<and> r = <abc_lm_v lm n> @ Bk\<up>rn)"
+
+fun mopup_bef_erase_a :: "mopup_type"
+  where
+  "mopup_bef_erase_a (s, l, r) lm n ires= 
+         (\<exists> ln m rn. l = Bk\<up>ln @ Bk # Bk # ires \<and> 
+                  r = Oc\<up>m@ Bk # <(drop ((s + 1) div 2) lm)> @ Bk\<up>rn)"
+
+fun mopup_bef_erase_b :: "mopup_type"
+  where
+  "mopup_bef_erase_b (s, l, r) lm n ires = 
+      (\<exists> ln m rn. l = Bk\<up>ln @ Bk # Bk # ires \<and> r = Bk # Oc\<up>m @ Bk # 
+                                      <(drop (s div 2) lm)> @ Bk\<up>rn)"
+
+fun mopup_jump_over1 :: "mopup_type"
+  where
+  "mopup_jump_over1 (s, l, r) lm n ires = 
+      (\<exists> ln m1 m2 rn. m1 + m2 = Suc (abc_lm_v lm n) \<and> 
+        l = Oc\<up>m1 @ Bk\<up>ln @ Bk # Bk # ires \<and> 
+     (r = Oc\<up>m2 @ Bk # <(drop (Suc n) lm)> @ Bk\<up>rn \<or> 
+     (r = Oc\<up>m2 \<and> (drop (Suc n) lm) = [])))"
+
+fun mopup_aft_erase_a :: "mopup_type"
+  where
+  "mopup_aft_erase_a (s, l, r) lm n ires = 
+      (\<exists> lnl lnr rn (ml::nat list) m. 
+          m = Suc (abc_lm_v lm n) \<and> l = Bk\<up>lnr @ Oc\<up>m @ Bk\<up>lnl @ Bk # Bk # ires \<and> 
+                                   (r = <ml> @ Bk\<up>rn))"
+
+fun mopup_aft_erase_b :: "mopup_type"
+  where
+  "mopup_aft_erase_b (s, l, r) lm n ires= 
+   (\<exists> lnl lnr rn (ml::nat list) m. 
+      m = Suc (abc_lm_v lm n) \<and> 
+      l = Bk\<up>lnr @ Oc\<up>m @ Bk\<up>lnl @ Bk # Bk # ires \<and> 
+     (r = Bk # <ml> @ Bk\<up>rn \<or>
+      r = Bk # Bk # <ml> @ Bk\<up>rn))"
+
+fun mopup_aft_erase_c :: "mopup_type"
+  where
+  "mopup_aft_erase_c (s, l, r) lm n ires = 
+ (\<exists> lnl lnr rn (ml::nat list) m. 
+     m = Suc (abc_lm_v lm n) \<and> 
+     l = Bk\<up>lnr @ Oc\<up>m @ Bk\<up>lnl @ Bk # Bk # ires \<and> 
+    (r = <ml> @ Bk\<up>rn \<or> r = Bk # <ml> @ Bk\<up>rn))"
+
+fun mopup_left_moving :: "mopup_type"
+  where
+  "mopup_left_moving (s, l, r) lm n ires = 
+  (\<exists> lnl lnr rn m.
+     m = Suc (abc_lm_v lm n) \<and> 
+   ((l = Bk\<up>lnr @ Oc\<up>m @ Bk\<up>lnl @ Bk # Bk # ires \<and> r = Bk\<up>rn) \<or>
+    (l = Oc\<up>(m - 1) @ Bk\<up>lnl @ Bk # Bk # ires \<and> r = Oc # Bk\<up>rn)))"
+
+fun mopup_jump_over2 :: "mopup_type"
+  where
+  "mopup_jump_over2 (s, l, r) lm n ires = 
+     (\<exists> ln rn m1 m2.
+          m1 + m2 = Suc (abc_lm_v lm n) 
+        \<and> r \<noteq> [] 
+        \<and> (hd r = Oc \<longrightarrow> (l = Oc\<up>m1 @ Bk\<up>ln @ Bk # Bk # ires \<and> r = Oc\<up>m2 @ Bk\<up>rn)) 
+        \<and> (hd r = Bk \<longrightarrow> (l = Bk\<up>ln @ Bk # ires \<and> r = Bk # Oc\<up>(m1+m2)@ Bk\<up>rn)))"
+
+
+fun mopup_inv :: "mopup_type"
+  where
+  "mopup_inv (s, l, r) lm n ires = 
+      (if s = 0 then mopup_stop (s, l, r) lm n ires
+       else if s \<le> 2*n then
+               if s mod 2 = 1 then mopup_bef_erase_a (s, l, r) lm n ires
+                   else mopup_bef_erase_b (s, l, r) lm n ires
+            else if s = 2*n + 1 then 
+                mopup_jump_over1 (s, l, r) lm n ires
+            else if s = 2*n + 2 then mopup_aft_erase_a (s, l, r) lm n ires
+            else if s = 2*n + 3 then mopup_aft_erase_b (s, l, r) lm n ires
+            else if s = 2*n + 4 then mopup_aft_erase_c (s, l, r) lm n ires
+            else if s = 2*n + 5 then mopup_left_moving (s, l, r) lm n ires
+            else if s = 2*n + 6 then mopup_jump_over2 (s, l, r) lm n ires
+            else False)"
+
+lemma mopup_fetch_0[simp]: 
+     "(fetch (mopup_a n @ shift mopup_b (2 * n)) 0 b) = (Nop, 0)"
+by(simp add: fetch.simps)
+
+lemma mop_bef_length[simp]: "length (mopup_a n) = 4 * n"
+apply(induct n, simp_all add: mopup_a.simps)
+done
+
+lemma mopup_a_nth: 
+  "\<lbrakk>q < n; x < 4\<rbrakk> \<Longrightarrow> mopup_a n ! (4 * q + x) = 
+                             mopup_a (Suc q) ! ((4 * q) + x)"
+apply(induct n, simp)
+apply(case_tac "q < n", simp add: mopup_a.simps, auto)
+apply(simp add: nth_append)
+apply(subgoal_tac "q = n", simp)
+apply(arith)
+done
+
+lemma fetch_bef_erase_a_o[simp]: 
+ "\<lbrakk>0 < s; s \<le> 2 * n; s mod 2 = Suc 0\<rbrakk>
+  \<Longrightarrow> (fetch (mopup_a n @ shift mopup_b (2 * n)) s Oc) = (W0, s + 1)"
+apply(subgoal_tac "\<exists> q. s = 2*q + 1", auto)
+apply(subgoal_tac "length (mopup_a n) = 4*n")
+apply(auto simp: fetch.simps nth_of.simps nth_append)
+apply(subgoal_tac "mopup_a n ! (4 * q + 1) = 
+                      mopup_a (Suc q) ! ((4 * q) + 1)", 
+      simp add: mopup_a.simps nth_append)
+apply(rule mopup_a_nth, auto)
+apply arith
+done
+
+lemma fetch_bef_erase_a_b[simp]:
+  "\<lbrakk>0 < s; s \<le> 2 * n; s mod 2 = Suc 0\<rbrakk>
+   \<Longrightarrow>  (fetch (mopup_a n @ shift mopup_b (2 * n)) s Bk) = (R, s + 2)"
+apply(subgoal_tac "\<exists> q. s = 2*q + 1", auto)
+apply(subgoal_tac "length (mopup_a n) = 4*n")
+apply(auto simp: fetch.simps nth_of.simps nth_append)
+apply(subgoal_tac "mopup_a n ! (4 * q + 0) = 
+                       mopup_a (Suc q) ! ((4 * q + 0))", 
+      simp add: mopup_a.simps nth_append)
+apply(rule mopup_a_nth, auto)
+apply arith
+done
+
+lemma fetch_bef_erase_b_b: 
+  "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = 0\<rbrakk> \<Longrightarrow> 
+     (fetch (mopup_a n @ shift mopup_b (2 * n)) s Bk) = (R, s - 1)"
+apply(subgoal_tac "\<exists> q. s = 2 * q", auto)
+apply(case_tac qa, simp, simp)
+apply(auto simp: fetch.simps nth_of.simps nth_append)
+apply(subgoal_tac "mopup_a n ! (4 * nat + 2) = 
+                     mopup_a (Suc nat) ! ((4 * nat) + 2)", 
+      simp add: mopup_a.simps nth_append)
+apply(rule mopup_a_nth, auto)
+done
+
+lemma fetch_jump_over1_o: 
+ "fetch (mopup_a n @ shift mopup_b (2 * n)) (Suc (2 * n)) Oc
+  = (R, Suc (2 * n))"
+apply(subgoal_tac "length (mopup_a n) = 4 * n")
+apply(auto simp: fetch.simps nth_of.simps mopup_b_def nth_append 
+                 shift.simps)
+done
+
+lemma fetch_jump_over1_b: 
+ "fetch (mopup_a n @ shift mopup_b (2 * n)) (Suc (2 * n)) Bk 
+ = (R, Suc (Suc (2 * n)))"
+apply(subgoal_tac "length (mopup_a n) = 4 * n")
+apply(auto simp: fetch.simps nth_of.simps mopup_b_def 
+                 nth_append shift.simps)
+done
+
+lemma fetch_aft_erase_a_o: 
+ "fetch (mopup_a n @ shift mopup_b (2 * n)) (Suc (Suc (2 * n))) Oc 
+ = (W0, Suc (2 * n + 2))"
+apply(subgoal_tac "length (mopup_a n) = 4 * n")
+apply(auto simp: fetch.simps nth_of.simps mopup_b_def 
+                 nth_append shift.simps)
+done
+
+lemma fetch_aft_erase_a_b: 
+ "fetch (mopup_a n @ shift mopup_b (2 * n)) (Suc (Suc (2 * n))) Bk
+  = (L, Suc (2 * n + 4))"
+apply(subgoal_tac "length (mopup_a n) = 4 * n")
+apply(auto simp: fetch.simps nth_of.simps mopup_b_def 
+                 nth_append shift.simps)
+done
+
+lemma fetch_aft_erase_b_b: 
+ "fetch (mopup_a n @ shift mopup_b (2 * n)) (2*n + 3) Bk
+  = (R, Suc (2 * n + 3))"
+apply(subgoal_tac "length (mopup_a n) = 4 * n")
+apply(subgoal_tac "2*n + 3 = Suc (2*n + 2)", simp only: fetch.simps)
+apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps)
+done
+
+lemma fetch_aft_erase_c_o: 
+ "fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 4) Oc 
+ = (W0, Suc (2 * n + 2))"
+apply(subgoal_tac "length (mopup_a n) = 4 * n")
+apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps)
+apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps)
+done
+
+lemma fetch_aft_erase_c_b: 
+ "fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 4) Bk 
+ = (R, Suc (2 * n + 1))"
+apply(subgoal_tac "length (mopup_a n) = 4 * n")
+apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps)
+apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps)
+done
+
+lemma fetch_left_moving_o: 
+ "(fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 5) Oc) 
+ = (L, 2*n + 6)"
+apply(subgoal_tac "length (mopup_a n) = 4 * n")
+apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps)
+apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps)
+done
+
+lemma fetch_left_moving_b: 
+ "(fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 5) Bk)
+  = (L, 2*n + 5)"
+apply(subgoal_tac "length (mopup_a n) = 4 * n")
+apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps)
+apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps)
+done
+
+lemma fetch_jump_over2_b:
+  "(fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 6) Bk) 
+ = (R, 0)"
+apply(subgoal_tac "length (mopup_a n) = 4 * n")
+apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps)
+apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps)
+done
+
+lemma fetch_jump_over2_o: 
+"(fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 6) Oc) 
+ = (L, 2*n + 6)"
+apply(subgoal_tac "length (mopup_a n) = 4 * n")
+apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps)
+apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps)
+done
+
+lemmas mopupfetchs = 
+fetch_bef_erase_a_o fetch_bef_erase_a_b fetch_bef_erase_b_b 
+fetch_jump_over1_o fetch_jump_over1_b fetch_aft_erase_a_o 
+fetch_aft_erase_a_b fetch_aft_erase_b_b fetch_aft_erase_c_o 
+fetch_aft_erase_c_b fetch_left_moving_o fetch_left_moving_b 
+fetch_jump_over2_b fetch_jump_over2_o
+
+declare 
+  mopup_jump_over2.simps[simp del] mopup_left_moving.simps[simp del]
+  mopup_aft_erase_c.simps[simp del] mopup_aft_erase_b.simps[simp del] 
+  mopup_aft_erase_a.simps[simp del] mopup_jump_over1.simps[simp del]
+  mopup_bef_erase_a.simps[simp del] mopup_bef_erase_b.simps[simp del]
+  mopup_stop.simps[simp del]
+
+lemma [simp]: 
+  "\<lbrakk>mopup_bef_erase_a (s, l, Oc # xs) lm n ires\<rbrakk> \<Longrightarrow> 
+  mopup_bef_erase_b (Suc s, l, Bk # xs) lm n ires"
+apply(auto simp: mopup_bef_erase_a.simps mopup_bef_erase_b.simps )
+apply(rule_tac x = "m - 1" in exI, rule_tac x = rn in exI)
+apply(case_tac m, simp, simp)
+done
+
+lemma mopup_false1:
+  "\<lbrakk>0 < s; s \<le> 2 * n; s mod 2 = Suc 0;  \<not> Suc s \<le> 2 * n\<rbrakk> 
+  \<Longrightarrow> RR"
+apply(arith)
+done
+
+lemma [simp]: 
+ "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = Suc 0; 
+   mopup_bef_erase_a (s, l, Oc # xs) lm n ires; r = Oc # xs\<rbrakk>
+ \<Longrightarrow> (Suc s \<le> 2 * n \<longrightarrow> mopup_bef_erase_b (Suc s, l, Bk # xs) lm n ires)  \<and>
+     (\<not> Suc s \<le> 2 * n \<longrightarrow> mopup_jump_over1 (Suc s, l, Bk # xs) lm n ires) "
+apply(auto elim: mopup_false1)
+done
+
+lemma drop_tape_of_cons: 
+  "\<lbrakk>Suc q < length lm; x = lm ! q\<rbrakk> \<Longrightarrow> <drop q lm> = Oc # Oc \<up> x @ Bk # <drop (Suc q) lm>"
+by (metis Suc_lessD append_Cons list.simps(2) nth_drop' replicate_Suc tape_of_nl_cons)
+
+lemma erase2jumpover1:
+  "\<lbrakk>q < length list; 
+             \<forall>rn. <drop q list> \<noteq> Oc # Oc \<up> abc_lm_v (a # list) (Suc q) @ Bk # <drop (Suc q) list> @ Bk \<up> rn\<rbrakk>
+       \<Longrightarrow> <drop q list> = Oc # Oc \<up> abc_lm_v (a # list) (Suc q)"
+apply(erule_tac x = 0 in allE, simp)
+apply(case_tac "Suc q < length list")
+apply(erule_tac notE)
+apply(rule_tac drop_tape_of_cons, simp_all add: abc_lm_v.simps)
+apply(subgoal_tac "length list = Suc q", auto)
+apply(subgoal_tac "drop q list = [list ! q]")
+apply(simp add: tape_of_nl_abv tape_of_nat_abv)
+by (metis append_Nil2 append_eq_conv_conj drop_Suc_conv_tl lessI)
+
+lemma erase2jumpover2:
+  "\<lbrakk>q < length list; \<forall>rn. <drop q list> @ Bk # Bk \<up> n \<noteq>
+  Oc # Oc \<up> abc_lm_v (a # list) (Suc q) @ Bk # <drop (Suc q) list> @ Bk \<up> rn\<rbrakk>
+  \<Longrightarrow> RR"
+apply(case_tac "Suc q < length list")
+apply(erule_tac x = "Suc n" in allE, simp)
+apply(erule_tac notE)
+apply(rule_tac drop_tape_of_cons, simp_all add: abc_lm_v.simps)
+apply(subgoal_tac "length list = Suc q", auto)
+apply(erule_tac x = "n" in allE, simp add: tape_of_nl_abv)
+by (metis append_Nil2 append_eq_conv_conj drop_Suc_conv_tl lessI replicate_Suc tape_of_nl_abv tape_of_nl_cons)
+
+lemma mopup_bef_erase_a_2_jump_over[simp]: 
+ "\<lbrakk>n < length lm; 0 < s; s mod 2 = Suc 0;  s \<le> 2 * n;
+   mopup_bef_erase_a (s, l, Bk # xs) lm n ires; \<not> (Suc (Suc s) \<le> 2 * n)\<rbrakk> 
+\<Longrightarrow> mopup_jump_over1 (s', Bk # l, xs) lm n ires"
+apply(auto simp: mopup_bef_erase_a.simps mopup_jump_over1.simps)
+apply(case_tac m, auto simp: mod_ex1)
+apply(subgoal_tac "n = Suc q", auto)
+apply(rule_tac x = "Suc ln" in exI, rule_tac x = 0 in exI, auto)
+apply(case_tac [!] lm, simp_all)
+apply(case_tac [!] rn, auto elim: erase2jumpover1 erase2jumpover2)
+apply(erule_tac x = 0 in allE, simp)
+apply(rule_tac classical, simp)
+apply(erule_tac notE)
+apply(rule_tac drop_tape_of_cons, simp_all add: abc_lm_v.simps)
+done
+
+lemma Suc_Suc_div:  "\<lbrakk>0 < s; s mod 2 = Suc 0; Suc (Suc s) \<le> 2 * n\<rbrakk>
+           \<Longrightarrow> (Suc (Suc (s div 2))) \<le> n"
+apply(arith)
+done
+
+lemma mopup_bef_erase_a_2_a[simp]: 
+ "\<lbrakk>n < length lm; 0 < s; s mod 2 = Suc 0; 
+   mopup_bef_erase_a (s, l, Bk # xs) lm n ires; 
+   Suc (Suc s) \<le> 2 * n\<rbrakk> \<Longrightarrow> 
+   mopup_bef_erase_a (Suc (Suc s), Bk # l, xs) lm n ires"
+apply(auto simp: mopup_bef_erase_a.simps)
+apply(subgoal_tac "drop (Suc (Suc (s div 2))) lm \<noteq> []")
+apply(case_tac m, simp_all)
+apply(rule_tac x = "Suc (abc_lm_v lm (Suc (s div 2)))" in exI, 
+      rule_tac x = rn in exI, auto simp: mod_ex1)
+apply(rule_tac drop_tape_of_cons)
+apply arith
+apply(simp add: abc_lm_v.simps)
+done
+
+lemma mopup_false2: 
+ "\<lbrakk>0 < s; s \<le> 2 * n; 
+   s mod 2 = Suc 0; Suc s \<noteq> 2 * n;
+   \<not> Suc (Suc s) \<le> 2 * n\<rbrakk> \<Longrightarrow> RR"
+apply(arith)
+done
+
+lemma [simp]: "mopup_bef_erase_a (s, l, []) lm n ires \<Longrightarrow> 
+                        mopup_bef_erase_a (s, l, [Bk]) lm n ires"
+apply(auto simp: mopup_bef_erase_a.simps)
+done
+
+lemma [simp]:
+   "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = Suc 0; \<not> Suc (Suc s) \<le> 2 *n;
+     mopup_bef_erase_a (s, l, []) lm n ires\<rbrakk>
+    \<Longrightarrow>  mopup_jump_over1 (s', Bk # l, []) lm n ires"
+by auto
+
+lemma "mopup_bef_erase_b (s, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []"
+apply(auto simp: mopup_bef_erase_b.simps)
+done
+
+lemma [simp]: "mopup_bef_erase_b (s, l, Oc # xs) lm n ires = False"
+apply(auto simp: mopup_bef_erase_b.simps )
+done
+ 
+lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow> 
+                                      (s - Suc 0) mod 2 = Suc 0"
+apply(arith)
+done
+
+lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow>
+                                       s - Suc 0 \<le> 2 * n"
+apply(simp)
+done
+
+lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow> \<not> s \<le> Suc 0"
+apply(arith)
+done
+
+lemma [simp]: "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; 
+               s mod 2 \<noteq> Suc 0; 
+               mopup_bef_erase_b (s, l, Bk # xs) lm n ires; r = Bk # xs\<rbrakk> 
+           \<Longrightarrow> mopup_bef_erase_a (s - Suc 0, Bk # l, xs) lm n ires"
+apply(auto simp: mopup_bef_erase_b.simps mopup_bef_erase_a.simps)
+done
+
+lemma [simp]: "\<lbrakk>mopup_bef_erase_b (s, l, []) lm n ires\<rbrakk> \<Longrightarrow> 
+                   mopup_bef_erase_a (s - Suc 0, Bk # l, []) lm n ires"
+apply(auto simp: mopup_bef_erase_b.simps mopup_bef_erase_a.simps)
+done
+
+lemma [simp]: 
+   "\<lbrakk>n < length lm;
+    mopup_jump_over1 (Suc (2 * n), l, Oc # xs) lm n ires;
+    r = Oc # xs\<rbrakk>
+  \<Longrightarrow> mopup_jump_over1 (Suc (2 * n), Oc # l, xs) lm n ires"
+apply(auto simp: mopup_jump_over1.simps)
+apply(rule_tac x = ln in exI, rule_tac x = "Suc m1" in exI,
+       rule_tac x = "m2 - 1" in exI, simp)
+apply(case_tac "m2", simp, simp)
+apply(rule_tac x = ln in exI, rule_tac x = "Suc m1" in exI, 
+      rule_tac x = "m2 - 1" in exI)
+apply(case_tac m2, simp, simp)
+done
+
+lemma mopup_jump_over1_2_aft_erase_a[simp]:  
+ "\<lbrakk>n < length lm; mopup_jump_over1 (Suc (2 * n), l, Bk # xs) lm n ires\<rbrakk>
+  \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, xs) lm n ires"
+apply(simp only: mopup_jump_over1.simps mopup_aft_erase_a.simps)
+apply(erule_tac exE)+
+apply(rule_tac x = ln in exI, rule_tac x = "Suc 0" in exI)
+apply(case_tac m2, simp)
+apply(rule_tac x = rn in exI, rule_tac x = "drop (Suc n) lm" in exI, 
+      simp)
+apply(simp)
+done
+
+lemma [simp]: 
+ "\<lbrakk>n < length lm; mopup_jump_over1 (Suc (2 * n), l, []) lm n ires\<rbrakk> \<Longrightarrow> 
+    mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, []) lm n ires"
+apply(rule mopup_jump_over1_2_aft_erase_a, simp)
+apply(auto simp: mopup_jump_over1.simps)
+apply(rule_tac x = ln in exI, rule_tac x = "Suc (abc_lm_v lm n)" in exI, 
+      rule_tac x = 0 in exI, simp add: )
+done
+
+
+lemma [simp]: 
+ "\<lbrakk>n < length lm; 
+   mopup_aft_erase_a (Suc (Suc (2 * n)), l, Oc # xs) lm n ires\<rbrakk> 
+ \<Longrightarrow> mopup_aft_erase_b (Suc (Suc (Suc (2 * n))), l, Bk # xs) lm n ires"
+apply(auto simp: mopup_aft_erase_a.simps mopup_aft_erase_b.simps )
+apply(case_tac ml)
+apply(simp_all add: tape_of_nl_cons split: if_splits)
+apply(case_tac a, simp_all)
+apply(rule_tac x = rn in exI, rule_tac x = "[]" in exI, simp)
+apply(rule_tac x = rn in exI, rule_tac x = "[nat]" in exI, simp)
+apply(case_tac a, simp_all)
+apply(rule_tac x = rn in exI, rule_tac x = "list" in exI, simp)
+apply(rule_tac x = rn in exI)
+apply(rule_tac x = "nat # list" in exI, simp add: tape_of_nl_cons)
+done
+
+lemma [simp]:
+  "mopup_aft_erase_a (Suc (Suc (2 * n)), l, Bk # xs) lm n ires \<Longrightarrow> l \<noteq> []"
+apply(auto simp: mopup_aft_erase_a.simps)
+done
+
+lemma [simp]:
+  "\<lbrakk>n < length lm;
+    mopup_aft_erase_a (Suc (Suc (2 * n)), l, Bk # xs) lm n ires\<rbrakk>
+  \<Longrightarrow> mopup_left_moving (5 + 2 * n, tl l, hd l # Bk # xs) lm n ires"
+apply(simp only: mopup_aft_erase_a.simps mopup_left_moving.simps)
+apply(erule exE)+
+apply(case_tac lnr, simp)
+apply(case_tac ml, simp, simp add: tape_of_nl_cons split: if_splits)
+apply(auto)
+apply(case_tac ml, simp_all add: tape_of_nl_cons split: if_splits)
+apply(rule_tac x = "Suc rn" in exI, simp)
+done
+
+lemma [simp]:
+  "mopup_aft_erase_a (Suc (Suc (2 * n)), l, []) lm n ires \<Longrightarrow> l \<noteq> []"
+apply(simp only: mopup_aft_erase_a.simps)
+apply(erule exE)+
+apply(auto)
+done
+
+lemma [simp]:
+  "\<lbrakk>n < length lm; mopup_aft_erase_a (Suc (Suc (2 * n)), l, []) lm n ires\<rbrakk>
+  \<Longrightarrow> mopup_left_moving (5 + 2 * n, tl l, [hd l]) lm n ires"
+apply(simp only: mopup_aft_erase_a.simps mopup_left_moving.simps)
+apply(erule exE)+
+apply(subgoal_tac "ml = [] \<and> rn = 0", erule conjE, erule conjE, simp)
+apply(case_tac lnr, simp)
+apply(rule_tac x = lnl in exI, simp)
+apply(rule_tac x = 1 in exI, simp)
+apply(case_tac ml, simp, simp)
+done
+
+
+lemma [simp]: "mopup_aft_erase_b (2 * n + 3, l, Oc # xs) lm n ires = False"
+apply(auto simp: mopup_aft_erase_b.simps )
+done
+
+lemma tape_of_ex1[intro]: 
+  "\<exists>rna ml. Oc \<up> a @ Bk \<up> rn = <ml::nat list> @ Bk \<up> rna \<or> Oc \<up> a @ Bk \<up> rn = Bk # <ml> @ Bk \<up> rna"
+apply(case_tac a, simp_all)
+apply(rule_tac x = rn in exI, rule_tac x = "[]" in exI, simp)
+apply(rule_tac x = rn in exI, rule_tac x = "[nat]" in exI, simp)
+done
+
+lemma [intro]: "\<exists>rna ml. Oc \<up> a @ Bk # <list::nat list> @ Bk \<up> rn = 
+  <ml> @ Bk \<up> rna \<or> Oc \<up> a @ Bk # <list> @ Bk \<up> rn = Bk # <ml::nat list> @ Bk \<up> rna"
+apply(case_tac "list = []", simp add: replicate_Suc[THEN sym] del: replicate_Suc)
+apply(rule_tac rn = "Suc rn" in tape_of_ex1)
+apply(case_tac a, simp)
+apply(rule_tac x = rn in exI, rule_tac x = list in exI, simp)
+apply(rule_tac x = rn in exI, rule_tac x = "nat # list" in exI)
+apply(simp add: tape_of_nl_cons)
+done
+
+lemma [simp]: 
+ "\<lbrakk>n < length lm; 
+   mopup_aft_erase_c (2 * n + 4, l, Oc # xs) lm n ires\<rbrakk>
+  \<Longrightarrow> mopup_aft_erase_b (Suc (Suc (Suc (2 * n))), l, Bk # xs) lm n ires"
+apply(auto simp: mopup_aft_erase_c.simps mopup_aft_erase_b.simps )
+apply(case_tac ml, simp_all add: tape_of_nl_cons split: if_splits, auto)
+done
+
+lemma mopup_aft_erase_c_aft_erase_a[simp]: 
+ "\<lbrakk>n < length lm; mopup_aft_erase_c (2 * n + 4, l, Bk # xs) lm n ires\<rbrakk> 
+ \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, xs) lm n ires"
+apply(simp only: mopup_aft_erase_c.simps mopup_aft_erase_a.simps )
+apply(erule_tac exE)+
+apply(erule conjE, erule conjE, erule disjE)
+apply(subgoal_tac "ml = []", simp, case_tac rn, 
+      simp, simp, rule conjI)
+apply(rule_tac x = lnl in exI, rule_tac x = "Suc lnr" in exI, simp)
+apply(rule_tac x = nat in exI, rule_tac x = "[]" in exI, simp)
+apply(case_tac ml, simp, simp add: tape_of_nl_cons split: if_splits)
+apply(rule_tac x = lnl in exI, rule_tac x = "Suc lnr" in exI, simp)
+apply(rule_tac x = rn in exI, rule_tac x = "ml" in exI, simp)
+done
+
+lemma [simp]: 
+ "\<lbrakk>n < length lm; mopup_aft_erase_c (2 * n + 4, l, []) lm n ires\<rbrakk> 
+ \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, []) lm n ires"
+apply(rule mopup_aft_erase_c_aft_erase_a, simp)
+apply(simp only: mopup_aft_erase_c.simps)
+apply(erule exE)+
+apply(rule_tac x = lnl in exI, rule_tac x = lnr in exI, simp add: )
+apply(rule_tac x = 0 in exI, rule_tac x = "[]" in exI, simp)
+done
+
+lemma mopup_aft_erase_b_2_aft_erase_c[simp]:
+  "\<lbrakk>n < length lm; mopup_aft_erase_b (2 * n + 3, l, Bk # xs) lm n ires\<rbrakk>  
+ \<Longrightarrow> mopup_aft_erase_c (4 + 2 * n, Bk # l, xs) lm n ires"
+apply(auto simp: mopup_aft_erase_b.simps mopup_aft_erase_c.simps)
+apply(rule_tac x = "lnl" in exI, rule_tac x = "Suc lnr" in exI, simp)
+apply(rule_tac x = "lnl" in exI, rule_tac x = "Suc lnr" in exI, simp)
+done
+
+lemma [simp]: 
+ "\<lbrakk>n < length lm; mopup_aft_erase_b (2 * n + 3, l, []) lm n ires\<rbrakk> 
+ \<Longrightarrow> mopup_aft_erase_c (4 + 2 * n, Bk # l, []) lm n ires"
+apply(rule_tac mopup_aft_erase_b_2_aft_erase_c, simp)
+apply(simp add: mopup_aft_erase_b.simps)
+done
+
+lemma [simp]: 
+    "mopup_left_moving (2 * n + 5, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []"
+apply(auto simp: mopup_left_moving.simps)
+done
+
+lemma [simp]:  
+ "\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, Oc # xs) lm n ires\<rbrakk>
+  \<Longrightarrow> mopup_jump_over2 (2 * n + 6, tl l, hd l # Oc # xs) lm n ires"
+apply(simp only: mopup_left_moving.simps mopup_jump_over2.simps)
+apply(erule_tac exE)+
+apply(erule conjE, erule disjE, erule conjE)
+apply(case_tac rn, simp, simp add: )
+apply(case_tac "hd l", simp add:  )
+apply(case_tac "abc_lm_v lm n", simp)
+apply(rule_tac x = "lnl" in exI, rule_tac x = rn in exI, 
+      rule_tac x = "Suc 0" in exI, rule_tac x = 0 in exI)
+apply(case_tac lnl, simp, simp, simp add: exp_ind[THEN sym], simp)
+apply(case_tac "abc_lm_v lm n", simp)
+apply(case_tac lnl, simp, simp)
+apply(rule_tac x = lnl in exI, rule_tac x = rn in exI)
+apply(rule_tac x = nat in exI, rule_tac x = "Suc (Suc 0)" in exI, simp)
+done
+
+lemma [simp]: "mopup_left_moving (2 * n + 5, l, xs) lm n ires \<Longrightarrow> l \<noteq> []"
+apply(auto simp: mopup_left_moving.simps)
+done
+
+lemma [simp]:
+  "\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, Bk # xs) lm n ires\<rbrakk> 
+ \<Longrightarrow> mopup_left_moving (2 * n + 5, tl l, hd l # Bk # xs) lm n ires"
+apply(simp only: mopup_left_moving.simps)
+apply(erule exE)+
+apply(case_tac lnr, simp)
+apply(rule_tac x = lnl in exI, rule_tac x = nat in exI, simp)
+apply(rule_tac x = "Suc rn" in exI, simp)
+done
+
+lemma [simp]: 
+"\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, []) lm n ires\<rbrakk>
+    \<Longrightarrow> mopup_left_moving (2 * n + 5, tl l, [hd l]) lm n ires"
+apply(simp only: mopup_left_moving.simps)
+apply(erule exE)+
+apply(case_tac lnr, auto)
+done
+
+
+lemma [simp]: 
+ "mopup_jump_over2 (2 * n + 6, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []"
+apply(auto simp: mopup_jump_over2.simps )
+done
+
+lemma [simp]: 
+"\<lbrakk>n < length lm; mopup_jump_over2 (2 * n + 6, l, Oc # xs) lm n ires\<rbrakk>
+ \<Longrightarrow>  mopup_jump_over2 (2 * n + 6, tl l, hd l # Oc # xs) lm n ires"
+apply(simp only: mopup_jump_over2.simps)
+apply(erule_tac exE)+
+apply(simp add:  , erule conjE, erule_tac conjE)
+apply(case_tac m1, simp)
+apply(rule_tac x = ln in exI, rule_tac x = rn in exI, 
+      rule_tac x = 0 in exI, simp)
+apply(case_tac ln, simp, simp, simp only: exp_ind[THEN sym], simp)
+apply(rule_tac x = ln in exI, rule_tac x = rn in exI, 
+      rule_tac x = nat in exI, rule_tac x = "Suc m2" in exI, simp)
+done
+
+lemma [simp]: 
+ "\<lbrakk>n < length lm; mopup_jump_over2 (2 * n + 6, l, Bk # xs) lm n ires\<rbrakk> 
+  \<Longrightarrow> mopup_stop (0, Bk # l, xs) lm n ires"
+apply(auto simp: mopup_jump_over2.simps mopup_stop.simps)
+apply(simp_all add: tape_of_nat_abv exp_ind[THEN sym])
+done
+
+lemma [simp]: "mopup_jump_over2 (2 * n + 6, l, []) lm n ires = False"
+apply(simp only: mopup_jump_over2.simps, simp)
+done
+
+lemma mopup_inv_step:
+  "\<lbrakk>n < length lm; mopup_inv (s, l, r) lm n ires\<rbrakk>
+  \<Longrightarrow> mopup_inv (step (s, l, r) (mopup_a n @ shift mopup_b (2 * n), 0)) lm n ires"
+apply(case_tac r, case_tac [2] a)
+apply(auto split:if_splits simp add:step.simps)
+apply(simp_all add: mopupfetchs)
+done
+
+declare mopup_inv.simps[simp del]
+lemma mopup_inv_steps: 
+"\<lbrakk>n < length lm; mopup_inv (s, l, r) lm n ires\<rbrakk> \<Longrightarrow> 
+     mopup_inv (steps (s, l, r) (mopup_a n @ shift mopup_b (2 * n), 0)  stp) lm n ires"
+apply(induct_tac stp, simp add: steps.simps)
+apply(simp add: step_red)
+apply(case_tac "steps (s, l, r) 
+                (mopup_a n @ shift mopup_b (2 * n), 0) na", simp)
+apply(rule_tac mopup_inv_step, simp, simp)
+done
+
+fun abc_mopup_stage1 :: "config \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "abc_mopup_stage1 (s, l, r) n = 
+           (if s > 0 \<and> s \<le> 2*n then 6
+            else if s = 2*n + 1 then 4
+            else if s \<ge> 2*n + 2 \<and> s \<le> 2*n + 4 then 3
+            else if s = 2*n + 5 then 2
+            else if s = 2*n + 6 then 1
+            else 0)"
+
+fun abc_mopup_stage2 :: "config \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "abc_mopup_stage2 (s, l, r) n = 
+           (if s > 0 \<and> s \<le> 2*n then length r
+            else if s = 2*n + 1 then length r
+            else if s = 2*n + 5 then length l
+            else if s = 2*n + 6 then length l
+            else if s \<ge> 2*n + 2 \<and> s \<le> 2*n + 4 then length r
+            else 0)"
+
+fun abc_mopup_stage3 :: "config \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "abc_mopup_stage3 (s, l, r) n = 
+          (if s > 0 \<and> s \<le> 2*n then 
+              if hd r = Bk then 0
+              else 1
+           else if s = 2*n + 2 then 1 
+           else if s = 2*n + 3 then 0
+           else if s = 2*n + 4 then 2
+           else 0)"
+
+fun abc_mopup_measure :: "(config \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)"
+  where
+  "abc_mopup_measure (c, n) = 
+    (abc_mopup_stage1 c n, abc_mopup_stage2 c n, 
+                                       abc_mopup_stage3 c n)"
+
+definition abc_mopup_LE ::
+   "(((nat \<times> cell list \<times> cell list) \<times> nat) \<times> 
+    ((nat \<times> cell list \<times> cell list) \<times> nat)) set"
+  where
+  "abc_mopup_LE \<equiv> (inv_image lex_triple abc_mopup_measure)"
+
+lemma wf_abc_mopup_le[intro]: "wf abc_mopup_LE"
+by(auto intro:wf_inv_image simp:abc_mopup_LE_def lex_triple_def lex_pair_def)
+
+lemma [simp]: "mopup_bef_erase_a (a, aa, []) lm n ires = False"
+apply(auto simp: mopup_bef_erase_a.simps)
+done
+
+lemma [simp]: "mopup_bef_erase_b (a, aa, []) lm n ires = False"
+apply(auto simp: mopup_bef_erase_b.simps) 
+done
+
+lemma [simp]: "mopup_aft_erase_b (2 * n + 3, aa, []) lm n ires = False"
+apply(auto simp: mopup_aft_erase_b.simps)
+done
+
+declare mopup_inv.simps[simp del]
+term mopup_inv
+
+lemma [simp]: 
+  "\<lbrakk>0 < q; q \<le> n\<rbrakk> \<Longrightarrow> 
+     (fetch (mopup_a n @ shift mopup_b (2 * n)) (2*q) Bk) = (R, 2*q - 1)"
+apply(case_tac q, simp, simp)
+apply(auto simp: fetch.simps nth_of.simps nth_append)
+apply(subgoal_tac "mopup_a n ! (4 * nat + 2) = 
+                     mopup_a (Suc nat) ! ((4 * nat) + 2)", 
+      simp add: mopup_a.simps nth_append)
+apply(rule mopup_a_nth, auto)
+done
+
+(* FIXME: is also in uncomputable *)
+lemma halt_lemma: 
+  "\<lbrakk>wf LE; \<forall>n. (\<not> P (f n) \<longrightarrow> (f (Suc n), (f n)) \<in> LE)\<rbrakk> \<Longrightarrow> \<exists>n. P (f n)"
+by (metis wf_iff_no_infinite_down_chain)
+
+
+lemma mopup_halt:
+  assumes 
+  less: "n < length lm"
+  and inv: "mopup_inv (Suc 0, l, r) lm n ires"
+  and f: "f = (\<lambda> stp. (steps (Suc 0, l, r) (mopup_a n @ shift mopup_b (2 * n), 0) stp, n))"
+  and P: "P = (\<lambda> (c, n). is_final c)"
+  shows "\<exists> stp. P (f stp)"
+proof(rule_tac LE = abc_mopup_LE in halt_lemma)
+  show "wf abc_mopup_LE" by(auto)
+next
+  show "\<forall>n. \<not> P (f n) \<longrightarrow> (f (Suc n), f n) \<in> abc_mopup_LE"
+  proof(rule_tac allI, rule_tac impI)
+    fix na
+    assume h: "\<not> P (f na)"
+    show "(f (Suc na), f na) \<in> abc_mopup_LE"
+    proof(simp add: f)
+      obtain a b c where g:"steps (Suc 0, l, r) (mopup_a n @ shift mopup_b (2 * n), 0) na = (a, b, c)"
+        apply(case_tac "steps (Suc 0, l, r) (mopup_a n @ shift mopup_b (2 * n), 0) na", auto)
+        done
+      then have "mopup_inv (a, b, c) lm n ires"
+        thm mopup_inv_steps
+        using inv less mopup_inv_steps[of n lm "Suc 0" l r ires na]
+        apply(simp)
+        done
+      moreover have "a > 0"
+        using h g
+        apply(simp add: f P)
+        done
+      ultimately have "((step (a, b, c) (mopup_a n @ shift mopup_b (2 * n), 0), n), (a, b, c), n) \<in> abc_mopup_LE"
+        apply(case_tac c, case_tac [2] aa)
+        apply(auto split:if_splits simp add:step.simps mopup_inv.simps)
+        apply(simp_all add: mopupfetchs abc_mopup_LE_def lex_triple_def lex_pair_def )
+        done
+      thus "((step (steps (Suc 0, l, r) (mopup_a n @ shift mopup_b (2 * n), 0) na) 
+        (mopup_a n @ shift mopup_b (2 * n), 0), n),
+        steps (Suc 0, l, r) (mopup_a n @ shift mopup_b (2 * n), 0) na, n)
+        \<in> abc_mopup_LE"
+        using g by simp
+    qed
+  qed
+qed     
+
+lemma mopup_inv_start: 
+  "n < length am \<Longrightarrow> mopup_inv (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) am n ires"
+apply(auto simp: mopup_inv.simps mopup_bef_erase_a.simps mopup_jump_over1.simps)
+apply(case_tac [!] am, auto split: if_splits simp: tape_of_nl_cons)
+apply(rule_tac x = "Suc a" in exI, rule_tac x = k in exI, simp)
+apply(case_tac [!] n, simp_all add: abc_lm_v.simps)
+apply(case_tac k, simp, simp_all)
+done
+      
+lemma mopup_correct:
+  assumes less: "n < length (am::nat list)"
+  and rs: "abc_lm_v am n = rs"
+  shows "\<exists> stp i j. (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp)
+    = (0, Bk\<up>i @ Bk # Bk # ires, Oc # Oc\<up> rs @ Bk\<up>j)"
+using less
+proof -
+  have a: "mopup_inv (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) am n ires"
+    using less
+    apply(simp add: mopup_inv_start)
+    done    
+  then have "\<exists> stp. is_final (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp)"
+    using less mopup_halt[of n am  "Bk # Bk # ires" "<am> @ Bk \<up> k" ires
+      "(\<lambda>stp. (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp, n))"
+      "(\<lambda>(c, n). is_final c)"]
+    apply(simp)
+    done
+  from this obtain stp where b:
+    "is_final (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp)" ..
+  from a b have
+    "mopup_inv (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp)
+    am n ires"
+    apply(rule_tac mopup_inv_steps, simp_all add: less)
+    done    
+  from b and this show "?thesis"
+    apply(rule_tac x = stp in exI, simp)
+    apply(case_tac "steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) 
+      (mopup_a n @ shift mopup_b (2 * n), 0) stp")
+    apply(simp add: mopup_inv.simps mopup_stop.simps rs)
+    using rs
+    apply(simp add: tape_of_nat_abv)
+    done
+qed
+
+(*we can use Hoare_plus here*)
+
+lemma wf_mopup[intro]: "tm_wf (mopup n, 0)"
+apply(induct n, simp add: mopup.simps shift.simps mopup_b_def tm_wf.simps)
+apply(auto simp: mopup.simps shift.simps mopup_b_def tm_wf.simps)
+done
+
+lemma length_tp:
+  "\<lbrakk>ly = layout_of ap; tp = tm_of ap\<rbrakk> \<Longrightarrow> 
+  start_of ly (length ap) = Suc (length tp div 2)"
+apply(frule_tac length_tp', simp_all)
+apply(simp add: start_of.simps)
+done
+
+lemma compile_correct_halt: 
+  assumes layout: "ly = layout_of ap"
+  and compile: "tp = tm_of ap"
+  and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires"
+  and abc_halt: "abc_steps_l (0, lm) ap stp = (length ap, am)"
+  and rs_loc: "n < length am"
+  and rs: "abc_lm_v am n = rs"
+  and off: "off = length tp div 2"
+  shows "\<exists> stp i j. steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stp = (0, Bk\<up>i @ Bk # Bk # ires, Oc\<up>Suc rs @ Bk\<up>j)"
+proof -
+  have "\<exists> stp k. steps (Suc 0, l, r) (tp, 0) stp = (Suc off, Bk # Bk # ires, <am> @ Bk\<up>k)"
+    using assms tp_correct'[of ly ap tp lm l r ires stp am]
+    by(simp add: length_tp)
+  then obtain stp k where "steps (Suc 0, l, r) (tp, 0) stp = (Suc off, Bk # Bk # ires, <am> @ Bk\<up>k)"
+    by blast
+  then have a: "steps (Suc 0, l, r) (tp@shift (mopup n) off , 0) stp = (Suc off, Bk # Bk # ires, <am> @ Bk\<up>k)"
+    using assms
+    by(auto intro: tm_append_first_steps_eq)
+  have "\<exists> stp i j. (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp)
+    = (0, Bk\<up>i @ Bk # Bk # ires, Oc # Oc\<up> rs @ Bk\<up>j)"
+    using assms
+    by(auto intro: mopup_correct)
+  then obtain stpb i j where 
+    "steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stpb
+    = (0, Bk\<up>i @ Bk # Bk # ires, Oc # Oc\<up> rs @ Bk\<up>j)" by blast
+  then have b: "steps (Suc 0 + off, Bk # Bk # ires, <am> @ Bk \<up> k) (tp @ shift (mopup n) off, 0) stpb
+    = (0, Bk\<up>i @ Bk # Bk # ires, Oc # Oc\<up> rs @ Bk\<up>j)"
+    using assms wf_mopup
+   thm tm_append_second_halt_eq
+    apply(drule_tac tm_append_second_halt_eq, auto)
+    done
+  from a b show "?thesis"
+    by(rule_tac x = "stp + stpb" in exI, simp add: steps_add)
+qed
+ 
+declare mopup.simps[simp del]
+lemma abc_step_red2:
+  "abc_steps_l (s, lm) p (Suc n) = (let (as', am') = abc_steps_l (s, lm) p n in
+                                    abc_step_l (as', am') (abc_fetch as' p))"
+apply(case_tac "abc_steps_l (s, lm) p n", simp)
+apply(drule_tac abc_step_red, simp)
+done
+
+lemma crsp_steps2:
+  assumes 
+  layout: "ly = layout_of ap"
+  and compile: "tp = tm_of ap"
+  and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires"
+  and nothalt: "as < length ap"
+  and aexec: "abc_steps_l (0, lm) ap stp = (as, am)"
+  shows "\<exists>stpa\<ge>stp. crsp ly (as, am) (steps (Suc 0, l, r) (tp, 0) stpa) ires"
+using nothalt aexec
+proof(induct stp arbitrary: as am)
+  case 0
+  thus "?case"
+    using crsp
+    by(rule_tac x = 0 in exI, auto simp: abc_steps_l.simps steps.simps crsp)
+next
+  case (Suc stp as am)
+  have ind: 
+    "\<And> as am.  \<lbrakk>as < length ap; abc_steps_l (0, lm) ap stp = (as, am)\<rbrakk> 
+    \<Longrightarrow> \<exists>stpa\<ge>stp. crsp ly (as, am) (steps (Suc 0, l, r) (tp, 0) stpa) ires" by fact
+  have a: "as < length ap" by fact
+  have b: "abc_steps_l (0, lm) ap (Suc stp) = (as, am)" by fact
+  obtain as' am' where c: "abc_steps_l (0, lm) ap stp = (as', am')" 
+    by(case_tac "abc_steps_l (0, lm) ap stp", auto)
+  then have d: "as' < length ap"
+    using a b
+    by(simp add: abc_step_red2, case_tac "as' < length ap", simp,
+      simp add: abc_fetch.simps abc_steps_l.simps abc_step_l.simps)
+  have "\<exists>stpa\<ge>stp. crsp ly (as', am') (steps (Suc 0, l, r) (tp, 0) stpa) ires"
+    using d c ind by simp
+  from this obtain stpa where e: 
+    "stpa \<ge> stp \<and>  crsp ly (as', am') (steps (Suc 0, l, r) (tp, 0) stpa) ires"
+    by blast
+  obtain s' l' r' where f: "steps (Suc 0, l, r) (tp, 0) stpa = (s', l', r')"
+    by(case_tac "steps (Suc 0, l, r) (tp, 0) stpa", auto)
+  obtain ins where g: "abc_fetch as' ap = Some ins" using d 
+    by(case_tac "abc_fetch as' ap",auto simp: abc_fetch.simps)
+  then have "\<exists>stp> (0::nat). crsp ly (abc_step_l (as', am') (Some ins)) 
+    (steps (s', l', r') (tp, 0) stp) ires "
+    using layout compile e f 
+    by(rule_tac crsp_step, simp_all)
+  then obtain stpb where "stpb > 0 \<and> crsp ly (abc_step_l (as', am') (Some ins)) 
+    (steps (s', l', r') (tp, 0) stpb) ires" ..
+  from this show "?case" using b e g f c
+    by(rule_tac x = "stpa + stpb" in exI, simp add: steps_add abc_step_red2)
+qed
+    
+lemma compile_correct_unhalt: 
+  assumes layout: "ly = layout_of ap"
+  and compile: "tp = tm_of ap"
+  and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires"
+  and off: "off = length tp div 2"
+  and abc_unhalt: "\<forall> stp. (\<lambda> (as, am). as < length ap) (abc_steps_l (0, lm) ap stp)"
+  shows "\<forall> stp.\<not> is_final (steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stp)"
+using assms
+proof(rule_tac allI, rule_tac notI)
+  fix stp
+  assume h: "is_final (steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stp)"
+  obtain as am where a: "abc_steps_l (0, lm) ap stp = (as, am)"
+    by(case_tac "abc_steps_l (0, lm) ap stp", auto)
+  then have b: "as < length ap"
+    using abc_unhalt
+    by(erule_tac x = stp in allE, simp)
+  have "\<exists> stpa\<ge>stp. crsp ly (as, am) (steps (Suc 0, l, r) (tp, 0) stpa) ires "
+    using assms b a
+    apply(rule_tac crsp_steps2, simp_all)
+    done
+  then obtain stpa where
+    "stpa\<ge>stp \<and> crsp ly (as, am) (steps (Suc 0, l, r) (tp, 0) stpa) ires" ..
+  then obtain s' l' r' where b: "(steps (Suc 0, l, r) (tp, 0) stpa) = (s', l', r') \<and> 
+       stpa\<ge>stp \<and> crsp ly (as, am) (s', l', r') ires"
+    by(case_tac "steps (Suc 0, l, r) (tp, 0) stpa", auto)
+  hence c:
+    "(steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stpa) = (s', l', r')"
+    by(rule_tac tm_append_first_steps_eq, simp_all add: crsp.simps)
+  from b have d: "s' > 0 \<and> stpa \<ge> stp"
+    by(simp add: crsp.simps)
+  then obtain diff where e: "stpa = stp + diff"   by (metis le_iff_add)
+  obtain s'' l'' r'' where f:
+    "steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stp = (s'', l'', r'') \<and> is_final (s'', l'', r'')"
+    using h
+    by(case_tac "steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stp", auto)
+
+  then have "is_final (steps (s'', l'', r'') (tp @ shift (mopup n) off, 0) diff)"
+    by(auto intro: after_is_final)
+  then have "is_final (steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stpa)"
+    using e
+    by(simp add: steps_add f)
+  from this and c d show "False" by simp
+qed
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/Rec_Def.thy	Sun Feb 10 19:49:07 2013 +0000
@@ -0,0 +1,87 @@
+theory Rec_Def
+imports Main
+begin
+
+section {*
+  Recursive functions
+*}
+
+text {*
+  Datatype of recursive operators.
+*}
+
+datatype recf = 
+ -- {* The zero function, which always resturns @{text "0"} as result. *}
+  z | 
+ -- {* The successor function, which increments its arguments. *}
+  s | 
+ -- {*
+  The projection function, where @{text "id i j"} returns the @{text "j"}-th
+  argment out of the @{text "i"} arguments.
+  *}
+  id nat nat | 
+ -- {*
+  The compostion operator, where "@{text "Cn n f [g1; g2; \<dots> ;gm]"} 
+  computes @{text "f (g1(x1, x2, \<dots>, xn), g2(x1, x2, \<dots>, xn), \<dots> , 
+  gm(x1, x2, \<dots> , xn))"} for input argments @{text "x1, \<dots>, xn"}.
+  *}
+  Cn nat recf "recf list" | 
+-- {*
+  The primitive resursive operator, where @{text "Pr n f g"} computes:
+  @{text "Pr n f g (x1, x2, \<dots>, xn-1, 0) = f(x1, \<dots>, xn-1)"} 
+  and @{text "Pr n f g (x1, x2, \<dots>, xn-1, k') = g(x1, x2, \<dots>, xn-1, k, 
+                                                  Pr n f g (x1, \<dots>, xn-1, k))"}.
+  *}
+  Pr nat recf recf | 
+-- {*
+  The minimization operator, where @{text "Mn n f (x1, x2, \<dots> , xn)"} 
+  computes the first i such that @{text "f (x1, \<dots>, xn, i) = 0"} and for all
+  @{text "j"}, @{text "f (x1, x2, \<dots>, xn, j) > 0"}.
+  *}
+  Mn nat recf 
+
+text {* 
+  The semantis of recursive operators is given by an inductively defined
+  relation as follows, where  
+  @{text "rec_calc_rel R [x1, x2, \<dots>, xn] r"} means the computation of 
+  @{text "R"} over input arguments @{text "[x1, x2, \<dots>, xn"} terminates
+  and gives rise to a result @{text "r"}
+*}
+
+inductive rec_calc_rel :: "recf \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> bool"
+where
+  calc_z: "rec_calc_rel z [n] 0" |
+  calc_s: "rec_calc_rel s [n] (Suc n)" |
+  calc_id: "\<lbrakk>length args = i; j < i; args!j = r\<rbrakk> \<Longrightarrow> rec_calc_rel (id i j) args r" |
+  calc_cn: "\<lbrakk>length args = n;
+             \<forall> k < length gs. rec_calc_rel (gs ! k) args (rs ! k);
+             length rs = length gs; 
+             rec_calc_rel f rs r\<rbrakk> 
+            \<Longrightarrow> rec_calc_rel (Cn n f gs) args r" |
+  calc_pr_zero: 
+           "\<lbrakk>length args = n;
+             rec_calc_rel f args r0 \<rbrakk> 
+            \<Longrightarrow> rec_calc_rel (Pr n f g) (args @ [0]) r0" |
+  calc_pr_ind: "
+           \<lbrakk> length args = n;
+             rec_calc_rel (Pr n f g) (args @ [k]) rk; 
+             rec_calc_rel g (args @ [k] @ [rk]) rk'\<rbrakk>
+            \<Longrightarrow> rec_calc_rel (Pr n f g) (args @ [Suc k]) rk'"  |
+  calc_mn: "\<lbrakk>length args = n; 
+             rec_calc_rel f (args@[r]) 0; 
+             \<forall> i < r. (\<exists> ri. rec_calc_rel f (args@[i]) ri \<and> ri \<noteq> 0)\<rbrakk> 
+            \<Longrightarrow> rec_calc_rel (Mn n f) args r" 
+
+inductive_cases calc_pr_reverse:
+              "rec_calc_rel (Pr n f g) (lm) rSucy"
+
+inductive_cases calc_z_reverse: "rec_calc_rel z lm x"
+
+inductive_cases calc_s_reverse: "rec_calc_rel s lm x"
+
+inductive_cases calc_id_reverse: "rec_calc_rel (id m n) lm x"
+
+inductive_cases calc_cn_reverse: "rec_calc_rel (Cn n f gs) lm x"
+
+inductive_cases calc_mn_reverse:"rec_calc_rel (Mn n f) lm x"
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/Recursive.thy	Sun Feb 10 19:49:07 2013 +0000
@@ -0,0 +1,5134 @@
+theory Recursive
+imports Rec_Def Abacus
+begin
+
+section {* 
+  Compiling from recursive functions to Abacus machines
+  *}
+
+text {*
+  Some auxilliary Abacus machines used to construct the result Abacus machines.
+*}
+
+text {*
+  @{text "get_paras_num recf"} returns the arity of recursive function @{text "recf"}.
+*}
+fun get_paras_num :: "recf \<Rightarrow> nat"
+  where
+  "get_paras_num z = 1" |
+  "get_paras_num s = 1" |
+  "get_paras_num (id m n) = m" |
+  "get_paras_num (Cn n f gs) = n" |
+  "get_paras_num (Pr n f g) = Suc n"  |
+  "get_paras_num (Mn n f) = n"  
+
+fun addition :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
+  where
+  "addition m n p = [Dec m 4, Inc n, Inc p, Goto 0, Dec p 7, 
+                       Inc m, Goto 4]"
+
+fun mv_box :: "nat \<Rightarrow> nat \<Rightarrow> abc_prog"
+  where
+  "mv_box m n = [Dec m 3, Inc n, Goto 0]"
+
+fun abc_inst_shift :: "abc_inst \<Rightarrow> nat \<Rightarrow> abc_inst"
+  where
+  "abc_inst_shift (Inc m) n = Inc m" |
+  "abc_inst_shift (Dec m e) n = Dec m (e + n)" |
+  "abc_inst_shift (Goto m) n = Goto (m + n)"
+
+fun abc_shift :: "abc_inst list \<Rightarrow> nat \<Rightarrow> abc_inst list" 
+  where
+  "abc_shift xs n = map (\<lambda> x. abc_inst_shift x n) xs" 
+
+fun abc_append :: "abc_inst list \<Rightarrow> abc_inst list \<Rightarrow> 
+                           abc_inst list" (infixl "[+]" 60)
+  where
+  "abc_append al bl = (let al_len = length al in 
+                           al @ abc_shift bl al_len)"
+
+text {*
+  The compilation of @{text "z"}-operator.
+*}
+definition rec_ci_z :: "abc_inst list"
+  where
+  "rec_ci_z \<equiv> [Goto 1]"
+
+text {*
+  The compilation of @{text "s"}-operator.
+*}
+definition rec_ci_s :: "abc_inst list"
+  where
+  "rec_ci_s \<equiv> (addition 0 1 2 [+] [Inc 1])"
+
+
+text {*
+  The compilation of @{text "id i j"}-operator
+*}
+
+fun rec_ci_id :: "nat \<Rightarrow> nat \<Rightarrow> abc_inst list"
+  where
+  "rec_ci_id i j = addition j i (i + 1)"
+
+fun mv_boxes :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_inst list"
+  where
+  "mv_boxes ab bb 0 = []" |
+  "mv_boxes ab bb (Suc n) = mv_boxes ab bb n [+] mv_box (ab + n)
+  (bb + n)"
+
+fun empty_boxes :: "nat \<Rightarrow> abc_inst list"
+  where
+  "empty_boxes 0 = []" |
+  "empty_boxes (Suc n) = empty_boxes n [+] [Dec n 2, Goto 0]"
+
+fun cn_merge_gs ::
+  "(abc_inst list \<times> nat \<times> nat) list \<Rightarrow> nat \<Rightarrow> abc_inst list"
+  where
+  "cn_merge_gs [] p = []" |
+  "cn_merge_gs (g # gs) p = 
+      (let (gprog, gpara, gn) = g in 
+         gprog [+] mv_box gpara p [+] cn_merge_gs gs (Suc p))"
+
+
+text {*
+  The compiler of recursive functions, where @{text "rec_ci recf"} return 
+  @{text "(ap, arity, fp)"}, where @{text "ap"} is the Abacus program, @{text "arity"} is the 
+  arity of the recursive function @{text "recf"}, 
+@{text "fp"} is the amount of memory which is going to be
+  used by @{text "ap"} for its execution. 
+*}
+
+function rec_ci :: "recf \<Rightarrow> abc_inst list \<times> nat \<times> nat"
+  where
+  "rec_ci z = (rec_ci_z, 1, 2)" |
+  "rec_ci s = (rec_ci_s, 1, 3)" |
+  "rec_ci (id m n) = (rec_ci_id m n, m, m + 2)" |
+  "rec_ci (Cn n f gs) = 
+      (let cied_gs = map (\<lambda> g. rec_ci g) (f # gs) in
+       let (fprog, fpara, fn) = hd cied_gs in 
+       let pstr = 
+        Max (set (Suc n # fn # (map (\<lambda> (aprog, p, n). n) cied_gs))) in
+       let qstr = pstr + Suc (length gs) in 
+       (cn_merge_gs (tl cied_gs) pstr [+] mv_boxes 0 qstr n [+] 
+          mv_boxes pstr 0 (length gs) [+] fprog [+] 
+            mv_box fpara pstr [+] empty_boxes (length gs) [+] 
+             mv_box pstr n [+] mv_boxes qstr 0 n, n,  qstr + n))" |
+  "rec_ci (Pr n f g) = 
+         (let (fprog, fpara, fn) = rec_ci f in 
+          let (gprog, gpara, gn) = rec_ci g in 
+          let p = Max (set ([n + 3, fn, gn])) in 
+          let e = length gprog + 7 in 
+           (mv_box n p [+] fprog [+] mv_box n (Suc n) [+] 
+               (([Dec p e] [+] gprog [+] 
+                 [Inc n, Dec (Suc n) 3, Goto 1]) @
+                     [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gprog + 4)]),
+             Suc n, p + 1))" |
+  "rec_ci (Mn n f) =
+         (let (fprog, fpara, fn) = rec_ci f in 
+          let len = length (fprog) in 
+            (fprog @ [Dec (Suc n) (len + 5), Dec (Suc n) (len + 3),
+             Goto (len + 1), Inc n, Goto 0], n, max (Suc n) fn) )"
+  by pat_completeness auto
+termination 
+proof
+term size
+  show "wf (measure size)" by auto
+next
+  fix n f gs x
+  assume "(x::recf) \<in> set (f # gs)" 
+  thus "(x, Cn n f gs) \<in> measure size"
+    by(induct gs, auto)
+next
+  fix n f g
+  show "(f, Pr n f g) \<in> measure size" by auto
+next
+  fix n f g x xa y xb ya
+  show "(g, Pr n f g) \<in> measure size" by auto
+next
+  fix n f
+  show "(f, Mn n f) \<in> measure size" by auto
+qed
+
+declare rec_ci.simps [simp del] rec_ci_s_def[simp del] 
+        rec_ci_z_def[simp del] rec_ci_id.simps[simp del]
+        mv_boxes.simps[simp del] abc_append.simps[simp del]
+        mv_box.simps[simp del] addition.simps[simp del]
+  
+thm rec_calc_rel.induct
+
+declare abc_steps_l.simps[simp del] abc_fetch.simps[simp del] 
+        abc_step_l.simps[simp del] 
+
+lemma abc_steps_add: 
+  "abc_steps_l (as, lm) ap (m + n) = 
+         abc_steps_l (abc_steps_l (as, lm) ap m) ap n"
+apply(induct m arbitrary: n as lm, simp add: abc_steps_l.simps)
+proof -
+  fix m n as lm
+  assume ind: 
+    "\<And>n as lm. abc_steps_l (as, lm) ap (m + n) = 
+                   abc_steps_l (abc_steps_l (as, lm) ap m) ap n"
+  show "abc_steps_l (as, lm) ap (Suc m + n) = 
+             abc_steps_l (abc_steps_l (as, lm) ap (Suc m)) ap n"
+    apply(insert ind[of as lm "Suc n"], simp)
+    apply(insert ind[of as lm "Suc 0"], simp add: abc_steps_l.simps)
+    apply(case_tac "(abc_steps_l (as, lm) ap m)", simp)
+    apply(simp add: abc_steps_l.simps)
+    apply(case_tac "abc_step_l (a, b) (abc_fetch a ap)", 
+          simp add: abc_steps_l.simps)
+    done
+qed
+
+(*lemmas: rec_ci and rec_calc_rel*)
+
+lemma rec_calc_inj_case_z: 
+  "\<lbrakk>rec_calc_rel z l x; rec_calc_rel z l y\<rbrakk> \<Longrightarrow> x = y"
+apply(auto elim: calc_z_reverse)
+done
+
+lemma  rec_calc_inj_case_s: 
+  "\<lbrakk>rec_calc_rel s l x; rec_calc_rel s l y\<rbrakk> \<Longrightarrow> x = y"
+apply(auto elim: calc_s_reverse)
+done
+
+lemma rec_calc_inj_case_id:
+  "\<lbrakk>rec_calc_rel (recf.id nat1 nat2) l x;
+    rec_calc_rel (recf.id nat1 nat2) l y\<rbrakk> \<Longrightarrow> x = y"
+apply(auto elim: calc_id_reverse)
+done
+
+lemma rec_calc_inj_case_mn:
+  assumes ind: "\<And> l x y. \<lbrakk>rec_calc_rel f l x; rec_calc_rel f l y\<rbrakk> 
+           \<Longrightarrow> x = y" 
+  and h: "rec_calc_rel (Mn n f) l x" "rec_calc_rel (Mn n f) l y"
+  shows "x = y"
+  apply(insert h)
+  apply(elim  calc_mn_reverse)
+  apply(case_tac "x > y", simp)
+  apply(erule_tac x = "y" in allE, auto)
+proof -
+  fix v va
+  assume "rec_calc_rel f (l @ [y]) 0" 
+    "rec_calc_rel f (l @ [y]) v"  
+    "0 < v"
+  thus "False"
+    apply(insert ind[of "l @ [y]" 0 v], simp)
+    done
+next
+  fix v va
+  assume 
+    "rec_calc_rel f (l @ [x]) 0" 
+    "\<forall>x<y. \<exists>v. rec_calc_rel f (l @ [x]) v \<and> 0 < v" "\<not> y < x"
+  thus "x = y"
+    apply(erule_tac x = "x" in allE)
+    apply(case_tac "x = y", auto)
+    apply(drule_tac y = v in ind, simp, simp)
+    done
+qed 
+
+lemma rec_calc_inj_case_pr: 
+  assumes f_ind: 
+  "\<And>l x y. \<lbrakk>rec_calc_rel f l x; rec_calc_rel f l y\<rbrakk> \<Longrightarrow> x = y"
+  and g_ind:
+  "\<And>x xa y xb ya l xc yb. 
+  \<lbrakk>x = rec_ci f; (xa, y) = x; (xb, ya) = y; 
+  rec_calc_rel g l xc; rec_calc_rel g l yb\<rbrakk> \<Longrightarrow> xc = yb"
+  and h: "rec_calc_rel (Pr n f g) l x" "rec_calc_rel (Pr n f g) l y"  
+  shows "x = y"
+  apply(case_tac "rec_ci f")
+proof -
+  fix a b c
+  assume "rec_ci f = (a, b, c)"
+  hence ng_ind: 
+    "\<And> l xc yb. \<lbrakk>rec_calc_rel g l xc; rec_calc_rel g l yb\<rbrakk>
+    \<Longrightarrow> xc = yb"
+    apply(insert g_ind[of "(a, b, c)" "a" "(b, c)" b c], simp)
+    done
+  from h show "x = y"
+    apply(erule_tac calc_pr_reverse, erule_tac calc_pr_reverse)
+    apply(erule f_ind, simp, simp)
+    apply(erule_tac calc_pr_reverse, simp, simp)
+  proof -
+    fix la ya ry laa yaa rya
+    assume k1:  "rec_calc_rel g (la @ [ya, ry]) x" 
+      "rec_calc_rel g (la @ [ya, rya]) y"
+      and k2: "rec_calc_rel (Pr (length la) f g) (la @ [ya]) ry"
+              "rec_calc_rel (Pr (length la) f g) (la @ [ya]) rya"
+    from k2 have "ry = rya"
+      apply(induct ya arbitrary: ry rya)
+      apply(erule_tac calc_pr_reverse, 
+        erule_tac calc_pr_reverse, simp)
+      apply(erule f_ind, simp, simp, simp)
+      apply(erule_tac calc_pr_reverse, simp)
+      apply(erule_tac rSucy = rya in calc_pr_reverse, simp, simp)
+    proof -
+      fix ya ry rya l y ryb laa yb ryc
+      assume ind:
+        "\<And>ry rya. \<lbrakk>rec_calc_rel (Pr (length l) f g) (l @ [y]) ry; 
+                   rec_calc_rel (Pr (length l) f g) (l @ [y]) rya\<rbrakk> \<Longrightarrow> ry = rya"
+        and j: "rec_calc_rel (Pr (length l) f g) (l @ [y]) ryb"
+        "rec_calc_rel g (l @ [y, ryb]) ry" 
+        "rec_calc_rel (Pr (length l) f g) (l @ [y]) ryc" 
+        "rec_calc_rel g (l @ [y, ryc]) rya"
+      from j show "ry = rya"
+	apply(insert ind[of ryb ryc], simp)
+	apply(insert ng_ind[of "l @ [y, ryc]" ry rya], simp)
+	done
+    qed 
+    from k1 and this show "x = y"
+      apply(simp)
+      apply(insert ng_ind[of "la @ [ya, rya]" x y], simp)
+      done
+  qed  
+qed
+
+lemma Suc_nth_part_eq:
+  "\<forall>k<Suc (length list). (a # xs) ! k = (aa # list) ! k
+       \<Longrightarrow> \<forall>k<(length list). (xs) ! k = (list) ! k"
+apply(rule allI, rule impI)
+apply(erule_tac x = "Suc k" in allE, simp)
+done
+
+
+lemma list_eq_intro:  
+  "\<lbrakk>length xs = length ys; \<forall> k < length xs. xs ! k = ys ! k\<rbrakk> 
+  \<Longrightarrow> xs = ys"
+apply(induct xs arbitrary: ys, simp)
+apply(case_tac ys, simp, simp)
+proof -
+  fix a xs ys aa list
+  assume ind: 
+    "\<And>ys. \<lbrakk>length list = length ys; \<forall>k<length ys. xs ! k = ys ! k\<rbrakk>
+    \<Longrightarrow> xs = ys"
+    and h: "length xs = length list" 
+    "\<forall>k<Suc (length list). (a # xs) ! k = (aa # list) ! k"
+  from h show "a = aa \<and> xs = list"
+    apply(insert ind[of list], simp)
+    apply(frule Suc_nth_part_eq, simp)
+    apply(erule_tac x = "0" in allE, simp)
+    done
+qed
+
+lemma rec_calc_inj_case_cn: 
+  assumes ind: 
+  "\<And>x l xa y.
+  \<lbrakk>x = f \<or> x \<in> set gs; rec_calc_rel x l xa; rec_calc_rel x l y\<rbrakk>
+  \<Longrightarrow> xa = y"
+  and h: "rec_calc_rel (Cn n f gs) l x" 
+         "rec_calc_rel (Cn n f gs) l y"
+  shows "x = y"
+  apply(insert h, elim  calc_cn_reverse)
+  apply(subgoal_tac "rs = rsa")
+  apply(rule_tac x = f and l = rsa and xa = x and y = y in ind, 
+        simp, simp, simp)
+  apply(intro list_eq_intro, simp, rule allI, rule impI)
+  apply(erule_tac x = k in allE, rule_tac x = k in allE, simp, simp)
+  apply(rule_tac x = "gs ! k" in ind, simp, simp, simp)
+  done
+
+lemma rec_calc_inj:
+  "\<lbrakk>rec_calc_rel f l x; 
+    rec_calc_rel f l y\<rbrakk> \<Longrightarrow> x = y"
+apply(induct f arbitrary: l x y rule: rec_ci.induct)
+apply(simp add: rec_calc_inj_case_z)
+apply(simp add: rec_calc_inj_case_s)
+apply(simp add: rec_calc_inj_case_id, simp)
+apply(erule rec_calc_inj_case_cn,simp, simp)
+apply(erule rec_calc_inj_case_pr, auto)
+apply(erule rec_calc_inj_case_mn, auto)
+done
+
+
+lemma calc_rel_reverse_ind_step_ex: 
+  "\<lbrakk>rec_calc_rel (Pr n f g) (lm @ [Suc x]) rs\<rbrakk> 
+  \<Longrightarrow> \<exists> rs. rec_calc_rel (Pr n f g) (lm @ [x]) rs"
+apply(erule calc_pr_reverse, simp, simp)
+apply(rule_tac x = rk in exI, simp)
+done
+
+lemma [simp]: "Suc x \<le> y \<Longrightarrow> Suc (y - Suc x) = y - x"
+by arith
+
+lemma calc_pr_para_not_null: 
+  "rec_calc_rel (Pr n f g) lm rs \<Longrightarrow> lm \<noteq> []"
+apply(erule calc_pr_reverse, simp, simp)
+done
+
+lemma calc_pr_less_ex: 
+ "\<lbrakk>rec_calc_rel (Pr n f g) lm rs; x \<le> last lm\<rbrakk> \<Longrightarrow> 
+ \<exists>rs. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rs"
+apply(subgoal_tac "lm \<noteq> []")
+apply(induct x, rule_tac x = rs in exI, simp, simp, erule exE)
+apply(rule_tac rs = xa in calc_rel_reverse_ind_step_ex, simp)
+apply(simp add: calc_pr_para_not_null)
+done
+
+lemma calc_pr_zero_ex:
+  "rec_calc_rel (Pr n f g) lm rs \<Longrightarrow> 
+             \<exists>rs. rec_calc_rel f (butlast lm) rs"
+apply(drule_tac x = "last lm" in calc_pr_less_ex, simp,
+      erule_tac exE, simp)
+apply(erule_tac calc_pr_reverse, simp)
+apply(rule_tac x = rs in exI, simp, simp)
+done
+
+
+lemma abc_steps_ind: 
+  "abc_steps_l (as, am) ap (Suc stp) =
+          abc_steps_l (abc_steps_l (as, am) ap stp) ap (Suc 0)"
+apply(insert abc_steps_add[of as am ap stp "Suc 0"], simp)
+done
+
+lemma abc_steps_zero: "abc_steps_l asm ap 0 = asm"
+apply(case_tac asm, simp add: abc_steps_l.simps)
+done
+
+lemma abc_append_nth: 
+  "n < length ap + length bp \<Longrightarrow> 
+       (ap [+] bp) ! n =
+         (if n < length ap then ap ! n 
+          else abc_inst_shift (bp ! (n - length ap)) (length ap))"
+apply(simp add: abc_append.simps nth_append map_nth split: if_splits)
+done
+
+lemma abc_state_keep:  
+  "as \<ge> length bp \<Longrightarrow> abc_steps_l (as, lm) bp stp = (as, lm)"
+apply(induct stp, simp add: abc_steps_zero)
+apply(simp add: abc_steps_ind)
+apply(simp add: abc_steps_zero)
+apply(simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps)
+done
+
+lemma abc_halt_equal: 
+  "\<lbrakk>abc_steps_l (0, lm) bp stpa = (length bp, lm1); 
+    abc_steps_l (0, lm) bp stpb = (length bp, lm2)\<rbrakk> \<Longrightarrow> lm1 = lm2"
+apply(case_tac "stpa - stpb > 0")
+apply(insert abc_steps_add[of 0 lm bp stpb "stpa - stpb"], simp)
+apply(insert abc_state_keep[of bp "length bp" lm2 "stpa - stpb"], 
+      simp, simp add: abc_steps_zero)
+apply(insert abc_steps_add[of 0 lm bp stpa "stpb - stpa"], simp)
+apply(insert abc_state_keep[of bp "length bp" lm1 "stpb - stpa"], 
+      simp)
+done  
+
+lemma abc_halt_point_ex: 
+  "\<lbrakk>\<exists>stp. abc_steps_l (0, lm) bp stp = (bs, lm');
+    bs = length bp; bp \<noteq> []\<rbrakk> 
+  \<Longrightarrow> \<exists> stp. (\<lambda> (s, l). s < bs \<and> 
+              (abc_steps_l (s, l) bp (Suc 0)) = (bs, lm')) 
+      (abc_steps_l (0, lm) bp stp) "
+apply(erule_tac exE)
+proof -
+  fix stp
+  assume "bs = length bp" 
+         "abc_steps_l (0, lm) bp stp = (bs, lm')" 
+         "bp \<noteq> []"
+  thus 
+    "\<exists>stp. (\<lambda>(s, l). s < bs \<and> 
+      abc_steps_l (s, l) bp (Suc 0) = (bs, lm')) 
+                       (abc_steps_l (0, lm) bp stp)"
+    apply(induct stp, simp add: abc_steps_zero, simp)
+  proof -
+    fix stpa
+    assume ind: 
+     "abc_steps_l (0, lm) bp stpa = (length bp, lm')
+       \<Longrightarrow> \<exists>stp. (\<lambda>(s, l). s < length bp  \<and> abc_steps_l (s, l) bp 
+             (Suc 0) = (length bp, lm')) (abc_steps_l (0, lm) bp stp)"
+    and h: "abc_steps_l (0, lm) bp (Suc stpa) = (length bp, lm')" 
+           "abc_steps_l (0, lm) bp stp = (length bp, lm')" 
+           "bp \<noteq> []"
+    from h show 
+      "\<exists>stp. (\<lambda>(s, l). s < length bp \<and> abc_steps_l (s, l) bp (Suc 0)
+                    = (length bp, lm')) (abc_steps_l (0, lm) bp stp)"
+      apply(case_tac "abc_steps_l (0, lm) bp stpa", 
+            case_tac "a = length bp")
+      apply(insert ind, simp)
+      apply(subgoal_tac "b = lm'", simp)
+      apply(rule_tac abc_halt_equal, simp, simp)
+      apply(rule_tac x = stpa in exI, simp add: abc_steps_ind)
+      apply(simp add: abc_steps_zero)
+      apply(rule classical, simp add: abc_steps_l.simps 
+                             abc_fetch.simps abc_step_l.simps)
+      done
+  qed
+qed  
+
+
+lemma abc_append_empty_r[simp]: "[] [+] ab = ab"
+apply(simp add: abc_append.simps abc_inst_shift.simps)
+apply(induct ab, simp, simp)
+apply(case_tac a, simp_all add: abc_inst_shift.simps)
+done
+
+lemma abc_append_empty_l[simp]:  "ab [+] [] = ab"
+apply(simp add: abc_append.simps abc_inst_shift.simps)
+done
+
+
+lemma abc_append_length[simp]:  
+  "length (ap [+] bp) = length ap + length bp"
+apply(simp add: abc_append.simps)
+done
+
+declare Let_def[simp]
+
+lemma abc_append_commute: "as [+] bs [+] cs = as [+] (bs [+] cs)"
+apply(simp add: abc_append.simps abc_shift.simps abc_inst_shift.simps)
+apply(induct cs, simp, simp)
+apply(case_tac a, auto simp: abc_inst_shift.simps Let_def)
+done
+
+lemma abc_halt_point_step[simp]: 
+  "\<lbrakk>a < length bp; abc_steps_l (a, b) bp (Suc 0) = (length bp, lm')\<rbrakk>
+  \<Longrightarrow> abc_steps_l (length ap + a, b) (ap [+] bp [+] cp) (Suc 0) = 
+                                        (length ap + length bp, lm')"
+apply(simp add: abc_steps_l.simps abc_fetch.simps abc_append_nth)
+apply(case_tac "bp ! a", 
+                      auto simp: abc_steps_l.simps abc_step_l.simps)
+done
+
+lemma abc_step_state_in:
+  "\<lbrakk>bs < length bp;  abc_steps_l (a, b) bp (Suc 0) = (bs, l)\<rbrakk>
+  \<Longrightarrow> a < length bp"
+apply(simp add: abc_steps_l.simps abc_fetch.simps)
+apply(rule_tac classical, 
+      simp add: abc_step_l.simps abc_steps_l.simps)
+done
+
+
+lemma abc_append_state_in_exc: 
+  "\<lbrakk>bs < length bp; abc_steps_l (0, lm) bp stpa = (bs, l)\<rbrakk>
+ \<Longrightarrow> abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa = 
+                                             (length ap + bs, l)"
+apply(induct stpa arbitrary: bs l, simp add: abc_steps_zero)
+proof -
+  fix stpa bs l
+  assume ind: 
+    "\<And>bs l. \<lbrakk>bs < length bp; abc_steps_l (0, lm) bp stpa = (bs, l)\<rbrakk>
+    \<Longrightarrow> abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa = 
+                                                (length ap + bs, l)"
+    and h: "bs < length bp" 
+           "abc_steps_l (0, lm) bp (Suc stpa) = (bs, l)"
+  from h show 
+    "abc_steps_l (length ap, lm) (ap [+] bp [+] cp) (Suc stpa) = 
+                                                (length ap + bs, l)"
+    apply(simp add: abc_steps_ind)
+    apply(case_tac "(abc_steps_l (0, lm) bp stpa)", simp)
+  proof -
+    fix a b
+    assume g: "abc_steps_l (0, lm) bp stpa = (a, b)" 
+              "abc_steps_l (a, b) bp (Suc 0) = (bs, l)"
+    from h and g have k1: "a < length bp"
+      apply(simp add: abc_step_state_in)
+      done
+    from h and g and k1 show 
+   "abc_steps_l (abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa) 
+              (ap [+] bp [+] cp) (Suc 0) = (length ap + bs, l)"
+      apply(insert ind[of a b], simp)
+      apply(simp add: abc_steps_l.simps abc_fetch.simps 
+                      abc_append_nth)
+      apply(case_tac "bp ! a", auto simp: 
+                                 abc_steps_l.simps abc_step_l.simps)
+      done
+  qed
+qed
+
+lemma [simp]: "abc_steps_l (0, am) [] stp = (0, am)"
+apply(induct stp, simp add: abc_steps_zero)
+apply(simp add: abc_steps_ind)
+apply(simp add: abc_steps_zero abc_steps_l.simps 
+                abc_fetch.simps abc_step_l.simps)
+done
+
+lemma abc_append_exc1:
+  "\<lbrakk>\<exists> stp. abc_steps_l (0, lm) bp stp = (bs, lm');
+    bs = length bp; 
+    as = length ap\<rbrakk>
+    \<Longrightarrow> \<exists> stp. abc_steps_l (as, lm) (ap [+] bp [+] cp) stp 
+                                                 = (as + bs, lm')"
+apply(case_tac "bp = []", erule_tac exE, simp,
+      rule_tac x = 0 in exI, simp add: abc_steps_zero)
+apply(frule_tac abc_halt_point_ex, simp, simp,
+      erule_tac exE, erule_tac exE) 
+apply(rule_tac x = "stpa + Suc 0" in exI)
+apply(case_tac "(abc_steps_l (0, lm) bp stpa)", 
+      simp add: abc_steps_ind)
+apply(subgoal_tac 
+  "abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa 
+                                   = (length ap + a, b)", simp)
+apply(simp add: abc_steps_zero)
+apply(rule_tac abc_append_state_in_exc, simp, simp)
+done
+
+lemma abc_append_exc3: 
+  "\<lbrakk>\<exists> stp. abc_steps_l (0, am) bp stp = (bs, bm); ss = length ap\<rbrakk>
+   \<Longrightarrow>  \<exists> stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
+apply(erule_tac exE)
+proof -
+  fix stp
+  assume h: "abc_steps_l (0, am) bp stp = (bs, bm)" "ss = length ap"
+  thus " \<exists>stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
+  proof(induct stp arbitrary: bs bm)
+    fix bs bm
+    assume "abc_steps_l (0, am) bp 0 = (bs, bm)"
+    thus "\<exists>stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
+      apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
+      done
+  next
+    fix stp bs bm
+    assume ind: 
+      "\<And>bs bm. \<lbrakk>abc_steps_l (0, am) bp stp = (bs, bm);
+                 ss = length ap\<rbrakk> \<Longrightarrow> 
+          \<exists>stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
+    and g: "abc_steps_l (0, am) bp (Suc stp) = (bs, bm)"
+    from g show 
+      "\<exists>stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
+      apply(insert abc_steps_add[of 0 am bp stp "Suc 0"], simp)
+      apply(case_tac "(abc_steps_l (0, am) bp stp)", simp)
+    proof -
+      fix a b
+      assume "(bs, bm) = abc_steps_l (a, b) bp (Suc 0)" 
+             "abc_steps_l (0, am) bp (Suc stp) = 
+                       abc_steps_l (a, b) bp (Suc 0)" 
+              "abc_steps_l (0, am) bp stp = (a, b)"
+      thus "?thesis"
+	apply(insert ind[of a b], simp add: h, erule_tac exE)
+	apply(rule_tac x = "Suc stp" in exI)
+	apply(simp only: abc_steps_ind, simp add: abc_steps_zero)
+      proof -
+	fix stp
+	assume "(bs, bm) = abc_steps_l (a, b) bp (Suc 0)"
+	thus "abc_steps_l (a + length ap, b) (ap [+] bp) (Suc 0)
+                                              = (bs + length ap, bm)"
+	  apply(simp add: abc_steps_l.simps abc_steps_zero
+                          abc_fetch.simps split: if_splits)
+	  apply(case_tac "bp ! a", 
+                simp_all add: abc_inst_shift.simps abc_append_nth
+                   abc_steps_l.simps abc_steps_zero abc_step_l.simps)
+	  apply(auto)
+	  done
+      qed
+    qed
+  qed
+qed
+
+lemma abc_add_equal:
+  "\<lbrakk>ap \<noteq> []; 
+    abc_steps_l (0, am) ap astp = (a, b);
+    a < length ap\<rbrakk>
+     \<Longrightarrow> (abc_steps_l (0, am) (ap @ bp) astp) = (a, b)"
+apply(induct astp arbitrary: a b, simp add: abc_steps_l.simps, simp)
+apply(simp add: abc_steps_ind)
+apply(case_tac "(abc_steps_l (0, am) ap astp)")
+proof -
+  fix astp a b aa ba
+  assume ind: 
+    "\<And>a b. \<lbrakk>abc_steps_l (0, am) ap astp = (a, b); 
+             a < length ap\<rbrakk> \<Longrightarrow> 
+                  abc_steps_l (0, am) (ap @ bp) astp = (a, b)"
+  and h: "abc_steps_l (abc_steps_l (0, am) ap astp) ap (Suc 0)
+                                                            = (a, b)"
+        "a < length ap" 
+        "abc_steps_l (0, am) ap astp = (aa, ba)"
+  from h show "abc_steps_l (abc_steps_l (0, am) (ap @ bp) astp)
+                                        (ap @ bp) (Suc 0) = (a, b)"
+    apply(insert ind[of aa ba], simp)
+    apply(subgoal_tac "aa < length ap", simp)
+    apply(simp add: abc_steps_l.simps abc_fetch.simps
+                     nth_append abc_steps_zero)
+    apply(rule abc_step_state_in, auto)
+    done
+qed
+
+
+lemma abc_add_exc1: 
+  "\<lbrakk>\<exists> astp. abc_steps_l (0, am) ap astp = (as, bm); as = length ap\<rbrakk>
+  \<Longrightarrow> \<exists> stp. abc_steps_l (0, am) (ap @ bp) stp = (as, bm)"
+apply(case_tac "ap = []", simp, 
+      rule_tac x = 0 in exI, simp add: abc_steps_zero)
+apply(drule_tac abc_halt_point_ex, simp, simp)
+apply(erule_tac exE, case_tac "(abc_steps_l (0, am) ap astp)", simp)
+apply(rule_tac x = "Suc astp" in exI, simp add: abc_steps_ind, auto)
+apply(frule_tac bp = bp in abc_add_equal, simp, simp, simp)
+apply(simp add: abc_steps_l.simps abc_steps_zero 
+                abc_fetch.simps nth_append)
+done
+
+declare abc_shift.simps[simp del] 
+
+lemma abc_append_exc2: 
+  "\<lbrakk>\<exists> astp. abc_steps_l (0, am) ap astp = (as, bm); as = length ap; 
+    \<exists> bstp. abc_steps_l (0, bm) bp bstp = (bs, bm'); bs = length bp;
+    cs = as + bs; bp \<noteq> []\<rbrakk>
+  \<Longrightarrow> \<exists> stp. abc_steps_l (0, am) (ap [+] bp) stp = (cs, bm')"
+apply(insert abc_append_exc1[of bm bp bs bm' as ap "[]"], simp)
+apply(drule_tac bp = "abc_shift bp (length ap)" in abc_add_exc1, simp)
+apply(subgoal_tac "ap @ abc_shift bp (length ap) = ap [+] bp", 
+      simp, auto)
+apply(rule_tac x = "stpa + stp" in exI, simp add: abc_steps_add)
+apply(simp add: abc_append.simps)
+done
+lemma exponent_add_iff: "a\<up>b @ a\<up>c@ xs = a\<up>(b+c) @ xs"
+apply(auto simp: replicate_add)
+done
+
+lemma exponent_cons_iff: "a # a\<up>c @ xs = a\<up>(Suc c) @ xs"
+apply(auto simp: replicate_add)
+done
+
+lemma  [simp]: "length lm = n \<Longrightarrow>  
+  abc_steps_l (Suc 0, lm @ Suc x # 0 # suf_lm) 
+       [Inc n, Dec (Suc n) 3, Goto (Suc 0)] (Suc (Suc 0))
+                                  = (3, lm @ Suc x # 0 # suf_lm)"
+apply(simp add: abc_steps_l.simps abc_fetch.simps 
+                abc_step_l.simps abc_lm_v.simps abc_lm_s.simps 
+                nth_append list_update_append)
+done
+
+lemma [simp]: 
+  "length lm = n \<Longrightarrow> 
+  abc_steps_l (Suc 0, lm @ Suc x # Suc y # suf_lm) 
+     [Inc n, Dec (Suc n) 3, Goto (Suc 0)] (Suc (Suc 0))
+  = (Suc 0, lm @ Suc x # y # suf_lm)"
+apply(simp add: abc_steps_l.simps abc_fetch.simps 
+                abc_step_l.simps abc_lm_v.simps abc_lm_s.simps 
+                nth_append list_update_append)
+done
+
+lemma pr_cycle_part_middle_inv: 
+  "\<lbrakk>length lm = n\<rbrakk> \<Longrightarrow> 
+  \<exists> stp. abc_steps_l (0, lm @ x # y # suf_lm) 
+                         [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp 
+  = (3, lm @ Suc x # 0 # suf_lm)"
+proof -
+  assume h: "length lm = n"
+  hence k1: "\<exists> stp. abc_steps_l (0, lm @ x # y # suf_lm) 
+                           [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp 
+    = (Suc 0, lm @ Suc x # y # suf_lm)"
+    apply(rule_tac x = "Suc 0" in exI)
+    apply(simp add: abc_steps_l.simps abc_step_l.simps 
+                    abc_lm_v.simps abc_lm_s.simps nth_append 
+                    list_update_append abc_fetch.simps)
+    done
+  from h have k2: 
+    "\<exists> stp. abc_steps_l (Suc 0, lm @ Suc x # y # suf_lm)
+                      [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp 
+    = (3, lm @ Suc x # 0 # suf_lm)"
+    apply(induct y)
+    apply(rule_tac x = "Suc (Suc 0)" in exI, simp, simp, 
+          erule_tac exE)
+    apply(rule_tac x = "Suc (Suc 0) + stp" in exI, 
+          simp only: abc_steps_add, simp)
+    done      
+  from k1 and k2 show 
+    "\<exists> stp. abc_steps_l (0, lm @ x # y # suf_lm) 
+                       [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp 
+    = (3, lm @ Suc x # 0 # suf_lm)"
+    apply(erule_tac exE, erule_tac exE)
+    apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
+    done
+qed
+
+lemma [simp]: 
+  "length lm = Suc n \<Longrightarrow> 
+  (abc_steps_l (length ap, lm @ x # Suc y # suf_lm) 
+           (ap @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length ap)]) 
+                    (Suc (Suc (Suc 0))))
+  = (length ap, lm @ Suc x # y # suf_lm)"
+apply(simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps 
+         abc_lm_v.simps list_update_append nth_append abc_lm_s.simps)
+done
+
+lemma switch_para_inv:
+  assumes bp_def:"bp =  ap @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto ss]"
+  and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" 
+         "ss = length ap" 
+         "length lm = Suc n"
+  shows " \<exists>stp. abc_steps_l (ss, lm @ x # y # suf_lm) bp stp =
+                               (0, lm @ (x + y) # 0 # suf_lm)"
+apply(induct y arbitrary: x)
+apply(rule_tac x = "Suc 0" in exI,
+  simp add: bp_def mv_box.simps abc_steps_l.simps 
+            abc_fetch.simps h abc_step_l.simps 
+            abc_lm_v.simps list_update_append nth_append
+            abc_lm_s.simps)
+proof -
+  fix y x
+  assume ind: 
+    "\<And>x. \<exists>stp. abc_steps_l (ss, lm @ x # y # suf_lm) bp stp = 
+                                     (0, lm @ (x + y) # 0 # suf_lm)"
+  show "\<exists>stp. abc_steps_l (ss, lm @ x # Suc y # suf_lm) bp stp = 
+                                  (0, lm @ (x + Suc y) # 0 # suf_lm)"
+    apply(insert ind[of "Suc x"], erule_tac exE)
+    apply(rule_tac x = "Suc (Suc (Suc 0)) + stp" in exI, 
+          simp only: abc_steps_add bp_def h)
+    apply(simp add: h)
+    done
+qed
+
+lemma [simp]:
+  "length lm = rs_pos \<and> Suc (Suc rs_pos) < a_md \<and> 0 < rs_pos \<Longrightarrow> 
+      a_md - Suc 0 < Suc (Suc (Suc (a_md + length suf_lm - 
+                                         Suc (Suc (Suc 0)))))"
+apply(arith)
+done
+
+lemma [simp]: 
+  "Suc (Suc rs_pos) < a_md \<and> 0 < rs_pos \<Longrightarrow> 
+                           \<not> a_md - Suc 0 < rs_pos - Suc 0"
+apply(arith)
+done
+
+lemma [simp]: 
+  "Suc (Suc rs_pos) < a_md \<and> 0 < rs_pos \<Longrightarrow> 
+           \<not> a_md - rs_pos < Suc (Suc (a_md - Suc (Suc rs_pos)))"
+apply(arith)
+done
+
+lemma butlast_append_last: "lm \<noteq> [] \<Longrightarrow> lm = butlast lm @ [last lm]"
+apply(auto)
+done
+
+lemma [simp]: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)
+           \<Longrightarrow> (Suc (Suc rs_pos)) < a_md"
+apply(simp add: rec_ci.simps)
+apply(case_tac "rec_ci f", simp)
+apply(case_tac "rec_ci g", simp)
+apply(arith)
+done
+
+(*
+lemma pr_para_ge_suc0: "rec_calc_rel (Pr n f g) lm xs \<Longrightarrow> 0 < n"
+apply(erule calc_pr_reverse, simp, simp)
+done
+*)
+
+lemma ci_pr_para_eq: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)
+                  \<Longrightarrow> rs_pos = Suc n"
+apply(simp add: rec_ci.simps)
+apply(case_tac "rec_ci g",  case_tac "rec_ci f", simp)
+done
+
+lemma [intro]:  
+  "\<lbrakk>rec_ci z = (aprog, rs_pos, a_md); rec_calc_rel z lm xs\<rbrakk>
+  \<Longrightarrow> length lm = rs_pos"
+apply(simp add: rec_ci.simps rec_ci_z_def)
+apply(erule_tac calc_z_reverse, simp)
+done
+
+lemma [intro]: 
+  "\<lbrakk>rec_ci s = (aprog, rs_pos, a_md); rec_calc_rel s lm xs\<rbrakk>
+  \<Longrightarrow> length lm = rs_pos"
+apply(simp add: rec_ci.simps rec_ci_s_def)
+apply(erule_tac calc_s_reverse, simp)
+done
+
+lemma [intro]: 
+  "\<lbrakk>rec_ci (recf.id nat1 nat2) = (aprog, rs_pos, a_md); 
+    rec_calc_rel (recf.id nat1 nat2) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
+apply(simp add: rec_ci.simps rec_ci_id.simps)
+apply(erule_tac calc_id_reverse, simp)
+done
+
+lemma [intro]: 
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
+    rec_calc_rel (Cn n f gs) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
+apply(erule_tac calc_cn_reverse, simp)
+apply(simp add: rec_ci.simps)
+apply(case_tac "rec_ci f",  simp)
+done
+
+lemma [intro]:
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
+    rec_calc_rel (Pr n f g) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
+apply(erule_tac  calc_pr_reverse, simp)
+apply(drule_tac ci_pr_para_eq, simp, simp)
+apply(drule_tac ci_pr_para_eq, simp)
+done
+
+lemma [intro]: 
+  "\<lbrakk>rec_ci (Mn n f) = (aprog, rs_pos, a_md);
+    rec_calc_rel (Mn n f) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
+apply(erule_tac calc_mn_reverse)
+apply(simp add: rec_ci.simps)
+apply(case_tac "rec_ci f",  simp)
+done
+
+lemma para_pattern: 
+  "\<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm xs\<rbrakk>
+  \<Longrightarrow> length lm = rs_pos"
+apply(case_tac f, auto)
+done
+
+lemma ci_pr_g_paras:
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
+    rec_ci g = (a, aa, ba);
+    rec_calc_rel (Pr n f g) (lm @ [x]) rs; x > 0\<rbrakk> \<Longrightarrow> 
+    aa = Suc rs_pos "
+apply(erule calc_pr_reverse, simp)
+apply(subgoal_tac "length (args @ [k, rk]) = aa", simp)
+apply(subgoal_tac "rs_pos = Suc n", simp)
+apply(simp add: ci_pr_para_eq)
+apply(erule para_pattern, simp)
+done
+
+lemma ci_pr_g_md_less: 
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
+    rec_ci g = (a, aa, ba)\<rbrakk> \<Longrightarrow> ba < a_md"
+apply(simp add: rec_ci.simps)
+apply(case_tac "rec_ci f",  auto)
+done
+
+lemma [intro]: "rec_ci z = (ap, rp, ad) \<Longrightarrow> rp < ad"
+  by(simp add: rec_ci.simps)
+
+lemma [intro]: "rec_ci s = (ap, rp, ad) \<Longrightarrow> rp < ad"
+  by(simp add: rec_ci.simps)
+
+lemma [intro]: "rec_ci (recf.id nat1 nat2) = (ap, rp, ad) \<Longrightarrow> rp < ad"
+  by(simp add: rec_ci.simps)
+
+lemma [intro]: "rec_ci (Cn n f gs) = (ap, rp, ad) \<Longrightarrow> rp < ad"
+apply(simp add: rec_ci.simps)
+apply(case_tac "rec_ci f",  simp)
+done
+
+lemma [intro]: "rec_ci (Pr n f g) = (ap, rp, ad) \<Longrightarrow> rp < ad"
+apply(simp add: rec_ci.simps)
+by(case_tac "rec_ci f", case_tac "rec_ci g",  auto)
+
+lemma [intro]: "rec_ci (Mn n f) = (ap, rp, ad) \<Longrightarrow> rp < ad"
+apply(simp add: rec_ci.simps)
+apply(case_tac "rec_ci f", simp)
+apply(arith)
+done
+
+lemma ci_ad_ge_paras: "rec_ci f = (ap, rp, ad) \<Longrightarrow> ad > rp"
+apply(case_tac f, auto)
+done
+
+lemma [elim]: "\<lbrakk>a [+] b = []; a \<noteq> [] \<or> b \<noteq> []\<rbrakk> \<Longrightarrow> RR"
+apply(auto simp: abc_append.simps abc_shift.simps)
+done
+
+lemma [intro]: "rec_ci z = ([], aa, ba) \<Longrightarrow> False"
+by(simp add: rec_ci.simps rec_ci_z_def)
+
+lemma [intro]: "rec_ci s = ([], aa, ba) \<Longrightarrow> False"
+by(auto simp: rec_ci.simps rec_ci_s_def addition.simps)
+
+lemma [intro]: "rec_ci (id m n) = ([], aa, ba) \<Longrightarrow> False"
+by(auto simp: rec_ci.simps rec_ci_id.simps addition.simps)
+
+lemma [intro]: "rec_ci (Cn n f gs) = ([], aa, ba) \<Longrightarrow> False"
+apply(case_tac "rec_ci f", auto simp: rec_ci.simps abc_append.simps)
+apply(simp add: abc_shift.simps mv_box.simps)
+done
+
+lemma [intro]: "rec_ci (Pr n f g) = ([], aa, ba) \<Longrightarrow> False"
+apply(simp add: rec_ci.simps)
+apply(case_tac "rec_ci f", case_tac "rec_ci g")
+by(auto)
+
+lemma [intro]: "rec_ci (Mn n f) = ([], aa, ba) \<Longrightarrow> False"
+apply(case_tac "rec_ci f", auto simp: rec_ci.simps)
+done
+
+lemma rec_ci_not_null:  "rec_ci g = (a, aa, ba) \<Longrightarrow> a \<noteq> []"
+by(case_tac g, auto)
+
+lemma calc_pr_g_def:
+ "\<lbrakk>rec_calc_rel (Pr rs_pos f g) (lm @ [Suc x]) rsa;
+   rec_calc_rel (Pr rs_pos f g) (lm @ [x]) rsxa\<rbrakk>
+ \<Longrightarrow> rec_calc_rel g (lm @ [x, rsxa]) rsa"
+apply(erule_tac calc_pr_reverse, simp, simp)
+apply(subgoal_tac "rsxa = rk", simp)
+apply(erule_tac rec_calc_inj, auto)
+done
+
+lemma ci_pr_md_def: 
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
+    rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\<rbrakk>
+  \<Longrightarrow> a_md = Suc (max (n + 3) (max bc ba))"
+by(simp add: rec_ci.simps)
+
+lemma  ci_pr_f_paras: 
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
+    rec_calc_rel (Pr n f g) lm rs;
+    rec_ci f = (ab, ac, bc)\<rbrakk>  \<Longrightarrow> ac = rs_pos - Suc 0"
+apply(subgoal_tac "\<exists>rs. rec_calc_rel f (butlast lm) rs", 
+      erule_tac exE)
+apply(drule_tac f = f and lm = "butlast lm" in para_pattern, 
+      simp, simp)
+apply(drule_tac para_pattern, simp)
+apply(subgoal_tac "lm \<noteq> []", simp)
+apply(erule_tac calc_pr_reverse, simp, simp)
+apply(erule calc_pr_zero_ex)
+done
+
+lemma ci_pr_md_ge_f:  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
+        rec_ci f = (ab, ac, bc)\<rbrakk> \<Longrightarrow> Suc bc \<le> a_md"
+apply(case_tac "rec_ci g")
+apply(simp add: rec_ci.simps, auto)
+done
+
+lemma ci_pr_md_ge_g:  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
+        rec_ci g = (ab, ac, bc)\<rbrakk> \<Longrightarrow> bc < a_md"
+apply(case_tac "rec_ci f")
+apply(simp add: rec_ci.simps, auto)
+done 
+
+lemma rec_calc_rel_def0: 
+  "\<lbrakk>rec_calc_rel (Pr n f g) lm rs; rec_calc_rel f (butlast lm) rsa\<rbrakk>
+  \<Longrightarrow> rec_calc_rel (Pr n f g) (butlast lm @ [0]) rsa"
+  apply(rule_tac calc_pr_zero, simp)
+apply(erule_tac calc_pr_reverse, simp, simp, simp)
+done
+
+lemma [simp]:  "length (mv_box m n) = 3"
+by (auto simp: mv_box.simps)
+
+
+lemma [simp]: "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); rec_calc_rel (Pr n f g) lm rs\<rbrakk>
+    \<Longrightarrow> rs_pos = Suc n"
+apply(simp add: ci_pr_para_eq)
+done
+
+
+lemma [simp]: "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); rec_calc_rel (Pr n f g) lm rs\<rbrakk>
+    \<Longrightarrow> length lm = Suc n"
+apply(subgoal_tac "rs_pos = Suc n", rule_tac para_pattern, simp, simp)
+apply(case_tac "rec_ci f", case_tac "rec_ci g", simp add: rec_ci.simps)
+done
+
+lemma [simp]: "rec_ci (Pr n f g) = (a, rs_pos, a_md) \<Longrightarrow> Suc (Suc n) < a_md"
+apply(case_tac "rec_ci f", case_tac "rec_ci g", simp add: rec_ci.simps)
+apply arith
+done
+
+lemma [simp]: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md) \<Longrightarrow> 0 < rs_pos"
+apply(case_tac "rec_ci f", case_tac "rec_ci g")
+apply(simp add: rec_ci.simps)
+done
+
+lemma [simp]: "Suc (Suc rs_pos) < a_md \<Longrightarrow> 
+       butlast lm @ (last lm - xa) # (rsa::nat) # 0 # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm =
+       butlast lm @ (last lm - xa) # rsa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm"
+apply(simp add: replicate_Suc[THEN sym])
+done
+
+lemma pr_cycle_part_ind: 
+  assumes g_ind: 
+  "\<And>lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow> 
+  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(ba - aa) @ suf_lm) a stp = 
+                    (length a, lm @ rs # 0\<up>(ba - Suc aa) @ suf_lm)"
+  and ap_def: 
+  "ap = ([Dec (a_md - Suc 0) (length a + 7)] [+]
+        (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)])) @
+         [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]"
+  and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" 
+         "rec_calc_rel (Pr n f g) 
+                   (butlast lm @ [last lm - Suc xa]) rsxa" 
+         "Suc xa \<le> last lm" 
+         "rec_ci g = (a, aa, ba)"
+         "rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rsa"
+         "lm \<noteq> []"
+  shows 
+  "\<exists>stp. abc_steps_l 
+     (0, butlast lm @ (last lm - Suc xa) # rsxa # 
+               0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) ap stp =
+     (0, butlast lm @ (last lm - xa) # rsa
+                 # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)"
+proof -
+  have k1: "\<exists>stp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) #
+    rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) ap stp =
+         (length a + 4, butlast lm @ (last lm - xa) # 0 # rsa #
+                           0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)"
+    apply(simp add: ap_def, rule_tac abc_add_exc1)
+    apply(rule_tac as = "Suc 0" and 
+      bm = "butlast lm @ (last lm - Suc xa) # 
+      rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm" in abc_append_exc2,
+      auto)
+  proof -
+    show 
+      "\<exists>astp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # rsxa 
+                   # 0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) 
+              [Dec (a_md - Suc 0)(length a + 7)] astp =
+      (Suc 0, butlast lm @ (last lm - Suc xa) # 
+             rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)"
+      apply(rule_tac x = "Suc 0" in exI, 
+          simp add: abc_steps_l.simps abc_step_l.simps
+                     abc_fetch.simps)
+      apply(subgoal_tac "length lm = Suc n \<and> rs_pos = Suc n \<and>
+                              a_md > Suc (Suc rs_pos)")
+      apply(simp add: abc_lm_v.simps nth_append abc_lm_s.simps)
+      apply(insert nth_append[of 
+                 "(last lm - Suc xa) # rsxa # 0\<up>(a_md - Suc (Suc rs_pos))" 
+                 "Suc xa # suf_lm" "(a_md - rs_pos)"], simp)
+      apply(simp add: list_update_append del: list_update.simps)
+      apply(insert list_update_append[of "(last lm - Suc xa) # rsxa # 
+                                          0\<up>(a_md - Suc (Suc rs_pos))" 
+                    "Suc xa # suf_lm" "a_md - rs_pos" "xa"], simp)
+      apply(case_tac a_md, simp, simp)
+      apply(insert h, simp)
+      apply(insert para_pattern[of "Pr n f g" aprog rs_pos a_md 
+                    "(butlast lm @ [last lm - Suc xa])" rsxa], simp)
+      done
+  next
+    show "\<exists>bstp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # 
+           rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm) (a [+] 
+            [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)]) bstp =
+         (3 + length a, butlast lm @ (last lm - xa) # 0 # rsa #
+                          0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)"
+      apply(rule_tac as = "length a" and
+               bm = "butlast lm @ (last lm - Suc xa) # rsxa # rsa #
+                     0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm" 
+        in abc_append_exc2, simp_all)
+    proof -
+      from h have j1: "aa = Suc rs_pos \<and> a_md > ba \<and> ba > Suc rs_pos"
+	apply(insert h)
+	apply(insert ci_pr_g_paras[of n f g aprog rs_pos
+                 a_md a aa ba "butlast lm" "last lm - xa" rsa], simp)
+	apply(drule_tac ci_pr_md_ge_g, auto)
+	apply(erule_tac ci_ad_ge_paras)
+	done
+      from h have j2: "rec_calc_rel g (butlast lm @ 
+                                  [last lm - Suc xa, rsxa]) rsa"
+	apply(rule_tac  calc_pr_g_def, simp, simp)
+	done
+      from j1 and j2 show 
+        "\<exists>astp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) #
+                rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm) a astp =
+        (length a, butlast lm @ (last lm - Suc xa) # rsxa # rsa 
+                         # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)"
+	apply(insert g_ind[of
+          "butlast lm @ (last lm - Suc xa) # [rsxa]" rsa 
+          "0\<up>(a_md - ba - Suc 0) @ xa # suf_lm"], simp, auto)
+	apply(simp add: exponent_add_iff)
+	apply(rule_tac x = stp in exI, simp add: numeral_3_eq_3)
+	done
+    next
+      from h have j3: "length lm = rs_pos \<and> rs_pos > 0"
+	apply(rule_tac conjI)
+	apply(drule_tac lm = "(butlast lm @ [last lm - Suc xa])"
+                          and xs = rsxa in para_pattern, simp, simp, simp)
+        done
+      from h have j4: "Suc (last lm - Suc xa) = last lm - xa"
+	apply(case_tac "last lm", simp, simp)
+	done
+      from j3 and j4 show
+      "\<exists>bstp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # rsxa #
+                     rsa # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)
+            [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)] bstp =
+        (3, butlast lm @ (last lm - xa) # 0 # rsa #
+                       0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)"
+	apply(insert pr_cycle_part_middle_inv[of "butlast lm" 
+          "rs_pos - Suc 0" "(last lm - Suc xa)" rsxa 
+          "rsa # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm"], simp)
+	done
+    qed
+  qed
+  from h have k2: 
+    "\<exists>stp. abc_steps_l (length a + 4, butlast lm @ (last lm - xa) # 0 
+           # rsa # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm) ap stp =
+    (0, butlast lm @ (last lm - xa) # rsa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)"
+    apply(insert switch_para_inv[of ap 
+      "([Dec (a_md - Suc 0) (length a + 7)] [+] 
+      (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)]))"
+      n "length a + 4" f g aprog rs_pos a_md 
+      "butlast lm @ [last lm - xa]" 0 rsa 
+      "0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm"])
+    apply(simp add: h ap_def)
+    apply(subgoal_tac "length lm = Suc n \<and> Suc (Suc rs_pos) < a_md", 
+          simp)
+    apply(insert h, simp)
+    apply(frule_tac lm = "(butlast lm @ [last lm - Suc xa])" 
+      and xs = rsxa in para_pattern, simp, simp)
+    done   
+  from k1 and k2 show "?thesis"
+    apply(auto)
+    apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
+    done
+qed
+
+lemma ci_pr_ex1: 
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
+    rec_ci g = (a, aa, ba);
+    rec_ci f = (ab, ac, bc)\<rbrakk>
+\<Longrightarrow> \<exists>ap bp. length ap = 6 + length ab \<and>
+    aprog = ap [+] bp \<and>
+    bp = ([Dec (a_md - Suc 0) (length a + 7)] [+] (a [+]
+         [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)])) @ 
+         [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "Recursive.mv_box n (max (Suc (Suc (Suc n)))
+    (max bc ba)) [+] ab [+] Recursive.mv_box n (Suc n)" in exI,
+     simp)
+apply(auto simp add: abc_append_commute numeral_3_eq_3)
+done
+
+lemma pr_cycle_part:
+  "\<lbrakk>\<And>lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow>
+     \<exists>stp. abc_steps_l (0, lm @ 0\<up>(ba - aa) @ suf_lm) a stp = 
+                        (length a, lm @ rs # 0\<up>(ba - Suc aa) @ suf_lm);
+  rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
+  rec_calc_rel (Pr n f g) lm rs;
+  rec_ci g = (a, aa, ba);
+  rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rsx;
+  rec_ci f = (ab, ac, bc);
+  lm \<noteq> [];
+  x \<le> last lm\<rbrakk> \<Longrightarrow> 
+  \<exists>stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - x) #
+              rsx # 0\<up>(a_md - Suc (Suc rs_pos)) @ x # suf_lm) aprog stp =
+  (6 + length ab, butlast lm @ last lm # rs #
+                                0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)"
+proof -
+  assume g_ind:
+    "\<And>lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow> 
+    \<exists>stp. abc_steps_l (0, lm @ 0\<up>(ba - aa) @ suf_lm) a stp =
+                      (length a, lm @ rs # 0\<up>(ba - Suc aa) @ suf_lm)"
+    and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" 
+           "rec_calc_rel (Pr n f g) lm rs" 
+           "rec_ci g = (a, aa, ba)"
+           "rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rsx" 
+           "lm \<noteq> []"
+           "x \<le> last lm" 
+           "rec_ci f = (ab, ac, bc)" 
+  from h show 
+    "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - x) # 
+            rsx # 0\<up>(a_md - Suc (Suc rs_pos)) @ x # suf_lm) aprog stp =
+    (6 + length ab, butlast lm @ last lm # rs #
+                               0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)" 
+  proof(induct x arbitrary: rsx, simp_all)
+    fix rsxa
+    assume "rec_calc_rel (Pr n f g) lm rsxa" 
+           "rec_calc_rel (Pr n f g) lm rs"
+    from h and this have "rs = rsxa"
+      apply(subgoal_tac "lm \<noteq> [] \<and> rs_pos = Suc n", simp)
+      apply(rule_tac rec_calc_inj, simp, simp)
+      apply(simp)
+      done
+    thus "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @  last lm # 
+             rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm) aprog stp =
+      (6 + length ab, butlast lm @ last lm # rs #
+                               0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)"
+      by(rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
+  next
+    fix xa rsxa
+    assume ind:
+   "\<And>rsx. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rsx 
+  \<Longrightarrow> \<exists>stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - xa) #
+             rsx # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm) aprog stp =
+      (6 + length ab, butlast lm @ last lm # rs # 
+                               0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)"
+      and g: "rec_calc_rel (Pr n f g) 
+                      (butlast lm @ [last lm - Suc xa]) rsxa"
+      "Suc xa \<le> last lm"
+      "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" 
+      "rec_calc_rel (Pr n f g) lm rs"
+      "rec_ci g = (a, aa, ba)" 
+      "rec_ci f = (ab, ac, bc)" "lm \<noteq> []"
+    from g have k1: 
+      "\<exists> rs. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rs"
+      apply(rule_tac rs = rs in  calc_pr_less_ex, simp, simp)
+      done
+    from g and this show 
+      "\<exists>stp. abc_steps_l (6 + length ab, 
+           butlast lm @ (last lm - Suc xa) # rsxa # 
+              0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) aprog stp =
+              (6 + length ab, butlast lm @ last lm # rs # 
+                                0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)"
+    proof(erule_tac exE)
+      fix rsa
+      assume k2: "rec_calc_rel (Pr n f g) 
+                           (butlast lm @ [last lm - xa]) rsa"
+      from g and k2 have
+      "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @ 
+       (last lm - Suc xa) # rsxa # 
+               0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) aprog stp
+        = (6 + length ab, butlast lm @ (last lm - xa) # rsa # 
+                               0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)"
+	proof -
+	  from g have k2_1: 
+            "\<exists> ap bp. length ap = 6 + length ab \<and>
+                   aprog = ap [+] bp \<and> 
+                   bp = ([Dec (a_md - Suc 0) (length a + 7)] [+]
+                  (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, 
+                  Goto (Suc 0)])) @
+                  [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]"
+            apply(rule_tac ci_pr_ex1, auto)
+	    done
+	  from k2_1 and k2 and g show "?thesis"
+	    proof(erule_tac exE, erule_tac exE)
+	      fix ap bp
+	      assume 
+                "length ap = 6 + length ab \<and> 
+                 aprog = ap [+] bp \<and> bp =
+                ([Dec (a_md - Suc 0) (length a + 7)] [+] 
+                (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3,
+                Goto (Suc 0)])) @ 
+                [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" 
+	      from g and this and k2 and g_ind show "?thesis"
+		apply(insert abc_append_exc3[of 
+                  "butlast lm @ (last lm - Suc xa) # rsxa #
+                  0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm" bp 0
+                  "butlast lm @ (last lm - xa) # rsa #
+                0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm" "length ap" ap],
+                 simp)
+		apply(subgoal_tac 
+                "\<exists>stp. abc_steps_l (0, butlast lm @ (last lm - Suc xa)
+                           # rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # 
+                              suf_lm) bp stp =
+	          (0, butlast lm @ (last lm - xa) # rsa #
+                           0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)",
+                      simp, erule_tac conjE, erule conjE)
+		apply(erule pr_cycle_part_ind, auto)
+		done
+	    qed
+	  qed  
+      from g and k2 and this show "?thesis"
+	apply(erule_tac exE)
+	apply(insert ind[of rsa], simp)
+	apply(erule_tac exE)
+	apply(rule_tac x = "stp + stpa" in exI, 
+              simp add: abc_steps_add)
+	done
+    qed
+  qed
+qed
+
+lemma ci_pr_length: 
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
+    rec_ci g = (a, aa, ba);  
+    rec_ci f = (ab, ac, bc)\<rbrakk>
+    \<Longrightarrow>  length aprog = 13 + length ab + length a"
+apply(auto simp: rec_ci.simps)
+done
+
+fun mv_box_inv :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> bool"
+  where
+  "mv_box_inv (as, lm) m n initlm = 
+         (let plus = initlm ! m + initlm ! n in
+           length initlm > max m n \<and> m \<noteq> n \<and> 
+              (if as = 0 then \<exists> k l. lm = initlm[m := k, n := l] \<and> 
+                    k + l = plus \<and> k \<le> initlm ! m 
+              else if as = 1 then \<exists> k l. lm = initlm[m := k, n := l]
+                             \<and> k + l + 1 = plus \<and> k < initlm ! m 
+              else if as = 2 then \<exists> k l. lm = initlm[m := k, n := l] 
+                              \<and> k + l = plus \<and> k \<le> initlm ! m
+              else if as = 3 then lm = initlm[m := 0, n := plus]
+              else False))"
+
+fun mv_box_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "mv_box_stage1 (as, lm) m  = 
+            (if as = 3 then 0 
+             else 1)"
+
+fun mv_box_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "mv_box_stage2 (as, lm) m = (lm ! m)"
+
+fun mv_box_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "mv_box_stage3 (as, lm) m = (if as = 1 then 3 
+                                else if as = 2 then 2
+                                else if as = 0 then 1 
+                                else 0)"
+ 
+fun mv_box_measure :: "((nat \<times> nat list) \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)"
+  where
+  "mv_box_measure ((as, lm), m) = 
+     (mv_box_stage1 (as, lm) m, mv_box_stage2 (as, lm) m,
+      mv_box_stage3 (as, lm) m)"
+
+definition lex_pair :: "((nat \<times> nat) \<times> nat \<times> nat) set"
+  where
+  "lex_pair = less_than <*lex*> less_than"
+
+definition lex_triple :: 
+ "((nat \<times> (nat \<times> nat)) \<times> (nat \<times> (nat \<times> nat))) set"
+  where
+  "lex_triple \<equiv> less_than <*lex*> lex_pair"
+
+definition mv_box_LE :: 
+ "(((nat \<times> nat list) \<times> nat) \<times> ((nat \<times> nat list) \<times> nat)) set"
+  where 
+  "mv_box_LE \<equiv> (inv_image lex_triple mv_box_measure)"
+
+lemma wf_lex_triple: "wf lex_triple"
+  by (auto intro:wf_lex_prod simp:lex_triple_def lex_pair_def)
+
+lemma wf_mv_box_le[intro]: "wf mv_box_LE"
+by(auto intro:wf_inv_image wf_lex_triple simp: mv_box_LE_def)
+
+declare mv_box_inv.simps[simp del]
+
+lemma mv_box_inv_init:  
+"\<lbrakk>m < length initlm; n < length initlm; m \<noteq> n\<rbrakk> \<Longrightarrow> 
+  mv_box_inv (0, initlm) m n initlm"
+apply(simp add: abc_steps_l.simps mv_box_inv.simps)
+apply(rule_tac x = "initlm ! m" in exI, 
+      rule_tac x = "initlm ! n" in exI, simp)
+done
+
+lemma [simp]: "abc_fetch 0 (Recursive.mv_box m n) = Some (Dec m 3)"
+apply(simp add: mv_box.simps abc_fetch.simps)
+done
+
+lemma [simp]: "abc_fetch (Suc 0) (Recursive.mv_box m n) =
+               Some (Inc n)"
+apply(simp add: mv_box.simps abc_fetch.simps)
+done
+
+lemma [simp]: "abc_fetch 2 (Recursive.mv_box m n) = Some (Goto 0)"
+apply(simp add: mv_box.simps abc_fetch.simps)
+done
+
+lemma [simp]: "abc_fetch 3 (Recursive.mv_box m n) = None"
+apply(simp add: mv_box.simps abc_fetch.simps)
+done
+
+lemma [simp]: 
+  "\<lbrakk>m \<noteq> n; m < length initlm; n < length initlm;
+    k + l = initlm ! m + initlm ! n; k \<le> initlm ! m; 0 < k\<rbrakk>
+ \<Longrightarrow> \<exists>ka la. initlm[m := k, n := l, m := k - Suc 0] = 
+     initlm[m := ka, n := la] \<and>
+     Suc (ka + la) = initlm ! m + initlm ! n \<and> 
+     ka < initlm ! m"
+apply(rule_tac x = "k - Suc 0" in exI, rule_tac x = l in exI, 
+      simp, auto)
+apply(subgoal_tac 
+      "initlm[m := k, n := l, m := k - Suc 0] = 
+       initlm[n := l, m := k, m := k - Suc 0]")
+apply(simp add: list_update_overwrite )
+apply(simp add: list_update_swap)
+apply(simp add: list_update_swap)
+done
+
+lemma [simp]:
+  "\<lbrakk>m \<noteq> n; m < length initlm; n < length initlm; 
+    Suc (k + l) = initlm ! m + initlm ! n;
+    k < initlm ! m\<rbrakk>
+    \<Longrightarrow> \<exists>ka la. initlm[m := k, n := l, n := Suc l] = 
+                initlm[m := ka, n := la] \<and> 
+                ka + la = initlm ! m + initlm ! n \<and> 
+                ka \<le> initlm ! m"
+apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, auto)
+done
+
+lemma [simp]: 
+  "\<lbrakk>length initlm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
+   \<forall>na. \<not> (\<lambda>(as, lm) m. as = 3) 
+    (abc_steps_l (0, initlm) (Recursive.mv_box m n) na) m \<and> 
+  mv_box_inv (abc_steps_l (0, initlm) 
+           (Recursive.mv_box m n) na) m n initlm \<longrightarrow>
+  mv_box_inv (abc_steps_l (0, initlm) 
+           (Recursive.mv_box m n) (Suc na)) m n initlm \<and>
+  ((abc_steps_l (0, initlm) (Recursive.mv_box m n) (Suc na), m),
+   abc_steps_l (0, initlm) (Recursive.mv_box m n) na, m) \<in> mv_box_LE"
+apply(rule allI, rule impI, simp add: abc_steps_ind)
+apply(case_tac "(abc_steps_l (0, initlm) (Recursive.mv_box m n) na)",
+      simp)
+apply(auto split:if_splits simp add:abc_steps_l.simps mv_box_inv.simps)
+apply(auto simp add: mv_box_LE_def lex_triple_def lex_pair_def 
+                     abc_step_l.simps abc_steps_l.simps
+                     mv_box_inv.simps abc_lm_v.simps abc_lm_s.simps
+                split: if_splits )
+apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, simp)
+done
+
+lemma mv_box_inv_halt: 
+  "\<lbrakk>length initlm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
+  \<exists> stp. (\<lambda> (as, lm). as = 3 \<and> 
+  mv_box_inv (as, lm) m n initlm) 
+             (abc_steps_l (0::nat, initlm) (mv_box m n) stp)"
+thm halt_lemma2
+apply(insert halt_lemma2[of mv_box_LE
+    "\<lambda> ((as, lm), m). mv_box_inv (as, lm) m n initlm"
+    "\<lambda> stp. (abc_steps_l (0, initlm) (Recursive.mv_box m n) stp, m)"
+    "\<lambda> ((as, lm), m). as = (3::nat)"
+    ])
+apply(insert wf_mv_box_le)
+apply(simp add: mv_box_inv_init abc_steps_zero)
+apply(erule_tac exE)
+apply(rule_tac x = na in exI)
+apply(case_tac "(abc_steps_l (0, initlm) (Recursive.mv_box m n) na)",
+      simp, auto)
+done
+
+lemma mv_box_halt_cond:
+  "\<lbrakk>m \<noteq> n; mv_box_inv (a, b) m n lm; a = 3\<rbrakk> \<Longrightarrow> 
+  b = lm[n := lm ! m + lm ! n, m := 0]"
+apply(simp add: mv_box_inv.simps, auto)
+apply(simp add: list_update_swap)
+done
+
+lemma mv_box_ex:
+  "\<lbrakk>length lm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
+  \<exists> stp. abc_steps_l (0::nat, lm) (mv_box m n) stp
+  = (3, (lm[n := (lm ! m + lm ! n)])[m := 0::nat])"
+apply(drule mv_box_inv_halt, simp, erule_tac exE)
+apply(rule_tac x = stp in exI)
+apply(case_tac "abc_steps_l (0, lm) (Recursive.mv_box m n) stp",
+      simp)
+apply(erule_tac mv_box_halt_cond, auto)
+done
+
+lemma [simp]: 
+  "\<lbrakk>a_md = Suc (max (Suc (Suc n)) (max bc ba)); 
+   length lm = rs_pos \<and> rs_pos = n \<and> n > 0\<rbrakk>
+  \<Longrightarrow> n - Suc 0 < length lm + 
+  (Suc (max (Suc (Suc n)) (max bc ba)) - rs_pos + length suf_lm) \<and>
+   Suc (Suc n) < length lm + (Suc (max (Suc (Suc n)) (max bc ba)) -
+  rs_pos + length suf_lm) \<and> bc < length lm + (Suc (max (Suc (Suc n)) 
+ (max bc ba)) - rs_pos + length suf_lm) \<and> ba < length lm + 
+  (Suc (max (Suc (Suc n)) (max bc ba)) - rs_pos + length suf_lm)"
+apply(arith)
+done
+
+lemma [simp]:
+  "\<lbrakk>a_md = Suc (max (Suc (Suc n)) (max bc ba)); 
+   length lm = rs_pos \<and> rs_pos = n \<and> n > 0\<rbrakk>
+ \<Longrightarrow> n - Suc 0 < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba)) \<and>
+     Suc n < length suf_lm + max (Suc (Suc n)) (max bc ba) \<and> 
+     bc < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba)) \<and> 
+     ba < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba))"
+apply(arith)
+done
+
+lemma [simp]: "n - Suc 0 \<noteq> max (Suc (Suc n)) (max bc ba)"
+apply(arith)
+done
+
+lemma [simp]: 
+  "a_md \<ge> Suc bc \<and> rs_pos > 0 \<and> bc \<ge> rs_pos \<Longrightarrow> 
+ bc - (rs_pos - Suc 0) + a_md - Suc bc = Suc (a_md - rs_pos - Suc 0)"
+apply(arith)
+done
+
+lemma [simp]: "length lm = n \<and> rs_pos = n \<and> 0 < rs_pos \<and> 
+                                                  Suc rs_pos < a_md 
+       \<Longrightarrow> n - Suc 0 < Suc (Suc (a_md + length suf_lm - Suc (Suc 0))) 
+        \<and> n < Suc (Suc (a_md + length suf_lm - Suc (Suc 0)))"
+apply(arith)
+done
+     
+lemma [simp]: "length lm = n \<and> rs_pos = n \<and> 0 < rs_pos \<and> 
+               Suc rs_pos < a_md \<Longrightarrow> n - Suc 0 \<noteq> n"
+by arith
+
+lemma ci_pr_ex2: 
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
+    rec_calc_rel (Pr n f g) lm rs; 
+    rec_ci g = (a, aa, ba); 
+    rec_ci f = (ab, ac, bc)\<rbrakk>
+  \<Longrightarrow> \<exists>ap bp. aprog = ap [+] bp \<and> 
+         ap = mv_box n (max (Suc (Suc (Suc n))) (max bc ba))"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "(ab [+] (Recursive.mv_box n (Suc n) [+]
+              ([Dec (max (n + 3) (max bc ba)) (length a + 7)] 
+      [+] (a [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @ 
+      [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]))" in exI, auto)
+apply(simp add: abc_append_commute numeral_3_eq_3)
+done
+
+lemma [simp]: 
+  "max (Suc (Suc (Suc n))) (max bc ba) - n < 
+     Suc (max (Suc (Suc (Suc n))) (max bc ba)) - n"
+apply(arith)
+done
+
+thm nth_replicate
+(*
+lemma exp_nth[simp]: "n < m \<Longrightarrow> a\<up>m ! n = a"
+apply(sim)
+done
+*)
+lemma [simp]: "length lm = n \<and> rs_pos = n \<and> 0 < n \<Longrightarrow> 
+                      lm[n - Suc 0 := 0::nat] = butlast lm @ [0]"
+apply(auto)
+apply(insert list_update_append[of "butlast lm" "[last lm]" 
+                                   "length lm - Suc 0" "0"], simp)
+done
+
+lemma [simp]: "\<lbrakk>length lm = n; 0 < n\<rbrakk>  \<Longrightarrow> lm ! (n - Suc 0) = last lm"
+apply(insert nth_append[of "butlast lm" "[last lm]" "n - Suc 0"],
+      simp)
+apply(insert butlast_append_last[of lm], auto)
+done
+lemma exp_suc_iff: "a\<up>b @ [a] = a\<up>(b + Suc 0)"
+apply(simp add: exp_ind del: replicate.simps)
+done
+
+lemma less_not_less[simp]: "n > 0 \<Longrightarrow> \<not> n < n - Suc 0"
+by auto
+
+lemma [simp]:
+  "Suc n < length suf_lm + max (Suc (Suc n)) (max bc ba) \<and> 
+  bc < Suc (length suf_lm + max (Suc (Suc n)) 
+  (max bc ba)) \<and> 
+  ba < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba))"
+  by arith
+
+lemma [simp]: "length lm = n \<and> rs_pos = n \<and> n > 0 \<Longrightarrow> 
+(lm @ 0\<up>(Suc (max (Suc (Suc n)) (max bc ba)) - n) @ suf_lm) 
+  [max (Suc (Suc n)) (max bc ba) :=
+   (lm @ 0\<up>(Suc (max (Suc (Suc n)) (max bc ba)) - n) @ suf_lm) ! (n - Suc 0) + 
+       (lm @ 0\<up>(Suc (max (Suc (Suc n)) (max bc ba)) - n) @ suf_lm) ! 
+                   max (Suc (Suc n)) (max bc ba), n - Suc 0 := 0::nat]
+ = butlast lm @ 0 # 0\<up>(max (Suc (Suc n)) (max bc ba) - n) @ last lm # suf_lm"
+apply(simp add: nth_append nth_replicate list_update_append)
+apply(insert list_update_append[of "0\<up>((max (Suc (Suc n)) (max bc ba)) - n)"
+         "[0]" "max (Suc (Suc n)) (max bc ba) - n" "last lm"], simp)
+apply(simp add: exp_suc_iff Suc_diff_le del: list_update.simps)
+done
+
+lemma exp_eq: "(a = b) = (c\<up>a = c\<up>b)"
+apply(auto)
+done
+
+lemma [simp]:
+  "\<lbrakk>length lm = n; 0 < n;  Suc n < a_md\<rbrakk> \<Longrightarrow> 
+   (butlast lm @ rsa # 0\<up>(a_md - Suc n) @ last lm # suf_lm)
+    [n := (butlast lm @ rsa # 0\<up>(a_md - Suc n) @ last lm # suf_lm) ! 
+        (n - Suc 0) + (butlast lm @ rsa # (0::nat)\<up>(a_md - Suc n) @ 
+                                last lm # suf_lm) ! n, n - Suc 0 := 0]
+ = butlast lm @ 0 # rsa # 0\<up>(a_md - Suc (Suc n)) @ last lm # suf_lm"
+apply(simp add: nth_append list_update_append)
+apply(case_tac "a_md - Suc n", auto)
+done
+
+lemma [simp]: 
+  "Suc (Suc rs_pos) \<le> a_md \<and> length lm = rs_pos \<and> 0 < rs_pos
+  \<Longrightarrow> a_md - Suc 0 < 
+          Suc (Suc (Suc (a_md + length suf_lm - Suc (Suc (Suc 0)))))"
+by arith
+
+lemma [simp]: 
+  "Suc (Suc rs_pos) \<le> a_md \<and> length lm = rs_pos \<and> 0 < rs_pos \<Longrightarrow> 
+                                   \<not> a_md - Suc 0 < rs_pos - Suc 0"
+by arith
+
+lemma [simp]: "Suc (Suc rs_pos) \<le> a_md \<Longrightarrow> 
+                                \<not> a_md - Suc 0 < rs_pos - Suc 0"
+by arith
+
+lemma [simp]: "\<lbrakk>Suc (Suc rs_pos) \<le> a_md\<rbrakk> \<Longrightarrow> 
+               \<not> a_md - rs_pos < Suc (Suc (a_md - Suc (Suc rs_pos)))"
+by arith 
+
+lemma [simp]: 
+  "Suc (Suc rs_pos) \<le> a_md \<and> length lm = rs_pos \<and> 0 < rs_pos
+ \<Longrightarrow> (abc_lm_v (butlast lm @ last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos)) @
+        0 # suf_lm) (a_md - Suc 0) = 0 \<longrightarrow>
+      abc_lm_s (butlast lm @ last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos)) @ 
+        0 # suf_lm) (a_md - Suc 0) 0 = 
+         lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) \<and>
+     abc_lm_v (butlast lm @ last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos)) @ 
+               0 # suf_lm) (a_md - Suc 0) = 0"
+apply(simp add: abc_lm_v.simps nth_append abc_lm_s.simps)
+apply(insert nth_append[of "last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos))" 
+               "0 # suf_lm" "(a_md - rs_pos)"], auto)
+apply(simp only: exp_suc_iff)
+apply(subgoal_tac "a_md - Suc 0 < a_md + length suf_lm", simp)
+apply(case_tac "lm = []", auto)
+done
+
+lemma pr_prog_ex[simp]: "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
+      rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\<rbrakk>
+    \<Longrightarrow> \<exists>cp. aprog = Recursive.mv_box n (max (n + 3) 
+                    (max bc ba)) [+] cp"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "(ab [+] (Recursive.mv_box n (Suc n) [+]
+              ([Dec (max (n + 3) (max bc ba)) (length a + 7)] 
+             [+] (a [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]))
+             @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]))" in exI)
+apply(auto simp: abc_append_commute)
+done
+
+lemma [simp]: "mv_box m n \<noteq> []"
+by (simp add: mv_box.simps)
+(*
+lemma [simp]: "\<lbrakk>rs_pos = n; 0 < rs_pos ; Suc rs_pos < a_md\<rbrakk> \<Longrightarrow> 
+                        n - Suc 0 < a_md + length suf_lm"
+by arith
+*)
+lemma [intro]: 
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
+    rec_ci f = (ab, ac, bc)\<rbrakk> \<Longrightarrow> 
+   \<exists>ap. (\<exists>cp. aprog = ap [+] ab [+] cp) \<and> length ap = 3"
+apply(case_tac "rec_ci g", simp add: rec_ci.simps)
+apply(rule_tac x = "mv_box n 
+              (max (n + 3) (max bc c))" in exI, simp)
+apply(rule_tac x = "Recursive.mv_box n (Suc n) [+]
+                 ([Dec (max (n + 3) (max bc c)) (length a + 7)]
+                 [+] a [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])
+               @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" in exI, 
+      auto)
+apply(simp add: abc_append_commute)
+done
+
+lemma [intro]: 
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
+    rec_ci g = (a, aa, ba); 
+    rec_ci f = (ab, ac, bc)\<rbrakk> \<Longrightarrow> 
+    \<exists>ap. (\<exists>cp. aprog = ap [+] Recursive.mv_box n (Suc n) [+] cp)
+      \<and> length ap = 3 + length ab"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "Recursive.mv_box n (max (n + 3)
+                                (max bc ba)) [+] ab" in exI, simp)
+apply(rule_tac x = "([Dec (max (n + 3) (max bc ba))
+  (length a + 7)] [+] a [+] 
+  [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ 
+  [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" in exI)
+apply(auto simp: abc_append_commute)
+done
+
+lemma [intro]: 
+  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
+    rec_ci g = (a, aa, ba); 
+    rec_ci f = (ab, ac, bc)\<rbrakk>
+    \<Longrightarrow> \<exists>ap. (\<exists>cp. aprog = ap [+] ([Dec (a_md - Suc 0) (length a + 7)]
+             [+] (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, 
+             Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n),
+             Goto (length a + 4)] [+] cp) \<and>
+             length ap = 6 + length ab"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "Recursive.mv_box n
+    (max (n + 3) (max bc ba)) [+] ab [+] 
+     Recursive.mv_box n (Suc n)" in exI, simp)
+apply(rule_tac x = "[]" in exI, auto)
+apply(simp add: abc_append_commute)
+done
+
+lemma [simp]: 
+  "n < Suc (max (n + 3) (max bc ba) + length suf_lm) \<and> 
+   Suc (Suc n) < max (n + 3) (max bc ba) + length suf_lm \<and> 
+   bc < Suc (max (n + 3) (max bc ba) + length suf_lm) \<and> 
+   ba < Suc (max (n + 3) (max bc ba) + length suf_lm)"
+by arith
+
+lemma [simp]: "n \<noteq> max (n + (3::nat)) (max bc ba)"
+by arith
+
+lemma [simp]:"length lm = Suc n \<Longrightarrow> lm[n := (0::nat)] = butlast lm @ [0]"
+apply(subgoal_tac "\<exists> xs x. lm = xs @ [x]", auto simp: list_update_append)
+apply(rule_tac x = "butlast lm" in exI, rule_tac x = "last lm" in exI)
+apply(case_tac lm, auto)
+done
+
+lemma [simp]:  "length lm = Suc n \<Longrightarrow> lm ! n =last lm"
+apply(subgoal_tac "lm \<noteq> []")
+apply(simp add: last_conv_nth, case_tac lm, simp_all)
+done
+
+lemma [simp]: "length lm = Suc n \<Longrightarrow> 
+      (lm @ (0::nat)\<up>(max (n + 3) (max bc ba) - n) @ suf_lm)
+           [max (n + 3) (max bc ba) := (lm @ 0\<up>(max (n + 3) (max bc ba) - n) @ suf_lm) ! n + 
+                  (lm @ 0\<up>(max (n + 3) (max bc ba) - n) @ suf_lm) ! max (n + 3) (max bc ba), n := 0]
+       = butlast lm @ 0 # 0\<up>(max (n + 3) (max bc ba) - Suc n) @ last lm # suf_lm"
+apply(auto simp: list_update_append nth_append)
+apply(subgoal_tac "(0\<up>(max (n + 3) (max bc ba) - n)) = 0\<up>(max (n + 3) (max bc ba) - Suc n) @ [0::nat]")
+apply(simp add: list_update_append)
+apply(simp add: exp_suc_iff)
+done
+
+lemma [simp]: "Suc (Suc n) < a_md \<Longrightarrow>  
+      n < Suc (Suc (a_md + length suf_lm - 2)) \<and>
+        n < Suc (a_md + length suf_lm - 2)"
+by(arith)
+
+lemma [simp]: "\<lbrakk>length lm = Suc n; Suc (Suc n) < a_md\<rbrakk>
+        \<Longrightarrow>(butlast lm @ (rsa::nat) # 0\<up>(a_md - Suc (Suc n)) @ last lm # suf_lm)
+          [Suc n := (butlast lm @ rsa # 0\<up>(a_md - Suc (Suc n)) @ last lm # suf_lm) ! n +
+                  (butlast lm @ rsa # 0\<up>(a_md - Suc (Suc n)) @ last lm # suf_lm) ! Suc n, n := 0]
+    = butlast lm @ 0 # rsa # 0\<up>(a_md - Suc (Suc (Suc n))) @ last lm # suf_lm"
+apply(auto simp: list_update_append)
+apply(subgoal_tac "(0\<up>(a_md - Suc (Suc n))) = (0::nat) # (0\<up>(a_md - Suc (Suc (Suc n))))", simp add: nth_append)
+apply(simp add: replicate_Suc[THEN sym])
+done
+
+lemma pr_case:
+  assumes nf_ind:
+  "\<And> lm rs suf_lm. rec_calc_rel f lm rs \<Longrightarrow> 
+  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(bc - ac) @ suf_lm) ab stp = 
+                (length ab, lm @ rs # 0\<up>(bc - Suc ac) @ suf_lm)"
+  and ng_ind: "\<And> lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow> 
+        \<exists>stp. abc_steps_l (0, lm @ 0\<up>(ba - aa) @ suf_lm) a stp = 
+                       (length a, lm @ rs # 0\<up>(ba - Suc aa) @ suf_lm)"
+    and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)"  "rec_calc_rel (Pr n f g) lm rs" 
+           "rec_ci g = (a, aa, ba)" "rec_ci f = (ab, ac, bc)" 
+  shows "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+proof -
+  from h have k1: "\<exists> stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp
+    = (3, butlast lm @ 0 # 0\<up>(a_md - rs_pos - 1) @ last lm # suf_lm)"
+  proof -
+    have "\<exists>bp cp. aprog = bp [+] cp \<and> bp = mv_box n 
+                 (max (n + 3) (max bc ba))"
+      apply(insert h, simp)
+      apply(erule pr_prog_ex, auto)
+      done
+    thus "?thesis"
+      apply(erule_tac exE, erule_tac exE, simp)
+      apply(subgoal_tac 
+           "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm)
+              ([] [+] Recursive.mv_box n
+                  (max (n + 3) (max bc ba)) [+] cp) stp =
+             (0 + 3, butlast lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ 
+                                        last lm # suf_lm)", simp)
+      apply(rule_tac abc_append_exc1, simp_all)
+      apply(insert mv_box_ex[of "n" "(max (n + 3) 
+                 (max bc ba))" "lm @ 0\<up>(a_md - rs_pos) @ suf_lm"], simp)
+      apply(subgoal_tac "a_md = Suc (max (n + 3) (max bc ba))",
+            simp)
+      apply(subgoal_tac "length lm = Suc n \<and> rs_pos = Suc n", simp)
+      apply(insert h)
+      apply(simp add: para_pattern ci_pr_para_eq)
+      apply(rule ci_pr_md_def, auto)
+      done
+  qed
+  from h have k2: 
+  "\<exists> stp. abc_steps_l (3,  butlast lm @ 0 # 0\<up>(a_md - rs_pos - 1) @ 
+             last lm # suf_lm) aprog stp 
+    = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+  proof -
+    from h have k2_1: "\<exists> rs. rec_calc_rel f (butlast lm) rs"
+      apply(erule_tac calc_pr_zero_ex)
+      done
+    thus "?thesis"
+    proof(erule_tac exE)
+      fix rsa
+      assume k2_2: "rec_calc_rel f (butlast lm) rsa"
+      from h and k2_2 have k2_2_1: 
+       "\<exists> stp. abc_steps_l (3, butlast lm @ 0 # 0\<up>(a_md - rs_pos - 1) 
+                 @ last lm # suf_lm) aprog stp
+        = (3 + length ab, butlast lm @ rsa # 0\<up>(a_md - rs_pos - 1) @ 
+                                             last lm # suf_lm)"
+      proof -
+	from h have j1: "
+          \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = 3 \<and> 
+              bp = ab"
+	  apply(auto)
+	  done
+	from h have j2: "ac = rs_pos - 1"
+	  apply(drule_tac ci_pr_f_paras, simp, auto)
+	  done
+	from h and j2 have j3: "a_md \<ge> Suc bc \<and> rs_pos > 0 \<and> bc \<ge> rs_pos"
+	  apply(rule_tac conjI)
+	  apply(erule_tac ab = ab and ac = ac in ci_pr_md_ge_f, simp)
+	  apply(rule_tac context_conjI)
+          apply(simp_all add: rec_ci.simps)
+	  apply(drule_tac ci_ad_ge_paras, drule_tac ci_ad_ge_paras)
+	  apply(arith)
+	  done	  
+	from j1 and j2 show "?thesis"
+	  apply(auto simp del: abc_append_commute)
+	  apply(rule_tac abc_append_exc1, simp_all)
+	  apply(insert nf_ind[of "butlast lm" "rsa" 
+                "0\<up>(a_md - bc - Suc 0) @ last lm # suf_lm"], 
+               simp add: k2_2 j2, erule_tac exE)
+	  apply(simp add: exponent_add_iff j3)
+	  apply(rule_tac x = "stp" in exI, simp)
+	  done
+      qed
+      from h have k2_2_2: 
+      "\<exists> stp. abc_steps_l (3 + length ab, butlast lm @ rsa # 
+                  0\<up>(a_md - rs_pos - 1) @ last lm # suf_lm) aprog stp
+        = (6 + length ab, butlast lm @ 0 # rsa # 
+                       0\<up>(a_md - rs_pos - 2) @ last lm # suf_lm)"
+      proof -	     
+	from h have "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
+          length ap = 3 + length ab \<and> bp = Recursive.mv_box n (Suc n)"
+	  by auto
+	thus "?thesis"
+	proof(erule_tac exE, erule_tac exE, erule_tac exE, 
+              erule_tac exE)
+	  fix ap cp bp apa
+	  assume "aprog = ap [+] bp [+] cp \<and> length ap = 3 + 
+                    length ab \<and> bp = Recursive.mv_box n (Suc n)"
+	  thus "?thesis"
+	    apply(simp del: abc_append_commute)
+	    apply(subgoal_tac 
+              "\<exists>stp. abc_steps_l (3 + length ab, 
+               butlast lm @ rsa # 0\<up>(a_md - Suc rs_pos) @
+                 last lm # suf_lm) (ap [+] 
+                   Recursive.mv_box n (Suc n) [+] cp) stp =
+              ((3 + length ab) + 3, butlast lm @ 0 # rsa # 
+                  0\<up>(a_md - Suc (Suc rs_pos)) @ last lm # suf_lm)", simp)
+	    apply(rule_tac abc_append_exc1, simp_all)
+	    apply(insert mv_box_ex[of n "Suc n" 
+                    "butlast lm @ rsa # 0\<up>(a_md - Suc rs_pos) @ 
+                          last lm # suf_lm"], simp)
+	    apply(subgoal_tac "length lm = Suc n \<and> rs_pos = Suc n \<and> a_md > Suc (Suc n)", simp)
+	    apply(insert h, simp)
+            done
+	qed
+      qed
+      from h have k2_3: "lm \<noteq> []"
+	apply(rule_tac calc_pr_para_not_null, simp)
+	done
+      from h and k2_2 and k2_3 have k2_2_3: 
+      "\<exists> stp. abc_steps_l (6 + length ab, butlast lm @ 
+          (last lm - last lm) # rsa # 
+            0\<up>(a_md - (Suc (Suc rs_pos))) @ last lm # suf_lm) aprog stp
+        = (6 + length ab, butlast lm @ last lm # rs # 
+                        0\<up>(a_md - Suc (Suc (rs_pos))) @ 0 # suf_lm)"
+	apply(rule_tac x = "last lm" and g = g in pr_cycle_part, auto)
+	apply(rule_tac ng_ind, simp)
+	apply(rule_tac rec_calc_rel_def0, simp, simp)
+	done
+      from h  have k2_2_4: 
+       "\<exists> stp. abc_steps_l (6 + length ab,
+             butlast lm @ last lm # rs # 0\<up>(a_md - rs_pos - 2) @
+                  0 # suf_lm) aprog stp
+        = (13 + length ab + length a,
+                   lm @ rs # 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+      proof -
+	from h have 
+        "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
+                     length ap = 6 + length ab \<and> 
+                    bp = ([Dec (a_md - Suc 0) (length a + 7)] [+] 
+                         (a [+] [Inc (rs_pos - Suc 0), 
+                         Dec rs_pos 3, Goto (Suc 0)])) @ 
+                        [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]"
+	  by auto
+	thus "?thesis"
+	  apply(auto)
+	  apply(subgoal_tac  
+            "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @ 
+                last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)
+                (ap [+] ([Dec (a_md - Suc 0) (length a + 7)] [+] 
+                (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, 
+                Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), 
+                Goto (length a + 4)] [+] cp) stp =
+            (6 + length ab + (length a + 7) , 
+                 lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)", simp)
+	  apply(subgoal_tac "13 + (length ab + length a) = 
+                              13 + length ab + length a", simp)
+	  apply(arith)
+	  apply(rule abc_append_exc1, simp_all)
+	  apply(rule_tac x = "Suc 0" in exI, 
+                simp add: abc_steps_l.simps abc_fetch.simps
+                         nth_append abc_append_nth abc_step_l.simps)
+	  apply(subgoal_tac "a_md > Suc (Suc rs_pos) \<and> 
+                            length lm = rs_pos \<and> rs_pos > 0", simp)
+	  apply(insert h, simp)
+	  apply(subgoal_tac "rs_pos = Suc n", simp, simp)
+          done
+      qed
+      from h have k2_2_5: "length aprog = 13 + length ab + length a"
+	apply(rule_tac ci_pr_length, simp_all)
+	done
+      from k2_2_1 and k2_2_2 and k2_2_3 and k2_2_4 and k2_2_5 
+      show "?thesis"
+	apply(auto)
+	apply(rule_tac x = "stp + stpa + stpb + stpc" in exI, 
+              simp add: abc_steps_add)
+	done
+    qed
+  qed	
+  from k1 and k2 show 
+    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp 
+               = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+    apply(erule_tac exE)
+    apply(erule_tac exE)
+    apply(rule_tac x = "stp + stpa" in exI)
+    apply(simp add: abc_steps_add)
+    done
+qed
+
+thm rec_calc_rel.induct
+
+lemma eq_switch: "x = y \<Longrightarrow> y = x"
+by simp
+
+lemma [simp]: 
+  "\<lbrakk>rec_ci f = (a, aa, ba); 
+    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk> \<Longrightarrow> \<exists>bp. aprog = a @ bp"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "[Dec (Suc n) (length a + 5), 
+      Dec (Suc n) (length a + 3), Goto (Suc (length a)), 
+      Inc n, Goto 0]" in exI, auto)
+done
+
+lemma ci_mn_para_eq[simp]: 
+  "rec_ci (Mn n f) = (aprog, rs_pos, a_md) \<Longrightarrow> rs_pos = n"
+apply(case_tac "rec_ci f", simp add: rec_ci.simps)
+done
+(*
+lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md); rec_calc_rel (Mn n f) lm rs\<rbrakk> \<Longrightarrow> aa = Suc rs_pos"
+apply(rule_tac calc_mn_reverse, simp)
+apply(insert para_pattern [of f a aa ba "lm @ [rs]" 0], simp)
+apply(subgoal_tac "rs_pos = length lm", simp)
+apply(drule_tac ci_mn_para_eq, simp)
+done
+*)
+lemma [simp]: "rec_ci f = (a, aa, ba) \<Longrightarrow> aa < ba"
+apply(simp add: ci_ad_ge_paras)
+done
+
+lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); 
+                rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
+    \<Longrightarrow> ba \<le> a_md"
+apply(simp add: rec_ci.simps)
+by arith
+
+lemma mn_calc_f: 
+  assumes ind: 
+  "\<And>aprog a_md rs_pos rs suf_lm lm.
+  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\<rbrakk>  
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp    
+           = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+  and h: "rec_ci f = (a, aa, ba)" 
+         "rec_ci (Mn n f) = (aprog, rs_pos, a_md)"  
+         "rec_calc_rel f (lm @ [x]) rsx" 
+         "aa = Suc n"
+  shows "\<exists>stp. abc_steps_l (0, lm @ x # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
+                  aprog stp = (length a, 
+                   lm @ x # rsx # 0\<up>(a_md - Suc (Suc rs_pos)) @ suf_lm)"
+proof -
+  from h have k1: "\<exists> ap bp. aprog = ap @ bp \<and> ap = a"
+    by simp
+  from h have k2: "rs_pos = n"
+    apply(erule_tac ci_mn_para_eq)
+    done
+  from h and k1 and k2 show "?thesis"
+  
+  proof(erule_tac exE, erule_tac exE, simp, 
+        rule_tac abc_add_exc1, auto)
+    fix bp
+    show 
+      "\<exists>astp. abc_steps_l (0, lm @ x # 0\<up>(a_md - Suc n) @ suf_lm) a astp
+      = (length a, lm @ x # rsx # 0\<up>(a_md - Suc (Suc n)) @ suf_lm)"
+      apply(insert ind[of a "Suc n" ba  "lm @ [x]" rsx 
+             "0\<up>(a_md - ba) @ suf_lm"], simp add: exponent_add_iff h k2)
+      apply(subgoal_tac "ba > aa \<and> a_md \<ge> ba \<and> aa = Suc n", 
+            insert h, auto)
+      done
+  qed
+qed
+
+fun mn_ind_inv ::
+  "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> nat list \<Rightarrow> bool"
+  where
+  "mn_ind_inv (as, lm') ss x rsx suf_lm lm = 
+           (if as = ss then lm' = lm @ x # rsx # suf_lm
+            else if as = ss + 1 then 
+                 \<exists>y. (lm' = lm @ x # y # suf_lm) \<and> y \<le> rsx
+            else if as = ss + 2 then 
+                 \<exists>y. (lm' = lm @ x # y # suf_lm) \<and> y \<le> rsx
+            else if as = ss + 3 then lm' = lm @ x # 0 # suf_lm
+            else if as = ss + 4 then lm' = lm @ Suc x # 0 # suf_lm
+            else if as = 0 then lm' = lm @ Suc x # 0 # suf_lm
+            else False
+)"
+
+fun mn_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "mn_stage1 (as, lm) ss n = 
+            (if as = 0 then 0 
+             else if as = ss + 4 then 1
+             else if as = ss + 3 then 2
+             else if as = ss + 2 \<or> as = ss + 1 then 3
+             else if as = ss then 4
+             else 0
+)"
+
+fun mn_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "mn_stage2 (as, lm) ss n = 
+            (if as = ss + 1 \<or> as = ss + 2 then (lm ! (Suc n))
+             else 0)"
+
+fun mn_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "mn_stage3 (as, lm) ss n = (if as = ss + 2 then 1 else 0)"
+
+ 
+fun mn_measure :: "((nat \<times> nat list) \<times> nat \<times> nat) \<Rightarrow>
+                                                (nat \<times> nat \<times> nat)"
+  where
+  "mn_measure ((as, lm), ss, n) = 
+     (mn_stage1 (as, lm) ss n, mn_stage2 (as, lm) ss n,
+                                       mn_stage3 (as, lm) ss n)"
+
+definition mn_LE :: "(((nat \<times> nat list) \<times> nat \<times> nat) \<times>
+                     ((nat \<times> nat list) \<times> nat \<times> nat)) set"
+  where "mn_LE \<equiv> (inv_image lex_triple mn_measure)"
+
+thm halt_lemma2
+lemma wf_mn_le[intro]: "wf mn_LE"
+by(auto intro:wf_inv_image wf_lex_triple simp: mn_LE_def)
+
+declare mn_ind_inv.simps[simp del]
+
+lemma mn_inv_init: 
+  "mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 0)
+                                         (length a) x rsx suf_lm lm"
+apply(simp add: mn_ind_inv.simps abc_steps_zero)
+done
+
+lemma mn_halt_init: 
+  "rec_ci f = (a, aa, ba) \<Longrightarrow> 
+  \<not> (\<lambda>(as, lm') (ss, n). as = 0) 
+    (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 0) 
+                                                       (length a, n)"
+apply(simp add: abc_steps_zero)
+apply(erule_tac rec_ci_not_null)
+done
+
+thm rec_ci.simps
+lemma [simp]: 
+  "\<lbrakk>rec_ci f = (a, aa, ba); 
+    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
+    \<Longrightarrow> abc_fetch (length a) aprog =
+                      Some (Dec (Suc n) (length a + 5))"
+apply(simp add: rec_ci.simps abc_fetch.simps, 
+                erule_tac conjE, erule_tac conjE, simp)
+apply(drule_tac eq_switch, drule_tac eq_switch, simp)
+done
+
+lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
+    \<Longrightarrow> abc_fetch (Suc (length a)) aprog = Some (Dec (Suc n) (length a + 3))"
+apply(simp add: rec_ci.simps abc_fetch.simps, erule_tac conjE, erule_tac conjE, simp)
+apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append)
+done
+
+lemma [simp]:
+  "\<lbrakk>rec_ci f = (a, aa, ba);
+    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
+    \<Longrightarrow> abc_fetch (Suc (Suc (length a))) aprog = 
+                                     Some (Goto (length a + 1))"
+apply(simp add: rec_ci.simps abc_fetch.simps,
+      erule_tac conjE, erule_tac conjE, simp)
+apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append)
+done
+
+lemma [simp]: 
+  "\<lbrakk>rec_ci f = (a, aa, ba);
+    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
+    \<Longrightarrow> abc_fetch (length a + 3) aprog = Some (Inc n)"
+apply(simp add: rec_ci.simps abc_fetch.simps, 
+      erule_tac conjE, erule_tac conjE, simp)
+apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append)
+done
+
+lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
+    \<Longrightarrow> abc_fetch (length a + 4) aprog = Some (Goto 0)"
+apply(simp add: rec_ci.simps abc_fetch.simps, erule_tac conjE, erule_tac conjE, simp)
+apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append)
+done
+
+lemma [simp]: 
+  "0 < rsx
+   \<Longrightarrow> \<exists>y. (lm @ x # rsx # suf_lm)[Suc (length lm) := rsx - Suc 0]   
+    = lm @ x # y # suf_lm \<and> y \<le> rsx"
+apply(case_tac rsx, simp, simp)
+apply(rule_tac x = nat in exI, simp add: list_update_append)
+done
+
+lemma [simp]: 
+  "\<lbrakk>y \<le> rsx; 0 < y\<rbrakk>
+   \<Longrightarrow> \<exists>ya. (lm @ x # y # suf_lm)[Suc (length lm) := y - Suc 0] 
+          = lm @ x # ya # suf_lm \<and> ya \<le> rsx"
+apply(case_tac y, simp, simp)
+apply(rule_tac x = nat in exI, simp add: list_update_append)
+done
+
+lemma mn_halt_lemma: 
+  "\<lbrakk>rec_ci f = (a, aa, ba);
+    rec_ci (Mn n f) = (aprog, rs_pos, a_md);
+     0 < rsx; length lm = n\<rbrakk>
+    \<Longrightarrow>
+  \<forall>na. \<not> (\<lambda>(as, lm') (ss, n). as = 0)
+  (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog na) 
+                                                       (length a, n)
+ \<and> mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm)
+                       aprog na) (length a) x rsx suf_lm lm 
+\<longrightarrow> mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 
+                         (Suc na)) (length a) x rsx suf_lm lm
+ \<and> ((abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog (Suc na), 
+                                                    length a, n), 
+    abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog na,
+                              length a, n) \<in> mn_LE"
+apply(rule allI, rule impI, simp add: abc_steps_ind)
+apply(case_tac "(abc_steps_l (length a, lm @ x # rsx # suf_lm) 
+                                                   aprog na)", simp)
+apply(auto split:if_splits simp add:abc_steps_l.simps 
+                           mn_ind_inv.simps abc_steps_zero)
+apply(auto simp add: mn_LE_def lex_triple_def lex_pair_def 
+            abc_step_l.simps abc_steps_l.simps mn_ind_inv.simps
+            abc_lm_v.simps abc_lm_s.simps nth_append
+           split: if_splits)
+apply(drule_tac  rec_ci_not_null, simp)
+done
+
+lemma mn_halt:
+  "\<lbrakk>rec_ci f = (a, aa, ba);
+    rec_ci (Mn n f) = (aprog, rs_pos, a_md);
+    0 < rsx; length lm = n\<rbrakk>
+    \<Longrightarrow> \<exists> stp. (\<lambda> (as, lm'). (as = 0 \<and> 
+           mn_ind_inv (as, lm')  (length a) x rsx suf_lm lm))
+            (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog stp)"
+apply(insert wf_mn_le)	  
+apply(insert halt_lemma2[of mn_LE
+  "\<lambda> ((as, lm'), ss, n). mn_ind_inv (as, lm') ss x rsx suf_lm lm"
+  "\<lambda> stp. (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog stp, 
+  length a, n)"
+  "\<lambda> ((as, lm'), ss, n). as = 0"], 
+   simp)
+apply(simp add: mn_halt_init mn_inv_init)
+apply(drule_tac x = x and suf_lm = suf_lm in mn_halt_lemma, auto)
+apply(rule_tac x = n in exI, 
+      case_tac "(abc_steps_l (length a, lm @ x # rsx # suf_lm)
+                              aprog n)", simp)
+done
+
+lemma [simp]: "Suc rs_pos < a_md \<Longrightarrow> 
+                Suc (a_md - Suc (Suc rs_pos)) = a_md - Suc rs_pos"
+by arith
+
+term rec_ci
+(*
+lemma [simp]: "\<lbrakk>rec_ci (Mn n f) = (aprog, rs_pos, a_md); rec_calc_rel (Mn n f) lm rs\<rbrakk>  \<Longrightarrow> Suc rs_pos < a_md"
+apply(case_tac "rec_ci f")
+apply(subgoal_tac "c > b \<and> b = Suc rs_pos \<and> a_md \<ge> c")
+apply(arith, auto)
+done
+*)
+lemma mn_ind_step: 
+  assumes ind:  
+  "\<And>aprog a_md rs_pos rs suf_lm lm.
+  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md);
+   rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
+  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp
+            = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+  and h: "rec_ci f = (a, aa, ba)" 
+         "rec_ci (Mn n f) = (aprog, rs_pos, a_md)"  
+         "rec_calc_rel f (lm @ [x]) rsx" 
+         "rsx > 0" 
+         "Suc rs_pos < a_md" 
+         "aa = Suc rs_pos"
+  shows "\<exists>stp. abc_steps_l (0, lm @ x # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
+             aprog stp = (0, lm @ Suc x # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+thm abc_add_exc1
+proof -
+  have k1: 
+    "\<exists> stp. abc_steps_l (0, lm @ x #  0\<up>(a_md - Suc (rs_pos)) @ suf_lm)
+         aprog stp = 
+       (length a, lm @ x # rsx # 0\<up>(a_md  - Suc (Suc rs_pos)) @ suf_lm)"
+    apply(insert h)
+    apply(auto intro: mn_calc_f ind)
+    done
+  from h have k2: "length lm = n"
+    apply(subgoal_tac "rs_pos = n")
+    apply(drule_tac  para_pattern, simp, simp, simp)
+    done
+  from h have k3: "a_md > (Suc rs_pos)"
+    apply(simp)
+    done
+  from k2 and h and k3 have k4: 
+    "\<exists> stp. abc_steps_l (length a,
+       lm @ x # rsx # 0\<up>(a_md  - Suc (Suc rs_pos)) @ suf_lm) aprog stp = 
+        (0, lm @ Suc x # 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+    apply(frule_tac x = x and 
+       suf_lm = "0\<up>(a_md - Suc (Suc rs_pos)) @ suf_lm" in mn_halt, auto)
+    apply(rule_tac x = "stp" in exI, 
+          simp add: mn_ind_inv.simps rec_ci_not_null)
+    apply(simp only: replicate.simps[THEN sym], simp)
+    done
+  from k1 and k4 show "?thesis"
+    apply(auto)
+    apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
+    done
+qed
+
+lemma [simp]: 
+  "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md);
+    rec_calc_rel (Mn n f) lm rs\<rbrakk> \<Longrightarrow> aa = Suc rs_pos"
+apply(rule_tac calc_mn_reverse, simp)
+apply(insert para_pattern [of f a aa ba "lm @ [rs]" 0], simp)
+apply(subgoal_tac "rs_pos = length lm", simp)
+apply(drule_tac ci_mn_para_eq, simp)
+done
+
+lemma [simp]: "\<lbrakk>rec_ci (Mn n f) = (aprog, rs_pos, a_md);      
+                rec_calc_rel (Mn n f) lm rs\<rbrakk>  \<Longrightarrow> Suc rs_pos < a_md"
+apply(case_tac "rec_ci f")
+apply(subgoal_tac "c > b \<and> b = Suc rs_pos \<and> a_md \<ge> c")
+apply(arith, auto)
+done
+
+lemma mn_ind_steps:  
+  assumes ind:
+  "\<And>aprog a_md rs_pos rs suf_lm lm. 
+  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
+  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
+              (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+  and h: "rec_ci f = (a, aa, ba)" 
+  "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
+  "rec_calc_rel (Mn n f) lm rs"
+  "rec_calc_rel f (lm @ [rs]) 0" 
+  "\<forall>x<rs. (\<exists> v. rec_calc_rel f (lm @ [x]) v \<and> 0 < v)"
+  "n = length lm" 
+  "x \<le> rs"
+  shows "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
+                 aprog stp = (0, lm @ x # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+apply(insert h, induct x, 
+      rule_tac x = 0 in exI, simp add: abc_steps_zero, simp)
+proof -
+  fix x
+  assume k1: 
+    "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
+                aprog stp = (0, lm @ x # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+  and k2: "rec_ci (Mn (length lm) f) = (aprog, rs_pos, a_md)" 
+          "rec_calc_rel (Mn (length lm) f) lm rs" 
+          "rec_calc_rel f (lm @ [rs]) 0" 
+          "\<forall>x<rs.(\<exists> v. rec_calc_rel f (lm @ [x]) v \<and> v > 0)" 
+          "n = length lm" 
+          "Suc x \<le> rs" 
+          "rec_ci f = (a, aa, ba)"
+  hence k2:
+    "\<exists>stp. abc_steps_l (0, lm @ x # 0\<up>(a_md - rs_pos - 1) @ suf_lm) aprog
+               stp = (0, lm @ Suc x # 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+    apply(erule_tac x = x in allE)
+    apply(auto)
+    apply(rule_tac  x = x in mn_ind_step)
+    apply(rule_tac ind, auto)      
+    done
+  from k1 and k2 show 
+    "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
+          aprog stp = (0, lm @ Suc x # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+    apply(auto)
+    apply(rule_tac x = "stp + stpa" in exI, simp only: abc_steps_add)
+    done
+qed
+    
+lemma [simp]: 
+"\<lbrakk>rec_ci f = (a, aa, ba); 
+  rec_ci (Mn n f) = (aprog, rs_pos, a_md); 
+  rec_calc_rel (Mn n f) lm rs;
+  length lm = n\<rbrakk>
+ \<Longrightarrow> abc_lm_v (lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) (Suc n) = 0"
+apply(auto simp: abc_lm_v.simps nth_append)
+done
+
+lemma [simp]: 
+  "\<lbrakk>rec_ci f = (a, aa, ba); 
+    rec_ci (Mn n f) = (aprog, rs_pos, a_md); 
+    rec_calc_rel (Mn n f) lm rs;
+     length lm = n\<rbrakk>
+    \<Longrightarrow> abc_lm_s (lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) (Suc n) 0 =
+                           lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm"
+apply(auto simp: abc_lm_s.simps list_update_append)
+done
+
+lemma mn_length: 
+  "\<lbrakk>rec_ci f = (a, aa, ba);
+    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
+  \<Longrightarrow> length aprog = length a + 5"
+apply(simp add: rec_ci.simps, erule_tac conjE)
+apply(drule_tac eq_switch, drule_tac eq_switch, simp)
+done
+
+lemma mn_final_step:
+  assumes ind:
+  "\<And>aprog a_md rs_pos rs suf_lm lm.
+  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); 
+  rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
+  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
+              (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+  and h: "rec_ci f = (a, aa, ba)" 
+         "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
+         "rec_calc_rel (Mn n f) lm rs" 
+         "rec_calc_rel f (lm @ [rs]) 0" 
+  shows "\<exists>stp. abc_steps_l (0, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
+     aprog stp = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+proof -
+  from h and ind have k1:
+    "\<exists>stp.  abc_steps_l (0, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
+        aprog stp = (length a,  lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+    thm mn_calc_f
+    apply(insert mn_calc_f[of f a aa ba n aprog 
+                               rs_pos a_md lm rs 0 suf_lm], simp)
+    apply(subgoal_tac "aa = Suc n", simp add: exponent_cons_iff)
+    apply(subgoal_tac "rs_pos = n", simp, simp)
+    done
+  from h have k2: "length lm = n"
+    apply(subgoal_tac "rs_pos = n")
+    apply(drule_tac f = "Mn n f" in para_pattern, simp, simp, simp)
+    done
+  from h and k2 have k3: 
+  "\<exists>stp. abc_steps_l (length a, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
+    aprog stp = (length a + 5, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+    apply(rule_tac x = "Suc 0" in exI, 
+          simp add: abc_step_l.simps abc_steps_l.simps)
+    done
+  from h have k4: "length aprog = length a + 5"
+    apply(simp add: mn_length)
+    done
+  from k1 and k3 and k4 show "?thesis"
+    apply(auto)
+    apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
+    done
+qed
+
+lemma mn_case: 
+  assumes ind: 
+  "\<And>aprog a_md rs_pos rs suf_lm lm.
+  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
+  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
+               (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+  and h: "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
+         "rec_calc_rel (Mn n f) lm rs"
+  shows "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp 
+  = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+apply(case_tac "rec_ci f", simp)
+apply(insert h, rule_tac calc_mn_reverse, simp)
+proof -
+  fix a b c v
+  assume h: "rec_ci f = (a, b, c)" 
+            "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
+            "rec_calc_rel (Mn n f) lm rs" 
+            "rec_calc_rel f (lm @ [rs]) 0" 
+            "\<forall>x<rs. \<exists>v. rec_calc_rel f (lm @ [x]) v \<and> 0 < v"
+            "n = length lm"
+  hence k1:
+    "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm) aprog
+                  stp = (0, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+    thm mn_ind_steps
+    apply(auto intro: mn_ind_steps ind)
+    done
+  from h have k2: 
+    "\<exists>stp. abc_steps_l (0, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) aprog
+         stp = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+    apply(auto intro: mn_final_step ind)
+    done
+  from k1 and k2 show 
+    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
+  (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+    apply(auto, insert h)
+    apply(subgoal_tac "Suc rs_pos < a_md")
+    apply(rule_tac x = "stp + stpa" in exI, 
+      simp only: abc_steps_add exponent_cons_iff, simp, simp)
+    done
+qed
+
+lemma z_rs: "rec_calc_rel z lm rs \<Longrightarrow> rs = 0"
+apply(rule_tac calc_z_reverse, auto)
+done
+
+lemma z_case:
+  "\<lbrakk>rec_ci z = (aprog, rs_pos, a_md); rec_calc_rel z lm rs\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
+           (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+apply(simp add: rec_ci.simps rec_ci_z_def, auto)
+apply(rule_tac x = "Suc 0" in exI, simp add: abc_steps_l.simps 
+                               abc_fetch.simps abc_step_l.simps z_rs)
+done
+
+fun addition_inv :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow>     
+                     nat list \<Rightarrow> bool"
+  where
+  "addition_inv (as, lm') m n p lm = 
+        (let sn = lm ! n in
+         let sm = lm ! m in
+         lm ! p = 0 \<and>
+             (if as = 0 then \<exists> x. x \<le> lm ! m \<and> lm' = lm[m := x,
+                                    n := (sn + sm - x), p := (sm - x)]
+             else if as = 1 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x,
+                            n := (sn + sm - x - 1), p := (sm - x - 1)]
+             else if as = 2 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x, 
+                               n := (sn + sm - x), p := (sm - x - 1)]
+             else if as = 3 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x,
+                                   n := (sn + sm - x), p := (sm - x)]
+             else if as = 4 then \<exists> x. x \<le> lm ! m \<and> lm' = lm[m := x,
+                                       n := (sn + sm), p := (sm - x)] 
+             else if as = 5 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x, 
+                                  n := (sn + sm), p := (sm - x - 1)] 
+             else if as = 6 then \<exists> x. x < lm ! m \<and> lm' =
+                     lm[m := Suc x, n := (sn + sm), p := (sm - x - 1)]
+             else if as = 7 then lm' = lm[m := sm, n := (sn + sm)]
+             else False))"
+
+fun addition_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "addition_stage1 (as, lm) m p = 
+          (if as = 0 \<or> as = 1 \<or> as = 2 \<or> as = 3 then 2 
+           else if as = 4 \<or> as = 5 \<or> as = 6 then 1
+           else 0)"
+
+fun addition_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow>  nat \<Rightarrow> nat"
+  where
+  "addition_stage2 (as, lm) m p = 
+              (if 0 \<le> as \<and> as \<le> 3 then lm ! m
+               else if 4 \<le> as \<and> as \<le> 6 then lm ! p
+               else 0)"
+
+fun addition_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
+  where
+  "addition_stage3 (as, lm) m p = 
+             (if as = 1 then 4  
+              else if as = 2 then 3 
+              else if as = 3 then 2
+              else if as = 0 then 1 
+              else if as = 5 then 2
+              else if as = 6 then 1 
+              else if as = 4 then 0 
+              else 0)"
+
+fun addition_measure :: "((nat \<times> nat list) \<times> nat \<times> nat) \<Rightarrow> 
+                                                 (nat \<times> nat \<times> nat)"
+  where
+  "addition_measure ((as, lm), m, p) =
+     (addition_stage1 (as, lm) m p, 
+      addition_stage2 (as, lm) m p,
+      addition_stage3 (as, lm) m p)"
+
+definition addition_LE :: "(((nat \<times> nat list) \<times> nat \<times> nat) \<times> 
+                          ((nat \<times> nat list) \<times> nat \<times> nat)) set"
+  where "addition_LE \<equiv> (inv_image lex_triple addition_measure)"
+
+lemma [simp]: "wf addition_LE"
+by(simp add: wf_inv_image wf_lex_triple addition_LE_def)
+
+declare addition_inv.simps[simp del]
+
+lemma addition_inv_init: 
+  "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow>
+                                   addition_inv (0, lm) m n p lm"
+apply(simp add: addition_inv.simps)
+apply(rule_tac x = "lm ! m" in exI, simp)
+done
+
+thm addition.simps
+
+lemma [simp]: "abc_fetch 0 (addition m n p) = Some (Dec m 4)"
+by(simp add: abc_fetch.simps addition.simps)
+
+lemma [simp]: "abc_fetch (Suc 0) (addition m n p) = Some (Inc n)"
+by(simp add: abc_fetch.simps addition.simps)
+
+lemma [simp]: "abc_fetch 2 (addition m n p) = Some (Inc p)"
+by(simp add: abc_fetch.simps addition.simps)
+
+lemma [simp]: "abc_fetch 3 (addition m n p) = Some (Goto 0)"
+by(simp add: abc_fetch.simps addition.simps)
+
+lemma [simp]: "abc_fetch 4 (addition m n p) = Some (Dec p 7)"
+by(simp add: abc_fetch.simps addition.simps)
+
+lemma [simp]: "abc_fetch 5 (addition m n p) = Some (Inc m)"
+by(simp add: abc_fetch.simps addition.simps)
+
+lemma [simp]: "abc_fetch 6 (addition m n p) = Some (Goto 4)"
+by(simp add: abc_fetch.simps addition.simps)
+
+lemma [simp]:
+  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x \<le> lm ! m; 0 < x\<rbrakk>
+ \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m - x, 
+                    p := lm ! m - x, m := x - Suc 0] =
+                 lm[m := xa, n := lm ! n + lm ! m - Suc xa,
+                    p := lm ! m - Suc xa]"
+apply(case_tac x, simp, simp)
+apply(rule_tac x = nat in exI, simp add: list_update_swap 
+                                         list_update_overwrite)
+done
+
+lemma [simp]:
+  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
+   \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m - Suc x,
+                      p := lm ! m - Suc x, n := lm ! n + lm ! m - x]
+                 = lm[m := xa, n := lm ! n + lm ! m - xa, 
+                      p := lm ! m - Suc xa]"
+apply(rule_tac x = x in exI, 
+      simp add: list_update_swap list_update_overwrite)
+done
+
+lemma [simp]: 
+  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
+   \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m - x, 
+                          p := lm ! m - Suc x, p := lm ! m - x]
+                 = lm[m := xa, n := lm ! n + lm ! m - xa, 
+                          p := lm ! m - xa]"
+apply(rule_tac x = x in exI, simp add: list_update_overwrite)
+done
+
+lemma [simp]: 
+  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = (0::nat); m < p; n < p; x < lm ! m\<rbrakk>
+  \<Longrightarrow> \<exists>xa\<le>lm ! m. lm[m := x, n := lm ! n + lm ! m - x,
+                                   p := lm ! m - x] = 
+                  lm[m := xa, n := lm ! n + lm ! m - xa, 
+                                   p := lm ! m - xa]"
+apply(rule_tac x = x in exI, simp)
+done
+
+lemma [simp]: 
+  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p;
+    x \<le> lm ! m; lm ! m \<noteq> x\<rbrakk>
+  \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m, 
+                       p := lm ! m - x, p := lm ! m - Suc x] 
+               = lm[m := xa, n := lm ! n + lm ! m, 
+                       p := lm ! m - Suc xa]"
+apply(rule_tac x = x in exI, simp add: list_update_overwrite)
+done
+
+lemma [simp]:
+  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
+  \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m,
+                             p := lm ! m - Suc x, m := Suc x]
+                = lm[m := Suc xa, n := lm ! n + lm ! m, 
+                             p := lm ! m - Suc xa]"
+apply(rule_tac x = x in exI, 
+     simp add: list_update_swap list_update_overwrite)
+done
+
+lemma [simp]: 
+  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
+  \<Longrightarrow> \<exists>xa\<le>lm ! m. lm[m := Suc x, n := lm ! n + lm ! m, 
+                             p := lm ! m - Suc x] 
+               = lm[m := xa, n := lm ! n + lm ! m, p := lm ! m - xa]"
+apply(rule_tac x = "Suc x" in exI, simp)
+done
+
+lemma addition_halt_lemma: 
+  "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow>
+  \<forall>na. \<not> (\<lambda>(as, lm') (m, p). as = 7) 
+        (abc_steps_l (0, lm) (addition m n p) na) (m, p) \<and> 
+  addition_inv (abc_steps_l (0, lm) (addition m n p) na) m n p lm 
+\<longrightarrow> addition_inv (abc_steps_l (0, lm) (addition m n p) 
+                                 (Suc na)) m n p lm 
+  \<and> ((abc_steps_l (0, lm) (addition m n p) (Suc na), m, p), 
+     abc_steps_l (0, lm) (addition m n p) na, m, p) \<in> addition_LE"
+apply(rule allI, rule impI, simp add: abc_steps_ind)
+apply(case_tac "(abc_steps_l (0, lm) (addition m n p) na)", simp)
+apply(auto split:if_splits simp add: addition_inv.simps
+                                 abc_steps_zero)
+apply(simp_all add: abc_steps_l.simps abc_steps_zero)
+apply(auto simp add: addition_LE_def lex_triple_def lex_pair_def 
+                     abc_step_l.simps addition_inv.simps 
+                     abc_lm_v.simps abc_lm_s.simps nth_append
+                split: if_splits)
+apply(rule_tac x = x in exI, simp)
+done
+
+lemma  addition_ex: 
+  "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow> 
+  \<exists> stp. (\<lambda> (as, lm'). as = 7 \<and> addition_inv (as, lm') m n p lm) 
+                        (abc_steps_l (0, lm) (addition m n p) stp)"
+apply(insert halt_lemma2[of addition_LE
+  "\<lambda> ((as, lm'), m, p). addition_inv (as, lm') m n p lm"
+  "\<lambda> stp. (abc_steps_l (0, lm) (addition m n p) stp, m, p)"
+  "\<lambda> ((as, lm'), m, p). as = 7"], 
+  simp add: abc_steps_zero addition_inv_init)
+apply(drule_tac addition_halt_lemma, simp, simp, simp,
+      simp, erule_tac exE)
+apply(rule_tac x = na in exI, 
+      case_tac "(abc_steps_l (0, lm) (addition m n p) na)", auto)
+done
+
+lemma [simp]: "length (addition m n p) = 7"
+by (simp add: addition.simps)
+
+lemma [elim]: "addition 0 (Suc 0) 2 = [] \<Longrightarrow> RR"
+by(simp add: addition.simps)
+
+lemma [simp]: "(0\<up>2)[0 := n] = [n, 0::nat]"
+apply(subgoal_tac "2 = Suc 1", 
+      simp only: replicate.simps)
+apply(auto)
+done
+
+lemma [simp]: 
+  "\<exists>stp. abc_steps_l (0, n # 0\<up>2 @ suf_lm) 
+     (addition 0 (Suc 0) 2 [+] [Inc (Suc 0)]) stp = 
+                                      (8, n # Suc n # 0 # suf_lm)"
+apply(rule_tac bm = "n # n # 0 # suf_lm" in abc_append_exc2, auto)
+apply(insert addition_ex[of 0 "Suc 0" 2 "n # 0\<up>2 @ suf_lm"], 
+      simp add: nth_append numeral_2_eq_2, erule_tac exE)
+apply(rule_tac x = stp in exI,
+      case_tac "(abc_steps_l (0, n # 0\<up>2 @ suf_lm)
+                      (addition 0 (Suc 0) 2) stp)", 
+      simp add: addition_inv.simps nth_append list_update_append numeral_2_eq_2)
+apply(simp add: nth_append numeral_2_eq_2, erule_tac exE)
+apply(rule_tac x = "Suc 0" in exI,
+      simp add: abc_steps_l.simps abc_fetch.simps 
+      abc_steps_zero abc_step_l.simps abc_lm_s.simps abc_lm_v.simps)
+done
+
+lemma s_case:
+  "\<lbrakk>rec_ci s = (aprog, rs_pos, a_md); rec_calc_rel s lm rs\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
+               (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+apply(simp add: rec_ci.simps rec_ci_s_def, auto)
+apply(rule_tac calc_s_reverse, auto)
+done
+
+lemma [simp]: 
+  "\<lbrakk>n < length lm; lm ! n = rs\<rbrakk>
+    \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0 # 0 #suf_lm)
+                     (addition n (length lm) (Suc (length lm))) stp 
+             = (7, lm @ rs # 0 # suf_lm)"
+apply(insert addition_ex[of n "length lm"
+                           "Suc (length lm)" "lm @ 0 # 0 # suf_lm"])
+apply(simp add: nth_append, erule_tac exE)
+apply(rule_tac x = stp in exI)
+apply(case_tac "abc_steps_l (0, lm @ 0 # 0 # suf_lm) (addition n (length lm)
+                 (Suc (length lm))) stp", simp)
+apply(simp add: addition_inv.simps)
+apply(insert nth_append[of lm "0 # 0 # suf_lm" "n"], simp)
+done
+
+lemma [simp]: "0\<up>2 = [0, 0::nat]"
+apply(auto simp:numeral_2_eq_2)
+done
+
+lemma id_case: 
+  "\<lbrakk>rec_ci (id m n) = (aprog, rs_pos, a_md); 
+    rec_calc_rel (id m n) lm rs\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
+               (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+apply(simp add: rec_ci.simps rec_ci_id.simps, auto)
+apply(rule_tac calc_id_reverse, simp, simp)
+done   
+
+lemma list_tl_induct:
+  "\<lbrakk>P []; \<And>a list. P list \<Longrightarrow> P (list @ [a::'a])\<rbrakk> \<Longrightarrow> 
+                                            P ((list::'a list))"
+apply(case_tac "length list", simp)
+proof -
+  fix nat
+  assume ind: "\<And>a list. P list \<Longrightarrow> P (list @ [a])"
+  and h: "length list = Suc nat" "P []"
+  from h show "P list"
+  proof(induct nat arbitrary: list, case_tac lista, simp, simp)
+    fix lista a listaa
+    from h show "P [a]"
+      by(insert ind[of "[]"], simp add: h)
+  next
+    fix nat list
+    assume nind: "\<And>list. \<lbrakk>length list = Suc nat; P []\<rbrakk> \<Longrightarrow> P list" 
+    and g: "length (list:: 'a list) = Suc (Suc nat)"
+    from g show "P (list::'a list)"
+      apply(insert nind[of "butlast list"], simp add: h)
+      apply(insert ind[of "butlast list" "last list"], simp)
+      apply(subgoal_tac "butlast list @ [last list] = list", simp)
+      apply(case_tac "list::'a list", simp, simp)
+      done
+  qed
+qed      
+  
+lemma nth_eq_butlast_nth: "\<lbrakk>length ys > Suc k\<rbrakk> \<Longrightarrow> 
+                                        ys ! k = butlast ys ! k"
+apply(subgoal_tac "\<exists> xs y. ys = xs @ [y]", auto simp: nth_append)
+apply(rule_tac x = "butlast ys" in exI, rule_tac x = "last ys" in exI)
+apply(case_tac "ys = []", simp, simp)
+done
+
+lemma [simp]: 
+"\<lbrakk>\<forall>k<Suc (length list). rec_calc_rel ((list @ [a]) ! k) lm (ys ! k);
+  length ys = Suc (length list)\<rbrakk>
+   \<Longrightarrow> \<forall>k<length list. rec_calc_rel (list ! k) lm (butlast ys ! k)"
+apply(rule allI, rule impI)
+apply(erule_tac  x = k in allE, simp add: nth_append)
+apply(subgoal_tac "ys ! k = butlast ys ! k", simp)
+apply(rule_tac nth_eq_butlast_nth, arith)
+done
+
+lemma cn_merge_gs_tl_app: 
+  "cn_merge_gs (gs @ [g]) pstr = 
+        cn_merge_gs gs pstr [+] cn_merge_gs [g] (pstr + length gs)"
+apply(induct gs arbitrary: pstr, simp add: cn_merge_gs.simps, simp)
+apply(case_tac a, simp add: abc_append_commute)
+done
+
+lemma cn_merge_gs_length: 
+  "length (cn_merge_gs (map rec_ci list) pstr) = 
+      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci list. length ap) + 3 * length list "
+apply(induct list arbitrary: pstr, simp, simp)
+apply(case_tac "rec_ci a", simp)
+done
+
+lemma [simp]: "Suc n \<le> pstr \<Longrightarrow> pstr + x - n > 0"
+by arith
+
+lemma [simp]:
+  "\<lbrakk>Suc (pstr + length list) \<le> a_md; 
+    length ys = Suc (length list);
+    length lm = n;
+     Suc n \<le> pstr\<rbrakk>
+   \<Longrightarrow>  (ys ! length list # 0\<up>(pstr - Suc n) @ butlast ys @
+             0\<up>(a_md - (pstr + length list)) @ suf_lm) ! 
+                      (pstr + length list - n) = (0 :: nat)"
+apply(insert nth_append[of "ys ! length list # 0\<up>(pstr - Suc n) @
+     butlast ys" "0\<up>(a_md - (pstr + length list)) @ suf_lm"
+      "(pstr + length list - n)"], simp add: nth_append)
+done
+
+lemma [simp]:
+  "\<lbrakk>Suc (pstr + length list) \<le> a_md; 
+    length ys = Suc (length list);
+    length lm = n;
+     Suc n \<le> pstr\<rbrakk>
+    \<Longrightarrow> (lm @ last ys # 0\<up>(pstr - Suc n) @ butlast ys @
+         0\<up>(a_md - (pstr + length list)) @ suf_lm)[pstr + length list := 
+                                        last ys, n := 0] =
+        lm @ (0::nat)\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm"
+apply(insert list_update_length[of 
+   "lm @ last ys # 0\<up>(pstr - Suc n) @ butlast ys" 0 
+   "0\<up>(a_md - Suc (pstr + length list)) @ suf_lm" "last ys"], simp)
+apply(simp add: exponent_cons_iff)
+apply(insert list_update_length[of "lm" 
+        "last ys" "0\<up>(pstr - Suc n) @ butlast ys @ 
+      last ys # 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm" 0], simp)
+apply(simp add: exponent_cons_iff)
+apply(case_tac "ys = []", simp_all add: append_butlast_last_id)
+done
+
+lemma cn_merge_gs_ex: 
+  "\<lbrakk>\<And>x aprog a_md rs_pos rs suf_lm lm.
+    \<lbrakk>x \<in> set gs; rec_ci x = (aprog, rs_pos, a_md);
+     rec_calc_rel x lm rs\<rbrakk>
+     \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp 
+           = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm); 
+   pstr + length gs\<le> a_md;
+   \<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k);
+   length ys = length gs; length lm = n;
+   pstr \<ge> Max (set (Suc n # map (\<lambda>(aprog, p, n). n) (map rec_ci gs)))\<rbrakk>
+  \<Longrightarrow> \<exists> stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suf_lm)
+                   (cn_merge_gs (map rec_ci gs) pstr) stp 
+   = (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) gs) +
+  3 * length gs, lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - (pstr + length gs)) @ suf_lm)"
+apply(induct gs arbitrary: ys rule: list_tl_induct)
+apply(simp add: exponent_add_iff, simp)
+proof -
+  fix a list ys
+  assume ind: "\<And>x aprog a_md rs_pos rs suf_lm lm.
+    \<lbrakk>x = a \<or> x \<in> set list; rec_ci x = (aprog, rs_pos, a_md); 
+     rec_calc_rel x lm rs\<rbrakk>
+    \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
+                (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+  and ind2: 
+    "\<And>ys. \<lbrakk>\<And>x aprog a_md rs_pos rs suf_lm lm.
+    \<lbrakk>x \<in> set list; rec_ci x = (aprog, rs_pos, a_md);
+     rec_calc_rel x lm rs\<rbrakk>
+    \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp
+        = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm);
+    \<forall>k<length list. rec_calc_rel (list ! k) lm (ys ! k); 
+    length ys = length list\<rbrakk>
+    \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suf_lm) 
+                   (cn_merge_gs (map rec_ci list) pstr) stp =
+    (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) +
+     3 * length list,
+                lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm)"
+    and h: "Suc (pstr + length list) \<le> a_md" 
+            "\<forall>k<Suc (length list). 
+                   rec_calc_rel ((list @ [a]) ! k) lm (ys ! k)" 
+            "length ys = Suc (length list)" 
+            "length lm = n"
+            "Suc n \<le> pstr \<and> (\<lambda>(aprog, p, n). n) (rec_ci a) \<le> pstr \<and> 
+            (\<forall>a\<in>set list. (\<lambda>(aprog, p, n). n) (rec_ci a) \<le> pstr)"
+  from h have k1: 
+    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suf_lm)
+                     (cn_merge_gs (map rec_ci list) pstr) stp =
+    (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) +
+     3 * length list, lm @ 0\<up>(pstr - n) @ butlast ys @
+                               0\<up>(a_md - (pstr + length list)) @ suf_lm) "
+    apply(rule_tac ind2)
+    apply(rule_tac ind, auto)
+    done
+  from k1 and h show 
+    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suf_lm) 
+          (cn_merge_gs (map rec_ci list @ [rec_ci a]) pstr) stp =
+        (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) + 
+        (\<lambda>(ap, pos, n). length ap) (rec_ci a) + (3 + 3 * length list),
+             lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm)"
+    apply(simp add: cn_merge_gs_tl_app)
+    thm abc_append_exc2
+    apply(rule_tac as = 
+  "(\<Sum>(ap, pos, n)\<leftarrow>map rec_ci list. length ap) + 3 * length list"    
+      and bm = "lm @ 0\<up>(pstr - n) @ butlast ys @ 
+                              0\<up>(a_md - (pstr + length list)) @ suf_lm" 
+      and bs = "(\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3" 
+      and bm' = "lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ 
+                                  suf_lm" in abc_append_exc2, simp)
+    apply(simp add: cn_merge_gs_length)
+  proof -
+    from h show 
+      "\<exists>bstp. abc_steps_l (0, lm @ 0\<up>(pstr - n) @ butlast ys @ 
+                                  0\<up>(a_md - (pstr + length list)) @ suf_lm) 
+              ((\<lambda>(gprog, gpara, gn). gprog [+] Recursive.mv_box gpara 
+              (pstr + length list)) (rec_ci a)) bstp =
+              ((\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3, 
+             lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm)"
+      apply(case_tac "rec_ci a", simp)
+      apply(rule_tac as = "length aa" and 
+                     bm = "lm @ (ys ! (length list)) # 
+          0\<up>(pstr - Suc n) @ butlast ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm" 
+        and bs = "3" and bm' = "lm @ 0\<up>(pstr - n) @ ys @
+             0\<up>(a_md - Suc (pstr + length list)) @ suf_lm" in abc_append_exc2)
+    proof -
+      fix aa b c
+      assume g: "rec_ci a = (aa, b, c)"
+      from h and g have k2: "b = n"
+	apply(erule_tac x = "length list" in allE, simp)
+	apply(subgoal_tac "length lm = b", simp)
+	apply(rule para_pattern, simp, simp)
+	done
+      from h and g and this show 
+        "\<exists>astp. abc_steps_l (0, lm @ 0\<up>(pstr - n) @ butlast ys @ 
+                         0\<up>(a_md - (pstr + length list)) @ suf_lm) aa astp =
+        (length aa, lm @ ys ! length list # 0\<up>(pstr - Suc n) @ 
+                       butlast ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm)"
+	apply(subgoal_tac "c \<ge> Suc n")
+	apply(insert ind[of a aa b c lm "ys ! length list" 
+          "0\<up>(pstr - c) @ butlast ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm"], simp)
+	apply(erule_tac x = "length list" in allE, 
+              simp add: exponent_add_iff)
+	apply(rule_tac Suc_leI, rule_tac ci_ad_ge_paras, simp)
+	done
+    next
+      fix aa b c
+      show "length aa = length aa" by simp 
+    next
+      fix aa b c
+      assume "rec_ci a = (aa, b, c)"
+      from h and this show     
+      "\<exists>bstp. abc_steps_l (0, lm @ ys ! length list #
+          0\<up>(pstr - Suc n) @ butlast ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm)
+                 (Recursive.mv_box b (pstr + length list)) bstp =
+       (3, lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm)"
+	apply(insert mv_box_ex [of b "pstr + length list" 
+         "lm @ ys ! length list # 0\<up>(pstr - Suc n) @ butlast ys @ 
+         0\<up>(a_md - (pstr + length list)) @ suf_lm"], simp)
+        apply(subgoal_tac "b = n")
+	apply(simp add: nth_append split: if_splits)
+	apply(erule_tac x = "length list" in allE, simp)
+        apply(drule para_pattern, simp, simp)
+	done
+    next
+      fix aa b c
+      show "3 = length (Recursive.mv_box b (pstr + length list))" 
+        by simp
+    next
+      fix aa b aaa ba
+      show "length aa + 3 = length aa + 3" by simp
+    next
+      fix aa b c
+      show "mv_box b (pstr + length list) \<noteq> []" 
+        by(simp add: mv_box.simps)
+    qed
+  next
+    show "(\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3 = 
+        length ((\<lambda>(gprog, gpara, gn). gprog [+]
+           Recursive.mv_box gpara (pstr + length list)) (rec_ci a))"
+      by(case_tac "rec_ci a", simp)
+  next
+    show "listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) +
+      (\<lambda>(ap, pos, n). length ap) (rec_ci a) + (3 + 3 * length list)=
+      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci list. length ap) + 3 * length list + 
+                ((\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3)" by simp
+  next
+    show "(\<lambda>(gprog, gpara, gn). gprog [+] 
+      Recursive.mv_box gpara (pstr + length list)) (rec_ci a) \<noteq> []"
+      by(case_tac "rec_ci a", 
+         simp add: abc_append.simps abc_shift.simps)
+  qed
+qed
+ 
+lemma [simp]: "length (mv_boxes aa ba n) = 3*n"
+by(induct n, auto simp: mv_boxes.simps)
+
+lemma exp_suc: "a\<up>Suc b = a\<up>b @ [a]"
+by(simp add: exp_ind del: replicate.simps)
+
+lemma [simp]: 
+  "\<lbrakk>Suc n \<le> ba - aa;  length lm2 = Suc n;
+    length lm3 = ba - Suc (aa + n)\<rbrakk>
+  \<Longrightarrow> (last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba - aa) = (0::nat)"
+proof -
+  assume h: "Suc n \<le> ba - aa"
+  and g: "length lm2 = Suc n" "length lm3 = ba - Suc (aa + n)"
+  from h and g have k: "ba - aa = Suc (length lm3 + n)"
+    by arith
+  from  k show 
+    "(last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba - aa) = 0"
+    apply(simp, insert g)
+    apply(simp add: nth_append)
+    done
+qed
+
+lemma [simp]: "length lm1 = aa \<Longrightarrow>
+      (lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (aa + n) = last lm2"
+apply(simp add: nth_append)
+done
+
+lemma [simp]: "\<lbrakk>Suc n \<le> ba - aa; aa < ba\<rbrakk> \<Longrightarrow> 
+                    (ba < Suc (aa + (ba - Suc (aa + n) + n))) = False"
+apply arith
+done
+
+lemma [simp]: "\<lbrakk>Suc n \<le> ba - aa; aa < ba; length lm1 = aa; 
+       length lm2 = Suc n; length lm3 = ba - Suc (aa + n)\<rbrakk>
+     \<Longrightarrow> (lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba + n) = 0"
+using nth_append[of "lm1 @ (0\<Colon>'a)\<up>n @ last lm2 # lm3 @ butlast lm2" 
+                     "(0\<Colon>'a) # lm4" "ba + n"]
+apply(simp)
+done
+
+lemma [simp]: 
+ "\<lbrakk>Suc n \<le> ba - aa; aa < ba; length lm1 = aa; length lm2 = Suc n;
+                 length lm3 = ba - Suc (aa + n)\<rbrakk>
+  \<Longrightarrow> (lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ (0::nat) # lm4)
+  [ba + n := last lm2, aa + n := 0] = 
+  lm1 @ 0 # 0\<up>n @ lm3 @ lm2 @ lm4"
+using list_update_append[of "lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2" "0 # lm4" 
+                            "ba + n" "last lm2"]
+apply(simp)
+apply(simp add: list_update_append)
+apply(simp only: exponent_cons_iff exp_suc, simp)
+apply(case_tac lm2, simp, simp)
+done
+
+lemma mv_boxes_ex:
+  "\<lbrakk>n \<le> ba - aa; ba > aa; length lm1 = aa; 
+    length (lm2::nat list) = n; length lm3 = ba - aa - n\<rbrakk>
+     \<Longrightarrow> \<exists> stp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<up>n @ lm4)
+       (mv_boxes aa ba n) stp = (3 * n, lm1 @ 0\<up>n @ lm3 @ lm2 @ lm4)"
+apply(induct n arbitrary: lm2 lm3 lm4, simp)
+apply(rule_tac x = 0 in exI, simp add: abc_steps_zero, 
+              simp add: mv_boxes.simps del: exp_suc_iff)
+apply(rule_tac as = "3 *n" and bm = "lm1 @ 0\<up>n @ last lm2 # lm3 @
+               butlast lm2 @ 0 # lm4" in abc_append_exc2, simp_all)
+apply(simp only: exponent_cons_iff, simp only: exp_suc, simp)
+proof -
+  fix n lm2 lm3 lm4
+  assume ind:
+    "\<And>lm2 lm3 lm4. \<lbrakk>length lm2 = n; length lm3 = ba - (aa + n)\<rbrakk> \<Longrightarrow>
+    \<exists>stp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<up>n @ lm4) 
+       (mv_boxes aa ba n) stp = (3 * n, lm1 @ 0\<up>n @ lm3 @ lm2 @ lm4)"
+  and h: "Suc n \<le> ba - aa" "aa < ba" "length (lm1::nat list) = aa" 
+         "length (lm2::nat list) = Suc n" 
+         "length (lm3::nat list) = ba - Suc (aa + n)"
+  from h show 
+    "\<exists>astp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<up>n @ 0 # lm4) 
+                       (mv_boxes aa ba n) astp = 
+        (3 * n, lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4)"
+    apply(insert ind[of "butlast lm2" "last lm2 # lm3" "0 # lm4"], 
+          simp)
+    apply(subgoal_tac "lm1 @ butlast lm2 @ last lm2 # lm3 @ 0\<up>n @ 
+              0 # lm4 = lm1 @ lm2 @ lm3 @ 0\<up>n @ 0 # lm4", simp, simp)
+    apply(case_tac "lm2 = []", simp, simp)
+    done
+next
+  fix n lm2 lm3 lm4
+  assume h: "Suc n \<le> ba - aa"
+            "aa < ba" 
+            "length (lm1::nat list) = aa" 
+            "length (lm2::nat list) = Suc n" 
+            "length (lm3::nat list) = ba - Suc (aa + n)"
+  thus " \<exists>bstp. abc_steps_l (0, lm1 @ 0\<up>n @ last lm2 # lm3 @
+                       butlast lm2 @ 0 # lm4) 
+                         (Recursive.mv_box (aa + n) (ba + n)) bstp
+               = (3, lm1 @ 0 # 0\<up>n @ lm3 @ lm2 @ lm4)"
+    apply(insert mv_box_ex[of "aa + n" "ba + n" 
+       "lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4"], simp)
+    done
+qed
+(*    
+lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; 
+                ba < aa; 
+               length lm2 = aa - Suc (ba + n)\<rbrakk>
+      \<Longrightarrow> ((0::nat) # lm2 @ 0\<up>n @ last lm3 # lm4) ! (aa - ba)
+         = last lm3"
+proof -
+  assume h: "Suc n \<le> aa - ba"
+    and g: " ba < aa" "length lm2 = aa - Suc (ba + n)"
+  from h and g have k: "aa - ba = Suc (length lm2 + n)"
+    by arith
+  thus "((0::nat) # lm2 @ 0\<up>n @ last lm3 # lm4) ! (aa - ba) = last lm3"
+    apply(simp,  simp add: nth_append)
+    done
+qed
+*)
+
+lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; ba < aa; length lm1 = ba; 
+        length lm2 = aa - Suc (ba + n); length lm3 = Suc n\<rbrakk>
+   \<Longrightarrow> (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ last lm3 # lm4) ! (aa + n) = last lm3"
+using nth_append[of "lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n" "last lm3 # lm4" "aa + n"]
+apply(simp)
+done
+
+lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; ba < aa; length lm1 = ba; 
+        length lm2 = aa - Suc (ba + n); length lm3 = Suc n\<rbrakk>
+     \<Longrightarrow> (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ last lm3 # lm4) ! (ba + n) = 0"
+apply(simp add: nth_append)
+done
+
+lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; ba < aa; length lm1 = ba; 
+        length lm2 = aa - Suc (ba + n); length lm3 = Suc n\<rbrakk> 
+     \<Longrightarrow> (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ last lm3 # lm4)[ba + n := last lm3, aa + n := 0]
+      = lm1 @ lm3 @ lm2 @ 0 # 0\<up>n @ lm4"
+using list_update_append[of "lm1 @ butlast lm3" "(0\<Colon>'a) # lm2 @ (0\<Colon>'a)\<up>n @ last lm3 # lm4"]
+apply(simp)
+using list_update_append[of "lm1 @ butlast lm3 @ last lm3 # lm2 @ (0\<Colon>'a)\<up>n"
+                            "last lm3 # lm4" "aa + n" "0"]
+apply(simp)
+apply(simp only: replicate_Suc[THEN sym] exp_suc, simp)
+apply(case_tac lm3, simp, simp)
+done
+
+lemma mv_boxes_ex2:
+  "\<lbrakk>n \<le> aa - ba; 
+    ba < aa; 
+    length (lm1::nat list) = ba;
+    length (lm2::nat list) = aa - ba - n; 
+    length (lm3::nat list) = n\<rbrakk>
+     \<Longrightarrow> \<exists> stp. abc_steps_l (0, lm1 @ 0\<up>n @ lm2 @ lm3 @ lm4) 
+                (mv_boxes aa ba n) stp =
+                    (3 * n, lm1 @ lm3 @ lm2 @ 0\<up>n @ lm4)"
+apply(induct n arbitrary: lm2 lm3 lm4, simp)
+apply(rule_tac x = 0 in exI, simp add: abc_steps_zero, 
+                   simp add: mv_boxes.simps del: exp_suc_iff)
+apply(rule_tac as = "3 *n" and bm = "lm1 @ butlast lm3 @ 0 # lm2 @
+                  0\<up>n @ last lm3 # lm4" in abc_append_exc2, simp_all)
+apply(simp only: exponent_cons_iff, simp only: exp_suc, simp)
+proof -
+  fix n lm2 lm3 lm4
+  assume ind: 
+"\<And>lm2 lm3 lm4. \<lbrakk>length lm2 = aa - (ba + n); length lm3 = n\<rbrakk> \<Longrightarrow> 
+  \<exists>stp. abc_steps_l (0, lm1 @ 0\<up>n @ lm2 @ lm3 @ lm4) 
+                 (mv_boxes aa ba n) stp = 
+                            (3 * n, lm1 @ lm3 @ lm2 @ 0\<up>n @ lm4)"
+  and h: "Suc n \<le> aa - ba" 
+         "ba < aa"  
+         "length (lm1::nat list) = ba" 
+         "length (lm2::nat list) = aa - Suc (ba + n)" 
+         "length (lm3::nat list) = Suc n"
+  from h show
+    "\<exists>astp. abc_steps_l (0, lm1 @ 0\<up>n @ 0 # lm2 @ lm3 @ lm4)
+        (mv_boxes aa ba n) astp = 
+          (3 * n, lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ last lm3 # lm4)"
+    apply(insert ind[of "0 # lm2" "butlast lm3" "last lm3 # lm4"],
+          simp)
+    apply(subgoal_tac
+      "lm1 @ 0\<up>n @ 0 # lm2 @ butlast lm3 @ last lm3 # lm4 =
+           lm1 @ 0\<up>n @ 0 # lm2 @ lm3 @ lm4", simp, simp)
+    apply(case_tac "lm3 = []", simp, simp)
+    done
+next
+  fix n lm2 lm3 lm4
+  assume h:
+    "Suc n \<le> aa - ba" 
+    "ba < aa"
+    "length lm1 = ba"
+    "length (lm2::nat list) = aa - Suc (ba + n)" 
+    "length (lm3::nat list) = Suc n"
+  thus
+    "\<exists>bstp. abc_steps_l (0, lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ 
+                               last lm3 # lm4) 
+           (Recursive.mv_box (aa + n) (ba + n)) bstp =
+                 (3, lm1 @ lm3 @ lm2 @ 0 # 0\<up>n @ lm4)"
+    apply(insert mv_box_ex[of "aa + n" "ba + n" "lm1 @ butlast lm3 @ 
+                          0 # lm2 @ 0\<up>n @ last lm3 # lm4"], simp)
+    done
+qed
+
+lemma cn_merge_gs_len: 
+  "length (cn_merge_gs (map rec_ci gs) pstr) = 
+      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs"
+apply(induct gs arbitrary: pstr, simp, simp)
+apply(case_tac "rec_ci a", simp)
+done
+
+lemma [simp]: "n < pstr \<Longrightarrow>
+     Suc (pstr + length ys - n) = Suc (pstr + length ys) - n"
+by arith
+
+lemma save_paras':  
+  "\<lbrakk>length lm = n; pstr > n; a_md > pstr + length ys + n\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(pstr - n) @ ys @
+               0\<up>(a_md - pstr - length ys) @ suf_lm) 
+                 (mv_boxes 0 (pstr + Suc (length ys)) n) stp
+        = (3 * n, 0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+thm mv_boxes_ex
+apply(insert mv_boxes_ex[of n "pstr + Suc (length ys)" 0 "[]" "lm" 
+         "0\<up>(pstr - n) @ ys @ [0]" "0\<up>(a_md - pstr - length ys - n - Suc 0) @ suf_lm"], simp)
+apply(erule_tac exE, rule_tac x = stp in exI,
+                            simp add: exponent_add_iff)
+apply(simp only: exponent_cons_iff, simp)
+done
+
+lemma [simp]:
+ "(max ba (Max (insert ba (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs))))
+ = (Max (insert ba (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs)))"
+apply(rule min_max.sup_absorb2, auto)
+done
+
+lemma [simp]:
+  "((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs) = 
+                  (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs)"
+apply(induct gs)
+apply(simp, simp)
+done
+
+lemma ci_cn_md_def:  
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
+  rec_ci f = (a, aa, ba)\<rbrakk>
+    \<Longrightarrow> a_md = max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) o 
+  rec_ci) ` set gs))) + Suc (length gs) + n"
+apply(simp add: rec_ci.simps, auto)
+done
+
+lemma save_paras_prog_ex:
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
+    rec_ci f = (a, aa, ba); 
+    pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
+                                    (map rec_ci (f # gs))))\<rbrakk>
+    \<Longrightarrow> \<exists>ap bp cp. 
+      aprog = ap [+] bp [+] cp \<and>
+      length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+              3 * length gs \<and> bp = mv_boxes 0 (pstr + Suc (length gs)) n"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = 
+  "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba 
+      (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs))))" in exI,
+      simp add: cn_merge_gs_len)
+apply(rule_tac x = 
+  "mv_boxes (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
+   0 (length gs) [+] a [+]Recursive.mv_box aa (max (Suc n) 
+   (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
+   empty_boxes (length gs) [+] Recursive.mv_box (max (Suc n) 
+  (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
+   mv_boxes (Suc (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) 
+   ` set gs))) + length gs)) 0 n" in exI, auto)
+apply(simp add: abc_append_commute)
+done
+
+lemma save_paras: 
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
+    rs_pos = n;
+    \<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k);
+    length ys = length gs;
+    length lm = n;
+    rec_ci f = (a, aa, ba);
+    pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
+                                          (map rec_ci (f # gs))))\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+          3 * length gs, lm @ 0\<up>(pstr - n) @ ys @
+                 0\<up>(a_md - pstr - length ys) @ suf_lm) aprog stp = 
+           ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+                      3 * length gs + 3 * n, 
+             0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+proof -
+  assume h:
+    "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
+    "rs_pos = n" 
+    "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)" 
+    "length ys = length gs"  
+    "length lm = n"    
+    "rec_ci f = (a, aa, ba)"
+    and g: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                                        (map rec_ci (f # gs))))"
+  from h and g have k1: 
+    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
+    length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+                3 *length gs \<and> bp = mv_boxes 0 (pstr + Suc (length ys)) n"
+    apply(drule_tac save_paras_prog_ex, auto)
+    done
+  from h have k2: 
+    "\<exists> stp. abc_steps_l (0, lm @ 0\<up>(pstr - n) @ ys @ 
+                         0\<up>(a_md - pstr - length ys) @ suf_lm)
+         (mv_boxes 0 (pstr + Suc (length ys)) n) stp = 
+        (3 * n, 0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+    apply(rule_tac save_paras', simp, simp_all add: g)
+    apply(drule_tac a = a and aa = aa and ba = ba in 
+                                        ci_cn_md_def, simp, simp)
+    done
+  from k1 show 
+    "\<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+         3 * length gs, lm @ 0\<up>(pstr - n) @ ys @ 
+                 0\<up>(a_md - pstr - length ys) @ suf_lm) aprog stp =
+             ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+               3 * length gs + 3 * n, 
+                0\<up> pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+  proof(erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
+    fix ap bp apa cp
+    assume "aprog = ap [+] bp [+] cp \<and> length ap = 
+            (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs
+            \<and> bp = mv_boxes 0 (pstr + Suc (length ys)) n"
+    from this and k2 show "?thesis"
+      apply(simp)
+      apply(rule_tac abc_append_exc1, simp, simp, simp)
+      done
+  qed
+qed
+ 
+lemma ci_cn_para_eq:
+  "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md) \<Longrightarrow> rs_pos = n"
+apply(simp add: rec_ci.simps, case_tac "rec_ci f", simp)
+done
+
+lemma calc_gs_prog_ex: 
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
+    rec_ci f = (a, aa, ba);
+    Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                         (map rec_ci (f # gs)))) = pstr\<rbrakk>
+   \<Longrightarrow> \<exists>ap bp. aprog = ap [+] bp \<and> 
+                 ap = cn_merge_gs (map rec_ci gs) pstr"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "mv_boxes 0 (Suc (max (Suc n)  
+   (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n [+]
+   mv_boxes (max (Suc n) (Max (insert ba 
+  (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
+   a [+] Recursive.mv_box aa (max (Suc n)
+    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
+   empty_boxes (length gs) [+] Recursive.mv_box (max (Suc n)
+    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
+    mv_boxes (Suc (max (Suc n) (Max 
+    (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n"
+   in exI)
+apply(auto simp: abc_append_commute)
+done
+
+lemma cn_calc_gs: 
+  assumes ind: 
+  "\<And>x aprog a_md rs_pos rs suf_lm lm.
+  \<lbrakk>x \<in> set gs; 
+   rec_ci x = (aprog, rs_pos, a_md); 
+   rec_calc_rel x lm rs\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
+     (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+  and h:  "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"  
+          "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
+          "length ys = length gs" 
+          "length lm = n" 
+          "rec_ci f = (a, aa, ba)" 
+          "Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                               (map rec_ci (f # gs)))) = pstr"
+  shows  
+  "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
+  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs, 
+   lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md -pstr - length ys) @ suf_lm) "
+proof -
+  from h have k1:
+    "\<exists> ap bp. aprog = ap [+] bp \<and> ap = 
+                        cn_merge_gs (map rec_ci gs) pstr"
+    by(erule_tac calc_gs_prog_ex, auto)
+  from h have j1: "rs_pos = n"
+    by(simp add: ci_cn_para_eq)
+  from h have j2: "a_md \<ge> pstr"
+    by(drule_tac a = a and aa = aa and ba = ba in 
+                                ci_cn_md_def, simp, simp)
+  from h have j3: "pstr > n"
+    by(auto)    
+  from j1 and j2 and j3 and h have k2:
+    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) 
+                         (cn_merge_gs (map rec_ci gs) pstr) stp 
+    = ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs, 
+                  lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - pstr - length ys) @ suf_lm)"
+    apply(simp)
+    apply(rule_tac cn_merge_gs_ex, rule_tac ind, simp, simp, auto)
+    apply(drule_tac a = a and aa = aa and ba = ba in 
+                                 ci_cn_md_def, simp, simp)
+    apply(rule min_max.le_supI2, auto)
+    done
+  from k1 show "?thesis"
+  proof(erule_tac exE, erule_tac exE, simp)
+    fix ap bp
+    from k2 show 
+      "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm)
+           (cn_merge_gs (map rec_ci gs) pstr [+] bp) stp =
+      (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) gs) +
+         3 * length gs, 
+         lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - (pstr + length ys)) @ suf_lm)"
+      apply(insert abc_append_exc1[of 
+        "lm @ 0\<up>(a_md - rs_pos) @ suf_lm" 
+        "(cn_merge_gs (map rec_ci gs) pstr)" 
+        "length (cn_merge_gs (map rec_ci gs) pstr)" 
+        "lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - pstr - length ys) @ suf_lm" 0 
+        "[]" bp], simp add: cn_merge_gs_len)
+      done      
+  qed
+qed
+
+lemma reset_new_paras': 
+  "\<lbrakk>length lm = n; 
+    pstr > 0; 
+    a_md \<ge> pstr + length ys + n;
+     pstr > length ys\<rbrakk> \<Longrightarrow>
+   \<exists>stp. abc_steps_l (0, 0\<up>pstr @ ys @ 0 # lm @  0\<up>(a_md - Suc (pstr + length ys + n)) @
+          suf_lm) (mv_boxes pstr 0 (length ys)) stp =
+  (3 * length ys, ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+thm mv_boxes_ex2
+apply(insert mv_boxes_ex2[of "length ys" "pstr" 0 "[]"
+     "0\<up>(pstr - length ys)" "ys" 
+     "0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm"], 
+     simp add: exponent_add_iff)
+done
+
+lemma [simp]:  
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
+  rec_calc_rel f ys rs; rec_ci f = (a, aa, ba);
+  pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+               (map rec_ci (f # gs))))\<rbrakk>
+  \<Longrightarrow> length ys < pstr"
+apply(subgoal_tac "length ys = aa", simp)
+apply(subgoal_tac "aa < ba \<and> ba \<le> pstr", 
+      rule basic_trans_rules(22), auto)
+apply(rule min_max.le_supI2)
+apply(auto)
+apply(erule_tac para_pattern, simp)
+done
+
+lemma reset_new_paras_prog_ex: 
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
+   rec_ci f = (a, aa, ba);
+   Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+  (map rec_ci (f # gs)))) = pstr\<rbrakk>
+  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
+  length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+           3 *length gs + 3 * n \<and> bp = mv_boxes pstr 0 (length gs)"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) 
+          (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
+          mv_boxes 0 (Suc (max (Suc n) (Max (insert ba 
+           (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n" in exI, 
+       simp add: cn_merge_gs_len)
+apply(rule_tac x = "a [+]
+     Recursive.mv_box aa (max (Suc n) (Max (insert ba 
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
+     empty_boxes (length gs) [+] Recursive.mv_box 
+     (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n
+      [+] mv_boxes (Suc (max (Suc n) (Max (insert ba 
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI,
+       auto simp: abc_append_commute)
+done
+
+lemma reset_new_paras:
+       "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
+        rs_pos = n;
+        \<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k);
+        length ys = length gs;
+        length lm = n;
+        length ys = aa;
+        rec_ci f = (a, aa, ba);
+        pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
+                                    (map rec_ci (f # gs))))\<rbrakk>
+\<Longrightarrow> \<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+                                               3 * length gs + 3 * n,
+        0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
+  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n,
+           ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+proof -
+  assume h:
+    "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
+    "rs_pos = n" 
+    "length ys = aa"
+    "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
+    "length ys = length gs"  "length lm = n"    
+    "rec_ci f = (a, aa, ba)"
+    and g: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
+                                         (map rec_ci (f # gs))))"
+  thm rec_ci.simps
+  from h and g have k1:
+    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = 
+    (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+          3 *length gs + 3 * n \<and> bp = mv_boxes pstr 0 (length ys)"
+    by(drule_tac reset_new_paras_prog_ex, auto)
+  from h have k2:
+    "\<exists> stp. abc_steps_l (0, 0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @
+              suf_lm) (mv_boxes pstr 0 (length ys)) stp = 
+    (3 * (length ys), 
+     ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+    apply(rule_tac reset_new_paras', simp)
+    apply(simp add: g)
+    apply(drule_tac a = a and aa = aa and ba = ba in ci_cn_md_def,
+      simp, simp add: g, simp)
+    apply(subgoal_tac "length gs = aa \<and> aa < ba \<and> ba \<le> pstr", arith,
+          simp add: para_pattern)
+    apply(insert g, auto intro: min_max.le_supI2)
+    done
+  from k1 show 
+    "\<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3
+    * length gs + 3 * n, 0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ 
+     suf_lm) aprog stp =
+    ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs +
+      3 * n, ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+  proof(erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
+    fix ap bp apa cp
+    assume "aprog = ap [+] bp [+] cp \<and> length ap = 
+      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs +
+                  3 * n \<and> bp = mv_boxes pstr 0 (length ys)"
+    from this and k2 show "?thesis"
+      apply(simp)
+      apply(drule_tac as = 
+        "(\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs +
+        3 * n" and ap = ap and cp = cp in abc_append_exc1, auto)
+      apply(rule_tac x = stp in exI, simp add: h)
+      using h
+      apply(simp)
+      done
+  qed
+qed
+
+thm rec_ci.simps 
+
+lemma calc_f_prog_ex: 
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
+    rec_ci f = (a, aa, ba);
+    Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                   (map rec_ci (f # gs)))) = pstr\<rbrakk>
+   \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
+  length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+                                6 *length gs + 3 * n \<and> bp = a"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
+     mv_boxes 0 (Suc (max (Suc n) (Max (insert ba 
+            (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n [+]
+     mv_boxes (max (Suc n) (Max (insert ba 
+      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs)" in exI,
+     simp add: cn_merge_gs_len)
+apply(rule_tac x = "Recursive.mv_box aa (max (Suc n) (Max (insert ba 
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
+     empty_boxes (length gs) [+] Recursive.mv_box (max (Suc n) (
+     Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
+     mv_boxes (Suc (max (Suc n) (Max (insert ba 
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI,
+  auto simp: abc_append_commute)
+done
+
+lemma calc_cn_f:
+  assumes ind:
+  "\<And>x aprog a_md rs_pos rs suf_lm lm.
+  \<lbrakk>x \<in> set (f # gs);
+  rec_ci x = (aprog, rs_pos, a_md); 
+  rec_calc_rel x lm rs\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
+  (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+  and h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
+  "rec_calc_rel (Cn n f gs) lm rs"
+  "length ys = length gs"
+  "rec_calc_rel f ys rs"
+  "length lm = n"
+  "rec_ci f = (a, aa, ba)" 
+  and p: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
+                                (map rec_ci (f # gs))))"
+  shows "\<exists>stp. abc_steps_l   
+  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n,
+  ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
+  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 
+                3 * n + length a,
+  ys @ rs # 0\<up>pstr @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+proof -
+  from h have k1: 
+    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
+    length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+    6 *length gs + 3 * n \<and> bp = a"
+    by(drule_tac calc_f_prog_ex, auto)
+  from h and k1 show "?thesis"
+  proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
+    fix ap bp apa cp
+    assume
+      "aprog = ap [+] bp [+] cp \<and> 
+      length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+      6 * length gs + 3 * n \<and> bp = a"
+    from h and this show "?thesis"
+      apply(simp, rule_tac abc_append_exc1, simp_all)
+      apply(insert ind[of f "a" aa ba ys rs 
+        "0\<up>(pstr - ba + length gs) @ 0 # lm @ 
+        0\<up>(a_md - Suc (pstr + length gs + n)) @ suf_lm"], simp)
+      apply(subgoal_tac "ba > aa \<and> aa = length gs\<and> pstr \<ge> ba", simp)
+      apply(simp add: exponent_add_iff)
+      apply(case_tac pstr, simp add: p)
+      apply(simp only: exp_suc, simp)
+      apply(rule conjI, rule ci_ad_ge_paras, simp, rule conjI)
+      apply(subgoal_tac "length ys = aa", simp,
+        rule para_pattern, simp, simp)
+      apply(insert p, simp)
+      apply(auto intro: min_max.le_supI2)
+      done
+  qed
+qed
+(*
+lemma [simp]: 
+  "\<lbrakk>pstr + length ys + n \<le> a_md; ys \<noteq> []\<rbrakk> \<Longrightarrow> 
+                          pstr < a_md + length suf_lm"
+apply(case_tac "length ys", simp)
+apply(arith)
+done
+*)
+
+lemma [simp]: 
+  "pstr > length ys 
+  \<Longrightarrow> (ys @ rs # 0\<up>pstr @ lm @
+  0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) ! pstr = (0::nat)"
+apply(simp add: nth_append)
+done
+
+(*
+lemma [simp]: "\<lbrakk>length ys < pstr; pstr - length ys = Suc x\<rbrakk>
+  \<Longrightarrow> pstr - Suc (length ys) = x"
+by arith
+*)
+
+lemma [simp]: "pstr > length ys \<Longrightarrow> 
+      (ys @ rs # 0\<up>pstr @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)
+                                         [pstr := rs, length ys := 0] =
+       ys @ 0\<up>(pstr - length ys) @ (rs::nat) # 0\<up>length ys @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm"
+apply(auto simp: list_update_append)
+apply(case_tac "pstr - length ys",simp_all)
+using list_update_length[of 
+  "0\<up>(pstr - Suc (length ys))" "0" "0\<up>length ys @ lm @ 
+  0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm" rs]
+apply(simp only: exponent_cons_iff exponent_add_iff, simp)
+apply(subgoal_tac "pstr - Suc (length ys) = nat", simp, simp)
+done
+
+lemma save_rs': 
+  "\<lbrakk>pstr > length ys\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, ys @ rs # 0\<up>pstr @ lm @ 
+  0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) 
+  (Recursive.mv_box (length ys) pstr) stp =
+  (3, ys @ 0\<up>(pstr - (length ys)) @ rs # 
+  0\<up>length ys  @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+apply(insert mv_box_ex[of "length ys" pstr 
+  "ys @ rs # 0\<up>pstr @ lm @ 0\<up>(a_md - Suc(pstr + length ys + n)) @ suf_lm"], 
+  simp)
+done
+
+
+lemma save_rs_prog_ex:
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
+  rec_ci f = (a, aa, ba);
+  Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                        (map rec_ci (f # gs)))) = pstr\<rbrakk>
+  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
+  length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+              6 *length gs + 3 * n + length a
+  \<and> bp = mv_box aa pstr"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x =
+  "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba 
+   (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
+   [+] mv_boxes 0 (Suc (max (Suc n) (Max (insert ba 
+   (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n [+]
+   mv_boxes (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
+    0 (length gs) [+] a" 
+  in exI, simp add: cn_merge_gs_len)
+apply(rule_tac x = 
+  "empty_boxes (length gs) [+]
+   Recursive.mv_box (max (Suc n) (Max (insert ba 
+    (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
+   mv_boxes (Suc (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))
+    + length gs)) 0 n" in exI, 
+  auto simp: abc_append_commute)
+done
+
+lemma save_rs:  
+  assumes h: 
+  "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
+  "rec_calc_rel (Cn n f gs) lm rs"
+  "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
+  "length ys = length gs" 
+  "rec_calc_rel f ys rs" 
+  "rec_ci f = (a, aa, ba)"  
+  "length lm = n"
+  and pdef: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                                            (map rec_ci (f # gs))))"
+  shows "\<exists>stp. abc_steps_l
+           ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs
+          + 3 * n + length a, ys @ rs # 0\<up>pstr @ lm @
+             0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
+  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs 
+  + 3 * n + length a + 3,
+  ys @ 0\<up>(pstr - length ys) @ rs # 0\<up>length ys @ lm @ 
+                               0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+proof -
+  thm rec_ci.simps
+  from h and pdef have k1: 
+    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
+    length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+    6 *length gs + 3 * n + length a \<and> bp = mv_box (length ys) pstr "
+    apply(subgoal_tac "length ys = aa")
+    apply(drule_tac a = a and aa = aa and ba = ba in save_rs_prog_ex, 
+      simp, simp, simp)
+    by(rule_tac para_pattern, simp, simp)
+  from k1 show 
+    "\<exists>stp. abc_steps_l
+    ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n
+    + length a, ys @ rs # 0\<up>pstr @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) 
+    @ suf_lm) aprog stp =
+    ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n
+    + length a + 3, ys @ 0\<up>(pstr - length ys) @ rs # 
+    0\<up>length ys @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+  proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
+    fix ap bp apa cp
+    assume "aprog = ap [+] bp [+] cp \<and> length ap = 
+      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 
+      3 * n + length a \<and> bp = Recursive.mv_box (length ys) pstr"
+    thus"?thesis"
+      apply(simp, rule_tac abc_append_exc1, simp_all)
+      apply(rule_tac save_rs', insert h)
+      apply(subgoal_tac "length gs = aa \<and> pstr \<ge> ba \<and> ba > aa",
+            arith)
+      apply(simp add: para_pattern, insert pdef, auto)
+      apply(rule_tac min_max.le_supI2, simp)
+      done
+  qed
+qed
+
+lemma [simp]: "length (empty_boxes n) = 2*n"
+apply(induct n, simp, simp)
+done
+
+lemma mv_box_step_ex: "length lm = n \<Longrightarrow> 
+      \<exists>stp. abc_steps_l (0, lm @ Suc x # suf_lm) [Dec n 2, Goto 0] stp
+  = (0, lm @ x # suf_lm)"
+apply(rule_tac x = "Suc (Suc 0)" in exI, 
+  simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps 
+         abc_lm_v.simps abc_lm_s.simps nth_append list_update_append)
+done
+
+lemma mv_box_ex': 
+  "\<lbrakk>length lm = n\<rbrakk> \<Longrightarrow> 
+  \<exists> stp. abc_steps_l (0, lm @ x # suf_lm) [Dec n 2, Goto 0] stp =
+  (Suc (Suc 0), lm @ 0 # suf_lm)"
+apply(induct x)
+apply(rule_tac x = "Suc 0" in exI, 
+  simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps
+            abc_lm_v.simps nth_append abc_lm_s.simps, simp)
+apply(drule_tac x = x and suf_lm = suf_lm in mv_box_step_ex, 
+      erule_tac exE, erule_tac exE)
+apply(rule_tac x = "stpa + stp" in exI, simp add: abc_steps_add)
+done
+
+lemma [simp]: "drop n lm = a # list \<Longrightarrow> list = drop (Suc n) lm"
+apply(induct n arbitrary: lm a list, simp)
+apply(case_tac "lm", simp, simp)
+done
+
+lemma empty_boxes_ex: "\<lbrakk>length lm \<ge> n\<rbrakk>
+     \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm) (empty_boxes n) stp = 
+                                          (2*n, 0\<up>n @ drop n lm)"
+apply(induct n, simp, simp)
+apply(rule_tac abc_append_exc2, auto)
+apply(case_tac "drop n lm", simp, simp)
+proof -
+  fix n stp a list
+  assume h: "Suc n \<le> length lm"  "drop n lm = a # list"
+  thus "\<exists>bstp. abc_steps_l (0, 0\<up>n @ a # list) [Dec n 2, Goto 0] bstp =
+                       (Suc (Suc 0), 0 # 0\<up>n @ drop (Suc n) lm)"
+    apply(insert mv_box_ex'[of "0\<up>n" n a list], simp, erule_tac exE)
+    apply(rule_tac x = stp in exI, simp, simp only: exponent_cons_iff)
+    apply(simp add:exp_ind del: replicate.simps)
+    done
+qed
+
+
+lemma mv_box_paras_prog_ex:
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
+  rec_ci f = (a, aa, ba); 
+  Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                    (map rec_ci (f # gs)))) = pstr\<rbrakk>
+  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
+  length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+  6 *length gs + 3 * n + length a + 3 \<and> bp = empty_boxes (length gs)"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) 
+    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
+    mv_boxes 0 (Suc (max (Suc n) (Max 
+     (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n
+    [+] mv_boxes (max (Suc n) (Max (insert ba 
+    (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
+     a [+] Recursive.mv_box aa (max (Suc n) 
+   (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))" 
+    in exI, simp add: cn_merge_gs_len)
+apply(rule_tac x = " Recursive.mv_box (max (Suc n) (Max (insert ba
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
+     mv_boxes (Suc (max (Suc n) (Max (insert ba 
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI, 
+  auto simp: abc_append_commute)
+done
+
+lemma mv_box_paras: 
+ assumes h: 
+  "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
+  "rec_calc_rel (Cn n f gs) lm rs" 
+  "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
+  "length ys = length gs" 
+  "rec_calc_rel f ys rs" 
+  "rec_ci f = (a, aa, ba)" 
+  and pdef: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                                             (map rec_ci (f # gs))))"
+  and starts: "ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+                              6 * length gs + 3 * n + length a + 3"
+  shows "\<exists>stp. abc_steps_l
+           (ss, ys @ 0\<up>(pstr - length ys) @ rs # 0\<up>length ys 
+               @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
+   (ss + 2 * length gs, 0\<up>pstr @ rs # 0\<up>length ys  @ lm @ 
+                                0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+proof -
+  from h and pdef and starts have k1: 
+    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
+    length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+                               6 *length gs + 3 * n + length a + 3
+    \<and> bp = empty_boxes (length ys)"
+    by(drule_tac mv_box_paras_prog_ex, auto)
+  from h have j1: "aa < ba"
+    by(simp add: ci_ad_ge_paras)
+  from h have j2: "length gs = aa"
+    by(drule_tac f = f in para_pattern, simp, simp)
+  from h and pdef have j3: "ba \<le> pstr"
+    apply simp 
+    apply(rule_tac min_max.le_supI2, simp)
+    done
+  from k1 show "?thesis"
+  proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
+    fix ap bp apa cp
+    assume "aprog = ap [+] bp [+] cp \<and> 
+      length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+      6 * length gs + 3 * n + length a + 3 \<and> 
+      bp = empty_boxes (length ys)"
+    thus"?thesis"
+      apply(simp, rule_tac abc_append_exc1, simp_all add: starts h)
+      apply(insert empty_boxes_ex[of 
+        "length gs" "ys @ 0\<up>(pstr - (length gs)) @ rs #
+        0\<up>length gs @ lm @ 0\<up>(a_md - Suc (pstr + length gs + n)) @ suf_lm"], 
+        simp add: h)
+      apply(erule_tac exE, rule_tac x = stp in exI, 
+        simp add: replicate.simps[THEN sym]
+        replicate_add[THEN sym] del: replicate.simps)
+      apply(subgoal_tac "pstr >(length gs)", simp)
+      apply(subgoal_tac "ba > aa \<and> length gs = aa \<and> pstr \<ge> ba", simp)
+      apply(simp add: j1 j2 j3)
+      done     
+  qed
+qed
+
+lemma restore_rs_prog_ex:
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
+  rec_ci f = (a, aa, ba);
+  Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
+  (map rec_ci (f # gs)))) = pstr;
+  ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+  8 * length gs + 3 * n + length a + 3\<rbrakk>
+  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
+                                           bp = mv_box pstr n"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) 
+      (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
+      mv_boxes 0 (Suc (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n)
+        \<circ> rec_ci) ` set gs))) + length gs)) n [+]
+     mv_boxes (max (Suc n) (Max (insert ba 
+      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
+     a [+] Recursive.mv_box aa (max (Suc n)
+       (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
+     empty_boxes (length gs)" in exI, simp add: cn_merge_gs_len)
+apply(rule_tac x = "mv_boxes (Suc (max (Suc n) 
+       (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) 
+        + length gs)) 0 n" 
+  in exI, auto simp: abc_append_commute)
+done
+
+lemma exp_add: "a\<up>(b+c) = a\<up>b @ a\<up>c"
+apply(simp add:replicate_add)
+done
+
+lemma [simp]: "n < pstr \<Longrightarrow> (0\<up>pstr)[n := rs] @ [0::nat] = 0\<up>n @ rs # 0\<up>(pstr - n)"
+using list_update_length[of "0\<up>n" "0::nat" "0\<up>(pstr - Suc n)" rs]
+apply(simp add: replicate_Suc[THEN sym] exp_add[THEN sym] exp_suc[THEN sym])
+done
+
+lemma restore_rs:
+  assumes h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
+  "rec_calc_rel (Cn n f gs) lm rs" 
+  "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
+  "length ys = length gs"
+  "rec_calc_rel f ys rs" 
+  "rec_ci f = (a, aa, ba)" 
+  and pdef: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
+                                        (map rec_ci (f # gs))))"
+  and starts: "ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+                              8 * length gs + 3 * n + length a + 3" 
+  shows "\<exists>stp. abc_steps_l
+           (ss, 0\<up>pstr @ rs # 0\<up>length ys  @ lm @
+                    0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
+  (ss + 3, 0\<up>n @ rs # 0\<up>(pstr - n) @ 0\<up>length ys  @ lm @ 
+                                   0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
+proof -
+ from h and pdef and starts have k1:
+   "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
+                                            bp = mv_box pstr n"
+   by(drule_tac restore_rs_prog_ex, auto)
+ from k1 show "?thesis"
+ proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
+   fix ap bp apa cp
+   assume "aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
+                                 bp = Recursive.mv_box pstr n"
+   thus"?thesis"
+     apply(simp, rule_tac abc_append_exc1, simp_all add: starts h)
+     apply(insert mv_box_ex[of pstr n "0\<up>pstr @ rs # 0\<up>length gs @
+                     lm @ 0\<up>(a_md - Suc (pstr + length gs + n)) @ suf_lm"], simp)
+     apply(subgoal_tac "pstr > n", simp)
+     apply(erule_tac exE, rule_tac x = stp in exI, 
+                         simp add: nth_append list_update_append)
+     apply(simp add: pdef)
+     done
+  qed
+qed
+
+lemma [simp]:"xs \<noteq> [] \<Longrightarrow> length xs \<ge> Suc 0"
+by(case_tac xs, auto)
+
+lemma  [simp]: "n < max (Suc n) (max ba (Max (((\<lambda>(aprog, p, n). n) o 
+                                                  rec_ci) ` set gs)))"
+by(simp)
+
+lemma restore_paras_prog_ex: 
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
+  rec_ci f = (a, aa, ba);
+  Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
+                          (map rec_ci (f # gs)))) = pstr;
+  ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+                         8 * length gs + 3 * n + length a + 6\<rbrakk>
+  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
+                      bp = mv_boxes (pstr + Suc (length gs)) (0::nat) n"
+apply(simp add: rec_ci.simps)
+apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) 
+      (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
+      [+] mv_boxes 0 (Suc (max (Suc n) 
+       (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) 
+     + length gs)) n [+] mv_boxes (max (Suc n) 
+    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
+     a [+] Recursive.mv_box aa (max (Suc n) 
+      (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
+     empty_boxes (length gs) [+]
+     Recursive.mv_box (max (Suc n) (Max (insert ba 
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n" in exI, simp add: cn_merge_gs_len)
+apply(rule_tac x = "[]" in exI, auto simp: abc_append_commute)
+done
+
+lemma restore_paras: 
+  assumes h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
+  "rec_calc_rel (Cn n f gs) lm rs" 
+  "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
+  "length ys = length gs"
+  "rec_calc_rel f ys rs" 
+  "rec_ci f = (a, aa, ba)"
+  and pdef: 
+  "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
+                         (map rec_ci (f # gs))))"
+  and starts: "ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
+                              8 * length gs + 3 * n + length a + 6" 
+  shows "\<exists>stp. abc_steps_l (ss, 0\<up>n @ rs # 0\<up>(pstr - n+ length ys) @
+                         lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)
+  aprog stp = (ss + 3 * n, lm @ rs # 0\<up>(a_md - Suc n) @ suf_lm)"
+proof -
+  thm rec_ci.simps
+  from h and pdef and starts have k1:
+    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and>
+                     bp = mv_boxes (pstr + Suc (length gs)) (0::nat) n"
+    by(drule_tac restore_paras_prog_ex, auto)
+  from k1 show "?thesis"
+  proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
+    fix ap bp apa cp
+    assume "aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
+                              bp = mv_boxes (pstr + Suc (length gs)) 0 n"
+    thus"?thesis"
+      apply(simp, rule_tac abc_append_exc1, simp_all add: starts h)
+      apply(insert mv_boxes_ex2[of n "pstr + Suc (length gs)" 0 "[]" 
+        "rs # 0\<up>(pstr - n + length gs)" "lm" 
+        "0\<up>(a_md - Suc (pstr + length gs + n)) @ suf_lm"], simp)
+      apply(subgoal_tac "pstr > n \<and> 
+        a_md > pstr + length gs + n \<and> length lm = n" , simp add: exponent_add_iff h)
+      using h pdef
+      apply(simp)     
+      apply(frule_tac a = a and 
+        aa = aa and ba = ba in ci_cn_md_def, simp, simp)
+      apply(subgoal_tac "length lm = rs_pos",
+        simp add: ci_cn_para_eq, erule_tac para_pattern, simp)
+      done
+  qed
+qed
+
+lemma ci_cn_length:
+  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
+  rec_calc_rel (Cn n f gs) lm rs;
+  rec_ci f = (a, aa, ba)\<rbrakk>
+  \<Longrightarrow> length aprog = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
+                             8 * length gs + 6 * n + length a + 6"
+apply(simp add: rec_ci.simps, auto simp: cn_merge_gs_len)
+done
+
+lemma  cn_case: 
+  assumes ind:
+  "\<And>x aprog a_md rs_pos rs suf_lm lm.
+  \<lbrakk>x \<in> set (f # gs);
+  rec_ci x = (aprog, rs_pos, a_md);
+  rec_calc_rel x lm rs\<rbrakk>
+  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
+               (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+  and h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
+         "rec_calc_rel (Cn n f gs) lm rs"
+  shows "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp 
+  = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+apply(insert h, case_tac "rec_ci f",  rule_tac calc_cn_reverse, simp)
+proof -
+  fix a b c ys
+  let ?pstr = "Max (set (Suc n # c # (map (\<lambda>(aprog, p, n). n) 
+                                         (map rec_ci (f # gs)))))"  
+  let ?gs_len = "listsum (map (\<lambda> (ap, pos, n). length ap) 
+                                                (map rec_ci (gs)))"
+  assume g: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
+    "rec_calc_rel (Cn n f gs) lm rs"
+    "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)" 
+    "length ys = length gs" 
+    "rec_calc_rel f ys rs"
+    "n = length lm"
+    "rec_ci f = (a, b, c)"  
+  hence k1:
+    "\<exists> stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
+    (?gs_len + 3 * length gs, lm @ 0\<up>(?pstr - n) @ ys @
+                               0\<up>(a_md - ?pstr - length ys) @ suf_lm)"	
+    apply(rule_tac a = a and aa = b and ba = c in cn_calc_gs)
+    apply(rule_tac ind, auto)
+    done  
+  thm rec_ci.simps
+  from g have k2: 
+    "\<exists> stp. abc_steps_l (?gs_len + 3 * length gs,  lm @ 
+        0\<up>(?pstr - n) @ ys @ 0\<up>(a_md - ?pstr - length ys) @ suf_lm) aprog stp = 
+    (?gs_len + 3 * length gs + 3 * n, 0\<up>?pstr @ ys @ 0 # lm @ 
+                              0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm)"
+    thm save_paras
+    apply(erule_tac ba = c in save_paras, auto intro: ci_cn_para_eq)
+    done
+  from g have k3: 
+    "\<exists> stp. abc_steps_l (?gs_len + 3 * length gs + 3 * n,
+    0\<up>?pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm) aprog stp =
+    (?gs_len + 6 * length gs + 3 * n,  
+           ys @ 0\<up>?pstr @ 0 # lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm)"
+    apply(erule_tac ba = c in reset_new_paras, 
+          auto intro: ci_cn_para_eq)
+    using para_pattern[of f a b c ys rs]
+    apply(simp)
+    done
+  from g have k4: 
+    "\<exists>stp. abc_steps_l  (?gs_len + 6 * length gs + 3 * n,  
+    ys @ 0\<up>?pstr @ 0 # lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm) aprog stp =
+    (?gs_len + 6 * length gs + 3 * n + length a, 
+   ys @ rs # 0\<up>?pstr  @ lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm)"
+    apply(rule_tac ba = c in calc_cn_f, rule_tac ind, auto)
+    done
+thm rec_ci.simps
+  from g h have k5:
+    "\<exists> stp. abc_steps_l (?gs_len + 6 * length gs + 3 * n + length a,
+    ys @ rs # 0\<up>?pstr @ lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm)
+    aprog stp =
+    (?gs_len + 6 * length gs + 3 * n + length a + 3,
+    ys @ 0\<up>(?pstr - length ys) @ rs # 0\<up>length ys @ lm @ 
+    0\<up>(a_md  - Suc (?pstr + length ys + n)) @ suf_lm)"
+    apply(rule_tac save_rs, auto simp: h)
+    done
+  from g have k6: 
+    "\<exists> stp. abc_steps_l (?gs_len + 6 * length gs + 3 * n + 
+    length a + 3, ys @ 0\<up>(?pstr - length ys) @ rs # 0\<up>length ys @ lm @ 
+    0\<up>(a_md  - Suc (?pstr + length ys + n)) @ suf_lm) 
+    aprog stp =
+    (?gs_len + 8 * length gs + 3 *n + length a + 3,
+    0\<up>?pstr @ rs # 0\<up>length ys @ lm @ 
+                        0\<up>(a_md -Suc (?pstr + length ys + n)) @ suf_lm)"
+    apply(drule_tac suf_lm = suf_lm in mv_box_paras, auto)
+    apply(rule_tac x = stp in exI, simp)
+    done
+  from g have k7: 
+    "\<exists> stp. abc_steps_l (?gs_len + 8 * length gs + 3 *n + 
+    length a + 3, 0\<up>?pstr  @ rs # 0\<up>length ys @ lm @ 
+    0\<up>(a_md -Suc (?pstr + length ys + n)) @ suf_lm) aprog stp =
+    (?gs_len + 8 * length gs + 3 * n + length a + 6, 
+    0\<up>n @ rs # 0\<up>(?pstr  - n) @ 0\<up>length ys @ lm @
+                        0\<up>(a_md -Suc (?pstr + length ys + n)) @ suf_lm)"
+    apply(drule_tac suf_lm = suf_lm in restore_rs, auto)
+    apply(rule_tac x = stp in exI, simp)
+    done
+  from g have k8: "\<exists> stp. abc_steps_l (?gs_len + 8 * length gs + 
+    3 * n + length a + 6,
+    0\<up>n @ rs # 0\<up>(?pstr  - n) @ 0\<up>length ys @ lm @
+                      0\<up>(a_md -Suc (?pstr + length ys + n)) @ suf_lm) aprog stp =
+    (?gs_len + 8 * length gs + 6 * n + length a + 6,
+                           lm @ rs # 0\<up>(a_md - Suc n) @ suf_lm)"
+    apply(drule_tac suf_lm = suf_lm in restore_paras, auto)
+    apply(simp add: exponent_add_iff)
+    apply(rule_tac x = stp in exI, simp)
+    done
+  from g have j1: 
+    "length aprog = ?gs_len + 8 * length gs + 6 * n + length a + 6"
+    by(drule_tac a = a and aa = b and ba = c in ci_cn_length,
+      simp, simp, simp)
+  from g have j2: "rs_pos = n"
+    by(simp add: ci_cn_para_eq)
+  from k1 and k2 and k3 and k4 and k5 and k6 and k7 and k8
+    and j1 and j2 show 
+    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
+    (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
+    apply(auto)
+    apply(rule_tac x = "stp + stpa + stpb + stpc +
+      stpd + stpe + stpf + stpg" in exI, simp add: abc_steps_add)
+    done
+qed
+
+text {*
+  Correctness of the complier (terminate case), which says if the execution of 
+  a recursive function @{text "recf"} terminates and gives result, then 
+  the Abacus program compiled from @{text "recf"} termintes and gives the same result.
+  Additionally, to facilitate induction proof, we append @{text "anything"} to the
+  end of Abacus memory.
+*}
+
+lemma recursive_compile_correct:
+  "\<lbrakk>rec_ci recf = (ap, arity, fp);
+    rec_calc_rel recf args r\<rbrakk>
+  \<Longrightarrow> (\<exists> stp. (abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp) = 
+              (length ap, args@[r]@0\<up>(fp - arity - 1) @ anything))"
+apply(induct arbitrary: ap fp arity r anything args
+  rule: rec_ci.induct)
+prefer 5
+proof(case_tac "rec_ci g", case_tac "rec_ci f", simp)
+  fix n f g ap fp arity r anything args  a b c aa ba ca
+  assume f_ind:
+    "\<And>ap fp arity r anything args.
+    \<lbrakk>aa = ap \<and> ba = arity \<and> ca = fp; rec_calc_rel f args r\<rbrakk> \<Longrightarrow> 
+    \<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
+    (length ap, args @ r # 0\<up>(fp - Suc arity) @ anything)"
+    and g_ind:
+    "\<And>x xa y xb ya ap fp arity r anything args.
+    \<lbrakk>x = (aa, ba, ca); xa = aa \<and> y = (ba, ca); xb = ba \<and> ya = ca; 
+    a = ap \<and> b = arity \<and> c = fp; rec_calc_rel g args r\<rbrakk>
+    \<Longrightarrow> \<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
+    (length ap, args @ r # 0\<up>(fp - Suc arity) @ anything)"
+    and h: "rec_ci (Pr n f g) = (ap, arity, fp)" 
+    "rec_calc_rel (Pr n f g) args r" 
+    "rec_ci g = (a, b, c)" 
+    "rec_ci f = (aa, ba, ca)"
+  from h have nf_ind: 
+    "\<And> args r anything. rec_calc_rel f args r \<Longrightarrow> 
+    \<exists>stp. abc_steps_l (0, args @ 0\<up>(ca - ba) @ anything) aa stp = 
+    (length aa, args @ r # 0\<up>(ca - Suc ba) @ anything)"
+    and ng_ind: 
+    "\<And> args r anything. rec_calc_rel g args r \<Longrightarrow> 
+    \<exists>stp. abc_steps_l (0, args @ 0\<up>(c - b) @ anything) a stp = 
+         (length a, args @ r # 0\<up>(c - Suc b)  @ anything)"
+    apply(insert f_ind[of aa ba ca], simp)
+    apply(insert g_ind[of "(aa, ba, ca)" aa "(ba, ca)" ba ca a b c],
+      simp)
+    done
+  from nf_ind and ng_ind and h show 
+    "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp = 
+    (length ap, args @ r # 0\<up>(fp - Suc arity) @ anything)"
+    apply(auto intro: nf_ind ng_ind pr_case)
+    done
+next
+  fix ap fp arity r anything args
+  assume h:
+    "rec_ci z = (ap, arity, fp)" "rec_calc_rel z args r"
+  thus "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
+    (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
+    by (rule_tac z_case)    
+next
+  fix ap fp arity r anything args
+  assume h: 
+    "rec_ci s = (ap, arity, fp)" 
+    "rec_calc_rel s args r"
+  thus 
+    "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
+    (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
+    by(erule_tac s_case, simp)
+next
+  fix m n ap fp arity r anything args
+  assume h: "rec_ci (id m n) = (ap, arity, fp)" 
+    "rec_calc_rel (id m n) args r"
+  thus "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp 
+    = (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
+    by(erule_tac id_case)
+next
+  fix n f gs ap fp arity r anything args
+  assume ind: "\<And>x ap fp arity r anything args.
+    \<lbrakk>x \<in> set (f # gs); 
+    rec_ci x = (ap, arity, fp); 
+    rec_calc_rel x args r\<rbrakk>
+    \<Longrightarrow> \<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
+    (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
+  and h: "rec_ci (Cn n f gs) = (ap, arity, fp)" 
+    "rec_calc_rel (Cn n f gs) args r"
+  from h show
+    "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) 
+       ap stp = (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
+    apply(rule_tac cn_case, rule_tac ind, auto)
+    done
+next
+  fix n f ap fp arity r anything args
+  assume ind:
+    "\<And>ap fp arity r anything args.
+    \<lbrakk>rec_ci f = (ap, arity, fp); rec_calc_rel f args r\<rbrakk> \<Longrightarrow> 
+    \<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
+    (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
+  and h: "rec_ci (Mn n f) = (ap, arity, fp)" 
+    "rec_calc_rel (Mn n f) args r"
+  from h show 
+    "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp = 
+              (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
+    apply(rule_tac mn_case, rule_tac ind, auto)
+    done    
+qed
+
+lemma abc_append_uhalt1:
+  "\<lbrakk>\<forall> stp. (\<lambda> (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp);
+    p = ap [+] bp [+] cp\<rbrakk>
+  \<Longrightarrow> \<forall> stp. (\<lambda> (ss, e). ss < length p) 
+                     (abc_steps_l (length ap, lm) p stp)"
+apply(auto)
+apply(erule_tac x = stp in allE, auto)
+apply(frule_tac ap = ap and cp = cp in abc_append_state_in_exc, auto)
+done
+
+
+lemma abc_append_unhalt2:
+  "\<lbrakk>abc_steps_l (0, am) ap stp = (length ap, lm); bp \<noteq> [];
+  \<forall> stp. (\<lambda> (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp);
+  p = ap [+] bp [+] cp\<rbrakk>
+  \<Longrightarrow> \<forall> stp. (\<lambda> (ss, e). ss < length p) (abc_steps_l (0, am) p stp)"
+proof -
+  assume h: 
+    "abc_steps_l (0, am) ap stp = (length ap, lm)" 
+    "bp \<noteq> []"
+    "\<forall> stp. (\<lambda> (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp)"
+    "p = ap [+] bp [+] cp"
+  have "\<exists> stp. (abc_steps_l (0, am) p stp) = (length ap, lm)"
+    using h
+    thm abc_add_exc1
+    apply(simp add: abc_append.simps)
+    apply(rule_tac abc_add_exc1, auto)
+    done
+  from this obtain stpa where g1: 
+    "(abc_steps_l (0, am) p stpa) = (length ap, lm)" ..
+  moreover have g2: "\<forall> stp. (\<lambda> (ss, e). ss < length p) 
+                          (abc_steps_l (length ap, lm) p stp)"
+    using h
+    apply(erule_tac abc_append_uhalt1, simp)
+    done
+  moreover from g1 and g2 have
+    "\<forall> stp. (\<lambda> (ss, e). ss < length p) 
+                    (abc_steps_l (0, am) p (stpa + stp))"
+    apply(simp add: abc_steps_add)
+    done
+  thus "\<forall> stp. (\<lambda> (ss, e). ss < length p) 
+                           (abc_steps_l (0, am) p stp)"
+    apply(rule_tac allI, auto)
+    apply(case_tac "stp \<ge>  stpa")
+    apply(erule_tac x = "stp - stpa" in allE, simp)
+  proof - 	
+    fix stp a b
+    assume g3:  "abc_steps_l (0, am) p stp = (a, b)" 
+                "\<not> stpa \<le> stp"
+    thus "a < length p"
+      using g1 h
+      apply(case_tac "a < length p", simp, simp)
+      apply(subgoal_tac "\<exists> d. stpa = stp + d")
+      using  abc_state_keep[of p a b "stpa - stp"]
+      apply(erule_tac exE, simp add: abc_steps_add)
+      apply(rule_tac x = "stpa - stp" in exI, simp)
+      done
+  qed
+qed
+
+text {*
+  Correctness of the complier (non-terminating case for Mn). There are many cases when a 
+  recursive function does not terminate. For the purpose of Uiversal Turing Machine, we only 
+  need to prove the case for @{text "Mn"} and @{text "Cn"}.
+  This lemma is for @{text "Mn"}. For @{text "Mn n f"}, this lemma describes what 
+  happens when @{text "f"} always terminates but always does not return zero, so that
+  @{text "Mn"} has to loop forever.
+  *}
+
+lemma Mn_unhalt:
+  assumes mn_rf: "rf = Mn n f"
+  and compiled_mnrf: "rec_ci rf = (aprog, rs_pos, a_md)"
+  and compiled_f: "rec_ci f = (aprog', rs_pos', a_md')"
+  and args: "length lm = n"
+  and unhalt_condition: "\<forall> y. (\<exists> rs. rec_calc_rel f (lm @ [y]) rs \<and> rs \<noteq> 0)"
+  shows "\<forall> stp. case abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm)
+               aprog stp of (ss, e) \<Rightarrow> ss < length aprog"
+  using mn_rf compiled_mnrf compiled_f args unhalt_condition
+proof(rule_tac allI)
+  fix stp
+  assume h: "rf = Mn n f" 
+            "rec_ci rf = (aprog, rs_pos, a_md)"
+            "rec_ci f = (aprog', rs_pos', a_md')" 
+            "\<forall>y. \<exists>rs. rec_calc_rel f (lm @ [y]) rs \<and> rs \<noteq> 0" "length lm = n"
+  thm mn_ind_step
+  have "\<exists>stpa \<ge> stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm) aprog stpa 
+         = (0, lm @ stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+  proof(induct stp, auto)
+    show "\<exists>stpa. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
+          aprog stpa = (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+      apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
+      done
+  next
+    fix stp stpa
+    assume g1: "stp \<le> stpa"
+      and g2: "abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
+                            aprog stpa
+               = (0, lm @ stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+    have "\<exists>rs. rec_calc_rel f (lm @ [stp]) rs \<and> rs \<noteq> 0"
+      using h
+      apply(erule_tac x = stp in allE, simp)
+      done
+    from this obtain rs where g3:
+      "rec_calc_rel f (lm @ [stp]) rs \<and> rs \<noteq> 0" ..
+    hence "\<exists> stpb. abc_steps_l (0, lm @ stp # 0\<up>(a_md - Suc rs_pos) @
+                     suf_lm) aprog stpb 
+      = (0, lm @ Suc stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+      using h
+      apply(rule_tac mn_ind_step)
+      apply(rule_tac recursive_compile_correct, simp, simp)
+    proof -
+      show "rec_ci f = ((aprog', rs_pos', a_md'))" using h by simp
+    next
+      show "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" using h by simp
+    next
+      show "rec_calc_rel f (lm @ [stp]) rs" using g3 by simp
+    next
+      show "0 < rs" using g3 by simp
+    next
+      show "Suc rs_pos < a_md"
+        using g3 h
+        apply(auto)
+        apply(frule_tac f = f in para_pattern, simp, simp)
+        apply(simp add: rec_ci.simps, auto)
+        apply(subgoal_tac "Suc (length lm) < a_md'")
+        apply(arith)
+        apply(simp add: ci_ad_ge_paras)
+        done
+    next
+      show "rs_pos' = Suc rs_pos"
+        using g3 h
+        apply(auto)
+        apply(frule_tac f = f in para_pattern, simp, simp)
+        apply(simp add: rec_ci.simps)
+        done
+    qed
+    thus "\<exists>stpa\<ge>Suc stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @
+                 suf_lm) aprog stpa 
+      = (0, lm @ Suc stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
+      using g2
+      apply(erule_tac exE)
+      apply(case_tac "stpb = 0", simp add: abc_steps_l.simps)
+      apply(rule_tac x = "stpa + stpb" in exI, simp add:
+        abc_steps_add)
+      using g1
+      apply(arith)
+      done
+  qed
+  from this obtain stpa where 
+    "stp \<le> stpa \<and> abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
+         aprog stpa = (0, lm @ stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)" ..
+  thus "case abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp
+    of (ss, e) \<Rightarrow> ss < length aprog"
+    apply(case_tac "abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog
+      stp", simp, case_tac "a \<ge> length aprog", 
+        simp, simp)
+    apply(subgoal_tac "\<exists> d. stpa = stp + d", erule_tac exE)
+    apply(subgoal_tac "lm @ 0\<up>(a_md - rs_pos) @ suf_lm = lm @ 0 # 
+             0\<up>(a_md - Suc rs_pos) @ suf_lm", simp add: abc_steps_add)
+    apply(frule_tac as = a and lm = b and stp = d in abc_state_keep, 
+          simp)
+    using h  
+    apply(simp add: rec_ci.simps, simp, 
+              simp only: replicate_Suc[THEN sym])
+    apply(case_tac rs_pos, simp, simp)
+    apply(rule_tac x = "stpa - stp" in exI, simp, simp)
+    done
+qed   
+
+lemma abc_append_cons_eq[intro!]: 
+  "\<lbrakk>ap = bp; cp = dp\<rbrakk> \<Longrightarrow> ap [+] cp = bp [+] dp"
+by simp 
+
+lemma cn_merge_gs_split: 
+  "\<lbrakk>i < length gs; rec_ci (gs!i) = (ga, gb, gc)\<rbrakk> \<Longrightarrow> 
+     cn_merge_gs (map rec_ci gs) p = 
+        cn_merge_gs (map rec_ci (take i gs)) p [+] ga [+] 
+       mv_box gb (p + i) [+] 
+      cn_merge_gs (map rec_ci (drop (Suc i) gs)) (p + Suc i)"
+apply(induct i arbitrary: gs p, case_tac gs, simp, simp)
+apply(case_tac gs, simp, case_tac "rec_ci a", 
+       simp add: abc_append_commute[THEN sym])
+done
+
+text {*
+  Correctness of the complier (non-terminating case for Mn). There are many cases when a 
+  recursive function does not terminate. For the purpose of Uiversal Turing Machine, we only 
+  need to prove the case for @{text "Mn"} and @{text "Cn"}.
+  This lemma is for @{text "Cn"}. For @{text "Cn f g1 g2 \<dots>gi, gi+1, \<dots> gn"}, this lemma describes what 
+  happens when every one of @{text "g1, g2, \<dots> gi"} terminates, but 
+  @{text "gi+1"} does not terminate, so that whole function @{text "Cn f g1 g2 \<dots>gi, gi+1, \<dots> gn"}
+  does not terminate.
+  *}
+
+lemma cn_gi_uhalt: 
+  assumes cn_recf: "rf = Cn n f gs"
+  and compiled_cn_recf: "rec_ci rf = (aprog, rs_pos, a_md)"
+  and args_length: "length lm = n"
+  and exist_unhalt_recf: "i < length gs" "gi = gs ! i"
+  and complied_unhalt_recf: "rec_ci gi = (ga, gb, gc)"  "gb = n"
+  and all_halt_before_gi: "\<forall> j < i. (\<exists> rs. rec_calc_rel (gs!j) lm rs)" 
+  and unhalt_condition: "\<And> slm. \<forall> stp. case abc_steps_l (0, lm @ 0\<up>(gc - gb) @ slm) 
+     ga stp of (se, e) \<Rightarrow> se < length ga"
+  shows " \<forall> stp. case abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suflm) aprog
+  stp of (ss, e) \<Rightarrow> ss < length aprog"
+  using cn_recf compiled_cn_recf args_length exist_unhalt_recf complied_unhalt_recf
+        all_halt_before_gi unhalt_condition
+proof(case_tac "rec_ci f", simp)
+  fix a b c
+  assume h1: "rf = Cn n f gs" 
+    "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
+    "length lm = n" 
+    "gi = gs ! i" 
+    "rec_ci (gs!i) = (ga, n, gc)" 
+    "gb = n" "rec_ci f = (a, b, c)"
+    and h2: "\<forall>j<i. \<exists>rs. rec_calc_rel (gs ! j) lm rs"
+    "i < length gs"
+  and ind:
+    "\<And> slm. \<forall> stp. case abc_steps_l (0, lm @ 0\<up>(gc - n) @ slm) ga stp of (se, e) \<Rightarrow> se < length ga"
+  have h3: "rs_pos = n"
+    using h1
+    by(rule_tac ci_cn_para_eq, simp)
+  let ?ggs = "take i gs"
+  have "\<exists> ys. (length ys = i \<and> 
+    (\<forall> k < i. rec_calc_rel (?ggs ! k) lm (ys ! k)))"
+    using h2
+    apply(induct i, simp, simp)
+    apply(erule_tac exE)
+    apply(erule_tac x = ia in allE, simp)
+    apply(erule_tac exE)
+    apply(rule_tac x = "ys @ [x]" in exI, simp add: nth_append, auto)
+    apply(subgoal_tac "k = length ys", simp, simp)
+    done
+  from this obtain ys where g1:
+    "(length ys = i \<and> (\<forall> k < i. rec_calc_rel (?ggs ! k)
+                        lm (ys ! k)))" ..
+  let ?pstr = "Max (set (Suc n # c # map (\<lambda>(aprog, p, n). n)
+    (map rec_ci (f # gs))))"
+  have "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suflm) 
+    (cn_merge_gs (map rec_ci ?ggs) ?pstr) stp =
+    (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) ?ggs) +
+    3 * length ?ggs, lm @ 0\<up>(?pstr - n) @ ys @ 0\<up>(a_md -(?pstr + length ?ggs)) @
+    suflm) "
+    apply(rule_tac  cn_merge_gs_ex)
+    apply(rule_tac  recursive_compile_correct, simp, simp)
+    using h1
+    apply(simp add: rec_ci.simps, auto)
+    using g1
+    apply(simp)
+    using h2 g1
+    apply(simp)
+    apply(rule_tac min_max.le_supI2)
+    apply(rule_tac Max_ge, simp, simp, rule_tac disjI2)
+    apply(subgoal_tac "aa \<in> set gs", simp)
+    using h2
+    apply(rule_tac A = "set (take i gs)" in subsetD, 
+      simp add: set_take_subset, simp)
+    done
+  thm cn_merge_gs.simps
+  from this obtain stpa where g2: 
+    "abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suflm) 
+    (cn_merge_gs (map rec_ci ?ggs) ?pstr) stpa =
+    (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) ?ggs) +
+    3 * length ?ggs, lm @ 0\<up>(?pstr - n) @ ys @ 0\<up>(a_md -(?pstr + length ?ggs)) @
+    suflm)" ..
+  moreover have 
+    "\<exists> cp. aprog = (cn_merge_gs
+    (map rec_ci ?ggs) ?pstr) [+] ga [+] cp"
+    using h1
+    apply(simp add: rec_ci.simps)
+    apply(rule_tac x = "mv_box n (?pstr + i) [+] 
+      (cn_merge_gs (map rec_ci (drop (Suc i) gs)) (?pstr + Suc i))
+      [+]mv_boxes 0 (Suc (max (Suc n) (Max (insert c 
+     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) +
+      length gs)) n [+] mv_boxes (max (Suc n) (Max (insert c 
+      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
+      a [+] Recursive.mv_box b (max (Suc n) 
+      (Max (insert c (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
+     empty_boxes (length gs) [+] Recursive.mv_box (max (Suc n) 
+      (Max (insert c (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
+      mv_boxes (Suc (max (Suc n) (Max (insert c 
+    (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI)
+    apply(simp add: abc_append_commute [THEN sym])
+    apply(auto)
+    using cn_merge_gs_split[of i gs ga "length lm" gc 
+      "(max (Suc (length lm))
+       (Max (insert c (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))"] 
+      h2
+    apply(simp)
+    done
+  from this obtain cp where g3: 
+    "aprog = (cn_merge_gs (map rec_ci ?ggs) ?pstr) [+] ga [+] cp" ..
+  show "\<forall> stp. case abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suflm) 
+    aprog stp of (ss, e) \<Rightarrow> ss < length aprog"
+  proof(rule_tac abc_append_unhalt2)
+    show "abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suflm) (
+      cn_merge_gs (map rec_ci ?ggs) ?pstr) stpa =
+         (length ((cn_merge_gs (map rec_ci ?ggs) ?pstr)),  
+          lm @ 0\<up>(?pstr - n) @ ys @ 0\<up>(a_md -(?pstr + length ?ggs)) @ suflm)"
+      using h3 g2
+      apply(simp add: cn_merge_gs_length)
+      done
+  next
+    show "ga \<noteq> []"
+      using h1
+      apply(simp add: rec_ci_not_null)
+      done
+  next
+    show "\<forall>stp. case abc_steps_l (0, lm @ 0\<up>(?pstr - n) @ ys
+      @ 0\<up>(a_md - (?pstr + length (take i gs))) @ suflm) ga  stp of
+          (ss, e) \<Rightarrow> ss < length ga"
+      using ind[of "0\<up>(?pstr - gc) @ ys @ 0\<up>(a_md - (?pstr + length (take i gs)))
+        @ suflm"]
+      apply(subgoal_tac "lm @ 0\<up>(?pstr - n) @ ys
+        @ 0\<up>(a_md - (?pstr + length (take i gs))) @ suflm
+                       = lm @ 0\<up>(gc - n) @ 
+        0\<up>(?pstr - gc) @ ys @ 0\<up>(a_md - (?pstr + length (take i gs))) @ suflm", simp)
+      apply(simp add: replicate_add[THEN sym])
+      apply(subgoal_tac "gc > n \<and> ?pstr \<ge> gc")
+      apply(erule_tac conjE)
+      apply(simp add: h1)
+      using h1
+      apply(auto)
+      apply(rule_tac min_max.le_supI2)
+      apply(rule_tac Max_ge, simp, simp)
+      apply(rule_tac disjI2)
+      using h2
+      thm rev_image_eqI
+      apply(rule_tac x = "gs!i" in rev_image_eqI, simp, simp)
+      done
+  next
+    show "aprog = cn_merge_gs (map rec_ci (take i gs)) 
+              ?pstr [+] ga [+] cp"
+      using g3 by simp
+  qed
+qed
+
+lemma recursive_compile_correct_spec: 
+  "\<lbrakk>rec_ci re = (ap, ary, fp); 
+    rec_calc_rel re args r\<rbrakk>
+  \<Longrightarrow> (\<exists> stp. (abc_steps_l (0, args @ 0\<up>(fp - ary)) ap stp) = 
+                     (length ap, args@[r]@0\<up>(fp - ary - 1)))"
+using recursive_compile_correct[of re ap ary fp args r "[]"]
+by simp
+
+definition dummy_abc :: "nat \<Rightarrow> abc_inst list"
+where
+"dummy_abc k = [Inc k, Dec k 0, Goto 3]"
+
+definition abc_list_crsp:: "nat list \<Rightarrow> nat list \<Rightarrow> bool"
+  where
+  "abc_list_crsp xs ys = (\<exists> n. xs = ys @ 0\<up>n \<or> ys = xs @ 0\<up>n)"
+
+lemma [intro]: "abc_list_crsp (lm @ 0\<up>m) lm"
+apply(auto simp: abc_list_crsp_def)
+done
+
+lemma abc_list_crsp_lm_v: 
+  "abc_list_crsp lma lmb \<Longrightarrow> abc_lm_v lma n = abc_lm_v lmb n"
+apply(auto simp: abc_list_crsp_def abc_lm_v.simps 
+                 nth_append)
+done
+
+lemma  rep_app_cons_iff: 
+  "k < n \<Longrightarrow> replicate n a[k:=b] = 
+          replicate k a @ b # replicate (n - k - 1) a"
+apply(induct n arbitrary: k, simp)
+apply(simp split:nat.splits)
+done
+
+lemma abc_list_crsp_lm_s: 
+  "abc_list_crsp lma lmb \<Longrightarrow> 
+      abc_list_crsp (abc_lm_s lma m n) (abc_lm_s lmb m n)"
+apply(auto simp: abc_list_crsp_def abc_lm_v.simps abc_lm_s.simps)
+apply(simp_all add: list_update_append, auto)
+proof -
+  fix na
+  assume h: "m < length lmb + na" " \<not> m < length lmb"
+  hence "m - length lmb < na" by simp
+  hence "replicate na 0[(m- length lmb):= n] = 
+           replicate (m - length lmb) 0 @ n # 
+              replicate (na - (m - length lmb) - 1) 0"
+    apply(erule_tac rep_app_cons_iff)
+    done
+  thus "\<exists>nb. replicate na 0[m - length lmb := n] =
+                 replicate (m - length lmb) 0 @ n # replicate nb 0 \<or>
+                 replicate (m - length lmb) 0 @ [n] =
+                 replicate na 0[m - length lmb := n] @ replicate nb 0"
+    apply(auto)
+    done
+next
+  fix na
+  assume h: "\<not> m < length lmb + na"
+  show 
+    "\<exists>nb. replicate na 0 @ replicate (m - (length lmb + na)) 0 @ [n] =
+           replicate (m - length lmb) 0 @ n # replicate nb 0 \<or>
+          replicate (m - length lmb) 0 @ [n] =
+            replicate na 0 @
+            replicate (m - (length lmb + na)) 0 @ n # replicate nb 0"
+    apply(rule_tac x = 0 in exI, simp, auto)
+    using h
+    apply(simp add: replicate_add[THEN sym])
+    done
+next
+  fix na
+  assume h: "\<not> m < length lma" "m < length lma + na"
+  hence "m - length lma < na" by simp
+  hence 
+    "replicate na 0[(m- length lma):= n] = replicate (m - length lma) 
+                  0 @ n # replicate (na - (m - length lma) - 1) 0"
+    apply(erule_tac rep_app_cons_iff)
+    done
+  thus "\<exists>nb. replicate (m - length lma) 0 @ [n] =
+                 replicate na 0[m - length lma := n] @ replicate nb 0 
+           \<or> replicate na 0[m - length lma := n] =
+                 replicate (m - length lma) 0 @ n # replicate nb 0"
+    apply(auto)
+    done
+next
+  fix na
+  assume "\<not> m < length lma + na"
+  thus " \<exists>nb. replicate (m - length lma) 0 @ [n] =
+            replicate na 0 @
+            replicate (m - (length lma + na)) 0 @ n # replicate nb 0 
+        \<or>   replicate na 0 @ 
+               replicate (m - (length lma + na)) 0 @ [n] =
+            replicate (m - length lma) 0 @ n # replicate nb 0"
+    apply(rule_tac x = 0 in exI, simp, auto)
+    apply(simp add: replicate_add[THEN sym])
+    done
+qed
+
+lemma abc_list_crsp_step: 
+  "\<lbrakk>abc_list_crsp lma lmb; abc_step_l (aa, lma) i = (a, lma'); 
+    abc_step_l (aa, lmb) i = (a', lmb')\<rbrakk>
+    \<Longrightarrow> a' = a \<and> abc_list_crsp lma' lmb'"
+apply(case_tac i, auto simp: abc_step_l.simps 
+       abc_list_crsp_lm_s abc_list_crsp_lm_v Let_def 
+                       split: abc_inst.splits if_splits)
+done
+
+lemma abc_list_crsp_steps: 
+  "\<lbrakk>abc_steps_l (0, lm @ 0\<up>m) aprog stp = (a, lm'); aprog \<noteq> []\<rbrakk> 
+      \<Longrightarrow> \<exists> lma. abc_steps_l (0, lm) aprog stp = (a, lma) \<and> 
+                                          abc_list_crsp lm' lma"
+apply(induct stp arbitrary: a lm', simp add: abc_steps_l.simps, auto)
+apply(case_tac "abc_steps_l (0, lm @ 0\<up>m) aprog stp", 
+      simp add: abc_step_red)
+proof -
+  fix stp a lm' aa b
+  assume ind:
+    "\<And>a lm'. aa = a \<and> b = lm' \<Longrightarrow> 
+     \<exists>lma. abc_steps_l (0, lm) aprog stp = (a, lma) \<and>
+                                          abc_list_crsp lm' lma"
+    and h: "abc_steps_l (0, lm @ 0\<up>m) aprog (Suc stp) = (a, lm')" 
+           "abc_steps_l (0, lm @ 0\<up>m) aprog stp = (aa, b)" 
+           "aprog \<noteq> []"
+  hence g1: "abc_steps_l (0, lm @ 0\<up>m) aprog (Suc stp)
+          = abc_step_l (aa, b) (abc_fetch aa aprog)"
+    apply(rule_tac abc_step_red, simp)
+    done
+  have "\<exists>lma. abc_steps_l (0, lm) aprog stp = (aa, lma) \<and> 
+              abc_list_crsp b lma"
+    apply(rule_tac ind, simp)
+    done
+  from this obtain lma where g2: 
+    "abc_steps_l (0, lm) aprog stp = (aa, lma) \<and> 
+     abc_list_crsp b lma"   ..
+  hence g3: "abc_steps_l (0, lm) aprog (Suc stp)
+          = abc_step_l (aa, lma) (abc_fetch aa aprog)"
+    apply(rule_tac abc_step_red, simp)
+    done
+  show "\<exists>lma. abc_steps_l (0, lm) aprog (Suc stp) = (a, lma) \<and> abc_list_crsp lm' lma"
+    using g1 g2 g3 h
+    apply(auto)
+    apply(case_tac "abc_step_l (aa, b) (abc_fetch aa aprog)",
+          case_tac "abc_step_l (aa, lma) (abc_fetch aa aprog)", simp)
+    apply(rule_tac abc_list_crsp_step, auto)
+    done
+qed
+
+lemma recursive_compile_correct_norm: 
+  "\<lbrakk>rec_ci re = (aprog, rs_pos, a_md);  
+   rec_calc_rel re lm rs\<rbrakk>
+  \<Longrightarrow> (\<exists> stp lm' m. (abc_steps_l (0, lm) aprog stp) = 
+  (length aprog, lm') \<and> abc_list_crsp lm' (lm @ rs # 0\<up>m))"
+apply(frule_tac recursive_compile_correct_spec, auto)
+apply(drule_tac abc_list_crsp_steps)
+apply(rule_tac rec_ci_not_null, simp)
+apply(erule_tac exE, rule_tac x = stp in exI, 
+  auto simp: abc_list_crsp_def)
+done
+
+lemma [simp]: "length (dummy_abc (length lm)) = 3"
+apply(simp add: dummy_abc_def)
+done
+
+lemma [simp]: "dummy_abc (length lm) \<noteq> []"
+apply(simp add: dummy_abc_def)
+done
+
+lemma dummy_abc_steps_ex: 
+  "\<exists>bstp. abc_steps_l (0, lm') (dummy_abc (length lm)) bstp = 
+  ((Suc (Suc (Suc 0))), abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)))"
+apply(rule_tac x = "Suc (Suc (Suc 0))" in exI)
+apply(auto simp: abc_steps_l.simps abc_step_l.simps 
+  dummy_abc_def abc_fetch.simps)
+apply(auto simp: abc_lm_s.simps abc_lm_v.simps nth_append)
+apply(simp add: butlast_append)
+done
+
+lemma [simp]: 
+  "\<lbrakk>Suc (length lm) - length lm' \<le> n; \<not> length lm < length lm'; lm @ rs # 0 \<up> m = lm' @ 0 \<up> n\<rbrakk> 
+  \<Longrightarrow> lm' @ 0 \<up> Suc (length lm - length lm') = lm @ [rs]"
+apply(subgoal_tac "n > m")
+apply(subgoal_tac "\<exists> d. n = d + m", erule_tac exE)
+apply(simp add: replicate_add)
+apply(drule_tac length_equal, simp)
+apply(simp add: replicate_Suc[THEN sym] del: replicate_Suc)
+apply(rule_tac x = "n - m" in exI, simp)
+apply(drule_tac length_equal, simp)
+done
+
+lemma [elim]: 
+  "lm @ rs # 0\<up>m = lm' @ 0\<up>n \<Longrightarrow> 
+  \<exists>m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) = 
+                            lm @ rs # 0\<up>m"
+proof(cases "length lm' > length lm")
+  case True 
+  assume h: "lm @ rs # 0\<up>m = lm' @ 0\<up>n" "length lm < length lm'"
+  hence "m \<ge> n"
+    apply(drule_tac length_equal)
+    apply(simp)
+    done
+  hence "\<exists> d. m = d + n"
+    apply(rule_tac x = "m - n" in exI, simp)
+    done
+  from this obtain d where "m = d + n" ..
+  from h and this show "?thesis"
+    apply(auto simp: abc_lm_s.simps abc_lm_v.simps 
+                     replicate_add)
+    done
+next
+  case False
+  assume h:"lm @ rs # 0\<up>m = lm' @ 0\<up>n" 
+    and    g: "\<not> length lm < length lm'"
+  have "take (Suc (length lm)) (lm @ rs # 0\<up>m) = 
+                        take (Suc (length lm)) (lm' @ 0\<up>n)"
+    using h by simp
+  moreover have "n \<ge> (Suc (length lm) - length lm')"
+    using h g
+    apply(drule_tac length_equal)
+    apply(simp)
+    done
+  ultimately show 
+    "\<exists>m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) =
+                                                       lm @ rs # 0\<up>m"
+    using g h
+    apply(simp add: abc_lm_s.simps abc_lm_v.simps min_def)
+    apply(rule_tac x = 0 in exI, 
+      simp add:replicate_append_same replicate_Suc[THEN sym]
+                                      del:replicate_Suc)
+    done    
+qed
+
+lemma [elim]: 
+  "abc_list_crsp lm' (lm @ rs # 0\<up>m)
+  \<Longrightarrow> \<exists>m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) 
+             = lm @ rs # 0\<up>m"
+apply(auto simp: abc_list_crsp_def)
+apply(simp add: abc_lm_v.simps abc_lm_s.simps)
+apply(rule_tac x =  "m + n" in exI, 
+      simp add: replicate_add)
+done
+
+lemma abc_append_dummy_complie:
+  "\<lbrakk>rec_ci recf = (ap, ary, fp);  
+    rec_calc_rel recf args r; 
+    length args = k\<rbrakk>
+  \<Longrightarrow> (\<exists> stp m. (abc_steps_l (0, args) (ap [+] dummy_abc k) stp) = 
+                  (length ap + 3, args @ r # 0\<up>m))"
+apply(drule_tac recursive_compile_correct_norm, auto simp: numeral_3_eq_3)
+proof -
+  fix stp lm' m
+  assume h: "rec_calc_rel recf args r"  
+    "abc_steps_l (0, args) ap stp = (length ap, lm')" 
+    "abc_list_crsp lm' (args @ r # 0\<up>m)"
+  thm abc_append_exc2
+  thm abc_lm_s.simps
+  have "\<exists>stp. abc_steps_l (0, args) (ap [+] 
+    (dummy_abc (length args))) stp = (length ap + 3, 
+    abc_lm_s lm' (length args) (abc_lm_v lm' (length args)))"
+    using h
+    apply(rule_tac bm = lm' in abc_append_exc2,
+          auto intro: dummy_abc_steps_ex simp: numeral_3_eq_3)
+    done
+  thus "\<exists>stp m. abc_steps_l (0, args) (ap [+] 
+    dummy_abc (length args)) stp = (Suc (Suc (Suc (length ap))), args @ r # 0\<up>m)"
+    using h
+    apply(erule_tac exE)
+    apply(rule_tac x = stpa in exI, auto)
+    done
+qed
+
+lemma [simp]: "length (dummy_abc k) = 3"
+apply(simp add: dummy_abc_def)
+done
+
+lemma [simp]: "length args = k \<Longrightarrow> abc_lm_v (args @ r # 0\<up>m) k = r "
+apply(simp add: abc_lm_v.simps nth_append)
+done
+
+lemma [simp]: "crsp (layout_of (ap [+] dummy_abc k)) (0, args)
+  (Suc 0, Bk # Bk # ires, <args> @ Bk \<up> rn) ires"
+apply(auto simp: crsp.simps start_of.simps)
+done
+
+(* cccc *)
+
+fun tm_of_rec :: "recf \<Rightarrow> instr list"
+where "tm_of_rec recf = (let (ap, k, fp) = rec_ci recf in
+                         let tp = tm_of (ap [+] dummy_abc k) in 
+                             tp @ (shift (mopup k) (length tp div 2)))"
+
+lemma recursive_compile_to_tm_correct: 
+  "\<lbrakk>rec_ci recf = (ap, ary, fp); 
+    rec_calc_rel recf args r;
+    length args = k;
+    ly = layout_of (ap [+] dummy_abc k);
+    tp = tm_of (ap [+] dummy_abc k)\<rbrakk>
+  \<Longrightarrow> \<exists> stp m l. steps0 (Suc 0, Bk # Bk # ires, <args> @ Bk\<up>rn)
+  (tp @ shift (mopup k) (length tp div 2)) stp
+  = (0, Bk\<up>m @ Bk # Bk # ires, Oc\<up>Suc r @ Bk\<up>l)"
+  using abc_append_dummy_complie[of recf ap ary fp args r k]
+apply(simp)
+apply(erule_tac exE)+
+apply(frule_tac tp = tp and n = k 
+               and ires = ires in compile_correct_halt, simp_all add: length_append)
+apply(simp_all add: length_append)
+done
+
+lemma recursive_compile_to_tm_correct2: 
+  assumes "rec_ci recf = (ap, ary, fp)" 
+  and     "rec_calc_rel recf args r"
+  and     "length args = k"
+  and     "tp = tm_of (ap [+] dummy_abc k)"
+  shows "\<exists> m n. {\<lambda>tp. tp = ([Bk, Bk], <args>)}
+             (tp @ (shift (mopup k) (length tp div 2)))
+             {\<lambda>tp. tp = (Bk \<up> m, Oc \<up> (Suc r) @ Bk \<up> n)}"
+using recursive_compile_to_tm_correct[where ires="[]" and rn="0", OF assms(1-3) _ assms(4)]
+apply(simp add: Hoare_halt_def)
+apply(drule_tac x="layout_of (ap [+] dummy_abc k)" in meta_spec)
+apply(auto)
+apply(rule_tac x="m + 2" in exI)
+apply(rule_tac x="l" in exI)
+apply(rule_tac x="stp" in exI)
+apply(auto)
+by (metis append_Nil2 replicate_app_Cons_same)
+
+lemma recursive_compile_to_tm_correct3: 
+  assumes "rec_calc_rel recf args r"
+  shows "{\<lambda>tp. tp = ([Bk, Bk], <args>)} tm_of_rec recf {\<lambda>tp. \<exists>k l. tp = (Bk \<up> k, <r> @ Bk \<up> l)}"
+using recursive_compile_to_tm_correct2[OF _ assms] 
+apply(auto)
+apply(case_tac "rec_ci recf")
+apply(auto)
+apply(drule_tac x="a" in meta_spec)
+apply(drule_tac x="b" in meta_spec)
+apply(drule_tac x="c" in meta_spec)
+apply(drule_tac x="length args" in meta_spec)
+apply(drule_tac x="tm_of (a [+] dummy_abc (length args))" in meta_spec)
+apply(auto)
+apply(simp add: tape_of_nat_abv)
+apply(subgoal_tac "b = length args")
+apply(simp add: Hoare_halt_def)
+apply(auto)[1]
+apply(rule_tac x="na" in exI)
+apply(auto)[1]
+apply(case_tac "steps0 (Suc 0, [Bk, Bk], <args>)
+                                   (tm_of (a [+] dummy_abc (length args)) @
+                                    shift (mopup (length args))
+                                     (listsum
+ (layout_of (a [+] dummy_abc (length args)))))
+                                   na")
+apply(simp)
+by (metis assms para_pattern)
+
+
+lemma [simp]:
+  "list_all (\<lambda>(acn, s). s \<le> Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))) xs \<Longrightarrow>
+  list_all (\<lambda>(acn, s). s \<le> Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))))) xs"
+apply(induct xs, simp, simp)
+apply(case_tac a, simp)
+done
+
+lemma shift_append: "shift (xs @ ys) n = shift xs n @ shift ys n"
+apply(simp add: shift.simps)
+done
+
+lemma [simp]: "length (shift (mopup n) ss) = 4 * n + 12"
+apply(auto simp: mopup.simps shift_append mopup_b_def)
+done
+
+lemma length_tm_even[intro]: "length (tm_of ap) mod 2 = 0"
+apply(simp add: tm_of.simps)
+done
+
+lemma [simp]: "k < length ap \<Longrightarrow> tms_of ap ! k  = 
+ ci (layout_of ap) (start_of (layout_of ap) k) (ap ! k)"
+apply(simp add: tms_of.simps tpairs_of.simps)
+done
+
+lemma start_of_suc_inc:
+  "\<lbrakk>k < length ap; ap ! k = Inc n\<rbrakk> \<Longrightarrow> start_of (layout_of ap) (Suc k) =
+                        start_of (layout_of ap) k + 2 * n + 9"
+apply(rule_tac start_of_Suc1, auto simp: abc_fetch.simps)
+done
+
+lemma start_of_suc_dec:
+  "\<lbrakk>k < length ap; ap ! k = (Dec n e)\<rbrakk> \<Longrightarrow> start_of (layout_of ap) (Suc k) =
+                        start_of (layout_of ap) k + 2 * n + 16"
+apply(rule_tac start_of_Suc2, auto simp: abc_fetch.simps)
+done
+
+lemma inc_state_all_le:
+  "\<lbrakk>k < length ap; ap ! k = Inc n; 
+       (a, b) \<in> set (shift (shift tinc_b (2 * n)) 
+                            (start_of (layout_of ap) k - Suc 0))\<rbrakk>
+       \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
+apply(subgoal_tac "b \<le> start_of (layout_of ap) (Suc k)")
+apply(subgoal_tac "start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap) ")
+apply(arith)
+apply(case_tac "Suc k = length ap", simp)
+apply(rule_tac start_of_less, simp)
+apply(auto simp: tinc_b_def shift.simps start_of_suc_inc length_of.simps startof_not0)
+done
+
+lemma findnth_le[elim]: 
+  "(a, b) \<in> set (shift (findnth n) (start_of (layout_of ap) k - Suc 0))
+  \<Longrightarrow> b \<le> Suc (start_of (layout_of ap) k + 2 * n)"
+apply(induct n, simp add: findnth.simps shift.simps)
+apply(simp add: findnth.simps shift_append, auto)
+apply(auto simp: shift.simps)
+done
+
+lemma findnth_state_all_le1:
+  "\<lbrakk>k < length ap; ap ! k = Inc n;
+  (a, b) \<in> 
+  set (shift (findnth n) (start_of (layout_of ap) k - Suc 0))\<rbrakk> 
+  \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
+apply(subgoal_tac "b \<le> start_of (layout_of ap) (Suc k)")
+apply(subgoal_tac "start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap) ")
+apply(arith)
+apply(case_tac "Suc k = length ap", simp)
+apply(rule_tac start_of_less, simp)
+apply(subgoal_tac "b \<le> start_of (layout_of ap) k + 2*n + 1 \<and> 
+     start_of (layout_of ap) k + 2*n + 1 \<le>  start_of (layout_of ap) (Suc k)", auto)
+apply(auto simp: tinc_b_def shift.simps length_of.simps startof_not0 start_of_suc_inc)
+done
+
+lemma start_of_eq: "length ap < as \<Longrightarrow> start_of (layout_of ap) as = start_of (layout_of ap) (length ap)"
+apply(induct as, simp)
+apply(case_tac "length ap < as", simp add: start_of.simps)
+apply(subgoal_tac "as = length ap")
+apply(simp add: start_of.simps)
+apply arith
+done
+
+lemma start_of_all_le: "start_of (layout_of ap) as \<le> start_of (layout_of ap) (length ap)"
+apply(subgoal_tac "as > length ap \<or> as = length ap \<or> as < length ap", 
+      auto simp: start_of_eq start_of_less)
+done
+
+lemma findnth_state_all_le2: 
+  "\<lbrakk>k < length ap; 
+  ap ! k = Dec n e;
+  (a, b) \<in> set (shift (findnth n) (start_of (layout_of ap) k - Suc 0))\<rbrakk>
+  \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
+apply(subgoal_tac "b \<le> start_of (layout_of ap) k + 2*n + 1 \<and> 
+     start_of (layout_of ap) k + 2*n + 1 \<le>  start_of (layout_of ap) (Suc k) \<and>
+      start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap)", auto)
+apply(subgoal_tac "start_of (layout_of ap) (Suc k) = 
+  start_of  (layout_of ap)  k + 2*n + 16", simp)
+apply(simp add: start_of_suc_dec)
+apply(rule_tac start_of_all_le)
+done
+
+lemma dec_state_all_le[simp]:
+  "\<lbrakk>k < length ap; ap ! k = Dec n e; 
+  (a, b) \<in> set (shift (shift tdec_b (2 * n))
+  (start_of (layout_of ap) k - Suc 0))\<rbrakk>
+       \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
+apply(subgoal_tac "2*n + start_of (layout_of ap) k + 16 \<le> start_of (layout_of ap) (length ap) \<and> start_of (layout_of ap) k > 0")
+prefer 2
+apply(subgoal_tac "start_of (layout_of ap) (Suc k) = start_of (layout_of ap) k + 2*n + 16
+                 \<and> start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap)")
+apply(simp add: startof_not0, rule_tac conjI)
+apply(simp add: start_of_suc_dec)
+apply(rule_tac start_of_all_le)
+apply(auto simp: tdec_b_def shift.simps)
+done
+
+lemma tms_any_less: 
+  "\<lbrakk>k < length ap; (a, b) \<in> set (tms_of ap ! k)\<rbrakk> \<Longrightarrow> 
+  b \<le> start_of (layout_of ap) (length ap)"
+apply(case_tac "ap!k", auto simp: tms_of.simps tpairs_of.simps ci.simps shift_append sete.simps)
+apply(erule_tac findnth_state_all_le1, simp_all)
+apply(erule_tac inc_state_all_le, simp_all)
+apply(erule_tac findnth_state_all_le2, simp_all)
+apply(rule_tac start_of_all_le)
+apply(rule_tac dec_state_all_le, simp_all)
+apply(rule_tac start_of_all_le)
+done
+
+lemma concat_in: "i < length (concat xs) \<Longrightarrow> \<exists>k < length xs. concat xs ! i \<in> set (xs ! k)"
+apply(induct xs rule: list_tl_induct, simp, simp)
+apply(case_tac "i < length (concat list)", simp)
+apply(erule_tac exE, rule_tac x = k in exI)
+apply(simp add: nth_append)
+apply(rule_tac x = "length list" in exI, simp)
+apply(simp add: nth_append)
+done 
+
+lemma [simp]: "length (tms_of ap) = length ap"
+apply(simp add: tms_of.simps tpairs_of.simps)
+done
+
+declare length_concat[simp]
+
+lemma in_tms: "i < length (tm_of ap) \<Longrightarrow> \<exists> k < length ap. (tm_of ap ! i) \<in> set (tms_of ap ! k)"
+apply(simp only: tm_of.simps)
+using concat_in[of i "tms_of ap"]
+apply(auto)
+done
+
+lemma all_le_start_of: "list_all (\<lambda>(acn, s). 
+  s \<le> start_of (layout_of ap) (length ap)) (tm_of ap)"
+apply(simp only: list_all_length)
+apply(rule_tac allI, rule_tac impI)
+apply(drule_tac in_tms, auto elim: tms_any_less)
+done
+
+lemma length_ci: 
+"\<lbrakk>k < length ap; length (ci ly y (ap ! k)) = 2 * qa\<rbrakk>
+      \<Longrightarrow> layout_of ap ! k = qa"
+apply(case_tac "ap ! k")
+apply(auto simp: layout_of.simps ci.simps 
+  length_of.simps tinc_b_def tdec_b_def length_findnth sete.simps)
+done
+
+lemma [intro]: "length (ci ly y i) mod 2 = 0"
+apply(case_tac i, auto simp: ci.simps length_findnth
+  tinc_b_def sete.simps tdec_b_def)
+done
+
+lemma [intro]: "listsum (map (length \<circ> (\<lambda>(x, y). ci ly x y)) zs) mod 2 = 0"
+apply(induct zs rule: list_tl_induct, simp)
+apply(case_tac a, simp)
+apply(subgoal_tac "length (ci ly aa b) mod 2 = 0")
+apply(auto)
+done
+
+lemma zip_pre:
+  "(length ys) \<le> length ap \<Longrightarrow>
+  zip ys ap = zip ys (take (length ys) (ap::'a list))"
+proof(induct ys arbitrary: ap, simp, case_tac ap, simp)
+  fix a ys ap aa list
+  assume ind: "\<And>(ap::'a list). length ys \<le> length ap \<Longrightarrow> 
+    zip ys ap = zip ys (take (length ys) ap)"
+  and h: "length (a # ys) \<le> length ap" "(ap::'a list) = aa # (list::'a list)"
+  from h show "zip (a # ys) ap = zip (a # ys) (take (length (a # ys)) ap)"
+    using ind[of list]
+    apply(simp)
+    done
+qed
+
+lemma length_start_of_tm: "start_of (layout_of ap) (length ap) = Suc (length (tm_of ap)  div 2)"
+using tpa_states[of "tm_of ap"  "length ap" ap]
+apply(simp add: tm_of.simps)
+done
+
+lemma [elim]: "list_all (\<lambda>(acn, s). s \<le> Suc q) xs
+        \<Longrightarrow> list_all (\<lambda>(acn, s). s \<le> q + (2 * n + 6)) xs"
+apply(simp add: list_all_length)
+apply(auto)
+done
+
+lemma [simp]: "length mopup_b = 12"
+apply(simp add: mopup_b_def)
+done
+(*
+lemma [elim]: "\<lbrakk>na < 4 * n; tshift (mop_bef n) q ! na = (a, b)\<rbrakk> \<Longrightarrow> 
+  b \<le> q + (2 * n + 6)"
+apply(induct n, simp, simp add: mop_bef.simps nth_append tshift_append shift_length)
+apply(case_tac "na < 4*n", simp, simp)
+apply(subgoal_tac "na = 4*n \<or> na = 1 + 4*n \<or> na = 2 + 4*n \<or> na = 3 + 4*n",
+  auto simp: shift_length)
+apply(simp_all add: tshift.simps)
+done
+*)
+
+lemma mp_up_all_le: "list_all  (\<lambda>(acn, s). s \<le> q + (2 * n + 6)) 
+  [(R, Suc (Suc (2 * n + q))), (R, Suc (2 * n + q)), 
+  (L, 5 + 2 * n + q), (W0, Suc (Suc (Suc (2 * n + q)))), (R, 4 + 2 * n + q),
+  (W0, Suc (Suc (Suc (2 * n + q)))), (R, Suc (Suc (2 * n + q))),
+  (W0, Suc (Suc (Suc (2 * n + q)))), (L, 5 + 2 * n + q),
+  (L, 6 + 2 * n + q), (R, 0),  (L, 6 + 2 * n + q)]"
+apply(auto)
+done
+
+lemma [simp]: "(a, b) \<in> set (mopup_a n) \<Longrightarrow> b \<le> 2 * n + 6"
+apply(induct n, auto simp: mopup_a.simps)
+done
+
+lemma [simp]: "(a, b) \<in> set (shift (mopup n) (listsum (layout_of ap)))
+  \<Longrightarrow> b \<le> (2 * listsum (layout_of ap) + length (mopup n)) div 2"
+apply(auto simp: mopup.simps shift_append shift.simps)
+apply(auto simp: mopup_a.simps mopup_b_def)
+done
+
+lemma [intro]: " 2 \<le> 2 * listsum (layout_of ap) + length (mopup n)"
+apply(simp add: mopup.simps)
+done
+
+lemma [intro]: " (2 * listsum (layout_of ap) + length (mopup n)) mod 2 = 0"
+apply(auto simp: mopup.simps)
+apply arith
+done
+
+lemma [simp]: "b \<le> Suc x
+          \<Longrightarrow> b \<le> (2 * x + length (mopup n)) div 2"
+apply(auto simp: mopup.simps)
+done
+
+lemma t_compiled_correct: 
+  "\<lbrakk>tp = tm_of ap; ly = layout_of ap; mop_ss = start_of ly (length ap)\<rbrakk> \<Longrightarrow> 
+    tm_wf (tp @ shift( mopup n) (length tp div 2), 0)"
+  using length_start_of_tm[of ap] all_le_start_of[of ap]
+apply(auto simp: tm_wf.simps List.list_all_iff)
+done
+
+end
+
+    
+  
+
+
+  
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/Turing.thy	Sun Feb 10 19:49:07 2013 +0000
@@ -0,0 +1,428 @@
+(* Title: Turing machines
+   Author: Xu Jian <xujian817@hotmail.com>
+   Maintainer: Xu Jian
+*)
+
+theory Turing
+imports Main
+begin
+
+section {* Basic definitions of Turing machine *}
+
+datatype action = W0 | W1 | L | R | Nop
+
+datatype cell = Bk | Oc
+
+type_synonym tape = "cell list \<times> cell list"
+
+type_synonym state = nat
+
+type_synonym instr = "action \<times> state"
+
+type_synonym tprog = "instr list \<times> nat"
+
+type_synonym tprog0 = "instr list"
+
+type_synonym config = "state \<times> tape"
+
+fun nth_of where
+  "nth_of xs i = (if i \<ge> length xs then None
+                  else Some (xs ! i))"
+
+lemma nth_of_map [simp]:
+  shows "nth_of (map f p) n = (case (nth_of p n) of None \<Rightarrow> None | Some x \<Rightarrow> Some (f x))"
+apply(induct p arbitrary: n)
+apply(auto)
+apply(case_tac n)
+apply(auto)
+done
+
+fun 
+  fetch :: "instr list \<Rightarrow> state \<Rightarrow> cell \<Rightarrow> instr"
+where
+  "fetch p 0 b = (Nop, 0)"
+| "fetch p (Suc s) Bk = 
+     (case nth_of p (2 * s) of
+        Some i \<Rightarrow> i
+      | None \<Rightarrow> (Nop, 0))"
+|"fetch p (Suc s) Oc = 
+     (case nth_of p ((2 * s) + 1) of
+         Some i \<Rightarrow> i
+       | None \<Rightarrow> (Nop, 0))"
+
+lemma fetch_Nil [simp]:
+  shows "fetch [] s b = (Nop, 0)"
+apply(case_tac s)
+apply(auto)
+apply(case_tac b)
+apply(auto)
+done
+
+fun 
+  update :: "action \<Rightarrow> tape \<Rightarrow> tape"
+where 
+  "update W0 (l, r) = (l, Bk # (tl r))" 
+| "update W1 (l, r) = (l, Oc # (tl r))"
+| "update L (l, r) = (if l = [] then ([], Bk # r) else (tl l, (hd l) # r))" 
+| "update R (l, r) = (if r = [] then (Bk # l, []) else ((hd r) # l, tl r))" 
+| "update Nop (l, r) = (l, r)"
+
+abbreviation 
+  "read r == if (r = []) then Bk else hd r"
+
+fun step :: "config \<Rightarrow> tprog \<Rightarrow> config"
+  where 
+  "step (s, l, r) (p, off) = 
+     (let (a, s') = fetch p (s - off) (read r) in (s', update a (l, r)))"
+
+fun steps :: "config \<Rightarrow> tprog \<Rightarrow> nat \<Rightarrow> config"
+  where
+  "steps c p 0 = c" |
+  "steps c p (Suc n) = steps (step c p) p n"
+
+
+abbreviation
+  "step0 c p \<equiv> step c (p, 0)"
+
+abbreviation
+  "steps0 c p n \<equiv> steps c (p, 0) n"
+
+lemma step_red [simp]: 
+  shows "steps c p (Suc n) = step (steps c p n) p"
+by (induct n arbitrary: c) (auto)
+
+lemma steps_add [simp]: 
+  shows "steps c p (m + n) = steps (steps c p m) p n"
+by (induct m arbitrary: c) (auto)
+
+lemma step_0 [simp]: 
+  shows "step (0, (l, r)) p = (0, (l, r))"
+by (case_tac p, simp)
+
+lemma steps_0 [simp]: 
+  shows "steps (0, (l, r)) p n = (0, (l, r))"
+by (induct n) (simp_all)
+
+
+
+fun
+  is_final :: "config \<Rightarrow> bool"
+where
+  "is_final (s, l, r) = (s = 0)"
+
+lemma is_final_eq: 
+  shows "is_final (s, tp) = (s = 0)"
+by (case_tac tp) (auto)
+
+lemma after_is_final:
+  assumes "is_final c"
+  shows "is_final (steps c p n)"
+using assms 
+apply(induct n) 
+apply(case_tac [!] c)
+apply(auto)
+done
+
+lemma not_is_final:
+  assumes a: "\<not> is_final (steps c p n1)"
+  and b: "n2 \<le> n1"
+  shows "\<not> is_final (steps c p n2)"
+proof (rule notI)
+  obtain n3 where eq: "n1 = n2 + n3" using b by (metis le_iff_add)
+  assume "is_final (steps c p n2)"
+  then have "is_final (steps c p n1)" unfolding eq
+    by (simp add: after_is_final)
+  with a show "False" by simp
+qed
+
+(* if the machine is in the halting state, there must have 
+   been a state just before the halting state *)
+lemma before_final: 
+  assumes "steps0 (1, tp) A n = (0, tp')"
+  shows "\<exists> n'. \<not> is_final (steps0 (1, tp) A n') \<and> steps0 (1, tp) A (Suc n') = (0, tp')"
+using assms
+proof(induct n arbitrary: tp')
+  case (0 tp')
+  have asm: "steps0 (1, tp) A 0 = (0, tp')" by fact
+  then show "\<exists>n'. \<not> is_final (steps0 (1, tp) A n') \<and> steps0 (1, tp) A (Suc n') = (0, tp')"
+    by simp
+next
+  case (Suc n tp')
+  have ih: "\<And>tp'. steps0 (1, tp) A n = (0, tp') \<Longrightarrow>
+    \<exists>n'. \<not> is_final (steps0 (1, tp) A n') \<and> steps0 (1, tp) A (Suc n') = (0, tp')" by fact
+  have asm: "steps0 (1, tp) A (Suc n) = (0, tp')" by fact
+  obtain s l r where cases: "steps0 (1, tp) A n = (s, l, r)"
+    by (auto intro: is_final.cases)
+  then show "\<exists>n'. \<not> is_final (steps0 (1, tp) A n') \<and> steps0 (1, tp) A (Suc n') = (0, tp')"
+  proof (cases "s = 0")
+    case True (* in halting state *)
+    then have "steps0 (1, tp) A n = (0, tp')"
+      using asm cases by (simp del: steps.simps)
+    then show ?thesis using ih by simp
+  next
+    case False (* not in halting state *)
+    then have "\<not> is_final (steps0 (1, tp) A n) \<and> steps0 (1, tp) A (Suc n) = (0, tp')"
+      using asm cases by simp
+    then show ?thesis by auto
+  qed
+qed
+
+(* well-formedness of Turing machine programs *)
+abbreviation "is_even n \<equiv> (n::nat) mod 2 = 0"
+
+fun 
+  tm_wf :: "tprog \<Rightarrow> bool"
+where
+  "tm_wf (p, off) = (length p \<ge> 2 \<and> is_even (length p) \<and> 
+                    (\<forall>(a, s) \<in> set p. s \<le> length p div 2 + off \<and> s \<ge> off))"
+
+abbreviation
+  "tm_wf0 p \<equiv> tm_wf (p, 0)"
+
+abbreviation exponent :: "'a \<Rightarrow> nat \<Rightarrow> 'a list" ("_ \<up> _" [100, 99] 100)
+  where "x \<up> n == replicate n x"
+
+consts tape_of :: "'a \<Rightarrow> cell list" ("<_>" 100)
+
+defs (overloaded)
+  tape_of_nat_abv: "<(n::nat)> \<equiv> Oc \<up> (Suc n)"
+
+fun tape_of_nat_list :: "'a list \<Rightarrow> cell list" 
+  where 
+  "tape_of_nat_list [] = []" |
+  "tape_of_nat_list [n] = <n>" |
+  "tape_of_nat_list (n#ns) = <n> @ Bk # (tape_of_nat_list ns)"
+
+fun tape_of_nat_pair :: "'a \<times> 'b \<Rightarrow> cell list" 
+  where 
+  "tape_of_nat_pair (n, m) = <n> @ [Bk] @ <m>" 
+
+
+defs (overloaded)
+  tape_of_nl_abv: "<(ns::'a list)> \<equiv> tape_of_nat_list ns"
+  tape_of_nat_pair: "<(np::'a\<times>'b)> \<equiv> tape_of_nat_pair np"
+
+fun 
+  shift :: "instr list \<Rightarrow> nat \<Rightarrow> instr list"
+where
+  "shift p n = (map (\<lambda> (a, s). (a, (if s = 0 then 0 else s + n))) p)"
+
+fun 
+  adjust :: "instr list \<Rightarrow> instr list"
+where
+  "adjust p = map (\<lambda> (a, s). (a, if s = 0 then (Suc (length p div 2)) else s)) p"
+
+lemma length_shift [simp]: 
+  shows "length (shift p n) = length p"
+by simp
+
+lemma length_adjust [simp]: 
+  shows "length (adjust p) = length p"
+by (induct p) (auto)
+
+
+(* composition of two Turing machines *)
+fun
+  tm_comp :: "instr list \<Rightarrow> instr list \<Rightarrow> instr list" ("_ |+| _" [0, 0] 100)
+where
+  "tm_comp p1 p2 = ((adjust p1) @ (shift p2 (length p1 div 2)))"
+
+lemma tm_comp_length:
+  shows "length (A |+| B) = length A + length B"
+by auto
+
+lemma tm_comp_wf[intro]: 
+  "\<lbrakk>tm_wf (A, 0); tm_wf (B, 0)\<rbrakk> \<Longrightarrow> tm_wf (A |+| B, 0)"
+by (auto simp: tm_wf.simps shift.simps adjust.simps tm_comp_length tm_comp.simps)
+
+
+lemma tm_comp_step: 
+  assumes unfinal: "\<not> is_final (step0 c A)"
+  shows "step0 c (A |+| B) = step0 c A"
+proof -
+  obtain s l r where eq: "c = (s, l, r)" by (metis is_final.cases) 
+  have "\<not> is_final (step0 (s, l, r) A)" using unfinal eq by simp
+  then have "case (fetch A s (read r)) of (a, s) \<Rightarrow> s \<noteq> 0"
+    by (auto simp add: is_final_eq)
+  then  have "fetch (A |+| B) s (read r) = fetch A s (read r)"
+    apply(case_tac [!] "read r")
+    apply(case_tac [!] s)
+    apply(auto simp: tm_comp_length nth_append)
+    done
+  then show "step0 c (A |+| B) = step0 c A" by (simp add: eq) 
+qed
+
+lemma tm_comp_steps:  
+  assumes "\<not> is_final (steps0 c A n)" 
+  shows "steps0 c (A |+| B) n = steps0 c A n"
+using assms
+proof(induct n)
+  case 0
+  then show "steps0 c (A |+| B) 0 = steps0 c A 0" by auto
+next 
+  case (Suc n)
+  have ih: "\<not> is_final (steps0 c A n) \<Longrightarrow> steps0 c (A |+| B) n = steps0 c A n" by fact
+  have fin: "\<not> is_final (steps0 c A (Suc n))" by fact
+  then have fin1: "\<not> is_final (step0 (steps0 c A n) A)" 
+    by (auto simp only: step_red)
+  then have fin2: "\<not> is_final (steps0 c A n)"
+    by (metis is_final_eq step_0 surj_pair) 
+ 
+  have "steps0 c (A |+| B) (Suc n) = step0 (steps0 c (A |+| B) n) (A |+| B)" 
+    by (simp only: step_red)
+  also have "... = step0 (steps0 c A n) (A |+| B)" by (simp only: ih[OF fin2])
+  also have "... = step0 (steps0 c A n) A" by (simp only: tm_comp_step[OF fin1])
+  finally show "steps0 c (A |+| B) (Suc n) = steps0 c A (Suc n)"
+    by (simp only: step_red)
+qed
+
+lemma tm_comp_fetch_in_A:
+  assumes h1: "fetch A s x = (a, 0)"
+  and h2: "s \<le> length A div 2" 
+  and h3: "s \<noteq> 0"
+  shows "fetch (A |+| B) s x = (a, Suc (length A div 2))"
+using h1 h2 h3
+apply(case_tac s)
+apply(case_tac [!] x)
+apply(auto simp: tm_comp_length nth_append)
+done
+
+lemma tm_comp_exec_after_first:
+  assumes h1: "\<not> is_final c" 
+  and h2: "step0 c A = (0, tp)"
+  and h3: "fst c \<le> length A div 2"
+  shows "step0 c (A |+| B) = (Suc (length A div 2), tp)"
+using h1 h2 h3
+apply(case_tac c)
+apply(auto simp del: tm_comp.simps)
+apply(case_tac "fetch A a Bk")
+apply(simp del: tm_comp.simps)
+apply(subst tm_comp_fetch_in_A)
+apply(auto)[4]
+apply(case_tac "fetch A a (hd c)")
+apply(simp del: tm_comp.simps)
+apply(subst tm_comp_fetch_in_A)
+apply(auto)[4]
+done
+
+lemma step_in_range: 
+  assumes h1: "\<not> is_final (step0 c A)"
+  and h2: "tm_wf (A, 0)"
+  shows "fst (step0 c A) \<le> length A div 2"
+using h1 h2
+apply(case_tac c)
+apply(case_tac a)
+apply(auto simp add: prod_case_unfold Let_def)
+apply(case_tac "hd c")
+apply(auto simp add: prod_case_unfold)
+done
+
+lemma steps_in_range: 
+  assumes h1: "\<not> is_final (steps0 (1, tp) A stp)"
+  and h2: "tm_wf (A, 0)"
+  shows "fst (steps0 (1, tp) A stp) \<le> length A div 2"
+using h1
+proof(induct stp)
+  case 0
+  then show "fst (steps0 (1, tp) A 0) \<le> length A div 2" using h2
+    by (auto simp add: steps.simps tm_wf.simps)
+next
+  case (Suc stp)
+  have ih: "\<not> is_final (steps0 (1, tp) A stp) \<Longrightarrow> fst (steps0 (1, tp) A stp) \<le> length A div 2" by fact
+  have h: "\<not> is_final (steps0 (1, tp) A (Suc stp))" by fact
+  from ih h h2 show "fst (steps0 (1, tp) A (Suc stp)) \<le> length A div 2"
+    by (metis step_in_range step_red)
+qed
+
+lemma tm_comp_pre_halt_same: 
+  assumes a_ht: "steps0 (1, tp) A n = (0, tp')"
+  and a_wf: "tm_wf (A, 0)"
+  obtains n' where "steps0 (1, tp) (A |+| B) n' = (Suc (length A div 2), tp')"
+proof -
+  assume a: "\<And>n. steps (1, tp) (A |+| B, 0) n = (Suc (length A div 2), tp') \<Longrightarrow> thesis"
+  obtain stp' where fin: "\<not> is_final (steps0 (1, tp) A stp')" and h: "steps0 (1, tp) A (Suc stp') = (0, tp')"
+  using before_final[OF a_ht] by blast
+  from fin have h1:"steps0 (1, tp) (A |+| B) stp' = steps0 (1, tp) A stp'"
+    by (rule tm_comp_steps)
+  from h have h2: "step0 (steps0 (1, tp) A stp') A = (0, tp')"
+    by (simp only: step_red)
+
+  have "steps0 (1, tp) (A |+| B) (Suc stp') = step0 (steps0 (1, tp) (A |+| B) stp') (A |+| B)" 
+    by (simp only: step_red)
+  also have "... = step0 (steps0 (1, tp) A stp') (A |+| B)" using h1 by simp
+  also have "... = (Suc (length A div 2), tp')" 
+    by (rule tm_comp_exec_after_first[OF fin h2 steps_in_range[OF fin a_wf]])
+  finally show thesis using a by blast
+qed
+
+lemma tm_comp_fetch_second_zero:
+  assumes h1: "fetch B s x = (a, 0)"
+  and hs: "tm_wf (A, 0)" "s \<noteq> 0"
+  shows "fetch (A |+| B) (s + (length A div 2)) x = (a, 0)"
+using h1 hs
+apply(case_tac x)
+apply(case_tac [!] s)
+apply(auto simp: tm_comp_length nth_append)
+done 
+
+lemma tm_comp_fetch_second_inst:
+  assumes h1: "fetch B sa x = (a, s)"
+  and hs: "tm_wf (A, 0)" "sa \<noteq> 0" "s \<noteq> 0"
+  shows "fetch (A |+| B) (sa + length A div 2) x = (a, s + length A div 2)"
+using h1 hs
+apply(case_tac x)
+apply(case_tac [!] sa)
+apply(auto simp: tm_comp_length nth_append)
+done 
+
+
+lemma tm_comp_second_same:
+  assumes a_wf: "tm_wf (A, 0)"
+  and steps: "steps0 (1, l, r) B stp = (s', l', r')"
+  shows "steps0 (Suc (length A div 2), l, r)  (A |+| B) stp 
+    = (if s' = 0 then 0 else s' + length A div 2, l', r')"
+using steps
+proof(induct stp arbitrary: s' l' r')
+  case 0
+  then show ?case by (simp add: steps.simps)
+next
+  case (Suc stp s' l' r')
+  obtain s'' l'' r'' where a: "steps0 (1, l, r) B stp = (s'', l'', r'')"
+    by (metis is_final.cases)
+  then have ih1: "s'' = 0 \<Longrightarrow> steps0 (Suc (length A div 2), l, r) (A |+| B) stp = (0, l'', r'')"
+  and ih2: "s'' \<noteq> 0 \<Longrightarrow> steps0 (Suc (length A div 2), l, r) (A |+| B) stp = (s'' + length A div 2, l'', r'')"
+    using Suc by (auto)
+  have h: "steps0 (1, l, r) B (Suc stp) = (s', l', r')" by fact
+
+  { assume "s'' = 0"
+    then have ?case using a h ih1 by (simp del: steps.simps) 
+  } moreover
+  { assume as: "s'' \<noteq> 0" "s' = 0"
+    from as a h 
+    have "step0 (s'', l'', r'') B = (0, l', r')" by (simp del: steps.simps)
+    with as have ?case
+    apply(simp add: ih2[OF as(1)] step.simps del: tm_comp.simps steps.simps)
+    apply(case_tac "fetch B s'' (read r'')")
+    apply(auto simp add: tm_comp_fetch_second_zero[OF _ a_wf] simp del: tm_comp.simps)
+    done
+  } moreover
+  { assume as: "s'' \<noteq> 0" "s' \<noteq> 0"
+    from as a h
+    have "step0 (s'', l'', r'') B = (s', l', r')" by (simp del: steps.simps)
+    with as have ?case
+    apply(simp add: ih2[OF as(1)] step.simps del: tm_comp.simps steps.simps)
+    apply(case_tac "fetch B s'' (read r'')")
+    apply(auto simp add: tm_comp_fetch_second_inst[OF _ a_wf as] simp del: tm_comp.simps)
+    done
+  }
+  ultimately show ?case by blast
+qed
+
+lemma tm_comp_second_halt_same:
+  assumes "tm_wf (A, 0)"  
+  and "steps0 (1, l, r) B stp = (0, l', r')"
+  shows "steps0 (Suc (length A div 2), l, r)  (A |+| B) stp = (0, l', r')"
+using tm_comp_second_same[OF assms] by (simp)
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/Turing_Hoare.thy	Sun Feb 10 19:49:07 2013 +0000
@@ -0,0 +1,159 @@
+theory Turing_Hoare
+imports Turing
+begin
+
+
+type_synonym assert = "tape \<Rightarrow> bool"
+
+definition 
+  assert_imp :: "assert \<Rightarrow> assert \<Rightarrow> bool" ("_ \<mapsto> _" [0, 0] 100)
+where
+  "P \<mapsto> Q \<equiv> \<forall>l r. P (l, r) \<longrightarrow> Q (l, r)"
+
+lemma [intro, simp]:
+  "P \<mapsto> P"
+unfolding assert_imp_def by simp
+
+fun 
+  holds_for :: "(tape \<Rightarrow> bool) \<Rightarrow> config \<Rightarrow> bool" ("_ holds'_for _" [100, 99] 100)
+where
+  "P holds_for (s, l, r) = P (l, r)"  
+
+lemma is_final_holds[simp]:
+  assumes "is_final c"
+  shows "Q holds_for (steps c p n) = Q holds_for c"
+using assms 
+apply(induct n)
+apply(auto)
+apply(case_tac [!] c)
+apply(auto)
+done
+
+(* Hoare Rules *)
+
+(* halting case *)
+definition
+  Hoare_halt :: "assert \<Rightarrow> tprog0 \<Rightarrow> assert \<Rightarrow> bool" ("({(1_)}/ (_)/ {(1_)})" 50)
+where
+  "{P} p {Q} \<equiv> \<forall>tp. P tp \<longrightarrow> (\<exists>n. is_final (steps0 (1, tp) p n) \<and> Q holds_for (steps0 (1, tp) p n))"
+
+
+(* not halting case *)
+definition
+  Hoare_unhalt :: "assert \<Rightarrow> tprog0 \<Rightarrow> bool" ("({(1_)}/ (_)) \<up>" 50)
+where
+  "{P} p \<up> \<equiv> \<forall>tp. P tp \<longrightarrow> (\<forall> n . \<not> (is_final (steps0 (1, tp) p n)))"
+
+
+lemma Hoare_haltI:
+  assumes "\<And>l r. P (l, r) \<Longrightarrow> \<exists>n. is_final (steps0 (1, (l, r)) p n) \<and> Q holds_for (steps0 (1, (l, r)) p n)"
+  shows "{P} p {Q}"
+unfolding Hoare_halt_def 
+using assms by auto
+
+lemma Hoare_unhaltI:
+  assumes "\<And>l r n. P (l, r) \<Longrightarrow> \<not> is_final (steps0 (1, (l, r)) p n)"
+  shows "{P} p \<up>"
+unfolding Hoare_unhalt_def 
+using assms by auto
+
+
+
+
+text {*
+  {P} A {Q}   {Q} B {S}  A well-formed
+  -----------------------------------------
+  {P} A |+| B {S}
+*}
+
+
+lemma Hoare_plus_halt [case_names A_halt B_halt A_wf]: 
+  assumes A_halt : "{P} A {Q}"
+  and B_halt : "{Q} B {S}"
+  and A_wf : "tm_wf (A, 0)"
+  shows "{P} A |+| B {S}"
+proof(rule Hoare_haltI)
+  fix l r
+  assume h: "P (l, r)"
+  then obtain n1 l' r' 
+    where "is_final (steps0 (1, l, r) A n1)"  
+      and a1: "Q holds_for (steps0 (1, l, r) A n1)"
+      and a2: "steps0 (1, l, r) A n1 = (0, l', r')"
+    using A_halt unfolding Hoare_halt_def
+    by (metis is_final_eq surj_pair) 
+  then obtain n2 
+    where "steps0 (1, l, r) (A |+| B) n2 = (Suc (length A div 2), l', r')"
+    using A_wf by (rule_tac tm_comp_pre_halt_same) 
+  moreover
+  from a1 a2 have "Q (l', r')" by (simp)
+  then obtain n3 l'' r''
+    where "is_final (steps0 (1, l', r') B n3)" 
+    and b1: "S holds_for (steps0 (1, l', r') B n3)"
+    and b2: "steps0 (1, l', r') B n3 = (0, l'', r'')"
+    using B_halt unfolding Hoare_halt_def 
+    by (metis is_final_eq surj_pair) 
+  then have "steps0 (Suc (length A div 2), l', r')  (A |+| B) n3 = (0, l'', r'')"
+    using A_wf by (rule_tac tm_comp_second_halt_same) 
+  ultimately show 
+    "\<exists>n. is_final (steps0 (1, l, r) (A |+| B) n) \<and> S holds_for (steps0 (1, l, r) (A |+| B) n)"
+    using b1 b2 by (rule_tac x = "n2 + n3" in exI) (simp)
+qed
+
+text {*
+  {P} A {Q}   {Q} B loops   A well-formed
+  ------------------------------------------
+          {P} A |+| B  loops
+*}
+
+lemma Hoare_plus_unhalt [case_names A_halt B_unhalt A_wf]:
+  assumes A_halt: "{P} A {Q}"
+  and B_uhalt: "{Q} B \<up>"
+  and A_wf : "tm_wf (A, 0)"
+  shows "{P} (A |+| B) \<up>"
+proof(rule_tac Hoare_unhaltI)
+  fix n l r 
+  assume h: "P (l, r)"
+  then obtain n1 l' r'
+    where a: "is_final (steps0 (1, l, r) A n1)" 
+    and b: "Q holds_for (steps0 (1, l, r) A n1)"
+    and c: "steps0 (1, l, r) A n1 = (0, l', r')"
+    using A_halt unfolding Hoare_halt_def 
+    by (metis is_final_eq surj_pair) 
+  then obtain n2 where eq: "steps0 (1, l, r) (A |+| B) n2 = (Suc (length A div 2), l', r')"
+    using A_wf by (rule_tac tm_comp_pre_halt_same)
+  then show "\<not> is_final (steps0 (1, l, r) (A |+| B) n)"
+  proof(cases "n2 \<le> n")
+    case True
+    from b c have "Q (l', r')" by simp
+    then have "\<forall> n. \<not> is_final (steps0 (1, l', r') B n)  "
+      using B_uhalt unfolding Hoare_unhalt_def by simp
+    then have "\<not> is_final (steps0 (1, l', r') B (n - n2))" by auto
+    then obtain s'' l'' r'' 
+      where "steps0 (1, l', r') B (n - n2) = (s'', l'', r'')" 
+      and "\<not> is_final (s'', l'', r'')" by (metis surj_pair)
+    then have "steps0 (Suc (length A div 2), l', r') (A |+| B) (n - n2) = (s''+ length A div 2, l'', r'')"
+      using A_wf by (auto dest: tm_comp_second_same simp del: tm_wf.simps)
+    then have "\<not> is_final (steps0 (1, l, r) (A |+| B) (n2 + (n  - n2)))"
+      using A_wf by (simp only: steps_add eq) (simp add: tm_wf.simps)
+    then show "\<not> is_final (steps0 (1, l, r) (A |+| B) n)" 
+      using `n2 \<le> n` by simp
+  next 
+    case False
+    then obtain n3 where "n = n2 - n3"
+      by (metis diff_le_self le_imp_diff_is_add nat_add_commute nat_le_linear)
+    moreover
+    with eq show "\<not> is_final (steps0 (1, l, r) (A |+| B) n)"
+      by (simp add: not_is_final[where ?n1.0="n2"])
+  qed
+qed
+
+lemma Hoare_consequence:
+  assumes "P' \<mapsto> P" "{P} p {Q}" "Q \<mapsto> Q'"
+  shows "{P'} p {Q'}"
+using assms
+unfolding Hoare_halt_def assert_imp_def
+by (metis holds_for.simps surj_pair)
+
+
+
+end
\ No newline at end of file
--- a/thys/UF.thy	Thu Feb 07 06:39:06 2013 +0000
+++ b/thys/UF.thy	Sun Feb 10 19:49:07 2013 +0000
@@ -1,5 +1,5 @@
 theory UF
-imports Main rec_def turing_basic GCD abacus
+imports Rec_Def GCD Abacus
 begin
 
 text {*
--- a/thys/UTM.thy	Thu Feb 07 06:39:06 2013 +0000
+++ b/thys/UTM.thy	Sun Feb 10 19:49:07 2013 +0000
@@ -1,5 +1,5 @@
 theory UTM
-imports Main recursive abacus UF GCD turing_hoare
+imports Recursive Abacus UF GCD Turing_Hoare
 begin
 
 section {* Wang coding of input arguments *}
@@ -1247,7 +1247,7 @@
 
 lemma adjust_fetch_norm: 
   "\<lbrakk>st > 0;  st \<le> length tp div 2; fetch ap st b = (aa, ns); ns \<noteq> 0\<rbrakk>
- \<Longrightarrow>  fetch (turing_basic.adjust ap) st b = (aa, ns)"
+ \<Longrightarrow>  fetch (Turing.adjust ap) st b = (aa, ns)"
  apply(case_tac b, auto simp: fetch.simps nth_of.simps nth_map
                        split: if_splits)
 apply(case_tac [!] st, auto simp: fetch.simps nth_of.simps)
@@ -1299,7 +1299,7 @@
 next
   case (Suc stp st' l' r')
   have ind: "\<And>st' l' r'. \<lbrakk>steps0 (st, l, r) ap stp = (st', l', r'); 0 < st'\<rbrakk> 
-    \<Longrightarrow> steps0 (st, l, r) (turing_basic.adjust ap) stp = (st', l', r')" by fact
+    \<Longrightarrow> steps0 (st, l, r) (Turing.adjust ap) stp = (st', l', r')" by fact
   have h: "steps0 (st, l, r) ap (Suc stp) = (st', l', r')" by fact
   have g:   "0 < st'" by fact
   obtain st'' l'' r'' where a: "steps0 (st, l, r) ap stp = (st'', l'', r'')"
@@ -1309,7 +1309,7 @@
     apply(simp add: step_red)
     apply(case_tac st'', auto)
     done
-  hence b: "steps0 (st, l, r) (turing_basic.adjust ap) stp = (st'', l'', r'')"
+  hence b: "steps0 (st, l, r) (Turing.adjust ap) stp = (st'', l'', r'')"
     using a
     by(rule_tac ind, simp_all)
   thus "?case"
@@ -3373,28 +3373,28 @@
 apply(auto simp: mopup.simps)
 done
 
-lemma [elim]: "(a, b) \<in> set (shift (turing_basic.adjust t_twice_compile) 12) \<Longrightarrow> 
+lemma [elim]: "(a, b) \<in> set (shift (Turing.adjust t_twice_compile) 12) \<Longrightarrow> 
   b \<le> (28 + (length t_twice_compile + length t_fourtimes_compile)) div 2"
 apply(simp add: t_twice_compile_def t_fourtimes_compile_def)
 proof -
-  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)"
+  assume g: "(a, b) \<in> set (shift (Turing.adjust (tm_of abc_twice @ shift (mopup (Suc 0)) (length (tm_of abc_twice) div 2))) 12)"
   moreover have "length (tm_of abc_twice) mod 2 = 0" by auto
   moreover have "length (tm_of abc_fourtimes) mod 2 = 0" by auto
   ultimately have "list_all (\<lambda>(acn, st). (st \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2)) 
-    (shift (turing_basic.adjust t_twice_compile) 12)"
+    (shift (Turing.adjust t_twice_compile) 12)"
   proof(auto simp: mod_ex1)
     fix q qa
     assume h: "length (tm_of abc_twice) = 2 * q" "length (tm_of abc_fourtimes) = 2 * qa"
-    hence "list_all (\<lambda>(acn, st). st \<le> (18 + (q + qa)) + 12) (shift (turing_basic.adjust t_twice_compile) 12)"
+    hence "list_all (\<lambda>(acn, st). st \<le> (18 + (q + qa)) + 12) (shift (Turing.adjust t_twice_compile) 12)"
     proof(rule_tac tm_wf_shift t_twice_compile_def)
       have "list_all (\<lambda>(acn, st). st \<le> Suc (length t_twice_compile div 2)) (adjust t_twice_compile)"
         by(rule_tac tm_wf_change_termi, auto)
-      thus "list_all (\<lambda>(acn, st). st \<le> 18 + (q + qa)) (turing_basic.adjust t_twice_compile)"
+      thus "list_all (\<lambda>(acn, st). st \<le> 18 + (q + qa)) (Turing.adjust t_twice_compile)"
         using h
         apply(simp add: t_twice_compile_def, auto simp: List.list_all_length)
         done
     qed
-    thus "list_all (\<lambda>(acn, st). st \<le> 30 + (q + qa)) (shift (turing_basic.adjust t_twice_compile) 12)"
+    thus "list_all (\<lambda>(acn, st). st \<le> 30 + (q + qa)) (shift (Turing.adjust t_twice_compile) 12)"
       by simp
   qed
   thus "b \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2"
@@ -3405,22 +3405,22 @@
     done
 qed 
 
-lemma [elim]: "(a, b) \<in> set (shift (turing_basic.adjust t_fourtimes_compile) (t_twice_len + 13)) 
+lemma [elim]: "(a, b) \<in> set (shift (Turing.adjust t_fourtimes_compile) (t_twice_len + 13)) 
   \<Longrightarrow> b \<le> (28 + (length t_twice_compile + length t_fourtimes_compile)) div 2"
 apply(simp add: t_twice_compile_def t_fourtimes_compile_def t_twice_len_def)
 proof -
-  assume g: "(a, b) \<in> set (shift (turing_basic.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0))
+  assume g: "(a, b) \<in> set (shift (Turing.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0))
     (length (tm_of abc_fourtimes) div 2))) (length t_twice div 2 + 13))"
   moreover have "length (tm_of abc_twice) mod 2 = 0" by auto
   moreover have "length (tm_of abc_fourtimes) mod 2 = 0" by auto
   ultimately have "list_all (\<lambda>(acn, st). (st \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2)) 
-    (shift (turing_basic.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0))
+    (shift (Turing.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0))
     (length (tm_of abc_fourtimes) div 2))) (length t_twice div 2 + 13))"
   proof(auto simp: mod_ex1 t_twice_def t_twice_compile_def)
     fix q qa
     assume h: "length (tm_of abc_twice) = 2 * q" "length (tm_of abc_fourtimes) = 2 * qa"
     hence "list_all (\<lambda>(acn, st). st \<le> (9 + qa + (21 + q)))
-      (shift (turing_basic.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa)) (21 + q))"
+      (shift (Turing.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa)) (21 + q))"
     proof(rule_tac tm_wf_shift t_twice_compile_def)
       have "list_all (\<lambda>(acn, st). st \<le> Suc (length (tm_of abc_fourtimes @ shift 
         (mopup (Suc 0)) qa) div 2)) (adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa))"
@@ -3428,12 +3428,12 @@
         using wf_fourtimes h
         apply(simp add: t_fourtimes_compile_def)
         done        
-      thus "list_all (\<lambda>(acn, st). st \<le> 9 + qa) ((turing_basic.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa)))"
+      thus "list_all (\<lambda>(acn, st). st \<le> 9 + qa) ((Turing.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa)))"
         using h
         apply(simp)
         done
     qed
-    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))"
+    thus "list_all (\<lambda>(acn, st). st \<le> 30 + (q + qa)) (shift (Turing.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa)) (21 + q))"
       apply(subgoal_tac "qa + q = q + qa")
       apply(simp, simp)
       done
@@ -3512,7 +3512,7 @@
     apply(auto simp: Hoare_halt_def)
     apply(rule_tac x = n in exI)
     apply(case_tac "(steps0 (Suc 0, [], <m # args>)
-      (turing_basic.adjust t_wcode_prepare @ shift t_wcode_main (length t_wcode_prepare div 2)) n)")
+      (Turing.adjust t_wcode_prepare @ shift t_wcode_main (length t_wcode_prepare div 2)) n)")
     apply(auto simp: tm_comp.simps)
     done
 qed
@@ -3919,12 +3919,12 @@
   where "wadjust_le \<equiv> (inv_image lex_square wadjust_measure)"
 
 lemma [intro]: "wf lex_square"
-by(auto intro:wf_lex_prod simp: abacus.lex_pair_def lex_square_def 
-  abacus.lex_triple_def)
+by(auto intro:wf_lex_prod simp: Abacus.lex_pair_def lex_square_def 
+  Abacus.lex_triple_def)
 
 lemma wf_wadjust_le[intro]: "wf wadjust_le"
 by(auto intro:wf_inv_image simp: wadjust_le_def
-           abacus.lex_triple_def abacus.lex_pair_def)
+           Abacus.lex_triple_def Abacus.lex_pair_def)
 
 lemma [simp]: "wadjust_start m rs (c, []) = False"
 apply(auto simp: wadjust_start.simps)
@@ -4518,7 +4518,7 @@
       apply(simp_all only: wadjust_inv.simps split: if_splits)
       apply(simp_all)
       apply(simp_all add: wadjust_inv.simps wadjust_le_def
-            abacus.lex_triple_def abacus.lex_pair_def lex_square_def  split: if_splits)
+            Abacus.lex_triple_def Abacus.lex_pair_def lex_square_def  split: if_splits)
       done
   next
     show "?Q (?f 0)"
@@ -4667,7 +4667,7 @@
 
 (*
 lemma F_abc_halt_eq:
-  "\<lbrakk>turing_basic.t_correct tp; 
+  "\<lbrakk>Turing.t_correct tp; 
     length lm = k;
     steps (Suc 0, Bk\<up>(l), <lm>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(n));
     rs > 0\<rbrakk>
@@ -4701,7 +4701,7 @@
 declare tape_of_nl_abv_cons[simp del]
 
 lemma t_utm_halt_eq': 
-  "\<lbrakk>turing_basic.t_correct tp;
+  "\<lbrakk>Turing.t_correct tp;
    0 < rs;
   steps (Suc 0, Bk\<up>(l), <lm::nat list>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(n))\<rbrakk>
   \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp = 
@@ -4729,7 +4729,7 @@
 *)
 (*
 lemma t_utm_halt_eq'': 
-  "\<lbrakk>turing_basic.t_correct tp;
+  "\<lbrakk>Turing.t_correct tp;
    0 < rs;
    steps (Suc 0, Bk\<up>(l), <lm::nat list>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(n))\<rbrakk>
   \<Longrightarrow>  \<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>(i)) t_utm stp = 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/thys/Uncomputable.thy	Sun Feb 10 19:49:07 2013 +0000
@@ -0,0 +1,1178 @@
+(* Title: Turing machine's definition and its charater
+   Author: XuJian <xujian817@hotmail.com>
+   Maintainer: Xujian
+*)
+
+header {* Undeciablity of the {\em Halting problem} *}
+
+theory Uncomputable
+imports Turing_Hoare
+begin
+
+lemma numeral:
+  shows "1 = Suc 0"
+  and "2 = Suc 1"
+  and "3 = Suc 2"
+  and "4 = Suc 3" 
+  and "5 = Suc 4" 
+  and "6 = Suc 5" 
+  and "7 = Suc 6"
+  and "8 = Suc 7" 
+  and "9 = Suc 8" 
+  and "10 = Suc 9" 
+  by arith+
+
+text {*
+  The {\em Copying} TM, which duplicates its input. 
+*}
+
+definition 
+  tcopy_begin :: "instr list"
+where
+  "tcopy_begin \<equiv> [(W0, 0), (R, 2), (R, 3), (R, 2),
+                 (W1, 3), (L, 4), (L, 4), (L, 0)]"
+
+definition 
+  tcopy_loop :: "instr list"
+where
+  "tcopy_loop \<equiv> [(R, 0), (R, 2),  (R, 3), (W0, 2),
+                 (R, 3), (R, 4), (W1, 5), (R, 4),
+                 (L, 6), (L, 5), (L, 6), (L, 1)]"
+
+definition 
+  tcopy_end :: "instr list"
+where
+  "tcopy_end \<equiv> [(L, 0), (R, 2), (W1, 3), (L, 4),
+                (R, 2), (R, 2), (L, 5), (W0, 4),
+                (R, 0), (L, 5)]"
+
+definition 
+  tcopy :: "instr list"
+where
+  "tcopy \<equiv> (tcopy_begin |+| tcopy_loop) |+| tcopy_end"
+
+
+(* tcopy_begin *)
+
+fun 
+  inv_begin0 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_begin1 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_begin2 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_begin3 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_begin4 :: "nat \<Rightarrow> tape \<Rightarrow> bool"
+where
+  "inv_begin0 n (l, r) = ((n > 1 \<and> (l, r) = (Oc \<up> (n - 2), [Oc, Oc, Bk, Oc])) \<or>   
+                          (n = 1 \<and> (l, r) = ([], [Bk, Oc, Bk, Oc])))"
+| "inv_begin1 n (l, r) = ((l, r) = ([], Oc \<up> n))"
+| "inv_begin2 n (l, r) = (\<exists> i j. i > 0 \<and> i + j = n \<and> (l, r) = (Oc \<up> i, Oc \<up> j))"
+| "inv_begin3 n (l, r) = (n > 0 \<and> (l, tl r) = (Bk # Oc \<up> n, []))"
+| "inv_begin4 n (l, r) = (n > 0 \<and> (l, r) = (Oc \<up> n, [Bk, Oc]) \<or> (l, r) = (Oc \<up> (n - 1), [Oc, Bk, Oc]))"
+
+fun inv_begin :: "nat \<Rightarrow> config \<Rightarrow> bool"
+  where
+  "inv_begin n (s, tp) = 
+        (if s = 0 then inv_begin0 n tp else
+         if s = 1 then inv_begin1 n tp else
+         if s = 2 then inv_begin2 n tp else
+         if s = 3 then inv_begin3 n tp else
+         if s = 4 then inv_begin4 n tp 
+         else False)"
+
+lemma [elim]: "\<lbrakk>0 < i; 0 < j\<rbrakk> \<Longrightarrow> 
+  \<exists>ia>0. ia + j - Suc 0 = i + j \<and> Oc # Oc \<up> i = Oc \<up> ia"
+by (rule_tac x = "Suc i" in exI, simp)
+
+lemma inv_begin_step: 
+  assumes "inv_begin n cf"
+  and "n > 0"
+  shows "inv_begin n (step0 cf tcopy_begin)"
+using assms
+unfolding tcopy_begin_def
+apply(case_tac cf)
+apply(auto simp: numeral split: if_splits)
+apply(case_tac "hd c")
+apply(auto)
+apply(case_tac c)
+apply(simp_all)
+done
+
+lemma inv_begin_steps: 
+  assumes "inv_begin n cf"
+  and "n > 0"
+  shows "inv_begin n (steps0 cf tcopy_begin stp)"
+apply(induct stp)
+apply(simp add: assms)
+apply(auto simp del: steps.simps)
+apply(rule_tac inv_begin_step)
+apply(simp_all add: assms)
+done
+
+lemma begin_partial_correctness:
+  assumes "is_final (steps0 (1, [], Oc \<up> n) tcopy_begin stp)"
+  shows "0 < n \<Longrightarrow> {inv_begin1 n} tcopy_begin {inv_begin0 n}"
+proof(rule_tac Hoare_haltI)
+  fix l r
+  assume h: "0 < n" "inv_begin1 n (l, r)"
+  have "inv_begin n (steps0 (1, [], Oc \<up> n) tcopy_begin stp)"
+    using h by (rule_tac inv_begin_steps) (simp_all add: inv_begin.simps)
+  then show
+    "\<exists>stp. is_final (steps0 (1, l, r) tcopy_begin stp) \<and> 
+    inv_begin0 n holds_for steps (1, l, r) (tcopy_begin, 0) stp"
+    using h assms
+    apply(rule_tac x = stp in exI)
+    apply(case_tac "(steps0 (1, [], Oc \<up> n) tcopy_begin stp)", simp add: inv_begin.simps)
+    done
+qed
+
+fun measure_begin_state :: "config \<Rightarrow> nat"
+  where
+  "measure_begin_state (s, l, r) = (if s = 0 then 0 else 5 - s)"
+
+fun measure_begin_step :: "config \<Rightarrow> nat"
+  where
+  "measure_begin_step (s, l, r) = 
+        (if s = 2 then length r else
+         if s = 3 then (if r = [] \<or> r = [Bk] then 1 else 0) else
+         if s = 4 then length l 
+         else 0)"
+
+definition
+  "measure_begin = measures [measure_begin_state, measure_begin_step]"
+
+lemma wf_measure_begin:
+  shows "wf measure_begin" 
+unfolding measure_begin_def 
+by auto
+
+lemma measure_begin_induct [case_names Step]: 
+  "\<lbrakk>\<And>n. \<not> P (f n) \<Longrightarrow> (f (Suc n), (f n)) \<in> measure_begin\<rbrakk> \<Longrightarrow> \<exists>n. P (f n)"
+using wf_measure_begin
+by (metis wf_iff_no_infinite_down_chain)
+
+lemma begin_halts: 
+  assumes h: "x > 0"
+  shows "\<exists> stp. is_final (steps0 (1, [], Oc \<up> x) tcopy_begin stp)"
+proof (induct rule: measure_begin_induct) 
+  case (Step n)
+  have "\<not> is_final (steps0 (1, [], Oc \<up> x) tcopy_begin n)" by fact
+  moreover
+  have "inv_begin x (steps0 (1, [], Oc \<up> x) tcopy_begin n)"
+    by (rule_tac inv_begin_steps) (simp_all add: inv_begin.simps h)
+  moreover
+  obtain s l r where eq: "(steps0 (1, [], Oc \<up> x) tcopy_begin n) = (s, l, r)"
+    by (metis measure_begin_state.cases)
+  ultimately 
+  have "(step0 (s, l, r) tcopy_begin, s, l, r) \<in> measure_begin"
+    apply(auto simp: measure_begin_def tcopy_begin_def numeral split: if_splits)
+    apply(subgoal_tac "r = [Oc]")
+    apply(auto)
+    by (metis cell.exhaust list.exhaust tl.simps(2))
+  then 
+  show "(steps0 (1, [], Oc \<up> x) tcopy_begin (Suc n), steps0 (1, [], Oc \<up> x) tcopy_begin n) \<in> measure_begin"
+    using eq by (simp only: step_red)
+qed
+
+lemma begin_correct: 
+  shows "0 < n \<Longrightarrow> {inv_begin1 n} tcopy_begin {inv_begin0 n}"
+using begin_partial_correctness begin_halts by blast
+
+declare tm_comp.simps [simp del] 
+declare adjust.simps[simp del] 
+declare shift.simps[simp del]
+declare tm_wf.simps[simp del]
+declare step.simps[simp del]
+declare steps.simps[simp del]
+
+(* tcopy_loop *)
+
+fun 
+  inv_loop1_loop :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_loop1_exit :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_loop5_loop :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_loop5_exit :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_loop6_loop :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_loop6_exit :: "nat \<Rightarrow> tape \<Rightarrow> bool"
+where
+  "inv_loop1_loop n (l, r) = (\<exists> i j. i + j + 1 = n \<and> (l, r) = (Oc\<up>i, Oc#Oc#Bk\<up>j @ Oc\<up>j) \<and> j > 0)"
+| "inv_loop1_exit n (l, r) = (0 < n \<and> (l, r) = ([], Bk#Oc#Bk\<up>n @ Oc\<up>n))"
+| "inv_loop5_loop x (l, r) = 
+     (\<exists> i j k t. i + j = Suc x \<and> i > 0 \<and> j > 0 \<and> k + t = j \<and> t > 0 \<and> (l, r) = (Oc\<up>k@Bk\<up>j@Oc\<up>i, Oc\<up>t))"
+| "inv_loop5_exit x (l, r) = 
+     (\<exists> i j. i + j = Suc x \<and> i > 0 \<and> j > 0 \<and> (l, r) = (Bk\<up>(j - 1)@Oc\<up>i, Bk # Oc\<up>j))"
+| "inv_loop6_loop x (l, r) = 
+     (\<exists> i j k t. i + j = Suc x \<and> i > 0 \<and> k + t + 1 = j \<and> (l, r) = (Bk\<up>k @ Oc\<up>i, Bk\<up>(Suc t) @ Oc\<up>j))"
+| "inv_loop6_exit x (l, r) = 
+     (\<exists> i j. i + j = x \<and> j > 0 \<and> (l, r) = (Oc\<up>i, Oc#Bk\<up>j @ Oc\<up>j))"
+
+fun 
+  inv_loop0 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_loop1 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_loop2 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_loop3 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_loop4 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_loop5 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_loop6 :: "nat \<Rightarrow> tape \<Rightarrow> bool"
+where
+  "inv_loop0 n (l, r) =  (0 < n \<and> (l, r) = ([Bk], Oc # Bk\<up>n @ Oc\<up>n))"
+| "inv_loop1 n (l, r) = (inv_loop1_loop n (l, r) \<or> inv_loop1_exit n (l, r))"
+| "inv_loop2 n (l, r) = (\<exists> i j any. i + j = n \<and> n > 0 \<and> i > 0 \<and> j > 0 \<and> (l, r) = (Oc\<up>i, any#Bk\<up>j@Oc\<up>j))"
+| "inv_loop3 n (l, r) = 
+     (\<exists> i j k t. i + j = n \<and> i > 0 \<and> j > 0 \<and>  k + t = Suc j \<and> (l, r) = (Bk\<up>k@Oc\<up>i, Bk\<up>t@Oc\<up>j))"
+| "inv_loop4 n (l, r) = 
+     (\<exists> i j k t. i + j = n \<and> i > 0 \<and> j > 0 \<and>  k + t = j \<and> (l, r) = (Oc\<up>k @ Bk\<up>(Suc j)@Oc\<up>i, Oc\<up>t))"
+| "inv_loop5 n (l, r) = (inv_loop5_loop n (l, r) \<or> inv_loop5_exit n (l, r))"
+| "inv_loop6 n (l, r) = (inv_loop6_loop n (l, r) \<or> inv_loop6_exit n (l, r))"
+
+fun inv_loop :: "nat \<Rightarrow> config \<Rightarrow> bool"
+  where
+  "inv_loop x (s, l, r) = 
+         (if s = 0 then inv_loop0 x (l, r)
+          else if s = 1 then inv_loop1 x (l, r)
+          else if s = 2 then inv_loop2 x (l, r)
+          else if s = 3 then inv_loop3 x (l, r)
+          else if s = 4 then inv_loop4 x (l, r)
+          else if s = 5 then inv_loop5 x (l, r)
+          else if s = 6 then inv_loop6 x (l, r)
+          else False)"
+       
+declare inv_loop.simps[simp del] inv_loop1.simps[simp del]
+        inv_loop2.simps[simp del] inv_loop3.simps[simp del] 
+        inv_loop4.simps[simp del] inv_loop5.simps[simp del] 
+        inv_loop6.simps[simp del]
+
+lemma [elim]: "Bk # list = Oc \<up> t \<Longrightarrow> RR"
+by (case_tac t, auto)
+
+lemma [simp]: "inv_loop1 x (b, []) = False"
+by (simp add: inv_loop1.simps)
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop2 x (b, [])\<rbrakk> \<Longrightarrow> inv_loop3 x (Bk # b, [])"
+by (auto simp: inv_loop2.simps inv_loop3.simps)
+
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop3 x (b, [])\<rbrakk> \<Longrightarrow> inv_loop3 x (Bk # b, [])"
+by (auto simp: inv_loop3.simps)
+
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop4 x (b, [])\<rbrakk> \<Longrightarrow> inv_loop5 x (b, [Oc])"
+apply(auto simp: inv_loop4.simps inv_loop5.simps)
+apply(rule_tac [!] x = i in exI, 
+      rule_tac [!] x  = "Suc j" in exI, simp_all)
+done
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop5 x ([], [])\<rbrakk> \<Longrightarrow> RR"
+by (auto simp: inv_loop4.simps inv_loop5.simps)
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop5 x (b, []); b \<noteq> []\<rbrakk> \<Longrightarrow> RR"
+by (auto simp: inv_loop4.simps inv_loop5.simps)
+
+lemma [elim]: "inv_loop6 x ([], []) \<Longrightarrow> RR"
+by (auto simp: inv_loop6.simps)
+
+lemma [elim]: "inv_loop6 x (b, []) \<Longrightarrow> RR" 
+by (auto simp: inv_loop6.simps)
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop1 x (b, Bk # list)\<rbrakk> \<Longrightarrow> b = []"
+by (auto simp: inv_loop1.simps)
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop1 x (b, Bk # list)\<rbrakk> \<Longrightarrow> list = Oc # Bk \<up> x @ Oc \<up> x"
+by (auto simp: inv_loop1.simps)
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop2 x (b, Bk # list)\<rbrakk> \<Longrightarrow> inv_loop3 x (Bk # b, list)"
+apply(auto simp: inv_loop2.simps inv_loop3.simps)
+apply(rule_tac [!] x = i  in exI, rule_tac [!] x = j in exI, simp_all)
+done
+
+lemma [elim]: "Bk # list = Oc \<up> j \<Longrightarrow> RR"
+by (case_tac j, simp_all)
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop3 x (b, Bk # list)\<rbrakk> \<Longrightarrow> inv_loop3 x (Bk # b, list)"
+apply(auto simp: inv_loop3.simps)
+apply(rule_tac [!] x = i in exI, 
+      rule_tac [!] x = j in exI, simp_all)
+apply(case_tac [!] t, auto)
+done
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop4 x (b, Bk # list)\<rbrakk> \<Longrightarrow> inv_loop5 x (b, Oc # list)"
+by (auto simp: inv_loop4.simps inv_loop5.simps)
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop5 x ([], Bk # list)\<rbrakk> \<Longrightarrow> inv_loop6 x ([], Bk # Bk # list)"
+by (auto simp: inv_loop6.simps inv_loop5.simps)
+
+lemma [simp]: "inv_loop5_loop x (b, Bk # list) = False"
+by (auto simp: inv_loop5.simps)
+
+lemma [simp]: "inv_loop6_exit x (b, Bk # list) = False"
+by (auto simp: inv_loop6.simps)
+
+declare inv_loop5_loop.simps[simp del]  inv_loop5_exit.simps[simp del]
+       inv_loop6_loop.simps[simp del]  inv_loop6_exit.simps[simp del]
+
+lemma [elim]:"\<lbrakk>0 < x; inv_loop5_exit x (b, Bk # list); b \<noteq> []; hd b = Bk\<rbrakk> 
+          \<Longrightarrow> inv_loop6_loop x (tl b, Bk # Bk # list)"
+apply(simp only: inv_loop5_exit.simps inv_loop6_loop.simps )
+apply(erule_tac exE)+
+apply(rule_tac x = i in exI, 
+      rule_tac x = j in exI,
+      rule_tac x = "j - Suc (Suc 0)" in exI, 
+      rule_tac x = "Suc 0" in exI, auto)
+apply(case_tac [!] j, simp_all)
+apply(case_tac [!] nat, simp_all)
+done
+
+lemma [simp]: "inv_loop6_loop x (b, Oc # Bk # list) = False"
+by (auto simp: inv_loop6_loop.simps)
+
+lemma [elim]: "\<lbrakk>x > 0; inv_loop5_exit x (b, Bk # list); b \<noteq> []; hd b = Oc\<rbrakk> \<Longrightarrow> 
+  inv_loop6_exit x (tl b, Oc # Bk # list)"
+apply(simp only: inv_loop5_exit.simps inv_loop6_exit.simps)
+apply(erule_tac exE)+
+apply(rule_tac x = "x - 1" in exI, rule_tac x = 1 in exI, simp)
+apply(case_tac j, auto)
+apply(case_tac [!] nat, auto)
+done
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop5 x (b, Bk # list); b \<noteq> []\<rbrakk> \<Longrightarrow> inv_loop6 x (tl b, hd b # Bk # list)"
+apply(simp add: inv_loop5.simps inv_loop6.simps)
+apply(case_tac "hd b", simp_all, auto)
+done
+
+lemma  [simp]: "inv_loop6 x ([], Bk # xs) = False"
+apply(simp add: inv_loop6.simps inv_loop6_loop.simps
+                inv_loop6_exit.simps)
+apply(auto)
+done
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop6 x ([], Bk # list)\<rbrakk> \<Longrightarrow> inv_loop6 x ([], Bk # Bk # list)"
+by (simp)
+
+lemma [simp]: "inv_loop6_exit x (b, Bk # list) = False"
+by (simp add: inv_loop6_exit.simps)
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop6_loop x (b, Bk # list); b \<noteq> []; hd b = Bk\<rbrakk>
+              \<Longrightarrow> inv_loop6_loop x (tl b, Bk # Bk # list)"
+apply(simp only: inv_loop6_loop.simps)
+apply(erule_tac exE)+
+apply(rule_tac x = i in exI, rule_tac x = j in exI, 
+      rule_tac x = "k - 1" in exI, rule_tac x = "Suc t" in exI, auto)
+apply(case_tac [!] k, auto)
+done
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop6_loop x (b, Bk # list); b \<noteq> []; hd b = Oc\<rbrakk> 
+        \<Longrightarrow> inv_loop6_exit x (tl b, Oc # Bk # list)"
+apply(simp only: inv_loop6_loop.simps inv_loop6_exit.simps)
+apply(erule_tac exE)+
+apply(rule_tac x = "i - 1" in exI, rule_tac x = j in exI, auto)
+apply(case_tac [!] k, auto)
+done
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop6 x (b, Bk # list); b \<noteq> []\<rbrakk> \<Longrightarrow> inv_loop6 x (tl b, hd b # Bk # list)"
+apply(simp add: inv_loop6.simps)
+apply(case_tac "hd b", simp_all, auto)
+done
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop1 x (b, Oc # list)\<rbrakk> \<Longrightarrow> inv_loop2 x (Oc # b, list)"
+apply(auto simp: inv_loop1.simps inv_loop2.simps)
+apply(rule_tac x = "Suc i" in exI, auto)
+done
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop2 x (b, Oc # list)\<rbrakk> \<Longrightarrow> inv_loop2 x (b, Bk # list)"
+by (auto simp: inv_loop2.simps)
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop3 x (b, Oc # list)\<rbrakk> \<Longrightarrow> inv_loop4 x (Oc # b, list)"
+apply(auto simp: inv_loop3.simps inv_loop4.simps)
+apply(rule_tac [!] x = i in exI, auto)
+apply(rule_tac [!] x = "Suc 0" in exI, rule_tac [!] x = "j - 1" in exI, auto)
+apply(case_tac [!] t, auto)
+apply(case_tac [!] j, auto)
+done
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop4 x (b, Oc # list)\<rbrakk> \<Longrightarrow> inv_loop4 x (Oc # b, list)"
+apply(auto simp: inv_loop4.simps)
+apply(rule_tac [!] x = "i" in exI, auto)
+apply(rule_tac [!] x = "Suc k" in exI, rule_tac [!] x = "t - 1" in exI, auto)
+apply(case_tac [!] t, simp_all)
+done
+
+lemma [simp]: "inv_loop5 x ([], list) = False"
+by (auto simp: inv_loop5.simps inv_loop5_exit.simps inv_loop5_loop.simps)
+
+lemma [simp]: "inv_loop5_exit x (b, Oc # list) = False"
+by (auto simp: inv_loop5_exit.simps)
+
+lemma [elim]: " \<lbrakk>inv_loop5_loop x (b, Oc # list); b \<noteq> []; hd b = Bk\<rbrakk>
+  \<Longrightarrow> inv_loop5_exit x (tl b, Bk # Oc # list)"
+apply(simp only: inv_loop5_loop.simps inv_loop5_exit.simps)
+apply(erule_tac exE)+
+apply(rule_tac x = i in exI, auto)
+apply(case_tac [!] k, auto)
+done
+
+lemma [elim]: "\<lbrakk>inv_loop5_loop x (b, Oc # list); b \<noteq> []; hd b = Oc\<rbrakk> 
+           \<Longrightarrow> inv_loop5_loop x (tl b, Oc # Oc # list)"
+apply(simp only:  inv_loop5_loop.simps)
+apply(erule_tac exE)+
+apply(rule_tac x = i in exI, rule_tac x = j in exI)
+apply(rule_tac x = "k - 1" in exI, rule_tac x = "Suc t" in exI, auto)
+apply(case_tac [!] k, auto)
+done
+
+lemma [elim]: "\<lbrakk>inv_loop5 x (b, Oc # list); b \<noteq> []\<rbrakk> \<Longrightarrow> inv_loop5 x (tl b, hd b # Oc # list)"
+apply(simp add: inv_loop5.simps)
+apply(case_tac "hd b", simp_all, auto)
+done
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop6 x ([], Oc # list)\<rbrakk> \<Longrightarrow> inv_loop1 x ([], Bk # Oc # list)"
+apply(auto simp: inv_loop6.simps inv_loop1.simps 
+  inv_loop6_loop.simps inv_loop6_exit.simps)
+done
+
+lemma [elim]: "\<lbrakk>0 < x; inv_loop6 x (b, Oc # list); b \<noteq> []\<rbrakk> 
+           \<Longrightarrow> inv_loop1 x (tl b, hd b # Oc # list)"
+apply(auto simp: inv_loop6.simps inv_loop1.simps 
+  inv_loop6_loop.simps inv_loop6_exit.simps)
+done
+
+lemma inv_loop_step: 
+  "\<lbrakk>inv_loop x cf; x > 0\<rbrakk> \<Longrightarrow> inv_loop x (step cf (tcopy_loop, 0))"
+apply(case_tac cf, case_tac c, case_tac [2] aa)
+apply(auto simp: inv_loop.simps step.simps tcopy_loop_def numeral split: if_splits)
+done
+
+lemma inv_loop_steps:
+  "\<lbrakk>inv_loop x cf; x > 0\<rbrakk> \<Longrightarrow> inv_loop x (steps cf (tcopy_loop, 0) stp)"
+apply(induct stp, simp add: steps.simps, simp)
+apply(erule_tac inv_loop_step, simp)
+done
+
+fun loop_stage :: "config \<Rightarrow> nat"
+  where
+  "loop_stage (s, l, r) = (if s = 0 then 0
+                           else (Suc (length (takeWhile (\<lambda>a. a = Oc) (rev l @ r)))))"
+
+fun loop_state :: "config \<Rightarrow> nat"
+  where
+  "loop_state (s, l, r) = (if s = 2 \<and> hd r = Oc then 0
+                           else if s = 1 then 1
+                           else 10 - s)"
+
+fun loop_step :: "config \<Rightarrow> nat"
+  where
+  "loop_step (s, l, r) = (if s = 3 then length r
+                          else if s = 4 then length r
+                          else if s = 5 then length l 
+                          else if s = 6 then length l
+                          else 0)"
+
+definition measure_loop :: "(config \<times> config) set"
+  where
+   "measure_loop = measures [loop_stage, loop_state, loop_step]"
+
+lemma wf_measure_loop: "wf measure_loop"
+unfolding measure_loop_def
+by (auto)
+
+lemma measure_loop_induct [case_names Step]: 
+  "\<lbrakk>\<And>n. \<not> P (f n) \<Longrightarrow> (f (Suc n), (f n)) \<in> measure_loop\<rbrakk> \<Longrightarrow> \<exists>n. P (f n)"
+using wf_measure_loop
+by (metis wf_iff_no_infinite_down_chain)
+
+
+
+lemma [simp]: "inv_loop2 x ([], b) = False"
+by (auto simp: inv_loop2.simps)
+
+lemma  [simp]: "inv_loop2 x (l', []) = False"
+by (auto simp: inv_loop2.simps)
+
+lemma [simp]: "inv_loop3 x (b, []) = False"
+by (auto simp: inv_loop3.simps)
+
+lemma [simp]: "inv_loop4 x ([], b) = False"
+by (auto simp: inv_loop4.simps)
+
+
+lemma [elim]: 
+  "\<lbrakk>inv_loop4 x (l', []); l' \<noteq> []; x > 0;
+  length (takeWhile (\<lambda>a. a = Oc) (rev l' @ [Oc])) \<noteq> 
+  length (takeWhile (\<lambda>a. a = Oc) (rev l'))\<rbrakk>
+  \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (rev l' @ [Oc])) < length (takeWhile (\<lambda>a. a = Oc) (rev l'))"
+apply(auto simp: inv_loop4.simps)
+apply(case_tac [!] j, simp_all add: List.takeWhile_tail)
+done
+
+
+lemma [elim]: 
+  "\<lbrakk>inv_loop4 x (l', Bk # list); l' \<noteq> []; 0 < x;
+   length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Oc # list)) \<noteq> 
+    length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Bk # list))\<rbrakk>
+    \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Oc # list)) < 
+    length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Bk # list))"
+by (auto simp: inv_loop4.simps)
+
+lemma takeWhile_replicate_append: 
+  "P a \<Longrightarrow> takeWhile P (a\<up>x @ ys) = a\<up>x @ takeWhile P ys"
+by (induct x, auto)
+
+lemma takeWhile_replicate: 
+  "P a \<Longrightarrow> takeWhile P (a\<up>x) = a\<up>x"
+by (induct x, auto)
+
+lemma [elim]: 
+   "\<lbrakk>inv_loop5 x (l', Bk # list); l' \<noteq> []; 0 < x;
+   length (takeWhile (\<lambda>a. a = Oc) (rev (tl l') @ hd l' # Bk # list)) \<noteq>
+   length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Bk # list))\<rbrakk>
+   \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (rev (tl l') @ hd l' # Bk # list)) < 
+   length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Bk # list))"
+apply(auto simp: inv_loop5.simps inv_loop5_exit.simps)
+apply(case_tac [!] j, simp_all)
+apply(case_tac [!] "nat", simp_all)
+apply(case_tac  nata, simp_all add: List.takeWhile_tail)
+apply(simp add: takeWhile_replicate_append takeWhile_replicate)
+apply(case_tac  nata, simp_all add: List.takeWhile_tail)
+done
+
+lemma [elim]: "\<lbrakk>inv_loop1 x (l', Oc # list)\<rbrakk> \<Longrightarrow> hd list = Oc"
+by (auto simp: inv_loop1.simps)
+
+lemma [elim]: 
+  "\<lbrakk>inv_loop6 x (l', Bk # list); l' \<noteq> []; 0 < x;
+  length (takeWhile (\<lambda>a. a = Oc) (rev (tl l') @ hd l' # Bk # list)) \<noteq> 
+  length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Bk # list))\<rbrakk>
+  \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (rev (tl l') @ hd l' # Bk # list)) < 
+  length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Bk # list))"
+apply(auto simp: inv_loop6.simps)
+apply(case_tac l', simp_all)
+done
+
+lemma [elim]:
+  "\<lbrakk>inv_loop2 x (l', Oc # list); l' \<noteq> []; 0 < x\<rbrakk> \<Longrightarrow> 
+  length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Bk # list)) <
+  length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Oc # list))"
+apply(auto simp: inv_loop2.simps)
+apply(simp_all add: takeWhile_tail takeWhile_replicate_append
+                takeWhile_replicate)
+done
+
+lemma [elim]: 
+  "\<lbrakk>inv_loop5 x (l', Oc # list); l' \<noteq> []; 0 < x;
+  length (takeWhile (\<lambda>a. a = Oc) (rev (tl l') @ hd l' # Oc # list)) \<noteq> 
+  length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Oc # list))\<rbrakk>
+  \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (rev (tl l') @ hd l' # Oc # list)) < 
+  length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Oc # list))"
+apply(auto simp: inv_loop5.simps)
+apply(case_tac l', auto)
+done
+
+lemma[elim]: 
+  "\<lbrakk>inv_loop6 x (l', Oc # list); l' \<noteq> []; 0 < x;
+  length (takeWhile (\<lambda>a. a = Oc) (rev (tl l') @ hd l' # Oc # list))
+  \<noteq> length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Oc # list))\<rbrakk>
+  \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (rev (tl l') @ hd l' # Oc # list)) < 
+      length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Oc # list))"
+apply(case_tac l')
+apply(auto simp: inv_loop6.simps)
+done
+
+lemma loop_halts: 
+  assumes h: "n > 0" "inv_loop n (1, l, r)"
+  shows "\<exists> stp. is_final (steps0 (1, l, r) tcopy_loop stp)"
+proof (induct rule: measure_loop_induct) 
+  case (Step stp)
+  have "\<not> is_final (steps0 (1, l, r) tcopy_loop stp)" by fact
+  moreover
+  have "inv_loop n (steps0 (1, l, r) tcopy_loop stp)"
+    by (rule_tac inv_loop_steps) (simp_all only: h)
+  moreover
+  obtain s l' r' where eq: "(steps0 (1, l, r) tcopy_loop stp) = (s, l', r')"
+    by (metis measure_begin_state.cases)
+  ultimately 
+  have "(step0 (s, l', r') tcopy_loop, s, l', r') \<in> measure_loop"
+    using h(1)
+    apply(case_tac r')
+    apply(case_tac [2] a)
+    apply(auto simp: inv_loop.simps step.simps tcopy_loop_def numeral measure_loop_def split: if_splits)
+    done
+  then 
+  show "(steps0 (1, l, r) tcopy_loop (Suc stp), steps0 (1, l, r) tcopy_loop stp) \<in> measure_loop"
+    using eq by (simp only: step_red)
+qed
+
+lemma loop_correct:
+  shows "0 < n \<Longrightarrow> {inv_loop1 n} tcopy_loop {inv_loop0 n}"
+  using assms
+proof(rule_tac Hoare_haltI)
+  fix l r
+  assume h: "0 < n" "inv_loop1 n (l, r)"
+  then obtain stp where k: "is_final (steps0 (1, l, r) tcopy_loop stp)" 
+    using loop_halts
+    apply(simp add: inv_loop.simps)
+    apply(blast)
+    done
+  moreover
+  have "inv_loop n (steps0 (1, l, r) tcopy_loop stp)"
+    using h 
+    by (rule_tac inv_loop_steps) (simp_all add: inv_loop.simps)
+  ultimately show
+    "\<exists>stp. is_final (steps0 (1, l, r) tcopy_loop stp) \<and> 
+    inv_loop0 n holds_for steps0 (1, l, r) tcopy_loop stp"
+    using h(1) 
+    apply(rule_tac x = stp in exI)
+    apply(case_tac "(steps0 (1, l, r) tcopy_loop stp)")
+    apply(simp add: inv_loop.simps)
+    done
+qed
+
+
+
+
+(* tcopy_end *)
+
+fun
+  inv_end5_loop :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_end5_exit :: "nat \<Rightarrow> tape \<Rightarrow> bool" 
+where  
+  "inv_end5_loop x (l, r) = 
+     (\<exists> i j. i + j = x \<and> x > 0 \<and> j > 0 \<and> l = Oc\<up>i @ [Bk] \<and> r = Oc\<up>j @ Bk # Oc\<up>x)"
+| "inv_end5_exit x (l, r) = (x > 0 \<and> l = [] \<and> r = Bk # Oc\<up>x @ Bk # Oc\<up>x)"
+
+fun 
+  inv_end0 :: "nat \<Rightarrow> tape \<Rightarrow>  bool" and
+  inv_end1 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_end2 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_end3 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
+  inv_end4 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and 
+  inv_end5 :: "nat \<Rightarrow> tape \<Rightarrow> bool" 
+where
+  "inv_end0 n (l, r) = (n > 0 \<and> (l, r) = ([Bk], Oc\<up>n @ Bk # Oc\<up>n))"
+| "inv_end1 n (l, r) = (n > 0 \<and> (l, r) = ([Bk], Oc # Bk\<up>n @ Oc\<up>n))"
+| "inv_end2 n (l, r) = (\<exists> i j. i + j = Suc n \<and> n > 0 \<and> l = Oc\<up>i @ [Bk] \<and> r = Bk\<up>j @ Oc\<up>n)"
+| "inv_end3 n (l, r) =
+     (\<exists> i j. n > 0 \<and> i + j = n \<and> l = Oc\<up>i @ [Bk] \<and> r = Oc # Bk\<up>j@ Oc\<up>n)"
+| "inv_end4 n (l, r) = (\<exists> any. n > 0 \<and> l = Oc\<up>n @ [Bk] \<and> r = any#Oc\<up>n)"
+| "inv_end5 n (l, r) = (inv_end5_loop n (l, r) \<or> inv_end5_exit n (l, r))"
+
+fun 
+  inv_end :: "nat \<Rightarrow> config \<Rightarrow> bool"
+where
+  "inv_end n (s, l, r) = (if s = 0 then inv_end0 n (l, r)
+                          else if s = 1 then inv_end1 n (l, r)
+                          else if s = 2 then inv_end2 n (l, r)
+                          else if s = 3 then inv_end3 n (l, r)
+                          else if s = 4 then inv_end4 n (l, r)
+                          else if s = 5 then inv_end5 n (l, r)
+                          else False)"
+
+declare inv_end.simps[simp del] inv_end1.simps[simp del]
+        inv_end0.simps[simp del] inv_end2.simps[simp del]
+        inv_end3.simps[simp del] inv_end4.simps[simp del]
+        inv_end5.simps[simp del]
+
+lemma [simp]:  "inv_end1 x (b, []) = False"
+by (auto simp: inv_end1.simps)
+
+lemma [simp]: "inv_end2 x (b, []) = False"
+by (auto simp: inv_end2.simps)
+
+lemma [simp]: "inv_end3 x (b, []) = False"
+by (auto simp: inv_end3.simps)
+
+lemma [simp]: "inv_end4 x (b, []) = False"
+by (auto simp: inv_end4.simps)
+
+lemma [simp]: "inv_end5 x (b, []) = False"
+by (auto simp: inv_end5.simps)
+
+lemma [simp]: "inv_end1 x ([], list) = False"
+by (auto simp: inv_end1.simps)
+
+lemma [elim]: "\<lbrakk>0 < x; inv_end1 x (b, Bk # list); b \<noteq> []\<rbrakk>
+  \<Longrightarrow> inv_end0 x (tl b, hd b # Bk # list)"
+by (auto simp: inv_end1.simps inv_end0.simps)
+
+lemma [elim]: "\<lbrakk>0 < x; inv_end2 x (b, Bk # list)\<rbrakk> 
+  \<Longrightarrow> inv_end3 x (b, Oc # list)"
+apply(auto simp: inv_end2.simps inv_end3.simps)
+apply(rule_tac x = "j - 1" in exI)
+apply(case_tac j, simp_all)
+apply(case_tac x, simp_all)
+done
+
+lemma [elim]: "\<lbrakk>0 < x; inv_end3 x (b, Bk # list)\<rbrakk> \<Longrightarrow> inv_end2 x (Bk # b, list)"
+by (auto simp: inv_end2.simps inv_end3.simps)
+  
+lemma [elim]: "\<lbrakk>0 < x; inv_end4 x ([], Bk # list)\<rbrakk> \<Longrightarrow> 
+  inv_end5 x ([], Bk # Bk # list)"
+by (auto simp: inv_end4.simps inv_end5.simps)
+ 
+lemma [elim]: "\<lbrakk>0 < x; inv_end4 x (b, Bk # list); b \<noteq> []\<rbrakk> \<Longrightarrow> 
+  inv_end5 x (tl b, hd b # Bk # list)"
+apply(auto simp: inv_end4.simps inv_end5.simps)
+apply(rule_tac x = 1 in exI, simp)
+done
+
+lemma [elim]: "\<lbrakk>0 < x; inv_end5 x (b, Bk # list)\<rbrakk> \<Longrightarrow> inv_end0 x (Bk # b, list)"
+apply(auto simp: inv_end5.simps inv_end0.simps)
+apply(case_tac [!] j, simp_all)
+done
+
+lemma [elim]: "\<lbrakk>0 < x; inv_end1 x (b, Oc # list)\<rbrakk> \<Longrightarrow> inv_end2 x (Oc # b, list)"
+by (auto simp: inv_end1.simps inv_end2.simps)
+
+lemma [elim]: "\<lbrakk>0 < x; inv_end2 x ([], Oc # list)\<rbrakk> \<Longrightarrow>
+               inv_end4 x ([], Bk # Oc # list)"
+by (auto simp: inv_end2.simps inv_end4.simps)
+
+lemma [elim]:  "\<lbrakk>0 < x; inv_end2 x (b, Oc # list); b \<noteq> []\<rbrakk> \<Longrightarrow>
+  inv_end4 x (tl b, hd b # Oc # list)"
+apply(auto simp: inv_end2.simps inv_end4.simps)
+apply(case_tac [!] j, simp_all)
+done
+
+lemma [elim]: "\<lbrakk>0 < x; inv_end3 x (b, Oc # list)\<rbrakk> \<Longrightarrow> inv_end2 x (Oc # b, list)"
+by (auto simp: inv_end2.simps inv_end3.simps)
+
+lemma [elim]: "\<lbrakk>0 < x; inv_end4 x (b, Oc # list)\<rbrakk> \<Longrightarrow> inv_end4 x (b, Bk # list)"
+by (auto simp: inv_end2.simps inv_end4.simps)
+
+lemma [elim]: "\<lbrakk>0 < x; inv_end5 x ([], Oc # list)\<rbrakk> \<Longrightarrow> inv_end5 x ([], Bk # Oc # list)"
+by (auto simp: inv_end2.simps inv_end5.simps)
+
+declare inv_end5_loop.simps[simp del]
+        inv_end5_exit.simps[simp del]
+
+lemma [simp]: "inv_end5_exit x (b, Oc # list) = False"
+by (auto simp: inv_end5_exit.simps)
+
+lemma [simp]: "inv_end5_loop x (tl b, Bk # Oc # list) = False"
+apply(auto simp: inv_end5_loop.simps)
+apply(case_tac [!] j, simp_all)
+done
+
+lemma [elim]:
+  "\<lbrakk>0 < x; inv_end5_loop x (b, Oc # list); b \<noteq> []; hd b = Bk\<rbrakk> \<Longrightarrow>
+  inv_end5_exit x (tl b, Bk # Oc # list)"
+apply(auto simp: inv_end5_loop.simps inv_end5_exit.simps)
+apply(case_tac [!] i, simp_all)
+done
+
+lemma [elim]: 
+  "\<lbrakk>0 < x; inv_end5_loop x (b, Oc # list); b \<noteq> []; hd b = Oc\<rbrakk> \<Longrightarrow>
+  inv_end5_loop x (tl b, Oc # Oc # list)"
+apply(simp only: inv_end5_loop.simps inv_end5_exit.simps)
+apply(erule_tac exE)+
+apply(rule_tac x = "i - 1" in exI, 
+      rule_tac x = "Suc j" in exI, auto)
+apply(case_tac [!] i, simp_all)
+done
+
+lemma [elim]: "\<lbrakk>0 < x; inv_end5 x (b, Oc # list); b \<noteq> []\<rbrakk> \<Longrightarrow> 
+  inv_end5 x (tl b, hd b # Oc # list)"
+apply(simp add: inv_end2.simps inv_end5.simps)
+apply(case_tac "hd b", simp_all, auto)
+done
+
+lemma inv_end_step:
+  "\<lbrakk>x > 0; inv_end x cf\<rbrakk> \<Longrightarrow> inv_end x (step cf (tcopy_end, 0))"
+apply(case_tac cf, case_tac c, case_tac [2] aa)
+apply(auto simp: inv_end.simps step.simps tcopy_end_def numeral split: if_splits)
+done
+
+lemma inv_end_steps:
+  "\<lbrakk>x > 0; inv_end x cf\<rbrakk> \<Longrightarrow> inv_end x (steps cf (tcopy_end, 0) stp)"
+apply(induct stp, simp add:steps.simps, simp)
+apply(erule_tac inv_end_step, simp)
+done
+
+fun end_state :: "config \<Rightarrow> nat"
+  where
+  "end_state (s, l, r) = 
+       (if s = 0 then 0
+        else if s = 1 then 5
+        else if s = 2 \<or> s = 3 then 4
+        else if s = 4 then 3 
+        else if s = 5 then 2
+        else 0)"
+
+fun end_stage :: "config \<Rightarrow> nat"
+  where
+  "end_stage (s, l, r) = 
+          (if s = 2 \<or> s = 3 then (length r) else 0)"
+
+fun end_step :: "config \<Rightarrow> nat"
+  where
+  "end_step (s, l, r) = 
+         (if s = 4 then (if hd r = Oc then 1 else 0)
+          else if s = 5 then length l
+          else if s = 2 then 1
+          else if s = 3 then 0
+          else 0)"
+
+definition end_LE :: "(config \<times> config) set"
+  where
+   "end_LE = measures [end_state, end_stage, end_step]"
+
+lemma wf_end_le: "wf end_LE"
+unfolding end_LE_def
+by (auto)
+
+lemma [simp]: "inv_end5 x ([], Oc # list) = False"
+by (auto simp: inv_end5.simps inv_end5_loop.simps)
+
+lemma halt_lemma: 
+  "\<lbrakk>wf LE; \<forall>n. (\<not> P (f n) \<longrightarrow> (f (Suc n), (f n)) \<in> LE)\<rbrakk> \<Longrightarrow> \<exists>n. P (f n)"
+by (metis wf_iff_no_infinite_down_chain)
+
+lemma end_halt: 
+  "\<lbrakk>x > 0; inv_end x (Suc 0, l, r)\<rbrakk> \<Longrightarrow> 
+      \<exists> stp. is_final (steps (Suc 0, l, r) (tcopy_end, 0) stp)"
+proof(rule_tac LE = end_LE in halt_lemma)
+  show "wf end_LE" by(intro wf_end_le)
+next
+  assume great: "0 < x"
+    and inv_start: "inv_end x (Suc 0, l, r)"
+  show "\<forall>n. \<not> is_final (steps (Suc 0, l, r) (tcopy_end, 0) n) \<longrightarrow> 
+    (steps (Suc 0, l, r) (tcopy_end, 0) (Suc n), steps (Suc 0, l, r) (tcopy_end, 0) n) \<in> end_LE"
+  proof(rule_tac allI, rule_tac impI)
+    fix n
+    assume notfinal: "\<not> is_final (steps (Suc 0, l, r) (tcopy_end, 0) n)"
+    obtain s' l' r' where d: "steps (Suc 0, l, r) (tcopy_end, 0) n = (s', l', r')"
+      apply(case_tac "steps (Suc 0, l, r) (tcopy_end, 0) n", auto)
+      done
+    hence "inv_end x (s', l', r') \<and> s' \<noteq> 0"
+      using great inv_start notfinal
+      apply(drule_tac stp = n in inv_end_steps, auto)
+      done
+    hence "(step (s', l', r') (tcopy_end, 0), s', l', r') \<in> end_LE"
+      apply(case_tac r', case_tac [2] a)
+      apply(auto simp: inv_end.simps step.simps tcopy_end_def numeral end_LE_def split: if_splits)
+      done
+    thus "(steps (Suc 0, l, r) (tcopy_end, 0) (Suc n), 
+      steps (Suc 0, l, r) (tcopy_end, 0) n) \<in> end_LE"
+      using d
+      by simp
+  qed
+qed
+
+lemma end_correct:
+  "n > 0 \<Longrightarrow> {inv_end1 n} tcopy_end {inv_end0 n}"
+proof(rule_tac Hoare_haltI)
+  fix l r
+  assume h: "0 < n"
+    "inv_end1 n (l, r)"
+  then have "\<exists> stp. is_final (steps0 (1, l, r) tcopy_end stp)"
+    by (simp add: end_halt inv_end.simps)
+  then obtain stp where "is_final (steps0 (1, l, r) tcopy_end stp)" ..
+  moreover have "inv_end n (steps0 (1, l, r) tcopy_end stp)"
+    apply(rule_tac inv_end_steps)
+    using h by(simp_all add: inv_end.simps)
+  ultimately show
+    "\<exists>stp. is_final (steps (1, l, r) (tcopy_end, 0) stp) \<and> 
+    inv_end0 n holds_for steps (1, l, r) (tcopy_end, 0) stp"        
+    using h
+    apply(rule_tac x = stp in exI)
+    apply(case_tac "(steps0 (1, l, r) tcopy_end stp)") 
+    apply(simp add: inv_end.simps)
+    done
+qed
+
+(* tcopy *)
+
+lemma [intro]: "tm_wf (tcopy_begin, 0)"
+by (auto simp: tm_wf.simps tcopy_begin_def)
+
+lemma [intro]: "tm_wf (tcopy_loop, 0)"
+by (auto simp: tm_wf.simps tcopy_loop_def)
+
+lemma [intro]: "tm_wf (tcopy_end, 0)"
+by (auto simp: tm_wf.simps tcopy_end_def)
+
+lemma tcopy_correct1: 
+  assumes "0 < x"
+  shows "{inv_begin1 x} tcopy {inv_end0 x}"
+proof -
+  have "{inv_begin1 x} tcopy_begin {inv_begin0 x}"
+    by (metis assms begin_correct) 
+  moreover 
+  have "inv_begin0 x \<mapsto> inv_loop1 x"
+    unfolding assert_imp_def
+    unfolding inv_begin0.simps inv_loop1.simps
+    unfolding inv_loop1_loop.simps inv_loop1_exit.simps
+    apply(auto simp add: numeral Cons_eq_append_conv)
+    by (rule_tac x = "Suc 0" in exI, auto)
+  ultimately have "{inv_begin1 x} tcopy_begin {inv_loop1 x}"
+    by (rule_tac Hoare_consequence) (auto)
+  moreover
+  have "{inv_loop1 x} tcopy_loop {inv_loop0 x}"
+    by (metis assms loop_correct) 
+  ultimately 
+  have "{inv_begin1 x} (tcopy_begin |+| tcopy_loop) {inv_loop0 x}"
+    by (rule_tac Hoare_plus_halt) (auto)
+  moreover 
+  have "{inv_end1 x} tcopy_end {inv_end0 x}"
+    by (metis assms end_correct) 
+  moreover 
+  have "inv_loop0 x = inv_end1 x"
+    by(auto simp: inv_end1.simps inv_loop1.simps assert_imp_def)
+  ultimately 
+  show "{inv_begin1 x} tcopy {inv_end0 x}"
+    unfolding tcopy_def
+    by (rule_tac Hoare_plus_halt) (auto)
+qed
+
+abbreviation (input)
+  "pre_tcopy n \<equiv> \<lambda>tp. tp = ([]::cell list, <(n::nat)>)"
+abbreviation (input)
+  "post_tcopy n \<equiv> \<lambda>tp. tp= ([Bk], <(n, n::nat)>)"
+
+lemma tcopy_correct:
+  shows "{pre_tcopy n} tcopy {post_tcopy n}"
+proof -
+  have "{inv_begin1 (Suc n)} tcopy {inv_end0 (Suc n)}"
+    by (rule tcopy_correct1) (simp)
+  moreover
+  have "pre_tcopy n = inv_begin1 (Suc n)" 
+    by (auto simp add: tape_of_nl_abv tape_of_nat_abv)
+  moreover
+  have "inv_end0 (Suc n) = post_tcopy n" 
+    by (auto simp add: inv_end0.simps tape_of_nat_abv tape_of_nat_pair)
+  ultimately
+  show "{pre_tcopy n} tcopy {post_tcopy n}" 
+    by simp
+qed
+
+
+section {* The {\em Dithering} Turing Machine *}
+
+text {*
+  The {\em Dithering} TM, when the input is @{text "1"}, it will loop forever, otherwise, it will
+  terminate.
+*}
+
+definition dither :: "instr list"
+  where
+  "dither \<equiv> [(W0, 1), (R, 2), (L, 1), (L, 0)] "
+
+(* invariants of dither *)
+abbreviation (input)
+  "dither_halt_inv \<equiv> \<lambda>tp. \<exists>k. tp = (Bk \<up> k, <1::nat>)"
+
+abbreviation (input)
+  "dither_unhalt_inv \<equiv> \<lambda>tp. \<exists>k. tp = (Bk \<up> k, <0::nat>)"
+
+lemma dither_loops_aux: 
+  "(steps0 (1, Bk \<up> m, [Oc]) dither stp = (1, Bk \<up> m, [Oc])) \<or> 
+   (steps0 (1, Bk \<up> m, [Oc]) dither stp = (2, Oc # Bk \<up> m, []))"
+  apply(induct stp)
+  apply(auto simp: steps.simps step.simps dither_def numeral tape_of_nat_abv)
+  done
+
+lemma dither_loops:
+  shows "{dither_unhalt_inv} dither \<up>" 
+apply(rule Hoare_unhaltI)
+using dither_loops_aux
+apply(auto simp add: numeral tape_of_nat_abv)
+by (metis Suc_neq_Zero is_final_eq)
+
+lemma dither_halts_aux: 
+  shows "steps0 (1, Bk \<up> m, [Oc, Oc]) dither 2 = (0, Bk \<up> m, [Oc, Oc])"
+unfolding dither_def
+by (simp add: steps.simps step.simps numeral)
+
+lemma dither_halts:
+  shows "{dither_halt_inv} dither {dither_halt_inv}" 
+apply(rule Hoare_haltI)
+using dither_halts_aux
+apply(auto simp add: tape_of_nat_abv)
+by (metis (lifting, mono_tags) holds_for.simps is_final_eq prod.cases)
+
+
+section {* The diagnal argument below shows the undecidability of Halting problem *}
+
+text {*
+  @{text "haltP tp x"} means TM @{text "tp"} terminates on input @{text "x"}
+  and the final configuration is standard.
+*}
+
+definition haltP :: "tprog0 \<Rightarrow> nat list \<Rightarrow> bool"
+  where
+  "haltP p ns \<equiv> {(\<lambda>tp. tp = ([], <ns>))} p {(\<lambda>tp. (\<exists>k n l. tp = (Bk \<up> k,  <n::nat> @ Bk \<up> l)))}"
+
+lemma [intro, simp]: "tm_wf0 tcopy"
+by (auto simp: tcopy_def)
+
+lemma [intro, simp]: "tm_wf0 dither"
+by (auto simp: tm_wf.simps dither_def)
+
+text {*
+  The following locale specifies that TM @{text "H"} can be used to solve 
+  the {\em Halting Problem} and @{text "False"} is going to be derived 
+  under this locale. Therefore, the undecidability of {\em Halting Problem}
+  is established. 
+*}
+
+locale uncomputable = 
+  -- {* The coding function of TM, interestingly, the detailed definition of this 
+  funciton @{text "code"} does not affect the final result. *}
+  fixes code :: "instr list \<Rightarrow> nat" 
+  -- {* 
+  The TM @{text "H"} is the one which is assummed being able to solve the Halting problem.
+  *}
+  and H :: "instr list"
+  assumes h_wf[intro]: "tm_wf0 H"
+  -- {*
+  The following two assumptions specifies that @{text "H"} does solve the Halting problem.
+  *}
+  and h_case: 
+  "\<And> M ns. haltP M ns \<Longrightarrow> {(\<lambda>tp. tp = ([Bk], <(code M, ns)>))} H {(\<lambda>tp. \<exists>k. tp = (Bk \<up> k, <0::nat>))}"
+  and nh_case: 
+  "\<And> M ns. \<not> haltP M ns \<Longrightarrow> {(\<lambda>tp. tp = ([Bk], <(code M, ns)>))} H {(\<lambda>tp. \<exists>k. tp = (Bk \<up> k, <1::nat>))}"
+begin
+
+(* invariants for H *)
+abbreviation (input)
+  "pre_H_inv M ns \<equiv> \<lambda>tp. tp = ([Bk], <(code M, ns::nat list)>)"
+
+abbreviation (input)
+  "post_H_halt_inv \<equiv> \<lambda>tp. \<exists>k. tp = (Bk \<up> k, <1::nat>)"
+
+abbreviation (input)
+  "post_H_unhalt_inv \<equiv> \<lambda>tp. \<exists>k. tp = (Bk \<up> k, <0::nat>)"
+
+
+lemma H_halt_inv:
+  assumes "\<not> haltP M ns" 
+  shows "{pre_H_inv M ns} H {post_H_halt_inv}"
+using assms nh_case by auto
+
+lemma H_unhalt_inv:
+  assumes "haltP M ns" 
+  shows "{pre_H_inv M ns} H {post_H_unhalt_inv}"
+using assms h_case by auto
+   
+(* TM that produces the contradiction and its code *)
+
+definition
+  "tcontra \<equiv> (tcopy |+| H) |+| dither"
+abbreviation
+  "code_tcontra \<equiv> code tcontra"
+
+(* assume tcontra does not halt on its code *)
+lemma tcontra_unhalt: 
+  assumes "\<not> haltP tcontra [code tcontra]"
+  shows "False"
+proof -
+  (* invariants *)
+  def P1 \<equiv> "\<lambda>tp. tp = ([]::cell list, <code_tcontra>)"
+  def P2 \<equiv> "\<lambda>tp. tp = ([Bk], <(code_tcontra, code_tcontra)>)"
+  def P3 \<equiv> "\<lambda>tp. \<exists>k. tp = (Bk \<up> k, <1::nat>)"
+
+  (*
+  {P1} tcopy {P2}  {P2} H {P3} 
+  ----------------------------
+     {P1} (tcopy |+| H) {P3}     {P3} dither {P3}
+  ------------------------------------------------
+                 {P1} tcontra {P3}
+  *)
+
+  have H_wf: "tm_wf0 (tcopy |+| H)" by auto
+
+  (* {P1} (tcopy |+| H) {P3} *)
+  have first: "{P1} (tcopy |+| H) {P3}" 
+  proof (cases rule: Hoare_plus_halt)
+    case A_halt (* of tcopy *)
+    show "{P1} tcopy {P2}" unfolding P1_def P2_def 
+      by (rule tcopy_correct)
+  next
+    case B_halt (* of H *)
+    show "{P2} H {P3}"
+      unfolding P2_def P3_def 
+      using H_halt_inv[OF assms]
+      by (simp add: tape_of_nat_pair tape_of_nl_abv)
+  qed (simp)
+
+  (* {P3} dither {P3} *)
+  have second: "{P3} dither {P3}" unfolding P3_def 
+    by (rule dither_halts)
+  
+  (* {P1} tcontra {P3} *)
+  have "{P1} tcontra {P3}" 
+    unfolding tcontra_def
+    by (rule Hoare_plus_halt[OF first second H_wf])
+
+  with assms show "False"
+    unfolding P1_def P3_def
+    unfolding haltP_def
+    unfolding Hoare_halt_def 
+    apply(auto)    
+    apply(drule_tac x = n in spec)
+    apply(case_tac "steps0 (Suc 0, [], <code tcontra>) tcontra n")
+    apply(auto simp add: tape_of_nl_abv)
+    by (metis append_Nil2 replicate_0)
+qed
+
+(* asumme tcontra halts on its code *)
+lemma tcontra_halt: 
+  assumes "haltP tcontra [code tcontra]"
+  shows "False"
+proof - 
+  (* invariants *)
+  def P1 \<equiv> "\<lambda>tp. tp = ([]::cell list, <code_tcontra>)"
+  def P2 \<equiv> "\<lambda>tp. tp = ([Bk], <(code_tcontra, code_tcontra)>)"
+  def Q3 \<equiv> "\<lambda>tp. \<exists>k. tp = (Bk \<up> k, <0::nat>)"
+
+  (*
+  {P1} tcopy {P2}  {P2} H {Q3} 
+  ----------------------------
+     {P1} (tcopy |+| H) {Q3}     {Q3} dither loops
+  ------------------------------------------------
+               {P1} tcontra loops
+  *)
+
+  have H_wf: "tm_wf0 (tcopy |+| H)" by auto
+
+  (* {P1} (tcopy |+| H) {Q3} *)
+  have first: "{P1} (tcopy |+| H) {Q3}" 
+  proof (cases rule: Hoare_plus_halt)
+    case A_halt (* of tcopy *)
+    show "{P1} tcopy {P2}" unfolding P1_def P2_def
+      by (rule tcopy_correct)
+  next
+    case B_halt (* of H *)
+    then show "{P2} H {Q3}"
+      unfolding P2_def Q3_def using H_unhalt_inv[OF assms]
+      by(simp add: tape_of_nat_pair tape_of_nl_abv)
+  qed (simp)
+
+  (* {P3} dither loops *)
+  have second: "{Q3} dither \<up>" unfolding Q3_def 
+    by (rule dither_loops)
+  
+  (* {P1} tcontra loops *)
+  have "{P1} tcontra \<up>" 
+    unfolding tcontra_def
+    by (rule Hoare_plus_unhalt[OF first second H_wf])
+
+  with assms show "False"
+    unfolding P1_def
+    unfolding haltP_def
+    unfolding Hoare_halt_def Hoare_unhalt_def
+    by (auto simp add: tape_of_nl_abv)
+qed
+
+      
+text {*
+  @{text "False"} can finally derived.
+*}
+
+lemma false: "False"
+using tcontra_halt tcontra_unhalt 
+by auto
+
+end
+
+declare replicate_Suc[simp del]
+
+
+end
+
--- a/thys/abacus.thy	Thu Feb 07 06:39:06 2013 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,4974 +0,0 @@
-header {* 
- {\em abacus} a kind of register machine
-*}
-
-theory abacus
-imports uncomputable
-begin
-
-(*
-declare tm_comp.simps [simp add] 
-declare adjust.simps[simp add] 
-declare shift.simps[simp add]
-declare tm_wf.simps[simp add]
-declare step.simps[simp add]
-declare steps.simps[simp add]
-*)
-declare replicate_Suc[simp add]
-
-text {*
-  {\em Abacus} instructions:
-*}
-
-datatype abc_inst =
-  -- {* @{text "Inc n"} increments the memory cell (or register) with address @{text "n"} by one.
-     *}
-     Inc nat
-  -- {*
-     @{text "Dec n label"} decrements the memory cell with address @{text "n"} by one. 
-      If cell @{text "n"} is already zero, no decrements happens and the executio jumps to
-      the instruction labeled by @{text "label"}.
-     *}
-   | Dec nat nat
-  -- {*
-  @{text "Goto label"} unconditionally jumps to the instruction labeled by @{text "label"}.
-  *}
-   | Goto nat
-  
-
-text {*
-  Abacus programs are defined as lists of Abacus instructions.
-*}
-type_synonym abc_prog = "abc_inst list"
-
-section {*
-  Sample Abacus programs
-  *}
-
-text {*
-  Abacus for addition and clearance.
-*}
-fun plus_clear :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"  
-  where
-  "plus_clear m n e = [Dec m e, Inc n, Goto 0]"
-
-text {*
-  Abacus for clearing memory untis.
-*}
-fun clear :: "nat \<Rightarrow> nat \<Rightarrow> abc_prog"
-  where
-  "clear n e = [Dec n e, Goto 0]"
-
-fun plus:: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
-  where
-  "plus m n p e = [Dec m 4, Inc n, Inc p,
-                   Goto 0, Dec p e, Inc m, Goto 4]"
-
-fun mult :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
-  where
-  "mult m1 m2 n p e = [Dec m1 e]@ plus m1 m2 p 1"
-
-fun expo :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
-  where
-  "expo n m1 m2 p e = [Inc n, Dec m1 e] @ mult m2 n n p 2"
-
-
-text {*
-  The state of Abacus machine.
-  *}
-type_synonym abc_state = nat
-
-(* text {*
-  The memory of Abacus machine is defined as a function from address to contents.
-*}
-type_synonym abc_mem = "nat \<Rightarrow> nat" *)
-
-text {*
-  The memory of Abacus machine is defined as a list of contents, with 
-  every units addressed by index into the list.
-  *}
-type_synonym abc_lm = "nat list"
-
-text {*
-  Fetching contents out of memory. Units not represented by list elements are considered
-  as having content @{text "0"}.
-*}
-fun abc_lm_v :: "abc_lm \<Rightarrow> nat \<Rightarrow> nat"
-  where 
-    "abc_lm_v lm n = (if (n < length lm) then (lm!n) else 0)"         
-
-
-text {*
-  Set the content of memory unit @{text "n"} to value @{text "v"}.
-  @{text "am"} is the Abacus memory before setting.
-  If address @{text "n"} is outside to scope of @{text "am"}, @{text "am"} 
-  is extended so that @{text "n"} becomes in scope.
-*}
-
-fun abc_lm_s :: "abc_lm \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_lm"
-  where
-    "abc_lm_s am n v = (if (n < length am) then (am[n:=v]) else 
-                           am@ (replicate (n - length am) 0) @ [v])"
-
-
-text {*
-  The configuration of Abaucs machines consists of its current state and its
-  current memory:
-*}
-type_synonym abc_conf = "abc_state \<times> abc_lm"
-
-text {*
-  Fetch instruction out of Abacus program:
-*}
-
-fun abc_fetch :: "nat \<Rightarrow> abc_prog \<Rightarrow> abc_inst option" 
-  where
-  "abc_fetch s p = (if (s < length p) then Some (p ! s)
-                    else None)"
-
-text {*
-  Single step execution of Abacus machine. If no instruction is feteched, 
-  configuration does not change.
-*}
-fun abc_step_l :: "abc_conf \<Rightarrow> abc_inst option \<Rightarrow> abc_conf"
-  where
-  "abc_step_l (s, lm) a = (case a of 
-               None \<Rightarrow> (s, lm) |
-               Some (Inc n)  \<Rightarrow> (let nv = abc_lm_v lm n in
-                       (s + 1, abc_lm_s lm n (nv + 1))) |
-               Some (Dec n e) \<Rightarrow> (let nv = abc_lm_v lm n in
-                       if (nv = 0) then (e, abc_lm_s lm n 0) 
-                       else (s + 1,  abc_lm_s lm n (nv - 1))) |
-               Some (Goto n) \<Rightarrow> (n, lm) 
-               )"
-
-text {*
-  Multi-step execution of Abacus machine.
-*}
-fun abc_steps_l :: "abc_conf \<Rightarrow> abc_prog \<Rightarrow> nat \<Rightarrow> abc_conf"
-  where
-  "abc_steps_l (s, lm) p 0 = (s, lm)" |
-  "abc_steps_l (s, lm) p (Suc n) = 
-      abc_steps_l (abc_step_l (s, lm) (abc_fetch s p)) p n"
-
-section {*
-  Compiling Abacus machines into Truing machines
-*}
-
-subsection {*
-  Compiling functions
-*}
-
-text {*
-  @{text "findnth n"} returns the TM which locates the represention of
-  memory cell @{text "n"} on the tape and changes representation of zero
-  on the way.
-*}
-
-fun findnth :: "nat \<Rightarrow> instr list"
-  where
-  "findnth 0 = []" |
-  "findnth (Suc n) = (findnth n @ [(W1, 2 * n + 1), 
-           (R, 2 * n + 2), (R, 2 * n + 3), (R, 2 * n + 2)])"
-
-text {*
-  @{text "tinc_b"} returns the TM which increments the representation 
-  of the memory cell under rw-head by one and move the representation 
-  of cells afterwards to the right accordingly.
-  *}
-
-definition tinc_b :: "instr list"
-  where
-  "tinc_b \<equiv> [(W1, 1), (R, 2), (W1, 3), (R, 2), (W1, 3), (R, 4), 
-             (L, 7), (W0, 5), (R, 6), (W0, 5), (W1, 3), (R, 6),
-             (L, 8), (L, 7), (R, 9), (L, 7), (R, 10), (W0, 9)]" 
-
-text {*
-  @{text "tinc ss n"} returns the TM which simulates the execution of 
-  Abacus instruction @{text "Inc n"}, assuming that TM is located at
-  location @{text "ss"} in the final TM complied from the whole
-  Abacus program.
-*}
-
-fun tinc :: "nat \<Rightarrow> nat \<Rightarrow> instr list"
-  where
-  "tinc ss n = shift (findnth n @ shift tinc_b (2 * n)) (ss - 1)"
-
-text {*
-  @{text "tinc_b"} returns the TM which decrements the representation 
-  of the memory cell under rw-head by one and move the representation 
-  of cells afterwards to the left accordingly.
-  *}
-
-definition tdec_b :: "instr list"
-  where
-  "tdec_b \<equiv>  [(W1, 1), (R, 2), (L, 14), (R, 3), (L, 4), (R, 3),
-              (R, 5), (W0, 4), (R, 6), (W0, 5), (L, 7), (L, 8),
-              (L, 11), (W0, 7), (W1, 8), (R, 9), (L, 10), (R, 9),
-              (R, 5), (W0, 10), (L, 12), (L, 11), (R, 13), (L, 11),
-              (R, 17), (W0, 13), (L, 15), (L, 14), (R, 16), (L, 14),
-              (R, 0), (W0, 16)]"
-
-text {*
-  @{text "sete tp e"} attaches the termination edges (edges leading to state @{text "0"}) 
-  of TM @{text "tp"} to the intruction labelled by @{text "e"}.
-  *}
-
-fun sete :: "instr list \<Rightarrow> nat \<Rightarrow> instr list"
-  where
-  "sete tp e = map (\<lambda> (action, state). (action, if state = 0 then e else state)) tp"
-
-text {*
-  @{text "tdec ss n label"} returns the TM which simulates the execution of 
-  Abacus instruction @{text "Dec n label"}, assuming that TM is located at
-  location @{text "ss"} in the final TM complied from the whole
-  Abacus program.
-*}
-
-fun tdec :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> instr list"
-  where
-  "tdec ss n e = shift (findnth n) (ss - 1) @ sete (shift (shift tdec_b (2 * n)) (ss - 1)) e"
- 
-text {*
-  @{text "tgoto f(label)"} returns the TM simulating the execution of Abacus instruction
-  @{text "Goto label"}, where @{text "f(label)"} is the corresponding location of
-  @{text "label"} in the final TM compiled from the overall Abacus program.
-*}
-
-fun tgoto :: "nat \<Rightarrow> instr list"
-  where
-  "tgoto n = [(Nop, n), (Nop, n)]"
-
-text {*
-  The layout of the final TM compiled from an Abacus program is represented
-  as a list of natural numbers, where the list element at index @{text "n"} represents the 
-  starting state of the TM simulating the execution of @{text "n"}-th instruction
-  in the Abacus program.
-*}
-
-type_synonym layout = "nat list"
-
-text {*
-  @{text "length_of i"} is the length of the 
-  TM simulating the Abacus instruction @{text "i"}.
-*}
-fun length_of :: "abc_inst \<Rightarrow> nat"
-  where
-  "length_of i = (case i of 
-                    Inc n   \<Rightarrow> 2 * n + 9 |
-                    Dec n e \<Rightarrow> 2 * n + 16 |
-                    Goto n  \<Rightarrow> 1)"
-
-text {*
-  @{text "layout_of ap"} returns the layout of Abacus program @{text "ap"}.
-*}
-fun layout_of :: "abc_prog \<Rightarrow> layout"
-  where "layout_of ap = map length_of ap"
-
-
-text {*
-  @{text "start_of layout n"} looks out the starting state of @{text "n"}-th
-  TM in the finall TM.
-*}
-thm listsum_def
-
-fun start_of :: "nat list \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "start_of ly x = (Suc (listsum (take x ly))) "
-
-text {*
-  @{text "ci lo ss i"} complies Abacus instruction @{text "i"}
-  assuming the TM of @{text "i"} starts from state @{text "ss"} 
-  within the overal layout @{text "lo"}.
-*}
-
-fun ci :: "layout \<Rightarrow> nat \<Rightarrow> abc_inst \<Rightarrow> instr list"
-  where
-  "ci ly ss (Inc n) = tinc ss n"
-| "ci ly ss (Dec n e) = tdec ss n (start_of ly e)"
-| "ci ly ss (Goto n) = tgoto (start_of ly n)"
-
-text {*
-  @{text "tpairs_of ap"} transfroms Abacus program @{text "ap"} pairing
-  every instruction with its starting state.
-*}
-
-fun tpairs_of :: "abc_prog \<Rightarrow> (nat \<times> abc_inst) list"
-  where "tpairs_of ap = (zip (map (start_of (layout_of ap)) 
-                         [0..<(length ap)]) ap)"
-
-text {*
-  @{text "tms_of ap"} returns the list of TMs, where every one of them simulates
-  the corresponding Abacus intruction in @{text "ap"}.
-*}
-
-fun tms_of :: "abc_prog \<Rightarrow> (instr list) list"
-  where "tms_of ap = map (\<lambda> (n, tm). ci (layout_of ap) n tm) 
-                         (tpairs_of ap)"
-
-text {*
-  @{text "tm_of ap"} returns the final TM machine compiled from Abacus program @{text "ap"}.
-*}
-fun tm_of :: "abc_prog \<Rightarrow> instr list"
-  where "tm_of ap = concat (tms_of ap)"
-
-text {*
-  The following two functions specify the well-formedness of complied TM.
-*}
-(*
-fun t_ncorrect :: "tprog \<Rightarrow> bool"
-  where
-  "t_ncorrect tp = (length tp mod 2 = 0)"
-
-fun abc2t_correct :: "abc_prog \<Rightarrow> bool"
-  where 
-  "abc2t_correct ap = list_all (\<lambda> (n, tm). 
-             t_ncorrect (ci (layout_of ap) n tm)) (tpairs_of ap)"
-*)
-
-lemma length_findnth: 
-  "length (findnth n) = 4 * n"
-apply(induct n, auto)
-done
-
-lemma ci_length : "length (ci ns n ai) div 2 = length_of ai"
-apply(auto simp: ci.simps tinc_b_def tdec_b_def length_findnth
-                 split: abc_inst.splits)
-done
-
-subsection {*
-  Representation of Abacus memory by TM tape
-*}
-
-text {*
-  @{text "crsp acf tcf"} meams the abacus configuration @{text "acf"}
-  is corretly represented by the TM configuration @{text "tcf"}.
-*}
-
-fun crsp :: "layout \<Rightarrow> abc_conf \<Rightarrow> config \<Rightarrow> cell list \<Rightarrow> bool"
-  where 
-  "crsp ly (as, lm) (s, l, r) inres = 
-           (s = start_of ly as \<and> (\<exists> x. r = <lm> @ Bk\<up>x) \<and> 
-            l = Bk # Bk # inres)"
-
-declare crsp.simps[simp del]
-
-subsection {*
-  A more general definition of TM execution. 
-*}
-
-(*
-fun nnth_of :: "(taction \<times> nat) list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> (taction \<times> nat)"
-  where
-  "nnth_of p s b = (if 2*s < length p 
-                    then (p ! (2*s + b)) else (Nop, 0))"
-
-thm nth_of.simps
-
-fun nfetch :: "tprog \<Rightarrow> nat \<Rightarrow> block \<Rightarrow> taction \<times> nat"
-  where
-  "nfetch p 0 b = (Nop, 0)" |
-  "nfetch p (Suc s) b = 
-             (case b of 
-                Bk \<Rightarrow> nnth_of p s 0 |
-                Oc \<Rightarrow> nnth_of p s 1)"
-*)
-
-                    
-text {*
-  The type of invarints expressing correspondence between 
-  Abacus configuration and TM configuration.
-*}
-
-type_synonym inc_inv_t = "abc_conf \<Rightarrow> config \<Rightarrow> cell list \<Rightarrow> bool"
-
-declare tms_of.simps[simp del] tm_of.simps[simp del]
-        layout_of.simps[simp del] abc_fetch.simps [simp del]  
-        tpairs_of.simps[simp del] start_of.simps[simp del]
-        ci.simps [simp del] length_of.simps[simp del] 
-        layout_of.simps[simp del]
-
-(*
-subsection {* The compilation of @{text "Inc n"} *}
-*)
-
-text {*
-  The lemmas in this section lead to the correctness of 
-  the compilation of @{text "Inc n"} instruction.
-*}
-
-declare abc_step_l.simps[simp del] abc_steps_l.simps[simp del]
-lemma [simp]: "start_of ly as > 0"
-apply(simp add: start_of.simps)
-done
-
-lemma abc_steps_l_0: "abc_steps_l ac ap 0 = ac"
-by(case_tac ac, simp add: abc_steps_l.simps)
-
-lemma abc_step_red: 
- "abc_steps_l (as, am) ap stp = (bs, bm) \<Longrightarrow> 
-  abc_steps_l (as, am) ap (Suc stp) = abc_step_l (bs, bm) (abc_fetch bs ap) "
-proof(induct stp arbitrary: as am bs bm)
-  case 0
-  thus "?case"
-    by(simp add: abc_steps_l.simps abc_steps_l_0)
-next
-  case (Suc stp as am bs bm)
-  have ind: "\<And>as am bs bm. abc_steps_l (as, am) ap stp = (bs, bm) \<Longrightarrow> 
-    abc_steps_l (as, am) ap (Suc stp) = abc_step_l (bs, bm) (abc_fetch bs ap)"
-    by fact
-  have h:" abc_steps_l (as, am) ap (Suc stp) = (bs, bm)" by fact
-  obtain as' am' where g: "abc_step_l (as, am) (abc_fetch as ap) = (as', am')"
-    by(case_tac "abc_step_l (as, am) (abc_fetch as ap)", auto)
-  then have "abc_steps_l (as', am') ap (Suc stp) = abc_step_l (bs, bm) (abc_fetch bs ap)"
-    using h
-    by(rule_tac ind, simp add: abc_steps_l.simps)
-  thus "?case"
-    using g
-    by(simp add: abc_steps_l.simps)
-qed
-
-lemma tm_shift_fetch: 
-  "\<lbrakk>fetch A s b = (ac, ns); ns \<noteq> 0 \<rbrakk>
-  \<Longrightarrow> fetch (shift A off) s b = (ac, ns + off)"
-apply(case_tac b)
-apply(case_tac [!] s, auto simp: fetch.simps shift.simps)
-done
-
-lemma tm_shift_eq_step:
-  assumes exec: "step (s, l, r) (A, 0) = (s', l', r')"
-  and notfinal: "s' \<noteq> 0"
-  shows "step (s + off, l, r) (shift A off, off) = (s' + off, l', r')"
-using assms
-apply(simp add: step.simps)
-apply(case_tac "fetch A s (read r)", auto)
-apply(drule_tac [!] off = off in tm_shift_fetch, simp_all)
-done
-
-declare step.simps[simp del] steps.simps[simp del] shift.simps[simp del]
-
-lemma tm_shift_eq_steps: 
-  assumes exec: "steps (s, l, r) (A, 0) stp = (s', l', r')"
-  and notfinal: "s' \<noteq> 0"
-  shows "steps (s + off, l, r) (shift A off, off) stp = (s' + off, l', r')"
-  using exec notfinal
-proof(induct stp arbitrary: s' l' r', simp add: steps.simps)
-  fix stp s' l' r'
-  assume ind: "\<And>s' l' r'. \<lbrakk>steps (s, l, r) (A, 0) stp = (s', l', r'); s' \<noteq> 0\<rbrakk> 
-     \<Longrightarrow> steps (s + off, l, r) (shift A off, off) stp = (s' + off, l', r')"
-  and h: " steps (s, l, r) (A, 0) (Suc stp) = (s', l', r')" "s' \<noteq> 0"
-  obtain s1 l1 r1 where g: "steps (s, l, r) (A, 0) stp = (s1, l1, r1)" 
-    apply(case_tac "steps (s, l, r) (A, 0) stp") by blast
-  moreover then have "s1 \<noteq> 0"
-    using h
-    apply(simp add: step_red)
-    apply(case_tac "0 < s1", auto)
-    done
-  ultimately have "steps (s + off, l, r) (shift A off, off) stp =
-                   (s1 + off, l1, r1)"
-    apply(rule_tac ind, simp_all)
-    done
-  thus "steps (s + off, l, r) (shift A off, off) (Suc stp) = (s' + off, l', r')"
-    using h g assms
-    apply(simp add: step_red)
-    apply(rule_tac tm_shift_eq_step, auto)
-    done
-qed
-
-lemma startof_not0[simp]: "0 < start_of ly as"
-apply(simp add: start_of.simps)
-done
-
-lemma startof_ge1[simp]: "Suc 0 \<le> start_of ly as"
-apply(simp add: start_of.simps)
-done
-
-lemma start_of_Suc1: "\<lbrakk>ly = layout_of ap; 
-       abc_fetch as ap = Some (Inc n)\<rbrakk>
-       \<Longrightarrow> start_of ly (Suc as) = start_of ly as + 2 * n + 9"
-apply(auto simp: start_of.simps layout_of.simps  
-                 length_of.simps abc_fetch.simps 
-                 take_Suc_conv_app_nth split: if_splits)
-done
-
-lemma start_of_Suc2:
-  "\<lbrakk>ly = layout_of ap;
-  abc_fetch as ap = Some (Dec n e)\<rbrakk> \<Longrightarrow> 
-        start_of ly (Suc as) = 
-            start_of ly as + 2 * n + 16"
-apply(auto simp: start_of.simps layout_of.simps  
-                 length_of.simps abc_fetch.simps 
-                 take_Suc_conv_app_nth split: if_splits)
-done
-
-lemma start_of_Suc3:
-  "\<lbrakk>ly = layout_of ap;
-  abc_fetch as ap = Some (Goto n)\<rbrakk> \<Longrightarrow> 
-  start_of ly (Suc as) = start_of ly as + 1"
-apply(auto simp: start_of.simps layout_of.simps  
-                 length_of.simps abc_fetch.simps 
-                 take_Suc_conv_app_nth split: if_splits)
-done
-
-lemma length_ci_inc: 
-  "length (ci ly ss (Inc n)) = 4*n + 18"
-apply(auto simp: ci.simps length_findnth tinc_b_def)
-done
-
-lemma length_ci_dec: 
-  "length (ci ly ss (Dec n e)) = 4*n + 32"
-apply(auto simp: ci.simps length_findnth tdec_b_def)
-done
-
-lemma length_ci_goto: 
-  "length (ci ly ss (Goto n )) = 2"
-apply(auto simp: ci.simps length_findnth tdec_b_def)
-done
-
-lemma take_Suc_last[elim]: "Suc as \<le> length xs \<Longrightarrow> 
-            take (Suc as) xs = take as xs @ [xs ! as]"
-apply(induct xs arbitrary: as, simp, simp)
-apply(case_tac as, simp, simp)
-done
-
-lemma concat_suc: "Suc as \<le> length xs \<Longrightarrow> 
-       concat (take (Suc as) xs) = concat (take as xs) @ xs! as"
-apply(subgoal_tac "take (Suc as) xs = take as xs @ [xs ! as]", simp)
-by auto
-
-lemma concat_take_suc_iff: "Suc n \<le> length tps \<Longrightarrow> 
-       concat (take n tps) @ (tps ! n) = concat (take (Suc n) tps)"
-apply(drule_tac concat_suc, simp)
-done
-
-lemma concat_drop_suc_iff: 
-   "Suc n < length tps \<Longrightarrow> concat (drop (Suc n) tps) = 
-           tps ! Suc n @ concat (drop (Suc (Suc n)) tps)"
-apply(induct tps arbitrary: n, simp, simp)
-apply(case_tac tps, simp, simp)
-apply(case_tac n, simp, simp)
-done
-
-declare append_assoc[simp del]
-
-lemma  tm_append:
-  "\<lbrakk>n < length tps; tp = tps ! n\<rbrakk> \<Longrightarrow> 
-  \<exists> tp1 tp2. concat tps = tp1 @ tp @ tp2 \<and> tp1 = 
-  concat (take n tps) \<and> tp2 = concat (drop (Suc n) tps)"
-apply(rule_tac x = "concat (take n tps)" in exI)
-apply(rule_tac x = "concat (drop (Suc n) tps)" in exI)
-apply(auto)
-apply(induct n, simp)
-apply(case_tac tps, simp, simp, simp)
-apply(subgoal_tac "concat (take n tps) @ (tps ! n) = 
-               concat (take (Suc n) tps)")
-apply(simp only: append_assoc[THEN sym], simp only: append_assoc)
-apply(subgoal_tac " concat (drop (Suc n) tps) = tps ! Suc n @ 
-                  concat (drop (Suc (Suc n)) tps)", simp)
-apply(rule_tac concat_drop_suc_iff, simp)
-apply(rule_tac concat_take_suc_iff, simp)
-done
-
-declare append_assoc[simp]
-
-lemma map_of:  "n < length xs \<Longrightarrow> (map f xs) ! n = f (xs ! n)"
-by(auto)
-
-lemma [simp]: "length (tms_of aprog) = length aprog"
-apply(auto simp: tms_of.simps tpairs_of.simps)
-done
-
-lemma ci_nth: 
-  "\<lbrakk>ly = layout_of aprog; 
-  abc_fetch as aprog = Some ins\<rbrakk>
-  \<Longrightarrow> ci ly (start_of ly as) ins = tms_of aprog ! as"
-apply(simp add: tms_of.simps tpairs_of.simps 
-      abc_fetch.simps  map_of del: map_append split: if_splits)
-done
-
-lemma t_split:"\<lbrakk>
-        ly = layout_of aprog;
-        abc_fetch as aprog = Some ins\<rbrakk>
-      \<Longrightarrow> \<exists> tp1 tp2. concat (tms_of aprog) = 
-            tp1 @ (ci ly (start_of ly as) ins) @ tp2
-            \<and> tp1 = concat (take as (tms_of aprog)) \<and> 
-              tp2 = concat (drop (Suc as) (tms_of aprog))"
-apply(insert tm_append[of "as" "tms_of aprog" 
-                             "ci ly (start_of ly as) ins"], simp)
-apply(subgoal_tac "ci ly (start_of ly as) ins = (tms_of aprog) ! as")
-apply(subgoal_tac "length (tms_of aprog) = length aprog")
-apply(simp add: abc_fetch.simps split: if_splits, simp)
-apply(rule_tac ci_nth, auto)
-done
-
-lemma math_sub: "\<lbrakk>x >= Suc 0; x - 1 = z\<rbrakk> \<Longrightarrow> x + y - Suc 0 = z + y"
-by auto
-
-lemma start_more_one: "as \<noteq> 0 \<Longrightarrow> start_of ly as >= Suc 0"
-apply(induct as, simp add: start_of.simps)
-apply(case_tac as, auto simp: start_of.simps)
-done
-
-lemma div_apart: "\<lbrakk>x mod (2::nat) = 0; y mod 2 = 0\<rbrakk> 
-          \<Longrightarrow> (x + y) div 2 = x div 2 + y div 2"
-apply(drule mod_eqD)+
-apply(auto)
-done
-
-lemma div_apart_iff: "\<lbrakk>x mod (2::nat) = 0; y mod 2 = 0\<rbrakk> \<Longrightarrow> 
-           (x + y) mod 2 = 0"
-apply(auto)
-done
-
-lemma [simp]: "length (layout_of aprog) = length aprog"
-apply(auto simp: layout_of.simps)
-done
-
-lemma start_of_ind: "\<lbrakk>as < length aprog; ly = layout_of aprog\<rbrakk> \<Longrightarrow> 
-       start_of ly (Suc as) = start_of ly as + 
-                          length ((tms_of aprog) ! as) div 2"
-apply(simp only: start_of.simps, simp)
-apply(auto simp: start_of.simps tms_of.simps layout_of.simps 
-                 tpairs_of.simps)
-apply(simp add: ci_length take_Suc take_Suc_conv_app_nth)
-done
-
-lemma concat_take_suc: "Suc n \<le> length xs \<Longrightarrow>
-  concat (take (Suc n) xs) = concat (take n xs) @ (xs ! n)"
-apply(subgoal_tac "take (Suc n) xs =
-                   take n xs @ [xs ! n]")
-apply(auto)
-done
-
-lemma [simp]: 
-  "\<lbrakk>as < length aprog; (abc_fetch as aprog) = Some ins\<rbrakk>
-  \<Longrightarrow> ci (layout_of aprog) 
-  (start_of (layout_of aprog) as) (ins) \<in> set (tms_of aprog)"
-apply(insert ci_nth[of "layout_of aprog" aprog as], simp)
-done
-
-lemma [simp]: "length (tms_of ap) = length ap"
-by(auto simp: tms_of.simps tpairs_of.simps)
-
-lemma [intro]:  "n < length ap \<Longrightarrow> length (tms_of ap ! n) mod 2 = 0"
-apply(auto simp: tms_of.simps tpairs_of.simps)
-apply(case_tac "ap ! n", auto simp: ci.simps length_findnth tinc_b_def tdec_b_def)
-apply arith
-by arith
-
-lemma compile_mod2: "length (concat (take n (tms_of ap))) mod 2 = 0"
-apply(induct n, auto)
-apply(case_tac "n < length (tms_of ap)", simp add: take_Suc_conv_app_nth, auto)
-apply(subgoal_tac "length (tms_of ap ! n) mod 2 = 0")
-apply arith
-by auto
-
-lemma tpa_states:
-  "\<lbrakk>tp = concat (take as (tms_of ap));
-  as \<le> length ap\<rbrakk> \<Longrightarrow> 
-  start_of (layout_of ap) as = Suc (length tp div 2)"
-proof(induct as arbitrary: tp)
-  case 0
-  thus "?case"
-    by(simp add: start_of.simps)
-next
-  case (Suc as tp)
-  have ind: "\<And>tp. \<lbrakk>tp = concat (take as (tms_of ap)); as \<le> length ap\<rbrakk> \<Longrightarrow>
-    start_of (layout_of ap) as = Suc (length tp div 2)" by fact
-  have tp: "tp = concat (take (Suc as) (tms_of ap))" by fact
-  have le: "Suc as \<le> length ap" by fact
-  have a: "start_of (layout_of ap) as = Suc (length (concat (take as (tms_of ap))) div 2)"
-    using le
-    by(rule_tac ind, simp_all)
-  from a tp le show "?case"
-    apply(simp add: start_of.simps take_Suc_conv_app_nth)
-    apply(subgoal_tac "length (concat (take as (tms_of ap))) mod 2= 0")
-    apply(subgoal_tac " length (tms_of ap ! as) mod 2 = 0")
-    apply(simp add: abacus.div_apart) 
-    apply(simp add: layout_of.simps ci_length  tms_of.simps tpairs_of.simps)
-    apply(auto  intro: compile_mod2)
-    done
-qed
-
-lemma append_append_fetch: 
-    "\<lbrakk>length tp1 mod 2 = 0; length tp mod 2 = 0;
-      length tp1 div 2 < a \<and> a \<le> length tp1 div 2 + length tp div 2\<rbrakk>
-    \<Longrightarrow>fetch (tp1 @ tp @ tp2) a b = fetch tp (a - length tp1 div 2) b "
-apply(subgoal_tac "\<exists> x. a = length tp1 div 2 + x", erule exE, simp)
-apply(case_tac x, simp)
-apply(subgoal_tac "length tp1 div 2 + Suc nat = 
-             Suc (length tp1 div 2 + nat)")
-apply(simp only: fetch.simps nth_of.simps, auto)
-apply(case_tac b, simp)
-apply(subgoal_tac "2 * (length tp1 div 2) = length tp1", simp)
-apply(subgoal_tac "2 * nat < length tp", simp add: nth_append, simp)
-apply(subgoal_tac "2 * (length tp1 div 2) = length tp1", simp)
-apply(subgoal_tac "2 * nat < length tp", simp add: nth_append, auto)
-apply(auto simp: nth_append)
-apply(rule_tac x = "a - length tp1 div 2" in exI, simp)
-done
-
-lemma step_eq_fetch':
-  assumes layout: "ly = layout_of ap"
-  and compile: "tp = tm_of ap"
-  and fetch: "abc_fetch as ap = Some ins"
-  and range1: "s \<ge> start_of ly as"
-  and range2: "s < start_of ly (Suc as)"
-  shows "fetch tp s b = fetch (ci ly (start_of ly as) ins)
-       (Suc s - start_of ly as) b "
-proof -
-  have "\<exists>tp1 tp2. concat (tms_of ap) = tp1 @ ci ly (start_of ly as) ins @ tp2 \<and>
-    tp1 = concat (take as (tms_of ap)) \<and> tp2 = concat (drop (Suc as) (tms_of ap))"
-    using assms
-    by(rule_tac t_split, simp_all)
-  then obtain tp1 tp2 where a: "concat (tms_of ap) = tp1 @ ci ly (start_of ly as) ins @ tp2 \<and>
-    tp1 = concat (take as (tms_of ap)) \<and> tp2 = concat (drop (Suc as) (tms_of ap))" by blast
-  then have b: "start_of (layout_of ap) as = Suc (length tp1 div 2)"
-    using fetch
-    apply(rule_tac tpa_states, simp, simp add: abc_fetch.simps split: if_splits)
-    done
-  have "fetch (tp1 @ (ci ly (start_of ly as) ins) @ tp2)  s b = 
-        fetch (ci ly (start_of ly as) ins) (s - length tp1 div 2) b"
-  proof(rule_tac append_append_fetch)
-    show "length tp1 mod 2 = 0"
-      using a
-      by(auto, rule_tac compile_mod2)
-  next
-    show "length (ci ly (start_of ly as) ins) mod 2 = 0"
-      apply(case_tac ins, auto simp: ci.simps length_findnth tinc_b_def tdec_b_def)
-      by(arith, arith)
-  next
-    show "length tp1 div 2 < s \<and> s \<le> 
-      length tp1 div 2 + length (ci ly (start_of ly as) ins) div 2"
-    proof -
-      have "length (ci ly (start_of ly as) ins) div 2 = length_of ins"
-        using ci_length by simp
-      moreover have "start_of ly (Suc as) = start_of ly as + length_of ins"
-        using fetch layout
-        apply(simp add: start_of.simps abc_fetch.simps List.take_Suc_conv_app_nth 
-          split: if_splits)
-        apply(simp add: layout_of.simps)
-        done
-      ultimately show "?thesis"
-        using b layout range1 range2
-        apply(simp)
-        done
-    qed
-  qed
-  thus "?thesis"
-    using b layout a compile  
-    apply(simp add: tm_of.simps)
-    done
-qed
-
-lemma step_eq_fetch: 
-  assumes layout: "ly = layout_of ap"
-  and compile: "tp = tm_of ap"
-  and abc_fetch: "abc_fetch as ap = Some ins" 
-  and fetch: "fetch (ci ly (start_of ly as) ins)
-       (Suc s - start_of ly as) b = (ac, ns)"
-  and notfinal: "ns \<noteq> 0"
-  shows "fetch tp s b = (ac, ns)"
-proof -
-  have "s \<ge> start_of ly as"
-  proof(cases "s \<ge> start_of ly as")
-    case True thus "?thesis" by simp
-  next
-    case False 
-    have "\<not> start_of ly as \<le> s" by fact
-    then have "Suc s - start_of ly as = 0"
-      by arith
-    then have "fetch (ci ly (start_of ly as) ins)
-       (Suc s - start_of ly as) b = (Nop, 0)"
-      by(simp add: fetch.simps)
-    with notfinal fetch show "?thesis"
-      by(simp)
-  qed
-  moreover have "s < start_of ly (Suc as)"
-  proof(cases "s < start_of ly (Suc as)")
-    case True thus "?thesis" by simp
-  next
-    case False
-    have h: "\<not> s < start_of ly (Suc as)"
-      by fact
-    then have "s > start_of ly as"
-      using abc_fetch layout
-      apply(simp add: start_of.simps abc_fetch.simps split: if_splits)
-      apply(simp add: List.take_Suc_conv_app_nth, auto)
-      apply(subgoal_tac "layout_of ap ! as > 0") 
-      apply arith
-      apply(simp add: layout_of.simps)
-      apply(case_tac "ap!as", auto simp: length_of.simps)
-      done
-    from this and h have "fetch (ci ly (start_of ly as) ins) (Suc s - start_of ly as) b = (Nop, 0)"
-      using abc_fetch layout
-      apply(case_tac b, simp_all add: Suc_diff_le)
-      apply(case_tac [!] ins, simp_all add: start_of_Suc2 start_of_Suc1 start_of_Suc3)
-      apply(simp_all only: length_ci_inc length_ci_dec length_ci_goto, auto)
-      using layout
-      apply(subgoal_tac [!] "start_of ly (Suc as) = start_of ly as + 2*nat1 + 16", simp_all)
-      apply(rule_tac [!] start_of_Suc2, auto)
-      done
-    from fetch and notfinal this show "?thesis"by simp
-  qed
-  ultimately show "?thesis"
-    using assms
-    apply(drule_tac b= b and ins = ins in step_eq_fetch', auto)
-    done
-qed
-
-
-lemma step_eq_in:
-  assumes layout: "ly = layout_of ap"
-  and compile: "tp = tm_of ap"
-  and fetch: "abc_fetch as ap = Some ins"    
-  and exec: "step (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) 
-  = (s', l', r')"
-  and notfinal: "s' \<noteq> 0"
-  shows "step (s, l, r) (tp, 0) = (s', l', r')"
-  using assms
-  apply(simp add: step.simps)
-  apply(case_tac "fetch (ci (layout_of ap) (start_of (layout_of ap) as) ins)
-    (Suc s - start_of (layout_of ap) as) (read r)", simp)
-  using layout
-  apply(drule_tac s = s and b = "read r" and ac = a in step_eq_fetch, auto)
-  done
-
-lemma steps_eq_in:
-  assumes layout: "ly = layout_of ap"
-  and compile: "tp = tm_of ap"
-  and crsp: "crsp ly (as, lm) (s, l, r) ires"
-  and fetch: "abc_fetch as ap = Some ins"    
-  and exec: "steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp 
-  = (s', l', r')"
-  and notfinal: "s' \<noteq> 0"
-  shows "steps (s, l, r) (tp, 0) stp = (s', l', r')"
-  using exec notfinal
-proof(induct stp arbitrary: s' l' r', simp add: steps.simps)
-  fix stp s' l' r'
-  assume ind: 
-    "\<And>s' l' r'. \<lbrakk>steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp = (s', l', r'); s' \<noteq> 0\<rbrakk>
-              \<Longrightarrow> steps (s, l, r) (tp, 0) stp = (s', l', r')"
-  and h: "steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) (Suc stp) = (s', l', r')" "s' \<noteq> 0"
-  obtain s1 l1 r1 where g: "steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp = 
-                        (s1, l1, r1)"
-    apply(case_tac "steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp") by blast
-  moreover hence "s1 \<noteq> 0"
-    using h
-    apply(simp add: step_red)
-    apply(case_tac "0 < s1", simp_all)
-    done
-  ultimately have "steps (s, l, r) (tp, 0) stp = (s1, l1, r1)"
-    apply(rule_tac ind, auto)
-    done
-  thus "steps (s, l, r) (tp, 0) (Suc stp) = (s', l', r')"
-    using h g assms
-    apply(simp add: step_red)
-    apply(rule_tac step_eq_in, auto)
-    done
-qed
-
-lemma tm_append_fetch_first: 
-  "\<lbrakk>fetch A s b = (ac, ns); ns \<noteq> 0\<rbrakk> \<Longrightarrow> 
-    fetch (A @ B) s b = (ac, ns)"
-apply(case_tac b)
-apply(case_tac [!] s, auto simp: fetch.simps nth_append split: if_splits)
-done
-
-lemma tm_append_first_step_eq: 
-  assumes "step (s, l, r) (A, off) = (s', l', r')"
-  and "s' \<noteq> 0"
-  shows "step (s, l, r) (A @ B, off) = (s', l', r')"
-using assms
-apply(simp add: step.simps)
-apply(case_tac "fetch A (s - off) (read r)")
-apply(frule_tac  B = B and b = "read r" in tm_append_fetch_first, auto)
-done
-
-lemma tm_append_first_steps_eq: 
-  assumes "steps (s, l, r) (A, off) stp = (s', l', r')"
-  and "s' \<noteq> 0"
-  shows "steps (s, l, r) (A @ B, off) stp = (s', l', r')"
-using assms
-proof(induct stp arbitrary: s' l' r', simp add: steps.simps)
-  fix stp s' l' r'
-  assume ind: "\<And>s' l' r'. \<lbrakk>steps (s, l, r) (A, off) stp = (s', l', r'); s' \<noteq> 0\<rbrakk>
-    \<Longrightarrow> steps (s, l, r) (A @ B, off) stp = (s', l', r')"
-    and h: "steps (s, l, r) (A, off) (Suc stp) = (s', l', r')" "s' \<noteq> 0"
-  obtain sa la ra where a: "steps (s, l, r) (A, off) stp = (sa, la, ra)"
-    apply(case_tac "steps (s, l, r) (A, off) stp") by blast
-  hence "steps (s, l, r) (A @ B, off) stp = (sa, la, ra) \<and> sa \<noteq> 0"
-    using h ind[of sa la ra]
-    apply(case_tac sa, simp_all)
-    done
-  thus "steps (s, l, r) (A @ B, off) (Suc stp) = (s', l', r')"
-    using h a
-    apply(simp add: step_red)
-    apply(rule_tac tm_append_first_step_eq, simp_all)
-    done
-qed
-
-lemma tm_append_second_fetch_eq:
-  assumes
-  even: "length A mod 2 = 0"
-  and off: "off = length A div 2"
-  and fetch: "fetch B s b = (ac, ns)"
-  and notfinal: "ns \<noteq> 0"
-  shows "fetch (A @ shift B off) (s + off) b = (ac, ns + off)"
-using assms
-apply(case_tac b)
-apply(case_tac [!] s, auto simp: fetch.simps nth_append shift.simps 
-  split: if_splits)
-done
-
-
-lemma tm_append_second_step_eq: 
-  assumes 
-  exec: "step0 (s, l, r) B = (s', l', r')"
-  and notfinal: "s' \<noteq> 0"
-  and off: "off = length A div 2"
-  and even: "length A mod 2 = 0"
-  shows "step0 (s + off, l, r) (A @ shift B off) = (s' + off, l', r')"
-using assms
-apply(simp add: step.simps)
-apply(case_tac "fetch B s (read r)")
-apply(frule_tac tm_append_second_fetch_eq, simp_all, auto)
-done
-
-
-lemma tm_append_second_steps_eq: 
-  assumes 
-  exec: "steps (s, l, r) (B, 0) stp = (s', l', r')"
-  and notfinal: "s' \<noteq> 0"
-  and off: "off = length A div 2"
-  and even: "length A mod 2 = 0"
-  shows "steps (s + off, l, r) (A @ shift B off, 0) stp = (s' + off, l', r')"
-using exec notfinal
-proof(induct stp arbitrary: s' l' r')
-  case 0
-  thus "steps0 (s + off, l, r) (A @ shift B off) 0 = (s' + off, l', r')"
-    by(simp add: steps.simps)
-next
-  case (Suc stp s' l' r')
-  have ind: "\<And>s' l' r'. \<lbrakk>steps0 (s, l, r) B stp = (s', l', r'); s' \<noteq> 0\<rbrakk> \<Longrightarrow> 
-    steps0 (s + off, l, r) (A @ shift B off) stp = (s' + off, l', r')"
-    by fact
-  have h: "steps0 (s, l, r) B (Suc stp) = (s', l', r')" by fact
-  have k: "s' \<noteq> 0" by fact
-  obtain s'' l'' r'' where a: "steps0 (s, l, r) B stp = (s'', l'', r'')"
-    by (metis prod_cases3)
-  then have b: "s'' \<noteq> 0"
-    using h k
-    by(rule_tac notI, auto simp: step_red)
-  from a b have c: "steps0 (s + off, l, r) (A @ shift B off) stp = (s'' + off, l'', r'')"
-    by(erule_tac ind, simp)
-  from c b h a k assms show "?case"
-    thm tm_append_second_step_eq
-    apply(simp add: step_red) by(rule tm_append_second_step_eq, simp_all)
-qed
-
-lemma tm_append_second_fetch0_eq:
-  assumes
-  even: "length A mod 2 = 0"
-  and off: "off = length A div 2"
-  and fetch: "fetch B s b = (ac, 0)"
-  and notfinal: "s \<noteq> 0"
-  shows "fetch (A @ shift B off) (s + off) b = (ac, 0)"
-using assms
-apply(case_tac b)
-apply(case_tac [!] s, auto simp: fetch.simps nth_append shift.simps 
-  split: if_splits)
-done
-
-lemma tm_append_second_halt_eq:
-  assumes 
-  exec: "steps (Suc 0, l, r) (B, 0) stp = (0, l', r')"
-  and wf_B: "tm_wf (B, 0)"
-  and off: "off = length A div 2"
-  and even: "length A mod 2 = 0"
-  shows "steps (Suc off, l, r) (A @ shift B off, 0) stp = (0, l', r')"
-proof -
-  thm before_final
-  have "\<exists>n. \<not> is_final (steps0 (1, l, r) B n) \<and> steps0 (1, l, r) B (Suc n) = (0, l', r')"
-    using exec by(rule_tac before_final, simp)
- then obtain n where a: 
-   "\<not> is_final (steps0 (1, l, r) B n) \<and> steps0 (1, l, r) B (Suc n) = (0, l', r')" ..
- obtain s'' l'' r'' where b: "steps0 (1, l, r) B n = (s'', l'', r'') \<and> s'' >0"
-   using a
-   by(case_tac "steps0 (1, l, r) B n", auto)
- have c: "steps (Suc 0 + off, l, r) (A @ shift B off, 0) n = (s'' + off, l'', r'')"
-   using a b assms
-   by(rule_tac tm_append_second_steps_eq, simp_all)
- obtain ac where d: "fetch B s'' (read r'') = (ac, 0)"
-   using  b a
-   by(case_tac "fetch B s'' (read r'')", auto simp: step_red step.simps)
- then have "fetch (A @ shift B off) (s'' + off) (read r'') = (ac, 0)"
-   using assms b
-   by(rule_tac tm_append_second_fetch0_eq, simp_all)
- then have e: "steps (Suc 0 + off, l, r) (A @ shift B off, 0) (Suc n) = (0, l', r')"
-   using a b assms c d
-   by(simp add: step_red step.simps)
- from a have "n < stp"
-   using exec
- proof(cases "n < stp")
-   case  True thus "?thesis" by simp
- next
-   case False
-   have "\<not> n < stp" by fact
-   then obtain d where  "n = stp + d"
-     by (metis add.comm_neutral less_imp_add_positive nat_neq_iff)
-   thus "?thesis"
-     using a e exec
-     by(simp add: steps_add)
- qed
- then obtain d where "stp = Suc n + d"
-   by(metis add_Suc less_iff_Suc_add)
- thus "?thesis"
-   using e
-   by(simp only: steps_add, simp)
-qed
-
-lemma tm_append_steps: 
-  assumes 
-  aexec: "steps (s, l, r) (A, 0) stpa = (Suc (length A div 2), la, ra)"
-  and bexec: "steps (Suc 0, la, ra) (B, 0) stpb =  (sb, lb, rb)"
-  and notfinal: "sb \<noteq> 0"
-  and off: "off = length A div 2"
-  and even: "length A mod 2 = 0"
-  shows "steps (s, l, r) (A @ shift B off, 0) (stpa + stpb) = (sb + off, lb, rb)"
-proof -
-  have "steps (s, l, r) (A@shift B off, 0) stpa = (Suc (length A div 2), la, ra)"
-    apply(rule_tac tm_append_first_steps_eq)
-    apply(auto simp: assms)
-    done
-  moreover have "steps (1 + off, la, ra) (A @ shift B off, 0) stpb = (sb + off, lb, rb)"
-    apply(rule_tac tm_append_second_steps_eq)
-    apply(auto simp: assms bexec)
-    done
-  ultimately show "steps (s, l, r) (A @ shift B off, 0) (stpa + stpb) = (sb + off, lb, rb)"
-    apply(simp add: steps_add off)
-    done
-qed
-       
-subsection {* Crsp of Inc*}
-
-fun at_begin_fst_bwtn :: "inc_inv_t"
-  where
-  "at_begin_fst_bwtn (as, lm) (s, l, r) ires = 
-      (\<exists> lm1 tn rn. lm1 = (lm @ 0\<up>tn) \<and> length lm1 = s \<and> 
-          (if lm1 = [] then l = Bk # Bk # ires
-           else l = [Bk]@<rev lm1>@Bk#Bk#ires) \<and> r = Bk\<up>rn)" 
-
-
-fun at_begin_fst_awtn :: "inc_inv_t"
-  where
-  "at_begin_fst_awtn (as, lm) (s, l, r) ires = 
-      (\<exists> lm1 tn rn. lm1 = (lm @ 0\<up>tn) \<and> length lm1 = s \<and>
-         (if lm1 = []  then l = Bk # Bk # ires
-          else l = [Bk]@<rev lm1>@Bk#Bk#ires) \<and> r = [Oc]@Bk\<up>rn)"
-
-fun at_begin_norm :: "inc_inv_t"
-  where
-  "at_begin_norm (as, lm) (s, l, r) ires= 
-      (\<exists> lm1 lm2 rn. lm = lm1 @ lm2 \<and> length lm1 = s \<and> 
-        (if lm1 = [] then l = Bk # Bk # ires
-         else l = Bk # <rev lm1> @ Bk # Bk # ires ) \<and> r = <lm2>@Bk\<up>rn)"
-
-fun in_middle :: "inc_inv_t"
-  where
-  "in_middle (as, lm) (s, l, r) ires = 
-      (\<exists> lm1 lm2 tn m ml mr rn. lm @ 0\<up>tn = lm1 @ [m] @ lm2
-       \<and> length lm1 = s \<and> m + 1 = ml + mr \<and>  
-         ml \<noteq> 0 \<and> tn = s + 1 - length lm \<and> 
-       (if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires 
-        else l = Oc\<up>ml@[Bk]@<rev lm1>@
-                 Bk # Bk # ires) \<and> (r = Oc\<up>mr @ [Bk] @ <lm2>@ Bk\<up>rn \<or> 
-      (lm2 = [] \<and> r = Oc\<up>mr))
-      )"
-
-fun inv_locate_a :: "inc_inv_t"
-  where "inv_locate_a (as, lm) (s, l, r) ires = 
-     (at_begin_norm (as, lm) (s, l, r) ires \<or>
-      at_begin_fst_bwtn (as, lm) (s, l, r) ires \<or>
-      at_begin_fst_awtn (as, lm) (s, l, r) ires
-      )"
-
-fun inv_locate_b :: "inc_inv_t"
-  where "inv_locate_b (as, lm) (s, l, r) ires = 
-        (in_middle (as, lm) (s, l, r)) ires "
-
-fun inv_after_write :: "inc_inv_t"
-  where "inv_after_write (as, lm) (s, l, r) ires = 
-           (\<exists> rn m lm1 lm2. lm = lm1 @ m # lm2 \<and>
-             (if lm1 = [] then l = Oc\<up>m @ Bk # Bk # ires
-              else Oc # l = Oc\<up>Suc m@ Bk # <rev lm1> @ 
-                      Bk # Bk # ires) \<and> r = [Oc] @ <lm2> @ Bk\<up>rn)"
-
-fun inv_after_move :: "inc_inv_t"
-  where "inv_after_move (as, lm) (s, l, r) ires = 
-      (\<exists> rn m lm1 lm2. lm = lm1 @ m # lm2 \<and>
-        (if lm1 = [] then l = Oc\<up>Suc m @ Bk # Bk # ires
-         else l = Oc\<up>Suc m@ Bk # <rev lm1> @ Bk # Bk # ires) \<and> 
-        r = <lm2> @ Bk\<up>rn)"
-
-fun inv_after_clear :: "inc_inv_t"
-  where "inv_after_clear (as, lm) (s, l, r) ires =
-       (\<exists> rn m lm1 lm2 r'. lm = lm1 @ m # lm2 \<and> 
-        (if lm1 = [] then l = Oc\<up>Suc m @ Bk # Bk # ires
-         else l = Oc\<up>Suc m @ Bk # <rev lm1> @ Bk # Bk # ires) \<and> 
-          r = Bk # r' \<and> Oc # r' = <lm2> @ Bk\<up>rn)"
-
-fun inv_on_right_moving :: "inc_inv_t"
-  where "inv_on_right_moving (as, lm) (s, l, r) ires = 
-       (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
-            ml + mr = m \<and> 
-          (if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires
-          else l = Oc\<up>ml  @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
-         ((r = Oc\<up>mr @ [Bk] @ <lm2> @ Bk\<up>rn) \<or> 
-          (r = Oc\<up>mr \<and> lm2 = [])))"
-
-fun inv_on_left_moving_norm :: "inc_inv_t"
-  where "inv_on_left_moving_norm (as, lm) (s, l, r) ires =
-      (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and>  
-             ml + mr = Suc m \<and> mr > 0 \<and> (if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires
-                                         else l =  Oc\<up>ml @ Bk # <rev lm1> @ Bk # Bk # ires)
-        \<and> (r = Oc\<up>mr @ Bk # <lm2> @ Bk\<up>rn \<or> 
-           (lm2 = [] \<and> r = Oc\<up>mr)))"
-
-fun inv_on_left_moving_in_middle_B:: "inc_inv_t"
-  where "inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires =
-                (\<exists> lm1 lm2 rn. lm = lm1 @ lm2 \<and>  
-                     (if lm1 = [] then l = Bk # ires
-                      else l = <rev lm1> @ Bk # Bk # ires) \<and> 
-                      r = Bk # <lm2> @ Bk\<up>rn)"
-
-fun inv_on_left_moving :: "inc_inv_t"
-  where "inv_on_left_moving (as, lm) (s, l, r) ires = 
-       (inv_on_left_moving_norm  (as, lm) (s, l, r) ires \<or>
-        inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires)"
-
-
-fun inv_check_left_moving_on_leftmost :: "inc_inv_t"
-  where "inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires = 
-                (\<exists> rn. l = ires \<and> r = [Bk, Bk] @ <lm> @  Bk\<up>rn)"
-
-fun inv_check_left_moving_in_middle :: "inc_inv_t"
-  where "inv_check_left_moving_in_middle (as, lm) (s, l, r) ires = 
-              (\<exists> lm1 lm2 r' rn. lm = lm1 @ lm2 \<and>
-                 (Oc # l = <rev lm1> @ Bk # Bk # ires) \<and> r = Oc # Bk # r' \<and> 
-                           r' = <lm2> @  Bk\<up>rn)"
-
-fun inv_check_left_moving :: "inc_inv_t"
-  where "inv_check_left_moving (as, lm) (s, l, r) ires = 
-             (inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires \<or>
-             inv_check_left_moving_in_middle (as, lm) (s, l, r) ires)"
-
-fun inv_after_left_moving :: "inc_inv_t"
-  where "inv_after_left_moving (as, lm) (s, l, r) ires= 
-              (\<exists> rn. l = Bk # ires \<and> r = Bk # <lm> @  Bk\<up>rn)"
-
-fun inv_stop :: "inc_inv_t"
-  where "inv_stop (as, lm) (s, l, r) ires= 
-              (\<exists> rn. l = Bk # Bk # ires \<and> r = <lm> @  Bk\<up>rn)"
-
-
-lemma halt_lemma2': 
-  "\<lbrakk>wf LE;  \<forall> n. ((\<not> P (f n) \<and> Q (f n)) \<longrightarrow> 
-    (Q (f (Suc n)) \<and> (f (Suc n), (f n)) \<in> LE)); Q (f 0)\<rbrakk> 
-      \<Longrightarrow> \<exists> n. P (f n)"
-apply(intro exCI, simp)
-apply(subgoal_tac "\<forall> n. Q (f n)", simp)
-apply(drule_tac f = f in wf_inv_image)
-apply(simp add: inv_image_def)
-apply(erule wf_induct, simp)
-apply(erule_tac x = x in allE)
-apply(erule_tac x = n in allE, erule_tac x = n in allE)
-apply(erule_tac x = "Suc x" in allE, simp)
-apply(rule_tac allI)
-apply(induct_tac n, simp)
-apply(erule_tac x = na in allE, simp)
-done
-
-lemma halt_lemma2'': 
-  "\<lbrakk>P (f n); \<not> P (f (0::nat))\<rbrakk> \<Longrightarrow> 
-         \<exists> n. (P (f n) \<and> (\<forall> i < n. \<not> P (f i)))"
-apply(induct n rule: nat_less_induct, auto)
-done
-
-lemma halt_lemma2''':
- "\<lbrakk>\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> LE;
-                 Q (f 0);  \<forall>i<na. \<not> P (f i)\<rbrakk> \<Longrightarrow> Q (f na)"
-apply(induct na, simp, simp)
-done
-
-lemma halt_lemma2: 
-  "\<lbrakk>wf LE;  
-    Q (f 0); \<not> P (f 0);
-    \<forall> n. ((\<not> P (f n) \<and> Q (f n)) \<longrightarrow> (Q (f (Suc n)) \<and> (f (Suc n), (f n)) \<in> LE))\<rbrakk> 
-  \<Longrightarrow> \<exists> n. P (f n) \<and> Q (f n)"
-apply(insert halt_lemma2' [of LE P f Q], simp, erule_tac exE)
-apply(subgoal_tac "\<exists> n. (P (f n) \<and> (\<forall> i < n. \<not> P (f i)))")
-apply(erule_tac exE)+
-apply(rule_tac x = na in exI, auto)
-apply(rule halt_lemma2''', simp, simp, simp)
-apply(erule_tac halt_lemma2'', simp)
-done
-
-
-fun findnth_inv :: "layout \<Rightarrow> nat \<Rightarrow> inc_inv_t"
-  where
-  "findnth_inv ly n (as, lm) (s, l, r) ires =
-              (if s = 0 then False
-               else if s \<le> Suc (2*n) then 
-                  if s mod 2 = 1 then inv_locate_a (as, lm) ((s - 1) div 2, l, r) ires
-                  else inv_locate_b (as, lm) ((s - 1) div 2, l, r) ires
-               else False)"
-
-
-fun findnth_state :: "config \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "findnth_state (s, l, r) n = (Suc (2*n) - s)"
-  
-fun findnth_step :: "config \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "findnth_step (s, l, r) n = 
-           (if s mod 2 = 1 then
-                   (if (r \<noteq> [] \<and> hd r = Oc) then 0
-                    else 1)
-            else length r)"
-
-fun findnth_measure :: "config \<times> nat \<Rightarrow> nat \<times> nat"
-  where
-  "findnth_measure (c, n) = 
-     (findnth_state c n, findnth_step c n)"
-
-definition lex_pair :: "((nat \<times> nat) \<times> nat \<times> nat) set"
-  where
-  "lex_pair \<equiv> less_than <*lex*> less_than"
-
-definition findnth_LE :: "((config \<times> nat) \<times> (config \<times> nat)) set"
-  where
-   "findnth_LE \<equiv> (inv_image lex_pair findnth_measure)"
-
-lemma wf_findnth_LE: "wf findnth_LE"
-by(auto intro:wf_inv_image simp: findnth_LE_def lex_pair_def)
-
-declare findnth_inv.simps[simp del]
-
-lemma [simp]: 
-  "\<lbrakk>x < Suc (Suc (2 * n)); Suc x mod 2 = Suc 0; \<not> x < 2 * n\<rbrakk>
- \<Longrightarrow> x = 2*n"
-by arith
-
-lemma [simp]: 
-  "\<lbrakk>0 < a; a < Suc (2 * n); a mod 2 = Suc 0\<rbrakk> 
-      \<Longrightarrow> fetch (findnth n) a Bk = (W1, a)"
-apply(case_tac a, simp_all)
-apply(induct n, auto simp: findnth.simps length_findnth nth_append)
-apply arith
-done
-
-lemma [simp]: 
-  "\<lbrakk>0 < a; a < Suc (2 * n); a mod 2 = Suc 0\<rbrakk> 
-      \<Longrightarrow> fetch (findnth n) a Oc = (R, Suc a)"
-apply(case_tac a, simp_all)
-apply(induct n, auto simp: findnth.simps length_findnth nth_append)
-apply(subgoal_tac "nat = 2 * n", simp)
-by arith
-
-lemma [simp]: 
-  "\<lbrakk>0 < a; a < Suc (2*n); a mod 2 \<noteq> Suc 0\<rbrakk>
-     \<Longrightarrow> fetch (findnth n) a Oc = (R, a)"
-apply(case_tac a, simp_all)
-apply(induct n, auto simp: findnth.simps length_findnth nth_append)
-apply(subgoal_tac "nat = Suc (2 * n)", simp)
-apply arith
-done
-
-lemma [simp]: 
-  "\<lbrakk>0 < a; a < Suc (2*n); a mod 2 \<noteq> Suc 0\<rbrakk>
-     \<Longrightarrow> fetch (findnth n) a Bk = (R, Suc a)"
-apply(case_tac a, simp_all)
-apply(induct n, auto simp: findnth.simps length_findnth nth_append)
-apply(subgoal_tac "nat = Suc (2 * n)", simp)
-by arith
-
-declare at_begin_norm.simps[simp del] at_begin_fst_bwtn.simps[simp del] 
-   at_begin_fst_awtn.simps[simp del] in_middle.simps[simp del] 
-   abc_lm_s.simps[simp del] abc_lm_v.simps[simp del]  
-   ci.simps[simp del] inv_after_move.simps[simp del] 
-   inv_on_left_moving_norm.simps[simp del] 
-   inv_on_left_moving_in_middle_B.simps[simp del]
-   inv_after_clear.simps[simp del] 
-   inv_after_write.simps[simp del] inv_on_left_moving.simps[simp del]
-   inv_on_right_moving.simps[simp del] 
-   inv_check_left_moving.simps[simp del] 
-   inv_check_left_moving_in_middle.simps[simp del]
-   inv_check_left_moving_on_leftmost.simps[simp del] 
-   inv_after_left_moving.simps[simp del]
-   inv_stop.simps[simp del] inv_locate_a.simps[simp del] 
-   inv_locate_b.simps[simp del]
-
-lemma [intro]: "\<exists>rn. [Bk] = Bk \<up> rn"
-by (metis replicate_0 replicate_Suc)
-
-lemma [intro]:  "at_begin_norm (as, am) (q, aaa, []) ires
-             \<Longrightarrow> at_begin_norm (as, am) (q, aaa, [Bk]) ires"
-apply(simp add: at_begin_norm.simps, erule_tac exE, erule_tac exE)
-apply(rule_tac x = lm1 in exI, simp, auto)
-done
-
-lemma [intro]: "at_begin_fst_bwtn (as, am) (q, aaa, []) ires 
-            \<Longrightarrow> at_begin_fst_bwtn (as, am) (q, aaa, [Bk]) ires"
-apply(simp only: at_begin_fst_bwtn.simps, erule_tac exE, erule_tac exE, erule_tac exE)
-apply(rule_tac x = "am @ 0\<up>tn" in exI, auto)
-done
-
-lemma [intro]: "at_begin_fst_awtn (as, am) (q, aaa, []) ires
-           \<Longrightarrow> at_begin_fst_awtn (as, am) (q, aaa, [Bk]) ires"
-apply(auto simp: at_begin_fst_awtn.simps)
-done 
-
-lemma [intro]: "inv_locate_a (as, am) (q, aaa, []) ires
-            \<Longrightarrow> inv_locate_a (as, am) (q, aaa, [Bk]) ires"
-apply(simp only: inv_locate_a.simps)
-apply(erule disj_forward)
-defer
-apply(erule disj_forward, auto)
-done
-
-lemma tape_of_nl_cons: "<m # lm> = (if lm = [] then Oc\<up>(Suc m)
-                    else Oc\<up>(Suc m) @ Bk # <lm>)"
-apply(case_tac lm, simp_all add: tape_of_nl_abv  tape_of_nat_abv split: if_splits)
-done
-
-
-lemma locate_a_2_locate_a[simp]: "inv_locate_a (as, am) (q, aaa, Bk # xs) ires
-       \<Longrightarrow> inv_locate_a (as, am) (q, aaa, Oc # xs) ires"
-apply(simp only: inv_locate_a.simps at_begin_norm.simps 
-                 at_begin_fst_bwtn.simps at_begin_fst_awtn.simps)
-apply(erule_tac disjE, erule exE, erule exE, erule exE, 
-      rule disjI2, rule disjI2)
-defer
-apply(erule_tac disjE, erule exE, erule exE, 
-      erule exE, rule disjI2, rule disjI2)
-prefer 2
-apply(simp)
-proof-
-  fix lm1 tn rn
-  assume k: "lm1 = am @ 0\<up>tn \<and> length lm1 = q \<and> (if lm1 = [] then aaa = Bk # Bk # 
-    ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> Bk # xs = Bk\<up>rn"
-  thus "\<exists>lm1 tn rn. lm1 = am @ 0 \<up> tn \<and> length lm1 = q \<and> 
-    (if lm1 = [] then aaa = Bk # Bk # ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> Oc # xs = [Oc] @ Bk \<up> rn"
-    (is "\<exists>lm1 tn rn. ?P lm1 tn rn")
-  proof -
-    from k have "?P lm1 tn (rn - 1)"
-      apply(auto simp: Oc_def)
-      by(case_tac [!] "rn::nat", auto)
-    thus ?thesis by blast
-  qed
-next
-  fix lm1 lm2 rn
-  assume h1: "am = lm1 @ lm2 \<and> length lm1 = q \<and> (if lm1 = [] 
-    then aaa = Bk # Bk # ires else aaa = Bk # <rev lm1> @ Bk # Bk # ires) \<and>
-    Bk # xs = <lm2> @ Bk\<up>rn"
-  from h1 have h2: "lm2 = []"
-    apply(auto split: if_splits)
-    apply(case_tac [!] lm2, simp_all add: tape_of_nl_cons split: if_splits)
-    done
-  from h1 and h2 show "\<exists>lm1 tn rn. lm1 = am @ 0\<up>tn \<and> length lm1 = q \<and> 
-    (if lm1 = [] then aaa = Bk # Bk # ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and>
-    Oc # xs = [Oc] @ Bk\<up>rn" 
-    (is "\<exists>lm1 tn rn. ?P lm1 tn rn")
-  proof -
-    from h1 and h2  have "?P lm1 0 (rn - 1)"
-      apply(auto simp: Oc_def
-                      tape_of_nl_abv tape_of_nat_list.simps)
-      by(case_tac "rn::nat", simp, simp)
-    thus ?thesis by blast
-  qed
-qed
-
-lemma [simp]: "inv_locate_a (as, am) (q, aaa, []) ires \<Longrightarrow> 
-               inv_locate_a (as, am) (q, aaa, [Oc]) ires"
-apply(insert locate_a_2_locate_a [of as am q aaa "[]"])
-apply(subgoal_tac "inv_locate_a (as, am) (q, aaa, [Bk]) ires", auto)
-done
-
-(*inv: from locate_b to locate_b*)
-lemma [simp]: "inv_locate_b (as, am) (q, aaa, Oc # xs) ires
-         \<Longrightarrow> inv_locate_b (as, am) (q, Oc # aaa, xs) ires"
-apply(simp only: inv_locate_b.simps in_middle.simps)
-apply(erule exE)+
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
-      rule_tac x = tn in exI, rule_tac x = m in exI)
-apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - 1" in exI,
-      rule_tac x = rn in exI)
-apply(case_tac mr, simp_all, auto)
-done
-
-(*
-lemma zero_and_nil[intro]: "(Bk # Bk\<^bsup>n\<^esup> = Oc\<^bsup>mr\<^esup> @ Bk # <lm::nat list> @ 
-                             Bk\<^bsup>rn \<^esup>) \<or> (lm2 = [] \<and> Bk # Bk\<^bsup>n\<^esup> = Oc\<^bsup>mr\<^esup>)
-       \<Longrightarrow> mr = 0 \<and> lm = []"
-apply(rule context_conjI)
-apply(case_tac mr, auto simp:exponent_def)
-apply(insert BkCons_nil[of "replicate (n - 1) Bk" lm rn])
-apply(case_tac n, auto simp: exponent_def Bk_def  tape_of_nl_nil_eq)
-done
-
-lemma tape_of_nat_def: "<[m::nat]> =  Oc # Oc\<^bsup>m\<^esup>"
-apply(simp add: tape_of_nl_abv tape_of_nat_list.simps)
-done
-*)
-lemma [simp]:  "<[x::nat]> = Oc\<up>(Suc x)"
-apply(simp add: tape_of_nat_abv tape_of_nl_abv)
-done
-
-lemma [simp]: " <([]::nat list)> = []"
-apply(simp add: tape_of_nl_abv)
-done
-
-lemma [simp]: "\<lbrakk>inv_locate_b (as, am) (q, aaa, Bk # xs) ires; \<exists>n. xs = Bk\<up>n\<rbrakk>
-            \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires"
-apply(simp add: inv_locate_b.simps inv_locate_a.simps)
-apply(rule_tac disjI2, rule_tac disjI1)
-apply(simp only: in_middle.simps at_begin_fst_bwtn.simps)
-apply(erule_tac exE)+
-apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = tn in exI, simp split: if_splits)
-apply(case_tac mr, simp_all)
-apply(case_tac "length am", simp_all, case_tac tn, simp_all)
-apply(case_tac lm2, simp_all add: tape_of_nl_cons split: if_splits)
-apply(case_tac am, simp_all)
-apply(case_tac n, simp_all)
-apply(case_tac n, simp_all)
-apply(case_tac mr, simp_all)
-apply(case_tac lm2, simp_all add: tape_of_nl_cons split: if_splits, auto)
-apply(case_tac [!] n, simp_all)
-done
-
-lemma [simp]: "(Oc # r = Bk \<up> rn) = False"
-apply(case_tac rn, simp_all)
-done
-
-lemma [simp]: "(\<exists>rna. Bk \<up> rn = Bk # Bk \<up> rna) \<or> rn = 0"
-apply(case_tac rn, auto)
-done
-
-lemma [simp]: "(\<forall> x. a \<noteq> x) = False"
-by auto
-
-lemma exp_ind: "a\<up>(Suc x) = a\<up>x @ [a]"
-apply(induct x, auto)
-done
-
-lemma [simp]: 
-      "inv_locate_a (as, lm) (q, l, Oc # r) ires
-       \<Longrightarrow> inv_locate_b (as, lm) (q, Oc # l, r) ires"
-apply(simp only: inv_locate_a.simps inv_locate_b.simps in_middle.simps
-          at_begin_norm.simps at_begin_fst_bwtn.simps
-          at_begin_fst_awtn.simps)
-apply(erule disjE, erule exE, erule exE, erule exE)
-apply(rule_tac x = lm1 in exI, rule_tac x = "tl lm2" in exI, simp)
-apply(rule_tac x = 0 in exI, rule_tac x = "hd lm2" in exI)
-apply(case_tac lm2, auto simp: tape_of_nl_cons )
-apply(rule_tac x = 1 in exI, rule_tac x = a in exI, auto)
-apply(case_tac list, simp_all)
-apply(case_tac rn, simp_all)
-apply(rule_tac x = "lm @ replicate tn 0" in exI, 
-      rule_tac x = "[]" in exI, 
-      rule_tac x = "Suc tn" in exI, 
-      rule_tac x = 0 in exI, auto)
-apply(simp only: replicate_Suc[THEN sym] exp_ind)
-apply(rule_tac x = "Suc 0" in exI, auto)
-done
-
-lemma length_equal: "xs = ys \<Longrightarrow> length xs = length ys"
-by auto
-
-lemma [simp]: "\<lbrakk>inv_locate_b (as, am) (q, aaa, Bk # xs) ires; 
-                \<not> (\<exists>n. xs = Bk\<up>n)\<rbrakk> 
-       \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires"
-apply(simp add: inv_locate_b.simps inv_locate_a.simps)
-apply(rule_tac disjI1)
-apply(simp only: in_middle.simps at_begin_norm.simps)
-apply(erule_tac exE)+
-apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = lm2 in exI, simp)
-apply(subgoal_tac "tn = 0", simp , auto split: if_splits)
-apply(case_tac [!] mr, simp_all, auto)
-apply(simp add: tape_of_nl_cons)
-apply(drule_tac length_equal, simp)
-apply(case_tac "length am", simp_all, erule_tac x = rn in allE, simp)
-apply(drule_tac length_equal, simp)
-apply(case_tac "(Suc (length lm1) - length am)", simp_all)
-apply(case_tac lm2, simp, simp)
-done
-
-lemma locate_b_2_a[intro]: 
-       "inv_locate_b (as, am) (q, aaa, Bk # xs) ires
-    \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires"
-apply(case_tac "\<exists> n. xs = Bk\<up>n", simp, simp)
-done
-
-
-lemma [simp]:  "inv_locate_b (as, am) (q, l, []) ires 
-           \<Longrightarrow>  inv_locate_b (as, am) (q, l, [Bk]) ires"
-apply(simp only: inv_locate_b.simps in_middle.simps)
-apply(erule exE)+
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
-      rule_tac x = tn in exI, rule_tac x = m in exI, 
-      rule_tac x = ml in exI, rule_tac x = mr in exI)
-apply(auto)
-done
-
-(*inv: from locate_b to after_write*)
-
-lemma [simp]: "(a mod 2 \<noteq> Suc 0) = (a mod 2 = 0)  "
-by arith
-
-lemma [simp]: "(a mod 2 \<noteq> 0) = (a mod 2 = Suc 0)  "
-by arith
-
-lemma mod_ex1: "(a mod 2 = Suc 0) = (\<exists> q. a = Suc (2 * q))"
-by arith
-
-lemma mod_ex2: "(a mod (2::nat) = 0) = (\<exists> q. a = 2 * q)"
-by arith
-
-lemma [simp]: "(2*q - Suc 0) div 2 = (q - 1)"
-by arith
-
-lemma [simp]: "(Suc (2*q)) div 2 = q"
-by arith
-
-lemma mod_2: "x mod 2  = 0 \<or>  x mod 2 = Suc 0"
-by arith
-
-lemma [simp]: "x mod 2 = 0 \<Longrightarrow> Suc x mod 2 = Suc 0"
-by arith
-
-lemma [simp]: "x mod 2 = Suc 0 \<Longrightarrow> Suc x mod 2 = 0"
-by arith
-
-lemma [simp]:  "inv_locate_b (as, am) (q, l, []) ires 
-           \<Longrightarrow>  inv_locate_b (as, am) (q, l, [Bk]) ires"
-apply(simp only: inv_locate_b.simps in_middle.simps)
-apply(erule exE)+
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
-      rule_tac x = tn in exI, rule_tac x = m in exI, 
-      rule_tac x = ml in exI, rule_tac x = mr in exI)
-apply(auto)
-done
-
-lemma locate_b_2_locate_a[simp]: 
-    "\<lbrakk>q > 0;  inv_locate_b (as, am) (q - Suc 0, aaa, Bk # xs) ires\<rbrakk>
-   \<Longrightarrow>  inv_locate_a (as, am) (q, Bk # aaa, xs) ires"
-apply(insert locate_b_2_a [of as am "q - 1" aaa xs ires], simp)
-done
-
-
-lemma [simp]:  "inv_locate_b (as, am) (q, l, []) ires 
-           \<Longrightarrow>  inv_locate_b (as, am) (q, l, [Bk]) ires"
-apply(simp only: inv_locate_b.simps in_middle.simps)
-apply(erule exE)+
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
-      rule_tac x = tn in exI, rule_tac x = m in exI, 
-      rule_tac x = ml in exI, rule_tac x = mr in exI)
-apply(auto)
-done
-
-(*inv: from locate_b to after_write*)
-
-lemma [simp]:
-  "crsp (layout_of ap) (as, lm) (s, l, r) ires
-  \<Longrightarrow> findnth_inv (layout_of ap) n (as, lm) (Suc 0, l, r) ires"
-apply(auto simp: crsp.simps findnth_inv.simps inv_locate_a.simps
-               at_begin_norm.simps at_begin_fst_awtn.simps at_begin_fst_bwtn.simps)
-done
-
-lemma findnth_correct_pre: 
-  assumes layout: "ly = layout_of ap"
-  and crsp: "crsp ly (as, lm) (s, l, r) ires"
-  and not0: "n > 0"
-  and f: "f = (\<lambda> stp. (steps (Suc 0, l, r) (findnth n, 0) stp, n))"
-  and P: "P = (\<lambda> ((s, l, r), n). s = Suc (2 * n))"
-  and Q: "Q = (\<lambda> ((s, l, r), n). findnth_inv ly n (as, lm) (s, l, r) ires)"
-  shows "\<exists> stp. P (f stp) \<and> Q (f stp)"
-thm halt_lemma2
-proof(rule_tac LE = findnth_LE in halt_lemma2)
-  show "wf findnth_LE"  by(intro wf_findnth_LE)
-next
-  show "Q (f 0)"
-    using crsp layout
-    apply(simp add: f P Q steps.simps)
-    done
-next
-  show "\<not> P (f 0)"
-    using not0
-    apply(simp add: f P steps.simps)
-    done
-next
-  show "\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) 
-        \<in> findnth_LE"
-  proof(rule_tac allI, rule_tac impI, simp add: f, 
-      case_tac "steps (Suc 0, l, r) (findnth n, 0) na", simp add: P)
-    fix na a b c
-    assume "a \<noteq> Suc (2 * n) \<and> Q ((a, b, c), n)"
-    thus  "Q (step (a, b, c) (findnth n, 0), n) \<and> 
-        ((step (a, b, c) (findnth n, 0), n), (a, b, c), n) \<in> findnth_LE"
-      apply(case_tac c, case_tac [2] aa)
-      apply(simp_all add: step.simps findnth_LE_def Q findnth_inv.simps mod_2  lex_pair_def split: if_splits)
-      apply(auto simp: mod_ex1 mod_ex2)
-      done
-  qed
-qed
-            
-lemma [intro]: "inv_locate_a (as, lm) (0, Bk # Bk # ires, <lm> @ Bk \<up> x) ires"
-apply(auto simp: crsp.simps inv_locate_a.simps at_begin_norm.simps)
-done
-lemma [simp]: "crsp ly (as, lm) (s, l, r) ires \<Longrightarrow> inv_locate_a (as, lm) (0, l, r) ires"
-apply(auto simp: crsp.simps inv_locate_a.simps at_begin_norm.simps)
-done
-
-lemma findnth_correct: 
-  assumes layout: "ly = layout_of ap"
-  and crsp: "crsp ly (as, lm) (s, l, r) ires"
-  shows "\<exists> stp l' r'. steps (Suc 0, l, r) (findnth n, 0) stp = (Suc (2 * n), l', r')
-              \<and> inv_locate_a (as, lm) (n, l', r') ires"
-  using crsp
-  apply(case_tac "n = 0")
-  apply(rule_tac x = 0 in exI, auto simp: steps.simps)
-  using assms
-  apply(drule_tac findnth_correct_pre, auto)
-  apply(rule_tac x = stp in exI, simp add: findnth_inv.simps)
-  done
-
-
-fun inc_inv :: "nat \<Rightarrow> inc_inv_t"
-  where
-  "inc_inv n (as, lm) (s, l, r) ires =
-              (let lm' = abc_lm_s lm n (Suc (abc_lm_v lm n)) in
-                if s = 0 then False
-                else if s = 1 then 
-                   inv_locate_a (as, lm) (n, l, r) ires
-                else if s = 2 then 
-                   inv_locate_b (as, lm) (n, l, r) ires
-                else if s = 3 then 
-                   inv_after_write (as, lm') (s, l, r) ires
-                else if s = Suc 3 then 
-                   inv_after_move (as, lm') (s, l, r) ires
-                else if s = Suc 4 then 
-                   inv_after_clear (as, lm') (s, l, r) ires
-                else if s = Suc (Suc 4) then 
-                   inv_on_right_moving (as, lm') (s, l, r) ires
-                else if s = Suc (Suc 5) then 
-                   inv_on_left_moving (as, lm') (s, l, r) ires
-                else if s = Suc (Suc (Suc 5)) then 
-                   inv_check_left_moving (as, lm') (s, l, r) ires
-                else if s = Suc (Suc (Suc (Suc 5))) then 
-                   inv_after_left_moving (as, lm') (s, l, r) ires
-                else if s = Suc (Suc (Suc (Suc (Suc 5)))) then 
-                   inv_stop (as, lm') (s, l, r) ires
-                else False)"
-
-
-fun abc_inc_stage1 :: "config \<Rightarrow> nat"
-  where
-  "abc_inc_stage1 (s, l, r) = 
-            (if s = 0 then 0
-             else if s \<le> 2 then 5
-             else if s \<le> 6 then 4
-             else if s \<le> 8 then 3
-             else if s = 9 then 2
-             else 1)"
-
-fun abc_inc_stage2 :: "config \<Rightarrow> nat"
-  where
-  "abc_inc_stage2 (s, l, r) =
-                (if s = 1 then 2
-                 else if s = 2 then 1
-                 else if s = 3 then length r
-                 else if s = 4 then length r
-                 else if s = 5 then length r
-                 else if s = 6 then 
-                                  if r \<noteq> [] then length r
-                                  else 1
-                 else if s = 7 then length l
-                 else if s = 8 then length l
-                 else 0)"
-
-fun abc_inc_stage3 :: "config \<Rightarrow>  nat"
-  where
-  "abc_inc_stage3 (s, l, r) = (
-              if s = 4 then 4
-              else if s = 5 then 3
-              else if s = 6 then 
-                   if r \<noteq> [] \<and> hd r = Oc then 2
-                   else 1
-              else if s = 3 then 0
-              else if s = 2 then length r
-              else if s = 1 then 
-                      if (r \<noteq> [] \<and> hd r = Oc) then 0
-                      else 1
-              else 10 - s)"
-
-
-definition inc_measure :: "config \<Rightarrow> nat \<times> nat \<times> nat"
-  where
-  "inc_measure c = 
-    (abc_inc_stage1 c, abc_inc_stage2 c, abc_inc_stage3 c)"
-
-definition lex_triple :: 
-   "((nat \<times> (nat \<times> nat)) \<times> (nat \<times> (nat \<times> nat))) set"
-  where "lex_triple \<equiv> less_than <*lex*> lex_pair"
-
-definition inc_LE :: "(config \<times> config) set"
-  where
-  "inc_LE \<equiv> (inv_image lex_triple inc_measure)"
-
-declare inc_inv.simps[simp del]
-
-lemma wf_inc_le[intro]: "wf inc_LE"
-by(auto intro:wf_inv_image simp: inc_LE_def lex_triple_def lex_pair_def)
-
-lemma numeral_5_eq_5: "5 = Suc (Suc (Suc (Suc (Suc 0))))"
-by arith
-
-lemma numeral_6_eq_6: "6 = Suc (Suc (Suc (Suc (Suc 1))))"
-by arith
-
-lemma numeral_7_eq_7: "7 = Suc (Suc (Suc (Suc (Suc 2))))"
-by arith
-
-lemma numeral_8_eq_8: "8 = Suc (Suc (Suc (Suc (Suc 3))))"
-by arith
-
-lemma numeral_9_eq_9: "9 = Suc (Suc (Suc (Suc (Suc (Suc 3)))))"
-by arith
-
-lemma numeral_10_eq_10: "10 = Suc (Suc (Suc (Suc (Suc (Suc (Suc 3))))))"
-by arith
-
-lemma inv_locate_b_2_after_write[simp]: 
-      "inv_locate_b (as, am) (n, aaa, Bk # xs) ires
-      \<Longrightarrow> inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n)))
-          (s, aaa, Oc # xs) ires"
-apply(auto simp: in_middle.simps inv_after_write.simps 
-                 abc_lm_v.simps abc_lm_s.simps  inv_locate_b.simps)
-apply(case_tac [!] mr, auto split: if_splits)
-apply(rule_tac x = rn in exI, rule_tac x = "Suc m" in exI,
-      rule_tac x = "lm1" in exI, simp)
-apply(rule_tac x = "lm2" in exI, simp)
-apply(simp only: Suc_diff_le exp_ind)
-apply(subgoal_tac "lm2 = []", simp)
-apply(drule_tac length_equal, simp)
-done
-
-lemma [simp]: "inv_locate_b (as, am) (n, aaa, []) ires \<Longrightarrow> 
-     inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) 
-                     (s, aaa, [Oc]) ires"
-apply(insert inv_locate_b_2_after_write [of as am n aaa "[]"])
-by(simp)
-
-
-
-(*inv: from after_write to after_move*)
-lemma [simp]: "inv_after_write (as, lm) (x, l, Oc # r) ires
-                \<Longrightarrow> inv_after_move (as, lm) (y, Oc # l, r) ires"
-apply(auto simp:inv_after_move.simps inv_after_write.simps split: if_splits)
-done
-
-lemma [simp]: "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n)
-                )) (x, aaa, Bk # xs) ires = False"
-apply(simp add: inv_after_write.simps )
-done
-
-lemma [simp]: 
- "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) 
-                        (x, aaa, []) ires = False"
-apply(simp add: inv_after_write.simps )
-done
-
-(*inv: from after_move to after_clear*)
-lemma [simp]: "inv_after_move (as, lm) (s, l, Oc # r) ires
-                \<Longrightarrow> inv_after_clear (as, lm) (s', l, Bk # r) ires"
-apply(auto simp: inv_after_move.simps inv_after_clear.simps split: if_splits)
-done
-
-(*inv: from after_move to on_leftmoving*)
-lemma [intro]: "Bk \<up> rn = Bk # Bk \<up> (rn - Suc 0) \<or> rn = 0"
-apply(case_tac rn, auto)
-done
-
-lemma inv_after_move_2_inv_on_left_moving[simp]:  
-   "inv_after_move (as, lm) (s, l, Bk # r) ires
-   \<Longrightarrow> (l = [] \<longrightarrow> 
-         inv_on_left_moving (as, lm) (s', [], Bk # Bk # r) ires) \<and>
-      (l \<noteq> [] \<longrightarrow> 
-         inv_on_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)"
-apply(simp only: inv_after_move.simps inv_on_left_moving.simps)
-apply(subgoal_tac "l \<noteq> []", rule conjI, simp, rule impI, 
-                rule disjI1, simp only: inv_on_left_moving_norm.simps)
-apply(erule exE)+
-apply(subgoal_tac "lm2 = []")
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,  
-    rule_tac x = m in exI, rule_tac x = m in exI, 
-    rule_tac x = 1 in exI,  
-    rule_tac x = "rn - 1" in exI, auto)
-apply(auto split: if_splits)
-apply(case_tac [1-2] rn, simp_all)
-apply(case_tac [!] lm2, simp_all add: tape_of_nl_cons split: if_splits)
-done
-
-
-lemma inv_after_move_2_inv_on_left_moving_B[simp]: 
-    "inv_after_move (as, lm) (s, l, []) ires
-      \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s', [], [Bk]) ires) \<and>
-          (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s', tl l, [hd l]) ires)"
-apply(simp only: inv_after_move.simps inv_on_left_moving.simps)
-apply(subgoal_tac "l \<noteq> []", rule conjI, simp, rule impI, rule disjI1,
-        simp only: inv_on_left_moving_norm.simps)
-apply(erule exE)+
-apply(subgoal_tac "lm2 = []")
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,  
-      rule_tac x = m in exI, rule_tac x = m in exI, 
-      rule_tac x = 1 in exI, rule_tac x = "rn - 1" in exI, simp, case_tac rn)
-apply(auto split: if_splits)
-apply(case_tac [!] lm2, auto simp: tape_of_nl_cons split: if_splits)
-done
-
-(*inv: from after_clear to on_right_moving*)
-lemma [simp]: "Oc # r = replicate rn Bk = False"
-apply(case_tac rn, simp, simp)
-done
-
-lemma inv_after_clear_2_inv_on_right_moving[simp]: 
-     "inv_after_clear (as, lm) (x, l, Bk # r) ires
-      \<Longrightarrow> inv_on_right_moving (as, lm) (y, Bk # l, r) ires"
-apply(auto simp: inv_after_clear.simps inv_on_right_moving.simps )
-apply(subgoal_tac "lm2 \<noteq> []")
-apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI, 
-      rule_tac x = "hd lm2" in exI, simp)
-apply(rule_tac x = 0 in exI, rule_tac x = "hd lm2" in exI)
-apply(simp, rule conjI)
-apply(case_tac [!] "lm2::nat list", auto)
-apply(case_tac rn, auto split: if_splits simp: tape_of_nl_cons)
-apply(case_tac [!] rn, simp_all)
-done
-
-lemma [simp]: "inv_after_clear (as, lm) (x, l, []) ires\<Longrightarrow> 
-               inv_after_clear (as, lm) (y, l, [Bk]) ires" 
-by(auto simp: inv_after_clear.simps)
-
-lemma [simp]: "inv_after_clear (as, lm) (x, l, []) ires
-             \<Longrightarrow> inv_on_right_moving (as, lm) (y, Bk # l, []) ires"
-by(insert 
-    inv_after_clear_2_inv_on_right_moving[of as lm n l "[]"], simp)
-
-(*inv: from on_right_moving to on_right_movign*)
-lemma [simp]: "inv_on_right_moving (as, lm) (x, l, Oc # r) ires
-      \<Longrightarrow> inv_on_right_moving (as, lm) (y, Oc # l, r) ires"
-apply(auto simp: inv_on_right_moving.simps)
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
-           rule_tac x = "ml + mr" in exI, simp)
-apply(rule_tac x = "Suc ml" in exI, 
-           rule_tac x = "mr - 1" in exI, simp)
-apply(case_tac mr, auto)
-apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, 
-      rule_tac x = "ml + mr" in exI, simp)
-apply(rule_tac x = "Suc ml" in exI, 
-      rule_tac x = "mr - 1" in exI, simp)
-apply(case_tac mr, auto split: if_splits)
-done
-
-lemma inv_on_right_moving_2_inv_on_right_moving[simp]: 
-     "inv_on_right_moving (as, lm) (x, l, Bk # r) ires
-     \<Longrightarrow> inv_after_write (as, lm) (y, l, Oc # r) ires"
-apply(auto simp: inv_on_right_moving.simps inv_after_write.simps )
-apply(case_tac mr, auto simp: split: if_splits)
-done
-      
-lemma [simp]: "inv_on_right_moving (as, lm) (x, l, []) ires\<Longrightarrow> 
-             inv_on_right_moving (as, lm) (y, l, [Bk]) ires"
-apply(auto simp: inv_on_right_moving.simps)
-apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, simp)
-done
-
-(*inv: from on_right_moving to after_write*)
-lemma [simp]: "inv_on_right_moving (as, lm) (x, l, []) ires
-       \<Longrightarrow> inv_after_write (as, lm) (y, l, [Oc]) ires"
-apply(rule_tac inv_on_right_moving_2_inv_on_right_moving, simp)
-done
-
-(*inv: from on_left_moving to on_left_moving*)
-lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) 
-               (s, l, Oc # r) ires = False"
-apply(auto simp: inv_on_left_moving_in_middle_B.simps )
-done
-
-lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, l, Bk # r) ires 
-             = False"
-apply(auto simp: inv_on_left_moving_norm.simps)
-apply(case_tac [!] mr, auto simp: )
-done
-
-lemma [simp]: 
-  "\<lbrakk>inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires;
-    hd l = Bk; l \<noteq> []\<rbrakk> \<Longrightarrow> 
-     inv_on_left_moving_in_middle_B (as, lm) (s, tl l, Bk # Oc # r) ires"
-apply(case_tac l, simp, simp)
-apply(simp only: inv_on_left_moving_norm.simps 
-                 inv_on_left_moving_in_middle_B.simps)
-apply(erule_tac exE)+
-apply(rule_tac x = lm1 in exI, rule_tac x = "m # lm2" in exI, auto)
-apply(case_tac [!] ml, auto)
-apply(auto simp: tape_of_nl_cons split: if_splits)
-apply(rule_tac [!] x = "Suc rn" in exI, simp_all)
-done
-
-lemma [simp]: "\<lbrakk>inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires; 
-                hd l = Oc; l \<noteq> []\<rbrakk>
-            \<Longrightarrow> inv_on_left_moving_norm (as, lm) 
-                                        (s, tl l, Oc # Oc # r) ires"
-apply(simp only: inv_on_left_moving_norm.simps)
-apply(erule exE)+
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
-      rule_tac x = m in exI, rule_tac x = "ml - 1" in exI,
-      rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI, simp)
-apply(case_tac ml, auto simp: split: if_splits)
-done
-
-lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, [], Oc # r) ires
-     \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # Oc # r) ires"
-apply(auto simp: inv_on_left_moving_norm.simps 
-                 inv_on_left_moving_in_middle_B.simps split: if_splits)
-done
-
-lemma [simp]:"inv_on_left_moving (as, lm) (s, l, Oc # r) ires
-    \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s, [], Bk # Oc # r) ires)
- \<and>  (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s, tl l, hd l # Oc # r) ires)"
-apply(simp add: inv_on_left_moving.simps)
-apply(case_tac "l \<noteq> []", rule conjI, simp, simp)
-apply(case_tac "hd l", simp, simp, simp)
-done
-
-(*inv: from on_left_moving to check_left_moving*)
-lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) 
-                                      (s, Bk # list, Bk # r) ires
-          \<Longrightarrow> inv_check_left_moving_on_leftmost (as, lm) 
-                                      (s', list, Bk # Bk # r) ires"
-apply(auto simp: inv_on_left_moving_in_middle_B.simps 
-                 inv_check_left_moving_on_leftmost.simps split: if_splits)
-apply(case_tac [!] "rev lm1", simp_all)
-apply(case_tac [!] lista, simp_all add: tape_of_nl_abv tape_of_nat_abv tape_of_nat_list.simps)
-done
-
-lemma [simp]:
-    "inv_check_left_moving_in_middle (as, lm) (s, l, Bk # r) ires= False"
-by(auto simp: inv_check_left_moving_in_middle.simps )
-
-lemma [simp]: 
- "inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # r) ires\<Longrightarrow> 
-  inv_check_left_moving_on_leftmost (as, lm) (s', [], Bk # Bk # r) ires"
-apply(auto simp: inv_on_left_moving_in_middle_B.simps 
-                 inv_check_left_moving_on_leftmost.simps split: if_splits)
-done
-
-lemma [simp]: "inv_check_left_moving_on_leftmost (as, lm) 
-                                       (s, list, Oc # r) ires= False"
-by(auto simp: inv_check_left_moving_on_leftmost.simps split: if_splits)
-
-lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) 
-                                         (s, Oc # list, Bk # r) ires
- \<Longrightarrow> inv_check_left_moving_in_middle (as, lm) (s', list, Oc # Bk # r) ires"
-apply(auto simp: inv_on_left_moving_in_middle_B.simps 
-                 inv_check_left_moving_in_middle.simps  split: if_splits)
-done
-
-lemma inv_on_left_moving_2_check_left_moving[simp]:
- "inv_on_left_moving (as, lm) (s, l, Bk # r) ires
- \<Longrightarrow> (l = [] \<longrightarrow> inv_check_left_moving (as, lm) (s', [], Bk # Bk # r) ires)
- \<and> (l \<noteq> [] \<longrightarrow> 
-      inv_check_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)"
-apply(simp add: inv_on_left_moving.simps inv_check_left_moving.simps)
-apply(case_tac l, simp, simp)
-apply(case_tac a, simp, simp)
-done
-
-lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, l, []) ires = False"
-apply(auto simp: inv_on_left_moving_norm.simps)
-done
-
-lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires\<Longrightarrow> 
-     inv_on_left_moving (as, lm) (6 + 2 * n, l, [Bk]) ires"
-apply(simp add: inv_on_left_moving.simps)
-apply(auto simp: inv_on_left_moving_in_middle_B.simps)
-done
-
-lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires = False"
-apply(simp add: inv_on_left_moving.simps)
-apply(simp add: inv_on_left_moving_in_middle_B.simps)
-done
-
-lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires
- \<Longrightarrow> (l = [] \<longrightarrow> inv_check_left_moving (as, lm) (s', [], [Bk]) ires) \<and>
-    (l \<noteq> [] \<longrightarrow> inv_check_left_moving (as, lm) (s', tl l, [hd l]) ires)"
-by simp
-
-lemma [intro]: "\<exists>rna. Bk # Bk \<up> rn = Bk \<up> rna"
-apply(rule_tac x = "Suc rn" in exI, simp)
-done
-
-lemma 
-inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[simp]:
-"inv_check_left_moving_in_middle (as, lm) (s, Bk # list, Oc # r) ires
-  \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s', list, Bk # Oc # r) ires"
-apply(simp only: inv_check_left_moving_in_middle.simps 
-                 inv_on_left_moving_in_middle_B.simps)
-apply(erule_tac exE)+
-apply(rule_tac x = "rev (tl (rev lm1))" in exI, 
-      rule_tac x = "[hd (rev lm1)] @ lm2" in exI, auto)
-apply(case_tac [!] "rev lm1",simp_all add: tape_of_nat_abv tape_of_nl_abv tape_of_nat_list.simps)
-apply(case_tac [!] a, simp_all)
-apply(case_tac [1] lm2, simp_all add: tape_of_nat_list.simps tape_of_nat_abv, auto)
-apply(case_tac [3] lm2, simp_all add: tape_of_nat_list.simps tape_of_nat_abv, auto)
-apply(case_tac [!] lista, simp_all add: tape_of_nat_abv tape_of_nat_list.simps)
-done
-
-lemma [simp]: 
- "inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires\<Longrightarrow>
-     inv_check_left_moving_in_middle (as, lm) (s', [Bk], Oc # r) ires"
-apply(auto simp: inv_check_left_moving_in_middle.simps )
-done
-
-lemma [simp]: 
- "inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires
-   \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s', [], Bk # Oc # r) ires"
-apply(insert 
-inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[of 
-                  as lm n "[]" r], simp)
-done 
-
-lemma [simp]: "inv_check_left_moving_in_middle (as, lm) 
-                       (s, Oc # list, Oc # r) ires
-   \<Longrightarrow> inv_on_left_moving_norm (as, lm) (s', list, Oc # Oc # r) ires"
-apply(auto simp: inv_check_left_moving_in_middle.simps 
-                 inv_on_left_moving_norm.simps)
-apply(rule_tac x = "rev (tl (rev lm1))" in exI, 
-      rule_tac x = lm2 in exI, rule_tac x = "hd (rev lm1)" in exI)
-apply(rule_tac conjI)
-apply(case_tac "rev lm1", simp, simp)
-apply(rule_tac x = "hd (rev lm1) - 1" in exI, auto)
-apply(rule_tac [!] x = "Suc (Suc 0)" in exI, simp)
-apply(case_tac [!] "rev lm1", simp_all)
-apply(case_tac [!] a, simp_all add: tape_of_nl_cons split: if_splits)
-done
-
-lemma [simp]: "inv_check_left_moving (as, lm) (s, l, Oc # r) ires
-\<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s', [], Bk # Oc # r) ires) \<and>
-   (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s', tl l, hd l # Oc # r) ires)"
-apply(case_tac l, 
-      auto simp: inv_check_left_moving.simps inv_on_left_moving.simps)
-apply(case_tac a, simp, simp)
-done
-
-(*inv: check_left_moving to after_left_moving*)
-lemma [simp]: "inv_check_left_moving (as, lm) (s, l, Bk # r) ires
-                \<Longrightarrow> inv_after_left_moving (as, lm) (s', Bk # l, r) ires"
-apply(auto simp: inv_check_left_moving.simps 
- inv_check_left_moving_on_leftmost.simps inv_after_left_moving.simps)
-done
-
-
-lemma [simp]:"inv_check_left_moving (as, lm) (s, l, []) ires
-      \<Longrightarrow> inv_after_left_moving (as, lm) (s', Bk # l, []) ires"
-by(simp add: inv_check_left_moving.simps  
-inv_check_left_moving_in_middle.simps 
-inv_check_left_moving_on_leftmost.simps)
-
-(*inv: after_left_moving to inv_stop*)
-lemma [simp]: "inv_after_left_moving (as, lm) (s, l, Bk # r) ires
-       \<Longrightarrow> inv_stop (as, lm) (s', Bk # l, r) ires"
-apply(auto simp: inv_after_left_moving.simps inv_stop.simps)
-done
-
-lemma [simp]: "inv_after_left_moving (as, lm) (s, l, []) ires
-             \<Longrightarrow> inv_stop (as, lm) (s', Bk # l, []) ires"
-by(auto simp: inv_after_left_moving.simps)
-
-(*inv: stop to stop*)
-lemma [simp]: "inv_stop (as, lm) (x, l, r) ires \<Longrightarrow> 
-               inv_stop (as, lm) (y, l, r) ires"
-apply(simp add: inv_stop.simps)
-done
-
-lemma [simp]: "inv_after_clear (as, lm) (s, aaa, Oc # xs) ires= False"
-apply(auto simp: inv_after_clear.simps )
-done
-
-lemma [simp]: 
-  "inv_after_left_moving (as, lm) (s, aaa, Oc # xs) ires = False"
-by(auto simp: inv_after_left_moving.simps  )
-
-lemma [simp]: "inv_after_clear (as, abc_lm_s lm n (Suc (abc_lm_v lm n))) (s, b, []) ires = False"
-apply(auto simp: inv_after_clear.simps)
-done
-
-lemma [simp]: "inv_on_left_moving (as, abc_lm_s lm n (Suc (abc_lm_v lm n))) 
-           (s, b, Oc # list) ires \<Longrightarrow> b \<noteq> []"
-apply(auto simp: inv_on_left_moving.simps inv_on_left_moving_norm.simps split: if_splits)
-done
-
-lemma [simp]: "inv_check_left_moving (as, abc_lm_s lm n (Suc (abc_lm_v lm n))) (s, b, Oc # list) ires \<Longrightarrow> b \<noteq> []"
-apply(auto simp: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps split: if_splits)
-done
-
-lemma tinc_correct_pre:
-  assumes layout: "ly = layout_of ap"
-  and inv_start: "inv_locate_a (as, lm) (n, l, r) ires"
-  and lm': "lm' = abc_lm_s lm n (Suc (abc_lm_v lm n))"
-  and f: "f = steps (Suc 0, l, r) (tinc_b, 0)"
-  and P: "P = (\<lambda> (s, l, r). s = 10)"
-  and Q: "Q = (\<lambda> (s, l, r). inc_inv n (as, lm) (s, l, r) ires)" 
-  shows "\<exists> stp. P (f stp) \<and> Q (f stp)"
-proof(rule_tac LE = inc_LE in halt_lemma2)
-  show "wf inc_LE" by(auto)
-next
-  show "Q (f 0)"
-    using inv_start
-    apply(simp add: f P Q steps.simps inc_inv.simps)
-    done
-next
-  show "\<not> P (f 0)"
-    apply(simp add: f P steps.simps)
-    done
-next
-  show "\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) 
-        \<in> inc_LE"
-  proof(rule_tac allI, rule_tac impI, simp add: f, 
-      case_tac "steps (Suc 0, l, r) (tinc_b, 0) n", simp add: P)
-    fix n a b c
-    assume "a \<noteq> 10 \<and> Q (a, b, c)"
-    thus  "Q (step (a, b, c) (tinc_b, 0)) \<and> (step (a, b, c) (tinc_b, 0), a, b, c) \<in> inc_LE"
-      apply(simp add:Q)
-      apply(simp add: inc_inv.simps)
-      apply(case_tac c, case_tac [2] aa)
-      apply(auto simp: Let_def step.simps tinc_b_def numeral_2_eq_2 numeral_3_eq_3  split: if_splits)
-      apply(simp_all add: inc_inv.simps inc_LE_def lex_triple_def lex_pair_def inc_measure_def numeral_5_eq_5
-                          numeral_6_eq_6 numeral_7_eq_7 numeral_8_eq_8 numeral_9_eq_9)         
-      done
-  qed
-qed
-         
-
-lemma tinc_correct: 
-  assumes layout: "ly = layout_of ap"
-  and inv_start: "inv_locate_a (as, lm) (n, l, r) ires"
-  and lm': "lm' = abc_lm_s lm n (Suc (abc_lm_v lm n))"
-  shows "\<exists> stp l' r'. steps (Suc 0, l, r) (tinc_b, 0) stp = (10, l', r')
-              \<and> inv_stop (as, lm') (10, l', r') ires"
-  using assms
-  apply(drule_tac tinc_correct_pre, auto)
-  apply(rule_tac x = stp in exI, simp)
-  apply(simp add: inc_inv.simps)
-  done
-
-declare inv_locate_a.simps[simp del] abc_lm_s.simps[simp del]
-        abc_lm_v.simps[simp del]
-
-lemma [simp]: "(4::nat) * n mod 2 = 0"
-apply(arith)
-done
-
-lemma crsp_step_inc_pre:
-  assumes layout: "ly = layout_of ap"
-  and crsp: "crsp ly (as, lm) (s, l, r) ires"
-  and aexec: "abc_step_l (as, lm) (Some (Inc n)) = (asa, lma)"
-  shows "\<exists> stp k. steps (Suc 0, l, r) (findnth n @ shift tinc_b (2 * n), 0) stp 
-        = (2*n + 10, Bk # Bk # ires, <lma> @ Bk\<up>k) \<and> stp > 0"
-proof -
-  thm tm_append_steps
-  have "\<exists> stp l' r'. steps (Suc 0, l, r) (findnth n, 0) stp = (Suc (2 * n), l', r')
-    \<and> inv_locate_a (as, lm) (n, l', r') ires"
-    using assms
-    apply(rule_tac findnth_correct, simp_all add: crsp layout)
-    done
-  from this obtain stp l' r' where a:
-    "steps (Suc 0, l, r) (findnth n, 0) stp = (Suc (2 * n), l', r')
-    \<and> inv_locate_a (as, lm) (n, l', r') ires" by blast
-  moreover have
-    "\<exists> stp la ra. steps (Suc 0, l', r') (tinc_b, 0) stp = (10, la, ra)
-                        \<and> inv_stop (as, lma) (10, la, ra) ires"
-    using assms a
-  proof(rule_tac lm' = lma and n = n and lm = lm and ly = ly and ap = ap in tinc_correct,
-      simp, simp)
-    show "lma = abc_lm_s lm n (Suc (abc_lm_v lm n))"
-      using aexec
-      apply(simp add: abc_step_l.simps)
-      done
-  qed
-  from this obtain stpa la ra where b:
-    "steps (Suc 0, l', r') (tinc_b, 0) stpa = (10, la, ra)
-    \<and> inv_stop (as, lma) (10, la, ra) ires" by blast
-  from a b show "\<exists>stp k. steps (Suc 0, l, r) (findnth n @ shift tinc_b (2 * n), 0) stp
-    = (2 * n + 10, Bk # Bk # ires, <lma> @ Bk \<up> k) \<and> stp > 0"
-    apply(rule_tac x = "stp + stpa" in exI)
-    using tm_append_steps[of "Suc 0" l r "findnth n" stp l' r' tinc_b stpa 10 la ra "length (findnth n) div 2"]
-    apply(simp add: length_findnth inv_stop.simps)
-    apply(case_tac stpa, simp_all add: steps.simps)
-    done
-qed 
-     
-lemma crsp_step_inc:
-  assumes layout: "ly = layout_of ap"
-  and crsp: "crsp ly (as, lm) (s, l, r) ires"
-  and fetch: "abc_fetch as ap = Some (Inc n)"
-  shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Inc n)))
-  (steps (s, l, r) (ci ly (start_of ly as) (Inc n), start_of ly as - Suc 0) stp) ires"
-proof(case_tac "(abc_step_l (as, lm) (Some (Inc n)))")
-  fix a b
-  assume aexec: "abc_step_l (as, lm) (Some (Inc n)) = (a, b)"
-  then have "\<exists> stp k. steps (Suc 0, l, r) (findnth n @ shift tinc_b (2 * n), 0) stp 
-        = (2*n + 10, Bk # Bk # ires, <b> @ Bk\<up>k) \<and> stp > 0"
-    using assms
-    apply(rule_tac crsp_step_inc_pre, simp_all)
-    done
-  thus "?thesis"
-    using assms aexec
-    apply(erule_tac exE)
-    apply(erule_tac exE)
-    apply(erule_tac conjE)
-    apply(rule_tac x = stp in exI, simp add: ci.simps tm_shift_eq_steps)
-    apply(drule_tac off = "(start_of (layout_of ap) as - Suc 0)" in tm_shift_eq_steps)
-    apply(auto simp: crsp.simps abc_step_l.simps fetch start_of_Suc1)
-    done
-qed
-    
-subsection{* Crsp of Dec n e*}
-declare sete.simps[simp del]
-
-type_synonym dec_inv_t = "(nat * nat list) \<Rightarrow> config \<Rightarrow> cell list \<Rightarrow>  bool"
-
-fun dec_first_on_right_moving :: "nat \<Rightarrow> dec_inv_t"
-  where
-  "dec_first_on_right_moving n (as, lm) (s, l, r) ires = 
-               (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
-         ml + mr = Suc m \<and> length lm1 = n \<and> ml > 0 \<and> m > 0 \<and>
-             (if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires
-                          else  l = Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
-    ((r = Oc\<up>mr @ [Bk] @ <lm2> @ Bk\<up>rn) \<or> (r = Oc\<up>mr \<and> lm2 = [])))"
-
-fun dec_on_right_moving :: "dec_inv_t"
-  where
-  "dec_on_right_moving (as, lm) (s, l, r) ires =  
-   (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
-                             ml + mr = Suc (Suc m) \<and>
-   (if lm1 = [] then l = Oc\<up>ml@ Bk # Bk # ires
-                else  l = Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
-   ((r = Oc\<up>mr @ [Bk] @ <lm2> @ Bk\<up>rn) \<or> (r = Oc\<up>mr \<and> lm2 = [])))"
-
-fun dec_after_clear :: "dec_inv_t"
-  where
-  "dec_after_clear (as, lm) (s, l, r) ires = 
-              (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
-                ml + mr = Suc m \<and> ml = Suc m \<and> r \<noteq> [] \<and> r \<noteq> [] \<and>
-               (if lm1 = [] then l = Oc\<up>ml@ Bk # Bk # ires
-                            else l = Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
-               (tl r = Bk # <lm2> @ Bk\<up>rn \<or> tl r = [] \<and> lm2 = []))"
-
-fun dec_after_write :: "dec_inv_t"
-  where
-  "dec_after_write (as, lm) (s, l, r) ires = 
-         (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
-       ml + mr = Suc m \<and> ml = Suc m \<and> lm2 \<noteq> [] \<and>
-       (if lm1 = [] then l = Bk # Oc\<up>ml @ Bk # Bk # ires
-                    else l = Bk # Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
-       tl r = <lm2> @ Bk\<up>rn)"
-
-fun dec_right_move :: "dec_inv_t"
-  where
-  "dec_right_move (as, lm) (s, l, r) ires = 
-        (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 
-            \<and> ml = Suc m \<and> mr = (0::nat) \<and> 
-              (if lm1 = [] then l = Bk # Oc\<up>ml @ Bk # Bk # ires
-                          else l = Bk # Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) 
-           \<and> (r = Bk # <lm2> @ Bk\<up>rn \<or> r = [] \<and> lm2 = []))"
-
-fun dec_check_right_move :: "dec_inv_t"
-  where
-  "dec_check_right_move (as, lm) (s, l, r) ires = 
-        (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> 
-           ml = Suc m \<and> mr = (0::nat) \<and> 
-           (if lm1 = [] then l = Bk # Bk # Oc\<up>ml @ Bk # Bk # ires
-                       else l = Bk # Bk # Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> 
-           r = <lm2> @ Bk\<up>rn)"
-
-fun dec_left_move :: "dec_inv_t"
-  where
-  "dec_left_move (as, lm) (s, l, r) ires = 
-    (\<exists> lm1 m rn. (lm::nat list) = lm1 @ [m::nat] \<and>   
-    rn > 0 \<and> 
-   (if lm1 = [] then l = Bk # Oc\<up>Suc m @ Bk # Bk # ires
-    else l = Bk # Oc\<up>Suc m @ Bk # <rev lm1> @ Bk # Bk # ires) \<and> r = Bk\<up>rn)"
-
-declare
-  dec_on_right_moving.simps[simp del] dec_after_clear.simps[simp del] 
-  dec_after_write.simps[simp del] dec_left_move.simps[simp del] 
-  dec_check_right_move.simps[simp del] dec_right_move.simps[simp del] 
-  dec_first_on_right_moving.simps[simp del]
-
-fun inv_locate_n_b :: "inc_inv_t"
-  where
-  "inv_locate_n_b (as, lm) (s, l, r) ires= 
-    (\<exists> lm1 lm2 tn m ml mr rn. lm @ 0\<up>tn = lm1 @ [m] @ lm2 \<and> 
-     length lm1 = s \<and> m + 1 = ml + mr \<and> 
-     ml = 1 \<and> tn = s + 1 - length lm \<and>
-     (if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires
-      else l = Oc\<up>ml @ Bk # <rev lm1> @ Bk # Bk # ires) \<and> 
-     (r = Oc\<up>mr @ [Bk] @ <lm2>@ Bk\<up>rn \<or> (lm2 = [] \<and> r = Oc\<up>mr))
-  )"
-(*
-fun dec_inv_1 :: "layout \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> dec_inv_t"
-  where
-  "dec_inv_1 ly n e (as, am) (s, l, r) ires = 
-           (let ss = start_of ly as in
-            let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in
-            let am'' = abc_lm_s am n (abc_lm_v am n) in
-              if s = start_of ly e then inv_stop (as, am'') (s, l, r) ires
-              else if s = ss then False
-              else if s = ss + 2 * n then 
-                  inv_locate_a (as, am) (n, l, r) ires
-                \<or> inv_locate_a (as, am'') (n, l, r) ires
-              else if s = ss + 2 * n + 1 then 
-                  inv_locate_b (as, am) (n, l, r) ires
-              else if s = ss + 2 * n + 13 then 
-                  inv_on_left_moving (as, am'') (s, l, r) ires
-              else if s = ss + 2 * n + 14 then 
-                  inv_check_left_moving (as, am'') (s, l, r) ires
-              else if s = ss + 2 * n + 15 then 
-                  inv_after_left_moving (as, am'') (s, l, r) ires
-              else False)"
-*)
-
-
-fun dec_inv_1 :: "layout \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> dec_inv_t"
-  where
-  "dec_inv_1 ly n e (as, am) (s, l, r) ires = 
-           (let ss = start_of ly as in
-            let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in
-            let am'' = abc_lm_s am n (abc_lm_v am n) in
-              if s = start_of ly e then inv_stop (as, am'') (s, l, r) ires
-              else if s = ss then False
-              else if s = ss + 2 * n + 1 then 
-                  inv_locate_b (as, am) (n, l, r) ires
-              else if s = ss + 2 * n + 13 then 
-                  inv_on_left_moving (as, am'') (s, l, r) ires
-              else if s = ss + 2 * n + 14 then 
-                  inv_check_left_moving (as, am'') (s, l, r) ires
-              else if s = ss + 2 * n + 15 then 
-                  inv_after_left_moving (as, am'') (s, l, r) ires
-              else False)"
-
-declare fetch.simps[simp del]
-lemma [simp]:
-  "fetch (ci ly (start_of ly as) (Dec n e)) (Suc (2 * n)) Bk = (W1,  start_of ly as + 2 *n)"
-apply(auto simp: fetch.simps length_ci_dec)
-apply(auto simp: ci.simps nth_append length_findnth sete.simps shift.simps tdec_b_def)
-using startof_not0[of ly as] by simp
-
-lemma [simp]:
-  "fetch (ci ly (start_of ly as) (Dec n e)) (Suc (2 * n)) Oc = (R,  Suc (start_of ly as) + 2 *n)"
-apply(auto simp: fetch.simps length_ci_dec)
-apply(auto simp: ci.simps nth_append length_findnth sete.simps shift.simps tdec_b_def)
-done
-
-lemma [simp]: 
-  "\<lbrakk>r = [] \<or> hd r = Bk; inv_locate_a (as, lm) (n, l, r) ires\<rbrakk>
-    \<Longrightarrow> \<exists>stp la ra.
-  steps (start_of ly as + 2 * n, l, r) (ci ly (start_of ly as) (Dec n e), 
-  start_of ly as - Suc 0) stp = (Suc (start_of ly as + 2 * n), la, ra) \<and>
-  inv_locate_b (as, lm) (n, la, ra) ires"
-apply(rule_tac x = "Suc (Suc 0)" in exI)
-apply(auto simp: steps.simps step.simps length_ci_dec)
-apply(case_tac r, simp_all)
-done
-
-lemma [simp]: 
-  "\<lbrakk>inv_locate_a (as, lm) (n, l, r) ires; r \<noteq> [] \<and> hd r \<noteq> Bk\<rbrakk>
-    \<Longrightarrow> \<exists>stp la ra.
-  steps (start_of ly as + 2 * n, l, r) (ci ly (start_of ly as) (Dec n e), 
-  start_of ly as - Suc 0) stp = (Suc (start_of ly as + 2 * n), la, ra) \<and>
-  inv_locate_b (as, lm) (n, la, ra) ires"
-apply(rule_tac x = "(Suc 0)" in exI, case_tac "hd r", simp_all)
-apply(auto simp: steps.simps step.simps length_ci_dec)
-apply(case_tac r, simp_all)
-done
-
-fun abc_dec_1_stage1:: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "abc_dec_1_stage1 (s, l, r) ss n = 
-       (if s > ss \<and> s \<le> ss + 2*n + 1 then 4
-        else if s = ss + 2 * n + 13 \<or> s = ss + 2*n + 14 then 3
-        else if s = ss + 2*n + 15 then 2
-        else 0)"
-
-fun abc_dec_1_stage2:: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "abc_dec_1_stage2 (s, l, r) ss n = 
-       (if s \<le> ss + 2 * n + 1 then (ss + 2 * n + 16 - s)
-        else if s = ss + 2*n + 13 then length l
-        else if s = ss + 2*n + 14 then length l
-        else 0)"
-
-fun abc_dec_1_stage3 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "abc_dec_1_stage3 (s, l, r) ss n  = 
-        (if s \<le> ss + 2*n + 1 then 
-             if (s - ss) mod 2 = 0 then 
-                         if r \<noteq> [] \<and> hd r = Oc then 0 else 1  
-                         else length r
-         else if s = ss + 2 * n + 13 then 
-             if r \<noteq> [] \<and> hd r = Oc then 2 
-             else 1
-         else if s = ss + 2 * n + 14 then 
-             if r \<noteq> [] \<and> hd r = Oc then 3 else 0 
-         else 0)"
-
-fun abc_dec_1_measure :: "(config \<times> nat \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)"
-  where
-  "abc_dec_1_measure (c, ss, n) = (abc_dec_1_stage1 c ss n, 
-                   abc_dec_1_stage2 c ss n, abc_dec_1_stage3 c ss n)"
-
-definition abc_dec_1_LE ::
-  "((config \<times> nat \<times>
-  nat) \<times> (config \<times> nat \<times> nat)) set"
-  where "abc_dec_1_LE \<equiv> (inv_image lex_triple abc_dec_1_measure)"
-
-lemma wf_dec_le: "wf abc_dec_1_LE"
-by(auto intro:wf_inv_image simp:abc_dec_1_LE_def lex_triple_def lex_pair_def)
-
-lemma startof_Suc2:
-  "abc_fetch as ap = Some (Dec n e) \<Longrightarrow> 
-        start_of (layout_of ap) (Suc as) = 
-            start_of (layout_of ap) as + 2 * n + 16"
-apply(auto simp: start_of.simps layout_of.simps  
-                 length_of.simps abc_fetch.simps 
-                 take_Suc_conv_app_nth split: if_splits)
-done
-
-lemma start_of_less_2: 
-  "start_of ly e \<le> start_of ly (Suc e)"
-thm take_Suc
-apply(case_tac "e < length ly")
-apply(auto simp: start_of.simps take_Suc take_Suc_conv_app_nth)
-done
-
-lemma start_of_less_1: "start_of ly e \<le> start_of ly (e + d)"
-proof(induct d)
-  case 0 thus "?case" by simp
-next
-  case (Suc d)
-  have "start_of ly e \<le> start_of ly (e + d)"  by fact
-  moreover have "start_of ly (e + d) \<le> start_of ly (Suc (e + d))"
-    by(rule_tac start_of_less_2)
-  ultimately show"?case"
-    by(simp)
-qed
-
-lemma start_of_less: 
-  assumes "e < as"
-  shows "start_of ly e \<le> start_of ly as"
-proof -
-  obtain d where " as = e + d"
-    using assms by (metis less_imp_add_positive)
-  thus "?thesis"
-    by(simp add: start_of_less_1)
-qed
-
-lemma start_of_ge: 
-  assumes fetch: "abc_fetch as ap = Some (Dec n e)"
-  and layout: "ly = layout_of ap"
-  and great: "e > as"
-  shows "start_of ly e \<ge> start_of ly as + 2*n + 16"
-proof(cases "e = Suc as")
-  case True
-  have "e = Suc as" by fact
-  moreover hence "start_of ly (Suc as) = start_of ly as + 2*n + 16"
-    using layout fetch
-    by(simp add: startof_Suc2)
-  ultimately show "?thesis" by (simp)
-next
-  case False
-  have "e \<noteq> Suc as" by fact
-  then have "e > Suc as" using great by arith
-  then have "start_of ly (Suc as) \<le> start_of ly e"
-    by(simp add: start_of_less)
-  moreover have "start_of ly (Suc as) = start_of ly as + 2*n + 16"
-    using layout fetch
-    by(simp add: startof_Suc2)
-  ultimately show "?thesis"
-    by arith
-qed
-    
-lemma [elim]: "\<lbrakk>abc_fetch as ap = Some (Dec n e); as < e; 
-  Suc (start_of (layout_of ap) as + 2 * n) = start_of (layout_of ap) e\<rbrakk> \<Longrightarrow> RR"
-apply(drule_tac start_of_ge, simp_all)
-apply(auto)
-done
-
-lemma [elim]: "\<lbrakk>abc_fetch as ap = Some (Dec n e); as > e;
-  Suc (start_of (layout_of ap) as + 2 * n) = start_of (layout_of ap) e\<rbrakk> \<Longrightarrow> RR"
-apply(drule_tac ly = "layout_of ap" in start_of_less[of])
-apply(arith)
-done
-
-lemma [elim]: "\<lbrakk>abc_fetch as ap = Some (Dec n e);
-  Suc (start_of (layout_of ap) as + 2 * n) = start_of (layout_of ap) e\<rbrakk> \<Longrightarrow> RR"
-apply(subgoal_tac "as = e \<or> as < e \<or> as > e", auto)
-done
-
-lemma [simp]:"fetch (ci (ly) (start_of ly as) (Dec n e)) (Suc (2 * n))  Oc
-  = (R, start_of ly as + 2*n + 1)"
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]: "(start_of ly as = 0) = False"
-apply(simp add: start_of.simps)
-done
-
-lemma [simp]: "fetch (ci (ly) 
-  (start_of ly as) (Dec n e)) (Suc (2 * n))  Bk
-  = (W1, start_of ly as + 2*n)"
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]: 
-  "fetch (ci (ly) 
-                (start_of ly as) (Dec n e)) (Suc (Suc (2 * n)))  Oc
-      = (R, start_of ly as + 2*n + 2)"
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-
-lemma [simp]: "fetch (ci (ly) 
-                  (start_of ly as) (Dec n e)) (Suc (Suc (2 * n))) Bk
-      = (L, start_of ly as + 2*n + 13)"
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-
-lemma [simp]: "fetch (ci (ly) 
-             (start_of ly as) (Dec n e)) (Suc (Suc (Suc (2 * n))))  Oc
-     = (R, start_of ly as + 2*n + 2)"
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-
-lemma [simp]: "fetch (ci (ly) (start_of ly as) (Dec n e)) 
-                             (Suc (Suc (Suc (2 * n))))  Bk
-     = (L, start_of ly as + 2*n + 3)"
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]:
-     "fetch (ci (ly) 
-                      (start_of ly as) (Dec n e)) (2 * n + 4) Oc
-    = (W0, start_of ly as + 2*n + 3)"
-apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]: "fetch (ci (ly) 
-                   (start_of ly as) (Dec n e)) (2 * n + 4) Bk
-    = (R, start_of ly as + 2*n + 4)"
-apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]:"fetch (ci (ly) 
-                          (start_of ly as) (Dec n e)) (2 * n + 5) Bk
-    = (R, start_of ly as + 2*n + 5)"
-apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-
-lemma [simp]: 
-  "fetch (ci (ly)
-                (start_of ly as) (Dec n e)) (2 * n + 6) Bk
-    = (L, start_of ly as + 2*n + 6)"
-apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]: 
-  "fetch (ci (ly) (start_of ly as) 
-                      (Dec n e)) (2 * n + 6) Oc
-    = (L, start_of ly as + 2*n + 7)"
-apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]:"fetch (ci (ly) 
-             (start_of ly as) (Dec n e)) (2 * n + 7) Bk
-    = (L, start_of ly as + 2*n + 10)"
-apply(subgoal_tac "2*n + 7 = Suc (2*n + 6)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]:
-  "fetch (ci (ly) 
-                   (start_of ly as) (Dec n e)) (2 * n + 8) Bk
-    = (W1, start_of ly as + 2*n + 7)"
-apply(subgoal_tac "2*n + 8 = Suc (2*n + 7)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-
-lemma [simp]: 
-  "fetch (ci (ly) 
-                   (start_of ly as) (Dec n e)) (2 * n + 8) Oc
-    = (R, start_of ly as + 2*n + 8)"
-apply(subgoal_tac "2*n + 8 = Suc (2*n + 7)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]: 
-  "fetch (ci (ly) 
-  (start_of ly as) (Dec n e)) (2 * n + 9) Bk
-  = (L, start_of ly as + 2*n + 9)"
-apply(subgoal_tac "2*n + 9 = Suc (2*n + 8)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]: 
-  "fetch (ci (ly) 
-  (start_of ly as) (Dec n e)) (2 * n + 9) Oc
-  = (R, start_of ly as + 2*n + 8)"
-apply(subgoal_tac "2*n + 9 = Suc (2*n + 8)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-
-lemma [simp]: 
-  "fetch (ci (ly) 
-  (start_of ly as) (Dec n e)) (2 * n + 10) Bk
-  = (R, start_of ly as + 2*n + 4)" 
-apply(subgoal_tac "2*n + 10 = Suc (2*n + 9)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]: "fetch (ci (ly) 
-             (start_of ly as) (Dec n e)) (2 * n + 10) Oc
-    = (W0, start_of ly as + 2*n + 9)"
-apply(subgoal_tac "2*n + 10 = Suc (2*n + 9)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-
-lemma [simp]: 
-  "fetch (ci (ly) 
-  (start_of ly as) (Dec n e)) (2 * n + 11) Oc
-  = (L, start_of ly as + 2*n + 10)"
-apply(subgoal_tac "2*n + 11 = Suc (2*n + 10)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-
-lemma [simp]: 
-  "fetch (ci (ly) 
-  (start_of ly as) (Dec n e)) (2 * n + 11) Bk
-  = (L, start_of ly as + 2*n + 11)"
-apply(subgoal_tac "2*n + 11 = Suc (2*n + 10)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]:
-  "fetch (ci (ly) 
-  (start_of ly as) (Dec n e)) (2 * n + 12) Oc
-  = (L, start_of ly as + 2*n + 10)"
-apply(subgoal_tac "2*n + 12 = Suc (2*n + 11)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-
-lemma [simp]: 
-  "fetch (ci (ly) 
-  (start_of ly as) (Dec n e)) (2 * n + 12) Bk
-  = (R, start_of ly as + 2*n + 12)"
-apply(subgoal_tac "2*n + 12 = Suc (2*n + 11)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]: 
-  "fetch (ci (ly) 
-  (start_of ly as) (Dec n e)) (2 * n + 13) Bk
-  = (R, start_of ly as + 2*n + 16)"
-apply(subgoal_tac "2*n + 13 = Suc (2*n + 12)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-
-lemma [simp]: 
-  "fetch (ci (ly) 
-  (start_of ly as) (Dec n e)) (14 + 2 * n) Oc
-  = (L, start_of ly as + 2*n + 13)"
-apply(subgoal_tac "14 + 2*n = Suc (2*n + 13)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]: 
-  "fetch (ci (ly) 
-  (start_of ly as) (Dec n e)) (14 + 2 * n) Bk
-  = (L, start_of ly as + 2*n + 14)"
-apply(subgoal_tac "14 + 2*n = Suc (2*n + 13)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]: 
-  "fetch (ci (ly) 
-  (start_of ly as) (Dec n e)) (15 + 2 * n)  Oc
-  = (L, start_of ly as + 2*n + 13)"
-apply(subgoal_tac "15 + 2*n = Suc (2*n + 14)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]: 
-  "fetch (ci (ly) 
-  (start_of ly as) (Dec n e)) (15 + 2 * n)  Bk
- = (R, start_of ly as + 2*n + 15)"
-apply(subgoal_tac "15 + 2*n = Suc (2*n + 14)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-lemma [simp]: 
-  "abc_fetch as aprog = Some (Dec n e) \<Longrightarrow> 
-     fetch (ci (ly) (start_of (ly) as) 
-              (Dec n e)) (16 + 2 * n)  Bk
- = (R, start_of (ly) e)"
-apply(subgoal_tac "16 + 2*n = Suc (2*n + 15)", simp only: fetch.simps)
-apply(auto simp: ci.simps findnth.simps fetch.simps
-                  nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps)
-done
-
-declare dec_inv_1.simps[simp del]
-
-
-lemma [simp]: 
- "\<lbrakk>abc_fetch as aprog = Some (Dec n e); ly = layout_of aprog\<rbrakk>
-   \<Longrightarrow> (start_of ly e \<noteq> Suc (start_of ly as + 2 * n) \<and>
-        start_of ly e \<noteq> Suc (Suc (start_of ly as + 2 * n)) \<and>  
-        start_of ly e \<noteq> start_of ly as + 2 * n + 3 \<and> 
-        start_of ly e \<noteq> start_of ly as + 2 * n + 4 \<and>
-        start_of ly e \<noteq> start_of ly as + 2 * n + 5 \<and> 
-        start_of ly e \<noteq> start_of ly as + 2 * n + 6 \<and> 
-        start_of ly e \<noteq> start_of ly as + 2 * n + 7 \<and>
-        start_of ly e \<noteq> start_of ly as + 2 * n + 8 \<and> 
-        start_of ly e \<noteq> start_of ly as + 2 * n + 9 \<and> 
-        start_of ly e \<noteq> start_of ly as + 2 * n + 10 \<and>
-        start_of ly e \<noteq> start_of ly as + 2 * n + 11 \<and> 
-        start_of ly e \<noteq> start_of ly as + 2 * n + 12 \<and> 
-        start_of ly e \<noteq> start_of ly as + 2 * n + 13 \<and>
-        start_of ly e \<noteq> start_of ly as + 2 * n + 14 \<and> 
-        start_of ly e \<noteq> start_of ly as + 2 * n + 15)"
-using start_of_ge[of as aprog n e ly] start_of_less[of e as ly]
-apply(case_tac "e < as", simp)
-apply(case_tac "e = as", simp, simp)
-done
-
-lemma [simp]: "\<lbrakk>abc_fetch as aprog = Some (Dec n e); ly = layout_of aprog\<rbrakk>
-      \<Longrightarrow> (Suc (start_of ly as + 2 * n) \<noteq> start_of ly e \<and>
-          Suc (Suc (start_of ly as + 2 * n)) \<noteq> start_of ly e \<and> 
-          start_of ly as + 2 * n + 3 \<noteq> start_of ly e \<and> 
-          start_of ly as + 2 * n + 4 \<noteq> start_of ly e \<and>
-          start_of ly as + 2 * n + 5 \<noteq>start_of ly e \<and> 
-          start_of ly as + 2 * n + 6 \<noteq> start_of ly e \<and>
-          start_of ly as + 2 * n + 7 \<noteq> start_of ly e \<and> 
-          start_of ly as + 2 * n + 8 \<noteq> start_of ly e \<and> 
-          start_of ly as + 2 * n + 9 \<noteq> start_of ly e \<and> 
-          start_of ly as + 2 * n + 10 \<noteq> start_of ly e \<and>
-          start_of ly as + 2 * n + 11 \<noteq> start_of ly e \<and> 
-          start_of ly as + 2 * n + 12 \<noteq> start_of ly e \<and> 
-          start_of ly as + 2 * n + 13 \<noteq> start_of ly e \<and> 
-          start_of ly as + 2 * n + 14 \<noteq> start_of ly e \<and> 
-          start_of ly as + 2 * n + 15 \<noteq> start_of ly e)"
-using start_of_ge[of as aprog n e ly] start_of_less[of e as ly]
-apply(case_tac "e < as", simp, simp)
-apply(case_tac "e = as", simp, simp)
-done
-
-lemma [simp]: "inv_locate_b (as, lm) (n, [], []) ires = False"
-apply(auto simp: inv_locate_b.simps in_middle.simps split: if_splits)
-done
-
-lemma [simp]: "inv_locate_b (as, lm) (n, [], Bk # list) ires = False"
-apply(auto simp: inv_locate_b.simps in_middle.simps split: if_splits)
-done
-
-(*
-
-lemma  inv_locate_b_2_on_left_moving_b[simp]: 
-   "inv_locate_b (as, am) (n, l, []) ires
-     \<Longrightarrow> inv_on_left_moving (as, 
-                  abc_lm_s am n (abc_lm_v am n)) (s, [], [Bk]) ires"
-apply(auto simp: inv_locate_b.simps inv_on_left_moving.simps inv_on_left_moving_in_middle_B.simps
-                 in_middle.simps split: if_splits)
-apply(drule_tac length_equal, simp)
-
-apply(insert inv_locate_b_2_on_left_moving[of as am n l "[]" ires s])
-apply(simp only: inv_on_left_moving.simps, simp)
-apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B 
-         (as, abc_lm_s am n (abc_lm_v am n)) (s, tl l, [hd l]) ires", simp)
-*)
-
-(*
-lemma [simp]:
-  "inv_locate_b (as, am) (n, l, []) ires; l \<noteq> []\<rbrakk>
- \<Longrightarrow> inv_on_left_moving (as, abc_lm_s am n 
-                  (abc_lm_v am n)) (s, tl l, [hd l]) ires"
-apply(auto simp: inv_locate_b.simps inv_on_left_moving.simps inv_on_left_moving_in_middle_B.simps
-                 in_middle.simps split: if_splits)
-apply(drule_tac length_equal, simp)
-
-apply(insert inv_locate_b_2_on_left_moving[of as am n l "[]" ires s])
-apply(simp only: inv_on_left_moving.simps, simp)
-apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B 
-         (as, abc_lm_s am n (abc_lm_v am n)) (s, tl l, [hd l]) ires", simp)
-
-apply(insert inv_locate_b_2_on_left_moving[of as am n l "[]" ires s])
-apply(simp only: inv_on_left_moving.simps, simp)
-apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B 
-         (as, abc_lm_s am n (abc_lm_v am n)) (s, tl l, [hd l]) ires", simp)
-apply(simp only: inv_on_left_moving_norm.simps)
-apply(erule_tac exE)+
-apply(erule_tac conjE)+
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
-      rule_tac x = m in exI, rule_tac x = ml in exI, 
-      rule_tac x = mr in exI, simp)
-apply(case_tac mr, simp, simp, case_tac nat, auto intro: nil_2_nil)
-done
-*)
-
-lemma [simp]: 
- "\<lbrakk>dec_first_on_right_moving n (as, am) (s, aaa, Oc # xs) ires\<rbrakk>
-   \<Longrightarrow> dec_first_on_right_moving n (as, am) (s', Oc # aaa, xs) ires"
-apply(simp only: dec_first_on_right_moving.simps)
-apply(erule exE)+
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
-      rule_tac x = m in exI, simp)
-apply(rule_tac x = "Suc ml" in exI, 
-      rule_tac x = "mr - 1" in exI, auto)
-apply(case_tac [!] mr, auto)
-done
-
-lemma [simp]: 
-  "dec_first_on_right_moving n (as, am) (s, l, Bk # xs) ires \<Longrightarrow> l \<noteq> []"
-apply(auto simp: dec_first_on_right_moving.simps split: if_splits)
-done
-
-lemma [elim]: 
-  "\<lbrakk>\<not> length lm1 < length am; 
-    am @ replicate (length lm1 - length am) 0 @ [0::nat] = 
-                                                lm1 @ m # lm2;
-    0 < m\<rbrakk>
-   \<Longrightarrow> RR"
-apply(subgoal_tac "lm2 = []", simp)
-apply(drule_tac length_equal, simp)
-done
-
-lemma [simp]: 
- "\<lbrakk>dec_first_on_right_moving n (as, 
-                   abc_lm_s am n (abc_lm_v am n)) (s, l, Bk # xs) ires\<rbrakk>
-\<Longrightarrow> dec_after_clear (as, abc_lm_s am n 
-                 (abc_lm_v am n - Suc 0)) (s', tl l, hd l # Bk # xs) ires"
-apply(simp only: dec_first_on_right_moving.simps 
-                 dec_after_clear.simps abc_lm_s.simps abc_lm_v.simps)
-apply(erule_tac exE)+
-apply(case_tac "n < length am")
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
-      rule_tac x = "m - 1" in exI, auto simp: )
-apply(case_tac [!] mr, auto)
-done
-
-lemma [simp]: 
- "\<lbrakk>dec_first_on_right_moving n (as, 
-                   abc_lm_s am n (abc_lm_v am n)) (s, l, []) ires\<rbrakk>
-\<Longrightarrow> (l = [] \<longrightarrow> dec_after_clear (as, 
-             abc_lm_s am n (abc_lm_v am n - Suc 0)) (s', [], [Bk]) ires) \<and>
-    (l \<noteq> [] \<longrightarrow> dec_after_clear (as, abc_lm_s am n 
-                      (abc_lm_v am n - Suc 0)) (s', tl l, [hd l]) ires)"
-apply(subgoal_tac "l \<noteq> []", 
-      simp only: dec_first_on_right_moving.simps 
-                 dec_after_clear.simps abc_lm_s.simps abc_lm_v.simps)
-apply(erule_tac exE)+
-apply(case_tac "n < length am", simp)
-apply(rule_tac x = lm1 in exI, rule_tac x = "m - 1" in exI, auto)
-apply(case_tac [1-2] m, auto)
-apply(auto simp: dec_first_on_right_moving.simps split: if_splits)
-done
-
-lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, Oc # r) ires\<rbrakk>
-                \<Longrightarrow> dec_after_clear (as, am) (s', l, Bk # r) ires"
-apply(auto simp: dec_after_clear.simps)
-done
-
-lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, Bk # r) ires\<rbrakk>
-                \<Longrightarrow> dec_right_move (as, am) (s', Bk # l, r) ires"
-apply(auto simp: dec_after_clear.simps dec_right_move.simps split: if_splits)
-done
-
-lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, []) ires\<rbrakk>
-             \<Longrightarrow> dec_right_move (as, am) (s', Bk # l, []) ires"
-apply(auto simp: dec_after_clear.simps dec_right_move.simps )
-done
-
-lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, []) ires\<rbrakk>
-             \<Longrightarrow> dec_right_move (as, am) (s', Bk # l, [Bk]) ires"
-apply(auto simp: dec_after_clear.simps dec_right_move.simps split: if_splits)
-done
-
-lemma [simp]:"dec_right_move (as, am) (s, l, Oc # r) ires = False"
-apply(auto simp: dec_right_move.simps)
-done
-              
-lemma dec_right_move_2_check_right_move[simp]:
-     "\<lbrakk>dec_right_move (as, am) (s, l, Bk # r) ires\<rbrakk>
-      \<Longrightarrow> dec_check_right_move (as, am) (s', Bk # l, r) ires"
-apply(auto simp: dec_right_move.simps dec_check_right_move.simps split: if_splits)
-done
-
-lemma [simp]: "(<lm::nat list> = []) = (lm = [])"
-apply(case_tac lm, simp_all add: tape_of_nl_cons)
-done
-
-lemma [simp]: 
- "dec_right_move (as, am) (s, l, []) ires= 
-  dec_right_move (as, am) (s, l, [Bk]) ires"
-apply(simp add: dec_right_move.simps)
-done
-
-lemma [simp]: "\<lbrakk>dec_right_move (as, am) (s, l, []) ires\<rbrakk>
-             \<Longrightarrow> dec_check_right_move (as, am) (s, Bk # l, []) ires"
-apply(insert dec_right_move_2_check_right_move[of as am s l "[]" s'], 
-      simp)
-done
-
-lemma [simp]: "dec_check_right_move (as, am) (s, l, r) ires\<Longrightarrow> l \<noteq> []"
-apply(auto simp: dec_check_right_move.simps split: if_splits)
-done
- 
-lemma [simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, Oc # r) ires\<rbrakk>
-             \<Longrightarrow> dec_after_write (as, am) (s', tl l, hd l # Oc # r) ires"
-apply(auto simp: dec_check_right_move.simps dec_after_write.simps)
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
-      rule_tac x = m in exI, auto)
-done
-
-
-
-lemma [simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, Bk # r) ires\<rbrakk>
-                \<Longrightarrow> dec_left_move (as, am) (s', tl l, hd l # Bk # r) ires"
-apply(auto simp: dec_check_right_move.simps 
-                 dec_left_move.simps inv_after_move.simps)
-apply(rule_tac x = lm1 in exI, rule_tac x = m in exI, auto split: if_splits)
-apply(case_tac [!] lm2, simp_all add: tape_of_nl_cons split: if_splits)
-apply(rule_tac [!] x = "(Suc rn)" in exI, simp_all)
-done
-
-lemma [simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, []) ires\<rbrakk>
-             \<Longrightarrow> dec_left_move (as, am) (s', tl l, [hd l]) ires"
-apply(auto simp: dec_check_right_move.simps 
-                 dec_left_move.simps inv_after_move.simps)
-apply(rule_tac x = lm1 in exI, rule_tac x = m in exI, auto)
-done
-
-lemma [simp]: "dec_left_move (as, am) (s, aaa, Oc # xs) ires = False"
-apply(auto simp: dec_left_move.simps inv_after_move.simps)
-done
-
-lemma [simp]: "dec_left_move (as, am) (s, l, r) ires
-             \<Longrightarrow> l \<noteq> []"
-apply(auto simp: dec_left_move.simps split: if_splits)
-done
-
-lemma [simp]: "inv_on_left_moving_in_middle_B (as, [m])
-  (s', Oc # Oc\<up>m @ Bk # Bk # ires, Bk # Bk\<up>rn) ires"
-apply(simp add: inv_on_left_moving_in_middle_B.simps)
-apply(rule_tac x = "[m]" in exI, auto)
-done
-
-lemma [simp]: "inv_on_left_moving_in_middle_B (as, [m])
-  (s', Oc # Oc\<up>m @ Bk # Bk # ires, [Bk]) ires"
-apply(simp add: inv_on_left_moving_in_middle_B.simps)
-done
-
-
-lemma [simp]: "lm1 \<noteq> [] \<Longrightarrow> 
-  inv_on_left_moving_in_middle_B (as, lm1 @ [m]) (s', 
-  Oc # Oc\<up>m @ Bk # <rev lm1> @ Bk # Bk # ires, Bk # Bk\<up>rn) ires"
-apply(simp only: inv_on_left_moving_in_middle_B.simps)
-apply(rule_tac x = "lm1 @ [m ]" in exI, rule_tac x = "[]" in exI, simp)
-apply(simp add: tape_of_nl_cons split: if_splits)
-done
-
-lemma [simp]: "lm1 \<noteq> [] \<Longrightarrow> 
-  inv_on_left_moving_in_middle_B (as, lm1 @ [m]) (s', 
-  Oc # Oc\<up> m @ Bk # <rev lm1> @ Bk # Bk # ires, [Bk]) ires"
-apply(simp only: inv_on_left_moving_in_middle_B.simps)
-apply(rule_tac x = "lm1 @ [m ]" in exI, rule_tac x = "[]" in exI, simp)
-apply(simp add: tape_of_nl_cons split: if_splits)
-done
-
-lemma [simp]: "dec_left_move (as, am) (s, l, Bk # r) ires
-       \<Longrightarrow> inv_on_left_moving (as, am) (s', tl l, hd l # Bk # r) ires"
-apply(auto simp: dec_left_move.simps inv_on_left_moving.simps split: if_splits)
-done
-
-(*
-lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm1 @ [m]) 
-                        (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # <rev lm1> @ Bk\<^bsup>ln\<^esup>, [Bk])  ires"
-apply(auto simp: inv_on_left_moving_in_middle_B.simps)
-apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "[]" in exI, auto)
-done
-*)
-
-lemma [simp]: "dec_left_move (as, am) (s, l, []) ires
-             \<Longrightarrow> inv_on_left_moving (as, am) (s', tl l, [hd l]) ires"
-apply(auto simp: dec_left_move.simps inv_on_left_moving.simps split: if_splits)
-done
-
-lemma [simp]: "dec_after_write (as, am) (s, l, Oc # r) ires
-       \<Longrightarrow> dec_on_right_moving (as, am) (s', Oc # l, r) ires"
-apply(auto simp: dec_after_write.simps dec_on_right_moving.simps)
-apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI, 
-      rule_tac x = "hd lm2" in exI, simp)
-apply(rule_tac x = "Suc 0" in exI,rule_tac x =  "Suc (hd lm2)" in exI)
-apply(case_tac lm2, auto split: if_splits simp: tape_of_nl_cons)
-done
-
-lemma [simp]: "dec_after_write (as, am) (s, l, Bk # r) ires
-       \<Longrightarrow> dec_after_write (as, am) (s', l, Oc # r) ires"
-apply(auto simp: dec_after_write.simps)
-done
-
-lemma [simp]: "dec_after_write (as, am) (s, aaa, []) ires
-             \<Longrightarrow> dec_after_write (as, am) (s', aaa, [Oc]) ires"
-apply(auto simp: dec_after_write.simps)
-done
-
-lemma [simp]: "dec_on_right_moving (as, am) (s, l, Oc # r) ires
-       \<Longrightarrow> dec_on_right_moving (as, am) (s', Oc # l, r) ires"
-apply(simp only: dec_on_right_moving.simps)
-apply(erule_tac exE)+
-apply(erule conjE)+
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,
-      rule_tac x = "m" in exI, rule_tac x = "Suc ml" in exI, 
-      rule_tac x = "mr - 1" in exI, simp)
-apply(case_tac mr, auto)
-done
-
-lemma [simp]: "dec_on_right_moving (as, am) (s, l, r) ires\<Longrightarrow>  l \<noteq> []"
-apply(auto simp: dec_on_right_moving.simps split: if_splits)
-done
-
-lemma [simp]: "dec_on_right_moving (as, am) (s, l, Bk # r) ires
-      \<Longrightarrow>  dec_after_clear (as, am) (s', tl l, hd l # Bk # r) ires"
-apply(auto simp: dec_on_right_moving.simps dec_after_clear.simps)
-apply(case_tac [!] mr, auto split: if_splits)
-done
-
-lemma [simp]: "dec_on_right_moving (as, am) (s, l, []) ires
-             \<Longrightarrow> dec_after_clear (as, am) (s', tl l, [hd l]) ires"
-apply(auto simp: dec_on_right_moving.simps dec_after_clear.simps)
-apply(simp_all split: if_splits)
-apply(rule_tac x = lm1 in exI, simp)
-done
-
-lemma [simp]: 
- "inv_stop (as, abc_lm_s am n (abc_lm_v am n)) (s, l, r) ires \<Longrightarrow> l \<noteq> []"
-apply(auto simp: inv_stop.simps)
-done
-
-lemma dec_false_1[simp]:
- "\<lbrakk>abc_lm_v am n = 0; inv_locate_b (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
-  \<Longrightarrow> False"
-apply(auto simp: inv_locate_b.simps in_middle.simps)
-apply(case_tac "length lm1 \<ge> length am", auto)
-apply(subgoal_tac "lm2 = []", simp, subgoal_tac "m = 0", simp)
-apply(case_tac mr, auto simp: )
-apply(subgoal_tac "Suc (length lm1) - length am = 
-                   Suc (length lm1 - length am)", 
-      simp add: exp_ind del: replicate.simps, simp)
-apply(drule_tac xs = "am @ replicate (Suc (length lm1) - length am) 0"
-                and ys = "lm1 @ m # lm2" in length_equal, simp)
-apply(case_tac mr, auto simp: abc_lm_v.simps)
-apply(case_tac "mr = 0", simp_all split: if_splits)
-apply(subgoal_tac "Suc (length lm1) - length am = 
-                       Suc (length lm1 - length am)", 
-      simp add: exp_ind del: replicate.simps, simp)
-done
-
-lemma [simp]: 
- "\<lbrakk>inv_locate_b (as, am) (n, aaa, Bk # xs) ires; 
-   abc_lm_v am n = 0\<rbrakk>
-   \<Longrightarrow> inv_on_left_moving (as, abc_lm_s am n 0) 
-                         (s, tl aaa, hd aaa # Bk # xs) ires"
-apply(simp add: inv_on_left_moving.simps)
-apply(simp only: inv_locate_b.simps in_middle.simps) 
-apply(erule_tac exE)+
-apply(simp add: inv_on_left_moving.simps)
-apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B 
-         (as, abc_lm_s am n 0) (s, tl aaa, hd aaa # Bk # xs) ires", simp)
-apply(simp only: inv_on_left_moving_norm.simps)
-apply(erule_tac conjE)+
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
-      rule_tac x =  m in exI, rule_tac x = m in exI, 
-      rule_tac x = "Suc 0" in exI, simp add: abc_lm_s.simps)
-apply(case_tac mr, simp_all, auto simp: abc_lm_v.simps)
-apply(simp only: exp_ind[THEN sym] replicate_Suc Nat.Suc_diff_le)
-apply(auto simp: inv_on_left_moving_in_middle_B.simps split: if_splits)
-done
-
-
-lemma [simp]:
- "\<lbrakk>abc_lm_v am n = 0; inv_locate_b (as, am) (n, aaa, []) ires\<rbrakk>
-   \<Longrightarrow> inv_on_left_moving (as, abc_lm_s am n 0) (s, tl aaa, [hd aaa]) ires"
-apply(simp add: inv_on_left_moving.simps)
-apply(simp only: inv_locate_b.simps in_middle.simps) 
-apply(erule_tac exE)+
-apply(simp add: inv_on_left_moving.simps)
-apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B 
-         (as, abc_lm_s am n 0) (s, tl aaa, [hd aaa]) ires", simp)
-apply(simp only: inv_on_left_moving_norm.simps)
-apply(erule_tac conjE)+
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
-      rule_tac x =  m in exI, rule_tac x = m in exI, 
-      rule_tac x = "Suc 0" in exI, simp add: abc_lm_s.simps)
-apply(case_tac mr, simp_all, auto simp: abc_lm_v.simps)
-apply(simp_all only: exp_ind Nat.Suc_diff_le del: replicate_Suc, simp_all)
-apply(auto simp: inv_on_left_moving_in_middle_B.simps split: if_splits)
-apply(case_tac [!] m, simp_all)
-done
-
-lemma [simp]: "\<lbrakk>am ! n = (0::nat); n < length am\<rbrakk> \<Longrightarrow> am[n := 0] = am"
-apply(simp add: list_update_same_conv)
-done
-
-lemma  [intro]: "\<lbrakk>abc_lm_v (a # list) 0 = 0\<rbrakk> \<Longrightarrow> a = 0"
-apply(simp add: abc_lm_v.simps split: if_splits)
-done
-
-lemma [simp]: 
- "inv_stop (as, abc_lm_s am n 0) 
-          (start_of (layout_of aprog) e, aaa, Oc # xs) ires
-  \<Longrightarrow> inv_locate_a (as, abc_lm_s am n 0) (0, aaa, Oc # xs) ires"
-apply(simp add: inv_locate_a.simps)
-apply(rule disjI1)
-apply(auto simp: inv_stop.simps at_begin_norm.simps)
-done
-
-lemma [simp]: 
- "\<lbrakk>inv_stop (as, abc_lm_s am n 0) 
-          (start_of (layout_of aprog) e, aaa, Oc # xs) ires\<rbrakk>
-  \<Longrightarrow> inv_locate_b (as, am) (0, Oc # aaa, xs) ires \<or> 
-      inv_locate_b (as, abc_lm_s am n 0) (0, Oc # aaa, xs) ires"
-apply(simp)
-done
-
-lemma dec_false2: 
- "inv_stop (as, abc_lm_s am n 0) 
-  (start_of (layout_of aprog) e, aaa, Bk # xs) ires = False"
-apply(auto simp: inv_stop.simps abc_lm_s.simps)
-apply(case_tac [!] am, auto)
-apply(case_tac [!] n, auto simp: tape_of_nl_cons split: if_splits)
-done
-
-lemma dec_false3:
-   "inv_stop (as, abc_lm_s am n 0) 
-              (start_of (layout_of aprog) e, aaa, []) ires = False"
-apply(auto simp: inv_stop.simps abc_lm_s.simps)
-done
-
-lemma [simp]:
-  "fetch (ci (layout_of aprog) 
-       (start_of (layout_of aprog) as) (Dec n e)) 0 b = (Nop, 0)"
-by(simp add: fetch.simps)
-
-declare dec_inv_1.simps[simp del]
-
-declare inv_locate_n_b.simps [simp del]
-
-lemma [simp]:
-  "\<lbrakk>0 < abc_lm_v am n; 0 < n; 
-    at_begin_fst_bwtn (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
- \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires"
-apply(simp add: at_begin_fst_bwtn.simps inv_locate_n_b.simps )
-done
- 
-lemma Suc_minus:"length am + tn = n
-       \<Longrightarrow> Suc tn = Suc n - length am "
-apply(arith)
-done
-
-lemma [simp]: 
- "\<lbrakk>0 < abc_lm_v am n; 0 < n; 
-   at_begin_fst_awtn (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
- \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires"
-apply(simp only: at_begin_fst_awtn.simps inv_locate_n_b.simps )
-apply(erule exE)+
-apply(erule conjE)+
-apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, 
-      rule_tac x = "Suc tn" in exI, rule_tac x = 0 in exI)
-apply(simp add: exp_ind del: replicate.simps)
-apply(rule conjI)+
-apply(auto)
-done
-
-lemma [simp]:
- "\<lbrakk>inv_locate_n_b (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
- \<Longrightarrow> dec_first_on_right_moving n (as, abc_lm_s am n (abc_lm_v am n))  
-                                      (s, Oc # aaa, xs) ires"
-apply(auto simp: inv_locate_n_b.simps dec_first_on_right_moving.simps 
-                 abc_lm_s.simps abc_lm_v.simps)
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
-      rule_tac x = m in exI, simp)
-apply(rule_tac x = "Suc (Suc 0)" in exI, 
-      rule_tac x = "m - 1" in exI, simp)
-apply(case_tac m, auto)
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
-      rule_tac x = m in exI, 
-      simp add: Suc_diff_le exp_ind del: replicate.simps)
-apply(rule_tac x = "Suc (Suc 0)" in exI, 
-      rule_tac x = "m - 1" in exI, simp)
-apply(case_tac m, auto)
-apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, 
-      rule_tac x = m in exI, simp)
-apply(rule_tac x = "Suc (Suc 0)" in exI, 
-      rule_tac x = "m - 1" in exI, simp)
-apply(case_tac m, auto)
-apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, 
-      rule_tac x = m in exI, 
-      simp add: Suc_diff_le exp_ind del: replicate.simps, simp)
-done
-
-lemma [simp]: "inv_on_left_moving (as, am) (s, [], r) ires 
-  = False"
-apply(simp add: inv_on_left_moving.simps inv_on_left_moving_norm.simps
-                inv_on_left_moving_in_middle_B.simps)
-done
-
-lemma [simp]: 
-  "inv_check_left_moving (as, abc_lm_s am n 0)
-  (start_of (layout_of aprog) as + 2 * n + 14, [], Oc # xs) ires
- = False"
-apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps)
-done
-
-lemma [simp]: "inv_check_left_moving (as, abc_lm_s lm n (abc_lm_v lm n)) (s, [], Oc # list) ires = False"
-apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps)
-done
-
-lemma [elim]: "\<lbrakk>abc_fetch as ap = Some (Dec n e);
-                start_of (layout_of ap) as < start_of (layout_of ap) e; 
-                start_of (layout_of ap) e \<le> Suc (start_of (layout_of ap) as + 2 * n)\<rbrakk>
-       \<Longrightarrow> RR"
-  using start_of_less[of e as "layout_of ap"] start_of_ge[of as ap n e "layout_of ap"]
-apply(case_tac "as < e", simp)
-apply(case_tac "as = e", simp, simp)
-done
-
-lemma crsp_step_dec_b_e_pre':
-  assumes layout: "ly = layout_of ap"
-  and inv_start: "inv_locate_b (as, lm) (n, la, ra) ires"
-  and fetch: "abc_fetch as ap = Some (Dec n e)"
-  and dec_0: "abc_lm_v lm n = 0"
-  and f: "f = (\<lambda> stp. (steps (Suc (start_of ly as) + 2 * n, la, ra) (ci ly (start_of ly as) (Dec n e), 
-            start_of ly as - Suc 0) stp, start_of ly as, n))"
-  and P: "P = (\<lambda> ((s, l, r), ss, x). s = start_of ly e)"
-  and Q: "Q = (\<lambda> ((s, l, r), ss, x). dec_inv_1 ly x e (as, lm) (s, l, r) ires)"
-  shows "\<exists> stp. P (f stp) \<and> Q (f stp)"
-proof(rule_tac LE = abc_dec_1_LE in halt_lemma2)
-  show "wf abc_dec_1_LE" by(intro wf_dec_le)
-next
-  show "Q (f 0)"
-    using layout fetch
-    apply(simp add: f steps.simps Q dec_inv_1.simps)
-    apply(subgoal_tac "e > as \<or> e = as \<or> e < as")
-    apply(auto simp: Let_def start_of_ge start_of_less inv_start)
-    done
-next
-  show "\<not> P (f 0)"
-    using layout fetch
-    apply(simp add: f steps.simps P)
-    done
-next
-  show "\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> abc_dec_1_LE"
-    using fetch
-  proof(rule_tac allI, rule_tac impI)
-    fix na
-    assume "\<not> P (f na) \<and> Q (f na)"
-    thus "Q (f (Suc na)) \<and> (f (Suc na), f na) \<in> abc_dec_1_LE"
-      apply(simp add: f)
-      apply(case_tac "steps (Suc (start_of ly as + 2 * n), la, ra)
-        (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) na", simp)
-    proof -
-      fix a b c 
-      assume "\<not> P ((a, b, c), start_of ly as, n) \<and> Q ((a, b, c), start_of ly as, n)"
-      thus "Q (step (a, b, c) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0), start_of ly as, n) \<and>
-               ((step (a, b, c) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0), start_of ly as, n), 
-                   (a, b, c), start_of ly as, n) \<in> abc_dec_1_LE"
-        apply(simp add: Q)
-        apply(case_tac c, case_tac [2] aa)
-        apply(simp_all add: dec_inv_1.simps Let_def split: if_splits)
-        using fetch layout dec_0
-        apply(auto simp: step.simps P dec_inv_1.simps Let_def abc_dec_1_LE_def lex_triple_def lex_pair_def)
-        using dec_0
-        apply(drule_tac dec_false_1, simp_all)
-        done
-    qed
-  qed
-qed
-
-lemma crsp_step_dec_b_e_pre:
-  assumes "ly = layout_of ap"
-  and inv_start: "inv_locate_b (as, lm) (n, la, ra) ires"
-  and dec_0: "abc_lm_v lm n  = 0"
-  and fetch: "abc_fetch as ap = Some (Dec n e)"
-  shows "\<exists>stp lb rb.
-       steps (Suc (start_of ly as) + 2 * n, la, ra) (ci ly (start_of ly as) (Dec n e), 
-       start_of ly as - Suc 0) stp = (start_of ly e, lb, rb) \<and>
-       dec_inv_1 ly n e (as, lm) (start_of ly e, lb, rb) ires"
-  using assms
-  apply(drule_tac crsp_step_dec_b_e_pre', auto)
-  apply(rule_tac x = stp in exI, simp)
-  done
-
-lemma [simp]:
-  "\<lbrakk>abc_lm_v lm n = 0;
-  inv_stop (as, abc_lm_s lm n (abc_lm_v lm n)) (start_of ly e, lb, rb) ires\<rbrakk>
-  \<Longrightarrow> crsp ly (abc_step_l (as, lm) (Some (Dec n e))) (start_of ly e, lb, rb) ires"
-apply(auto simp: crsp.simps abc_step_l.simps inv_stop.simps)
-done
-
-lemma crsp_step_dec_b_e:
-  assumes layout: "ly = layout_of ap"
-  and inv_start: "inv_locate_a (as, lm) (n, l, r) ires"
-  and dec_0: "abc_lm_v lm n = 0"
-  and fetch: "abc_fetch as ap = Some (Dec n e)"
-  shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
-  (steps (start_of ly as + 2 * n, l, r) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires"
-proof -
-  let ?P = "ci ly (start_of ly as) (Dec n e)"
-  let ?off = "start_of ly as - Suc 0"
-  have "\<exists> stp la ra. steps (start_of ly as + 2 * n, l, r) (?P, ?off) stp = (Suc (start_of ly as) + 2*n, la, ra)
-             \<and>  inv_locate_b (as, lm) (n, la, ra) ires"
-    using inv_start
-    apply(case_tac "r = [] \<or> hd r = Bk", simp_all)
-    done
-  from this obtain stpa la ra where a:
-    "steps (start_of ly as + 2 * n, l, r) (?P, ?off) stpa = (Suc (start_of ly as) + 2*n, la, ra)
-             \<and>  inv_locate_b (as, lm) (n, la, ra) ires" by blast
-  term dec_inv_1
-  have "\<exists> stp lb rb. steps (Suc (start_of ly as) + 2 * n, la, ra) (?P, ?off) stp = (start_of ly e, lb, rb)
-             \<and>  dec_inv_1 ly n e (as, lm) (start_of ly e, lb, rb) ires"
-    using assms a
-    apply(rule_tac crsp_step_dec_b_e_pre, auto)
-    done
-  from this obtain stpb lb rb where b:
-    "steps (Suc (start_of ly as) + 2 * n, la, ra) (?P, ?off) stpb = (start_of ly e, lb, rb)
-             \<and>  dec_inv_1 ly n e (as, lm) (start_of ly e, lb, rb) ires"  by blast
-  from a b show "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e))) 
-    (steps (start_of ly as + 2 * n, l, r) (?P, ?off) stp) ires"
-    apply(rule_tac x = "stpa + stpb" in exI)
-    apply(simp add: steps_add)
-    using dec_0
-    apply(simp add: dec_inv_1.simps)
-    apply(case_tac stpa, simp_all add: steps.simps)
-    done
-qed    
-  
-fun dec_inv_2 :: "layout \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> dec_inv_t"
-  where
-  "dec_inv_2 ly n e (as, am) (s, l, r) ires =
-           (let ss = start_of ly as in
-            let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in
-            let am'' = abc_lm_s am n (abc_lm_v am n) in
-              if s = 0 then False
-              else if s = ss + 2 * n then 
-                      inv_locate_a (as, am) (n, l, r) ires
-              else if s = ss + 2 * n + 1 then 
-                      inv_locate_n_b (as, am) (n, l, r) ires
-              else if s = ss + 2 * n + 2 then 
-                      dec_first_on_right_moving n (as, am'') (s, l, r) ires
-              else if s = ss + 2 * n + 3 then 
-                      dec_after_clear (as, am') (s, l, r) ires
-              else if s = ss + 2 * n + 4 then 
-                      dec_right_move (as, am') (s, l, r) ires
-              else if s = ss + 2 * n + 5 then 
-                      dec_check_right_move (as, am') (s, l, r) ires
-              else if s = ss + 2 * n + 6 then 
-                      dec_left_move (as, am') (s, l, r) ires
-              else if s = ss + 2 * n + 7 then 
-                      dec_after_write (as, am') (s, l, r) ires
-              else if s = ss + 2 * n + 8 then 
-                      dec_on_right_moving (as, am') (s, l, r) ires
-              else if s = ss + 2 * n + 9 then 
-                      dec_after_clear (as, am') (s, l, r) ires
-              else if s = ss + 2 * n + 10 then 
-                      inv_on_left_moving (as, am') (s, l, r) ires
-              else if s = ss + 2 * n + 11 then 
-                      inv_check_left_moving (as, am') (s, l, r) ires
-              else if s = ss + 2 * n + 12 then 
-                      inv_after_left_moving (as, am') (s, l, r) ires
-              else if s = ss + 2 * n + 16 then 
-                      inv_stop (as, am') (s, l, r) ires
-              else False)"
-
-declare dec_inv_2.simps[simp del]
-fun abc_dec_2_stage1 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "abc_dec_2_stage1 (s, l, r) ss n = 
-              (if s \<le> ss + 2*n + 1 then 7
-               else if s = ss + 2*n + 2 then 6 
-               else if s = ss + 2*n + 3 then 5
-               else if s \<ge> ss + 2*n + 4 \<and> s \<le> ss + 2*n + 9 then 4
-               else if s = ss + 2*n + 6 then 3
-               else if s = ss + 2*n + 10 \<or> s = ss + 2*n + 11 then 2
-               else if s = ss + 2*n + 12 then 1
-               else 0)"
-
-fun abc_dec_2_stage2 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "abc_dec_2_stage2 (s, l, r) ss n = 
-       (if s \<le> ss + 2 * n + 1 then (ss + 2 * n + 16 - s)
-        else if s = ss + 2*n + 10 then length l
-        else if s = ss + 2*n + 11 then length l
-        else if s = ss + 2*n + 4 then length r - 1
-        else if s = ss + 2*n + 5 then length r 
-        else if s = ss + 2*n + 7 then length r - 1
-        else if s = ss + 2*n + 8 then  
-              length r + length (takeWhile (\<lambda> a. a = Oc) l) - 1
-        else if s = ss + 2*n + 9 then 
-              length r + length (takeWhile (\<lambda> a. a = Oc) l) - 1
-        else 0)"
-
-fun abc_dec_2_stage3 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "abc_dec_2_stage3 (s, l, r) ss n  =
-        (if s \<le> ss + 2*n + 1 then 
-            if (s - ss) mod 2 = 0 then if r \<noteq> [] \<and> 
-                                          hd r = Oc then 0 else 1  
-            else length r
-         else if s = ss + 2 * n + 10 then 
-             if  r \<noteq> [] \<and> hd r = Oc then 2
-             else 1
-         else if s = ss + 2 * n + 11 then 
-             if r \<noteq> [] \<and> hd r = Oc then 3 
-             else 0 
-         else (ss + 2 * n + 16 - s))"
-
-fun abc_dec_2_stage4 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "abc_dec_2_stage4 (s, l, r) ss n = 
-          (if s = ss + 2*n + 2 then length r
-           else if s = ss + 2*n + 8 then length r
-           else if s = ss + 2*n + 3 then 
-               if r \<noteq> [] \<and> hd r = Oc then 1
-               else 0
-           else if s = ss + 2*n + 7 then 
-               if r \<noteq> [] \<and> hd r = Oc then 0 
-               else 1
-           else if s = ss + 2*n + 9 then 
-               if r \<noteq> [] \<and> hd r = Oc then 1
-               else 0 
-           else 0)"
-
-fun abc_dec_2_measure :: "(config \<times> nat \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat \<times> nat)"
-  where
-  "abc_dec_2_measure (c, ss, n) = 
-  (abc_dec_2_stage1 c ss n, 
-  abc_dec_2_stage2 c ss n, abc_dec_2_stage3 c ss n,  abc_dec_2_stage4 c ss n)"
-
-definition lex_square:: 
-   "((nat \<times> nat \<times> nat \<times> nat) \<times> (nat \<times> nat \<times> nat \<times> nat)) set"
-  where "lex_square \<equiv> less_than <*lex*> lex_triple"
-
-definition abc_dec_2_LE ::
-  "((config \<times> nat \<times>
-  nat) \<times> (config \<times> nat \<times> nat)) set"
-  where "abc_dec_2_LE \<equiv> (inv_image lex_square abc_dec_2_measure)"
-
-lemma wf_dec2_le: "wf abc_dec_2_LE"
-by(auto intro:wf_inv_image simp:abc_dec_2_LE_def lex_square_def lex_triple_def lex_pair_def)
-
-lemma fix_add: "fetch ap ((x::nat) + 2*n) b = fetch ap (2*n + x) b"
-by (metis Suc_1 mult_2 nat_add_commute)
-
-lemma [elim]: 
- "\<lbrakk>0 < abc_lm_v am n; inv_locate_n_b (as, am) (n, aaa, Bk # xs) ires\<rbrakk> 
- \<Longrightarrow> RR"
-apply(auto simp: inv_locate_n_b.simps abc_lm_v.simps split: if_splits)
-apply(case_tac [!] m, auto)
-done
- 
-lemma [elim]:
- "\<lbrakk>0 < abc_lm_v am n; inv_locate_n_b (as, am) 
-                                (n, aaa, []) ires\<rbrakk> \<Longrightarrow> RR"
-apply(auto simp: inv_locate_n_b.simps abc_lm_v.simps split: if_splits)
-done
-
-lemma [simp]: "dec_after_write (as, am) (s, aa, r) ires
-           \<Longrightarrow> takeWhile (\<lambda>a. a = Oc) aa = []"
-apply(simp only : dec_after_write.simps)
-apply(erule exE)+
-apply(erule_tac conjE)+
-apply(case_tac aa, simp)
-apply(case_tac a, simp only: takeWhile.simps , simp_all split: if_splits)
-done
-
-lemma [simp]: 
-     "\<lbrakk>dec_on_right_moving (as, lm) (s, aa, []) ires; 
-       length (takeWhile (\<lambda>a. a = Oc) (tl aa)) 
-           \<noteq> length (takeWhile (\<lambda>a. a = Oc) aa) - Suc 0\<rbrakk>
-    \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (tl aa)) < 
-                       length (takeWhile (\<lambda>a. a = Oc) aa) - Suc 0"
-apply(simp only: dec_on_right_moving.simps)
-apply(erule_tac exE)+
-apply(erule_tac conjE)+
-apply(case_tac mr, auto split: if_splits)
-done
-
-lemma [simp]: 
-  "dec_after_clear (as, abc_lm_s am n (abc_lm_v am n - Suc 0)) 
-             (start_of (layout_of aprog) as + 2 * n + 9, aa, Bk # xs) ires
- \<Longrightarrow> length xs - Suc 0 < length xs + 
-                             length (takeWhile (\<lambda>a. a = Oc) aa)"
-apply(simp only: dec_after_clear.simps)
-apply(erule_tac exE)+
-apply(erule conjE)+
-apply(simp split: if_splits )
-done
-
-lemma [simp]: 
- "\<lbrakk>dec_after_clear (as, abc_lm_s am n (abc_lm_v am n - Suc 0))
-       (start_of (layout_of aprog) as + 2 * n + 9, aa, []) ires\<rbrakk>
-    \<Longrightarrow> Suc 0 < length (takeWhile (\<lambda>a. a = Oc) aa)"
-apply(simp add: dec_after_clear.simps split: if_splits)
-done
-
-lemma [elim]: 
-  "inv_check_left_moving (as, lm)
-  (s, [], Oc # xs) ires
- \<Longrightarrow> RR"
-apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps)
-done
-
-lemma [simp]:
-"\<lbrakk>0 < abc_lm_v am n; 
-  at_begin_norm (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
-  \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires"
-apply(simp only: at_begin_norm.simps inv_locate_n_b.simps)
-apply(erule_tac exE)+
-apply(rule_tac x = lm1 in exI, simp)
-apply(case_tac "length lm2", simp)
-apply(case_tac "lm2", simp, simp)
-apply(case_tac "lm2", auto simp: tape_of_nl_cons split: if_splits)
-done
-
-lemma [simp]: 
- "\<lbrakk>0 < abc_lm_v am n; 
-   at_begin_fst_awtn (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
- \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires"
-apply(simp only: at_begin_fst_awtn.simps inv_locate_n_b.simps )
-apply(erule exE)+
-apply(erule conjE)+
-apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, 
-      rule_tac x = "Suc tn" in exI, rule_tac x = 0 in exI)
-apply(simp add: exp_ind del: replicate.simps)
-apply(rule conjI)+
-apply(auto)
-done
-
-lemma [simp]: 
- "\<lbrakk>0 < abc_lm_v am n; inv_locate_a (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
- \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc#aaa, xs) ires"
-apply(auto simp: inv_locate_a.simps at_begin_fst_bwtn.simps)
-done
-
-lemma [simp]: 
- "\<lbrakk>dec_on_right_moving (as, am) (s, aa, Bk # xs) ires; 
-   Suc (length (takeWhile (\<lambda>a. a = Oc) (tl aa)))
-   \<noteq> length (takeWhile (\<lambda>a. a = Oc) aa)\<rbrakk>
-  \<Longrightarrow> Suc (length (takeWhile (\<lambda>a. a = Oc) (tl aa))) 
-    < length (takeWhile (\<lambda>a. a = Oc) aa)"
-apply(simp only: dec_on_right_moving.simps)
-apply(erule exE)+
-apply(erule conjE)+
-apply(case_tac ml, auto split: if_splits )
-done
-
-lemma crsp_step_dec_b_suc_pre:
-  assumes layout: "ly = layout_of ap"
-  and crsp: "crsp ly (as, lm) (s, l, r) ires"
-  and inv_start: "inv_locate_a (as, lm) (n, la, ra) ires"
-  and fetch: "abc_fetch as ap = Some (Dec n e)"
-  and dec_suc: "0 < abc_lm_v lm n"
-  and f: "f = (\<lambda> stp. (steps (start_of ly as + 2 * n, la, ra) (ci ly (start_of ly as) (Dec n e), 
-            start_of ly as - Suc 0) stp, start_of ly as, n))"
-  and P: "P = (\<lambda> ((s, l, r), ss, x). s = start_of ly as + 2*n + 16)"
-  and Q: "Q = (\<lambda> ((s, l, r), ss, x). dec_inv_2 ly x e (as, lm) (s, l, r) ires)"
-  shows "\<exists> stp. P (f stp) \<and> Q(f stp)"
-  proof(rule_tac LE = abc_dec_2_LE in halt_lemma2)
-  show "wf abc_dec_2_LE" by(intro wf_dec2_le)
-next
-  show "Q (f 0)"
-    using layout fetch inv_start
-    apply(simp add: f steps.simps Q)
-    apply(simp only: dec_inv_2.simps)
-    apply(auto simp: Let_def start_of_ge start_of_less inv_start dec_inv_2.simps)
-    done
-next
-  show "\<not> P (f 0)"
-    using layout fetch
-    apply(simp add: f steps.simps P)
-    done
-next
-  show "\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> abc_dec_2_LE"
-    using fetch
-  proof(rule_tac allI, rule_tac impI)
-    fix na
-    assume "\<not> P (f na) \<and> Q (f na)"
-    thus "Q (f (Suc na)) \<and> (f (Suc na), f na) \<in> abc_dec_2_LE"
-      apply(simp add: f)
-      apply(case_tac "steps ((start_of ly as + 2 * n), la, ra)
-        (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) na", simp)
-    proof -
-      fix a b c 
-      assume "\<not> P ((a, b, c), start_of ly as, n) \<and> Q ((a, b, c), start_of ly as, n)"
-      thus "Q (step (a, b, c) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0), start_of ly as, n) \<and>
-               ((step (a, b, c) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0), start_of ly as, n), 
-                   (a, b, c), start_of ly as, n) \<in> abc_dec_2_LE"
-        apply(simp add: Q)
-        apply(erule_tac conjE)
-        apply(case_tac c, case_tac [2] aa)
-        apply(simp_all add: dec_inv_2.simps Let_def)
-        apply(simp_all split: if_splits)
-        using fetch layout dec_suc
-        apply(auto simp: step.simps P dec_inv_2.simps Let_def abc_dec_2_LE_def lex_triple_def lex_pair_def lex_square_def
-                         fix_add numeral_3_eq_3) 
-        done
-    qed
-  qed
-qed
-
-lemma [simp]: 
-  "\<lbrakk>inv_stop (as, abc_lm_s lm n (abc_lm_v lm n - Suc 0)) 
-  (start_of (layout_of ap) as + 2 * n + 16, a, b) ires;
-   abc_lm_v lm n > 0;
-   abc_fetch as ap = Some (Dec n e)\<rbrakk>
-  \<Longrightarrow> crsp (layout_of ap) (abc_step_l (as, lm) (Some (Dec n e))) 
-  (start_of (layout_of ap) as + 2 * n + 16, a, b) ires"
-apply(auto simp: inv_stop.simps crsp.simps  abc_step_l.simps startof_Suc2)
-apply(drule_tac startof_Suc2, simp)
-done
-  
-lemma crsp_step_dec_b_suc:
-  assumes layout: "ly = layout_of ap"
-  and crsp: "crsp ly (as, lm) (s, l, r) ires"
-  and inv_start: "inv_locate_a (as, lm) (n, la, ra) ires"
-  and fetch: "abc_fetch as ap = Some (Dec n e)"
-  and dec_suc: "0 < abc_lm_v lm n"
-  shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
-              (steps (start_of ly as + 2 * n, la, ra) (ci (layout_of ap) 
-                  (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires"
-  using assms
-  apply(drule_tac crsp_step_dec_b_suc_pre, auto)
-  apply(rule_tac x = stp in exI, simp)
-  apply(simp add: dec_inv_2.simps)
-  apply(case_tac stp, simp_all add: steps.simps)
-  done
-
-lemma crsp_step_dec_b:
-  assumes layout: "ly = layout_of ap"
-  and crsp: "crsp ly (as, lm) (s, l, r) ires"
-  and inv_start: "inv_locate_a (as, lm) (n, la, ra) ires"
-  and fetch: "abc_fetch as ap = Some (Dec n e)"
-  shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
-  (steps (start_of ly as + 2 * n, la, ra) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires"
-using assms
-apply(case_tac "abc_lm_v lm n = 0")
-apply(rule_tac crsp_step_dec_b_e, simp_all)
-apply(rule_tac crsp_step_dec_b_suc, simp_all)
-done
-
-lemma crsp_step_dec: 
-  assumes layout: "ly = layout_of ap"
-  and crsp: "crsp ly (as, lm) (s, l, r) ires"
-  and fetch: "abc_fetch as ap = Some (Dec n e)"
-  shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
-  (steps (s, l, r) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires"
-proof(simp add: ci.simps)
-  let ?off = "start_of ly as - Suc 0"
-  let ?A = "findnth n"
-  let ?B = "sete (shift (shift tdec_b (2 * n)) ?off) (start_of ly e)"
-  have "\<exists> stp la ra. steps (s, l, r) (shift ?A ?off @ ?B, ?off) stp = (start_of ly as + 2*n, la, ra)
-                    \<and> inv_locate_a (as, lm) (n, la, ra) ires"
-  proof -
-    have "\<exists>stp l' r'. steps (Suc 0, l, r) (?A, 0) stp = (Suc (2 * n), l', r') \<and> 
-                     inv_locate_a (as, lm) (n, l', r') ires"
-      using assms
-      apply(rule_tac findnth_correct, simp_all)
-      done
-    then obtain stp l' r' where a: 
-      "steps (Suc 0, l, r) (?A, 0) stp = (Suc (2 * n), l', r') \<and> 
-      inv_locate_a (as, lm) (n, l', r') ires" by blast
-    then have "steps (Suc 0 + ?off, l, r) (shift ?A ?off, ?off) stp = (Suc (2 * n) + ?off, l', r')"
-      apply(rule_tac tm_shift_eq_steps, simp_all)
-      done
-    moreover have "s = start_of ly as"
-      using crsp
-      apply(auto simp: crsp.simps)
-      done
-    ultimately show "\<exists> stp la ra. steps (s, l, r) (shift ?A ?off @ ?B, ?off) stp = (start_of ly as + 2*n, la, ra)
-                    \<and> inv_locate_a (as, lm) (n, la, ra) ires"
-      using a
-      apply(drule_tac B = ?B in tm_append_first_steps_eq, auto)
-      apply(rule_tac x = stp in exI, simp)
-      done
-  qed
-  from this obtain stpa la ra where a: 
-    "steps (s, l, r) (shift ?A ?off @ ?B, ?off) stpa = (start_of ly as + 2*n, la, ra)
-                    \<and> inv_locate_a (as, lm) (n, la, ra) ires" by blast
-  have "\<exists>stp. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
-           (steps (start_of ly as + 2*n, la, ra) (shift ?A ?off @ ?B, ?off) stp) ires \<and> stp > 0"
-    using assms a
-    apply(drule_tac crsp_step_dec_b, auto)
-    apply(rule_tac x = stp in exI, simp add: ci.simps)
-    done
-  then obtain stpb where b: 
-    "crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
-    (steps (start_of ly as + 2*n, la, ra) (shift ?A ?off @ ?B, ?off) stpb) ires \<and> stpb > 0" ..
-  from a b show "\<exists> stp>0. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
-    (steps (s, l, r) (shift ?A ?off @ ?B, ?off) stp) ires"
-    apply(rule_tac x = "stpa + stpb" in exI)
-    apply(simp add: steps_add)
-    done
-qed    
-  
-subsection{*Crsp of Goto*}
-
-lemma crsp_step_goto:
-  assumes layout: "ly = layout_of ap"
-  and crsp: "crsp ly (as, lm) (s, l, r) ires"
-  shows "\<exists>stp>0. crsp ly (abc_step_l (as, lm) (Some (Goto n)))
-  (steps (s, l, r) (ci ly (start_of ly as) (Goto n), 
-            start_of ly as - Suc 0) stp) ires"
-using crsp
-apply(rule_tac x = "Suc 0" in exI)
-apply(case_tac r, case_tac [2] a)
-apply(simp_all add: ci.simps steps.simps step.simps crsp.simps fetch.simps
-  crsp.simps abc_step_l.simps)
-done
-
-lemma crsp_step_in:
-  assumes layout: "ly = layout_of ap"
-  and compile: "tp = tm_of ap"
-  and crsp: "crsp ly (as, lm) (s, l, r) ires"
-  and fetch: "abc_fetch as ap = Some ins"
-  shows "\<exists> stp>0. crsp ly (abc_step_l (as, lm) (Some ins))
-                      (steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp) ires"
-  using assms
-  apply(case_tac ins, simp_all)
-  apply(rule crsp_step_inc, simp_all)
-  apply(rule crsp_step_dec, simp_all)
-  apply(rule_tac crsp_step_goto, simp_all)
-  done
-
-lemma crsp_step:
-  assumes layout: "ly = layout_of ap"
-  and compile: "tp = tm_of ap"
-  and crsp: "crsp ly (as, lm) (s, l, r) ires"
-  and fetch: "abc_fetch as ap = Some ins"
-  shows "\<exists> stp>0. crsp ly (abc_step_l (as, lm) (Some ins))
-                      (steps (s, l, r) (tp, 0) stp) ires"
-proof -
-  have "\<exists> stp>0. crsp ly (abc_step_l (as, lm) (Some ins))
-                      (steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp) ires"
-    using assms
-    apply(rule_tac crsp_step_in, simp_all)
-    done
-  from this obtain stp where d: "stp > 0 \<and> crsp ly (abc_step_l (as, lm) (Some ins))
-                      (steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp) ires" ..
-  obtain s' l' r' where e:
-    "(steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp) = (s', l', r')"
-    apply(case_tac "(steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp)")
-    by blast
-  then have "steps (s, l, r) (tp, 0) stp = (s', l', r')"
-    using assms d
-    apply(rule_tac steps_eq_in)
-    apply(simp_all)
-    apply(case_tac "(abc_step_l (as, lm) (Some ins))", simp add: crsp.simps)
-    done    
-  thus " \<exists>stp>0. crsp ly (abc_step_l (as, lm) (Some ins)) (steps (s, l, r) (tp, 0) stp) ires"
-    using d e
-    apply(rule_tac x = stp in exI, simp)
-    done
-qed
-
-lemma crsp_steps:
-  assumes layout: "ly = layout_of ap"
-  and compile: "tp = tm_of ap"
-  and crsp: "crsp ly (as, lm) (s, l, r) ires"
-  shows "\<exists> stp. crsp ly (abc_steps_l (as, lm) ap n)
-                      (steps (s, l, r) (tp, 0) stp) ires"
-(*
-proof(induct n)
-  case 0
-  have "crsp ly (abc_steps_l (as, lm) ap 0) (steps (s, l, r) (tp, 0) 0) ires"
-    using crsp by(simp add: steps.simps abc_steps_l.simps)
-  thus "?case"
-    by(rule_tac x = 0 in exI, simp)
-next
-  case (Suc n)
-  obtain as' lm' where a: "abc_steps_l (as, lm) ap n = (as', lm')"
-    by(case_tac "abc_steps_l (as, lm) ap n", auto) 
-  have "\<exists>stp\<ge>n. crsp ly (abc_steps_l (as, lm) ap n) (steps (s, l, r) (tp, 0) stp) ires"
-    by fact
-  from this a obtain stpa where b:
-    "stpa\<ge>n \<and> crsp ly (as', lm') (steps (s, l, r) (tp, 0) stpa) ires" by auto
-  obtain s' l' r' where "steps (s, l, r) (tp, 0) stpa = (s', l', r')"
-    by(case_tac "steps (s, l, r) (tp, 0) stpa")
-  then have "stpa\<ge>n \<and> crsp ly (as', lm') (s', l', r') ires" using b by simp
-  from a and this show "?case"
-  proof(cases "abc_fetch as' ap")
-    case None
-    
-  
-
-    have "crsp ly (abc_steps_l (as, lm) ap 0) (steps (s, l, r) (tp, 0) stp) ires"
-    apply(simp add: steps.simps abc_steps_l.simps)
-*)
-  using crsp
-  apply(induct n)
-  apply(rule_tac x = 0 in exI) 
-  apply(simp add: steps.simps abc_steps_l.simps, simp)
-  apply(case_tac "(abc_steps_l (as, lm) ap n)", auto)
-  apply(frule_tac abc_step_red, simp)
-  apply(case_tac "abc_fetch a ap", simp add: abc_step_l.simps, auto)
-  apply(case_tac "steps (s, l, r) (tp, 0) stp", simp)
-  using assms
-  apply(drule_tac s = ab and l = ba and r = c in crsp_step, auto)
-  apply(rule_tac x = "stp + stpa" in exI, simp add: steps_add)
-  done
-
-lemma tp_correct': 
-  assumes layout: "ly = layout_of ap"
-  and compile: "tp = tm_of ap"
-  and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires"
-  and abc_halt: "abc_steps_l (0, lm) ap stp = (length ap, am)"
-  shows "\<exists> stp k. steps (Suc 0, l, r) (tp, 0) stp = (start_of ly (length ap), Bk # Bk # ires, <am> @ Bk\<up>k)"
-  using assms
-  apply(drule_tac n = stp in crsp_steps, auto)
-  apply(rule_tac x = stpa in exI)
-  apply(case_tac "steps (Suc 0, l, r) (tm_of ap, 0) stpa", simp add: crsp.simps)
-  done
-
-text{*The tp @ [(Nop, 0), (Nop, 0)] is nomoral turing machines, so we can use Hoare_plus when composing with Mop machine*}
-
-thm layout_of.simps
-lemma layout_id_cons: "layout_of (ap @ [p]) = layout_of ap @ [length_of p]"
-apply(simp add: layout_of.simps)
-done
-
-lemma [simp]: "length (layout_of xs) = length xs"
-by(simp add: layout_of.simps)
-
-thm tms_of.simps
-term ci
-thm tms_of.simps
-thm tpairs_of.simps
-
-lemma [simp]:  
-  "map (start_of (layout_of xs @ [length_of x])) [0..<length xs] =  (map (start_of (layout_of xs)) [0..<length xs])"
-apply(auto)
-apply(simp add: layout_of.simps start_of.simps)
-done
-
-lemma tpairs_id_cons: 
-  "tpairs_of (xs @ [x]) = tpairs_of xs @ [(start_of (layout_of (xs @ [x])) (length xs), x)]"
-apply(auto simp: tpairs_of.simps layout_id_cons )
-done
-
-lemma map_length_ci:
-  "(map (length \<circ> (\<lambda>(xa, y). ci (layout_of xs @ [length_of x]) xa y)) (tpairs_of xs)) = 
-  (map (length \<circ> (\<lambda>(x, y). ci (layout_of xs) x y)) (tpairs_of xs)) "
-apply(auto)
-apply(case_tac b, auto simp: ci.simps sete.simps)
-done
-
-lemma length_tp'[simp]: 
-  "\<lbrakk>ly = layout_of ap; tp = tm_of ap\<rbrakk> \<Longrightarrow>
-       length tp = 2 * listsum (take (length ap) (layout_of ap))"
-proof(induct ap arbitrary: ly tp rule: rev_induct)
-  case Nil
-  thus "?case"
-    by(simp add: tms_of.simps tm_of.simps tpairs_of.simps)
-next
-  fix x xs ly tp
-  assume ind: "\<And>ly tp. \<lbrakk>ly = layout_of xs; tp = tm_of xs\<rbrakk> \<Longrightarrow> 
-    length tp = 2 * listsum (take (length xs) (layout_of xs))"
-  and layout: "ly = layout_of (xs @ [x])"
-  and tp: "tp = tm_of (xs @ [x])"
-  obtain ly' where a: "ly' = layout_of xs"
-    by metis
-  obtain tp' where b: "tp' = tm_of xs"
-    by metis
-  have c: "length tp' = 2 * listsum (take (length xs) (layout_of xs))"
-    using a b
-    by(erule_tac ind, simp)
-  thus "length tp = 2 * 
-    listsum (take (length (xs @ [x])) (layout_of (xs @ [x])))"
-    using tp b
-    apply(auto simp: layout_id_cons tm_of.simps tms_of.simps length_concat tpairs_id_cons map_length_ci)
-    apply(case_tac x)
-    apply(auto simp: ci.simps tinc_b_def tdec_b_def length_findnth sete.simps length_of.simps
-                 split: abc_inst.splits)
-    done
-qed
-
-lemma [simp]: 
-  "\<lbrakk>ly = layout_of ap; tp = tm_of ap\<rbrakk> \<Longrightarrow>
-        fetch (tp @ [(Nop, 0), (Nop, 0)]) (start_of ly (length ap)) b = 
-       (Nop, 0)"
-apply(case_tac b)
-apply(simp_all add: start_of.simps fetch.simps nth_append)
-done
-(*
-lemma tp_correct: 
-  assumes layout: "ly = layout_of ap"
-  and compile: "tp = tm_of ap"
-  and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires"
-  and abc_halt: "abc_steps_l (0, lm) ap stp = (length ap, am)"
-  shows "\<exists> stp k. steps (Suc 0, l, r) (tp @ [(Nop, 0), (Nop, 0)], 0) stp = (0, Bk # Bk # ires, <am> @ Bk\<up>k)"
-  using assms
-proof -
-  have "\<exists> stp k. steps (Suc 0, l, r) (tp @ [(Nop, 0), (Nop, 0)], 0) stp =
-    (start_of ly (length ap), Bk # Bk # ires, <am> @ Bk\<up>k)"
-  proof -
-    have "\<exists> stp k. steps (Suc 0, l, r) (tp, 0) stp =
-    (start_of ly (length ap), Bk # Bk # ires, <am> @ Bk\<up>k)"
-      using assms
-      apply(rule_tac tp_correct', simp_all)
-      done
-    from this obtain stp k where "steps (Suc 0, l, r) (tp, 0) stp =
-    (start_of ly (length ap), Bk # Bk # ires, <am> @ Bk\<up>k)" by blast
-    thus "?thesis"
-      apply(rule_tac x = stp in exI, rule_tac x = k in exI)
-      apply(drule_tac tm_append_first_steps_eq, simp_all)
-      done
-  qed
-  from this obtain stp k where 
-    "steps (Suc 0, l, r) (tp @ [(Nop, 0), (Nop, 0)], 0) stp =
-    (start_of ly (length ap), Bk # Bk # ires, <am> @ Bk\<up>k)"
-    by blast
-  thus "\<exists>stp k. steps (Suc 0, l, r) (tp @ [(Nop, 0), (Nop, 0)], 0) stp 
-    = (0, Bk # Bk # ires, <am> @ Bk \<up> k)"
-    using assms
-    apply(rule_tac x = "stp + Suc 0" in exI)
-    apply(simp add: steps_add)
-    apply(auto simp: step.simps)
-    done
-qed
- *)   
-(********for mopup***********)
-fun mopup_a :: "nat \<Rightarrow> instr list"
-  where
-  "mopup_a 0 = []" |
-  "mopup_a (Suc n) = mopup_a n @ 
-       [(R, 2*n + 3), (W0, 2*n + 2), (R, 2*n + 1), (W1, 2*n + 2)]"
-
-definition mopup_b :: "instr list"
-  where
-  "mopup_b \<equiv> [(R, 2), (R, 1), (L, 5), (W0, 3), (R, 4), (W0, 3),
-            (R, 2), (W0, 3), (L, 5), (L, 6), (R, 0), (L, 6)]"
-
-fun mopup :: "nat \<Rightarrow> instr list"
-  where 
-  "mopup n = mopup_a n @ shift mopup_b (2*n)"
-(****)
-
-type_synonym mopup_type = "config \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> cell list \<Rightarrow> bool"
-
-fun mopup_stop :: "mopup_type"
-  where
-  "mopup_stop (s, l, r) lm n ires= 
-        (\<exists> ln rn. l = Bk\<up>ln @ Bk # Bk # ires \<and> r = <abc_lm_v lm n> @ Bk\<up>rn)"
-
-fun mopup_bef_erase_a :: "mopup_type"
-  where
-  "mopup_bef_erase_a (s, l, r) lm n ires= 
-         (\<exists> ln m rn. l = Bk\<up>ln @ Bk # Bk # ires \<and> 
-                  r = Oc\<up>m@ Bk # <(drop ((s + 1) div 2) lm)> @ Bk\<up>rn)"
-
-fun mopup_bef_erase_b :: "mopup_type"
-  where
-  "mopup_bef_erase_b (s, l, r) lm n ires = 
-      (\<exists> ln m rn. l = Bk\<up>ln @ Bk # Bk # ires \<and> r = Bk # Oc\<up>m @ Bk # 
-                                      <(drop (s div 2) lm)> @ Bk\<up>rn)"
-
-fun mopup_jump_over1 :: "mopup_type"
-  where
-  "mopup_jump_over1 (s, l, r) lm n ires = 
-      (\<exists> ln m1 m2 rn. m1 + m2 = Suc (abc_lm_v lm n) \<and> 
-        l = Oc\<up>m1 @ Bk\<up>ln @ Bk # Bk # ires \<and> 
-     (r = Oc\<up>m2 @ Bk # <(drop (Suc n) lm)> @ Bk\<up>rn \<or> 
-     (r = Oc\<up>m2 \<and> (drop (Suc n) lm) = [])))"
-
-fun mopup_aft_erase_a :: "mopup_type"
-  where
-  "mopup_aft_erase_a (s, l, r) lm n ires = 
-      (\<exists> lnl lnr rn (ml::nat list) m. 
-          m = Suc (abc_lm_v lm n) \<and> l = Bk\<up>lnr @ Oc\<up>m @ Bk\<up>lnl @ Bk # Bk # ires \<and> 
-                                   (r = <ml> @ Bk\<up>rn))"
-
-fun mopup_aft_erase_b :: "mopup_type"
-  where
-  "mopup_aft_erase_b (s, l, r) lm n ires= 
-   (\<exists> lnl lnr rn (ml::nat list) m. 
-      m = Suc (abc_lm_v lm n) \<and> 
-      l = Bk\<up>lnr @ Oc\<up>m @ Bk\<up>lnl @ Bk # Bk # ires \<and> 
-     (r = Bk # <ml> @ Bk\<up>rn \<or>
-      r = Bk # Bk # <ml> @ Bk\<up>rn))"
-
-fun mopup_aft_erase_c :: "mopup_type"
-  where
-  "mopup_aft_erase_c (s, l, r) lm n ires = 
- (\<exists> lnl lnr rn (ml::nat list) m. 
-     m = Suc (abc_lm_v lm n) \<and> 
-     l = Bk\<up>lnr @ Oc\<up>m @ Bk\<up>lnl @ Bk # Bk # ires \<and> 
-    (r = <ml> @ Bk\<up>rn \<or> r = Bk # <ml> @ Bk\<up>rn))"
-
-fun mopup_left_moving :: "mopup_type"
-  where
-  "mopup_left_moving (s, l, r) lm n ires = 
-  (\<exists> lnl lnr rn m.
-     m = Suc (abc_lm_v lm n) \<and> 
-   ((l = Bk\<up>lnr @ Oc\<up>m @ Bk\<up>lnl @ Bk # Bk # ires \<and> r = Bk\<up>rn) \<or>
-    (l = Oc\<up>(m - 1) @ Bk\<up>lnl @ Bk # Bk # ires \<and> r = Oc # Bk\<up>rn)))"
-
-fun mopup_jump_over2 :: "mopup_type"
-  where
-  "mopup_jump_over2 (s, l, r) lm n ires = 
-     (\<exists> ln rn m1 m2.
-          m1 + m2 = Suc (abc_lm_v lm n) 
-        \<and> r \<noteq> [] 
-        \<and> (hd r = Oc \<longrightarrow> (l = Oc\<up>m1 @ Bk\<up>ln @ Bk # Bk # ires \<and> r = Oc\<up>m2 @ Bk\<up>rn)) 
-        \<and> (hd r = Bk \<longrightarrow> (l = Bk\<up>ln @ Bk # ires \<and> r = Bk # Oc\<up>(m1+m2)@ Bk\<up>rn)))"
-
-
-fun mopup_inv :: "mopup_type"
-  where
-  "mopup_inv (s, l, r) lm n ires = 
-      (if s = 0 then mopup_stop (s, l, r) lm n ires
-       else if s \<le> 2*n then
-               if s mod 2 = 1 then mopup_bef_erase_a (s, l, r) lm n ires
-                   else mopup_bef_erase_b (s, l, r) lm n ires
-            else if s = 2*n + 1 then 
-                mopup_jump_over1 (s, l, r) lm n ires
-            else if s = 2*n + 2 then mopup_aft_erase_a (s, l, r) lm n ires
-            else if s = 2*n + 3 then mopup_aft_erase_b (s, l, r) lm n ires
-            else if s = 2*n + 4 then mopup_aft_erase_c (s, l, r) lm n ires
-            else if s = 2*n + 5 then mopup_left_moving (s, l, r) lm n ires
-            else if s = 2*n + 6 then mopup_jump_over2 (s, l, r) lm n ires
-            else False)"
-
-lemma mopup_fetch_0[simp]: 
-     "(fetch (mopup_a n @ shift mopup_b (2 * n)) 0 b) = (Nop, 0)"
-by(simp add: fetch.simps)
-
-lemma mop_bef_length[simp]: "length (mopup_a n) = 4 * n"
-apply(induct n, simp_all add: mopup_a.simps)
-done
-
-lemma mopup_a_nth: 
-  "\<lbrakk>q < n; x < 4\<rbrakk> \<Longrightarrow> mopup_a n ! (4 * q + x) = 
-                             mopup_a (Suc q) ! ((4 * q) + x)"
-apply(induct n, simp)
-apply(case_tac "q < n", simp add: mopup_a.simps, auto)
-apply(simp add: nth_append)
-apply(subgoal_tac "q = n", simp)
-apply(arith)
-done
-
-lemma fetch_bef_erase_a_o[simp]: 
- "\<lbrakk>0 < s; s \<le> 2 * n; s mod 2 = Suc 0\<rbrakk>
-  \<Longrightarrow> (fetch (mopup_a n @ shift mopup_b (2 * n)) s Oc) = (W0, s + 1)"
-apply(subgoal_tac "\<exists> q. s = 2*q + 1", auto)
-apply(subgoal_tac "length (mopup_a n) = 4*n")
-apply(auto simp: fetch.simps nth_of.simps nth_append)
-apply(subgoal_tac "mopup_a n ! (4 * q + 1) = 
-                      mopup_a (Suc q) ! ((4 * q) + 1)", 
-      simp add: mopup_a.simps nth_append)
-apply(rule mopup_a_nth, auto)
-apply arith
-done
-
-lemma fetch_bef_erase_a_b[simp]:
-  "\<lbrakk>0 < s; s \<le> 2 * n; s mod 2 = Suc 0\<rbrakk>
-   \<Longrightarrow>  (fetch (mopup_a n @ shift mopup_b (2 * n)) s Bk) = (R, s + 2)"
-apply(subgoal_tac "\<exists> q. s = 2*q + 1", auto)
-apply(subgoal_tac "length (mopup_a n) = 4*n")
-apply(auto simp: fetch.simps nth_of.simps nth_append)
-apply(subgoal_tac "mopup_a n ! (4 * q + 0) = 
-                       mopup_a (Suc q) ! ((4 * q + 0))", 
-      simp add: mopup_a.simps nth_append)
-apply(rule mopup_a_nth, auto)
-apply arith
-done
-
-lemma fetch_bef_erase_b_b: 
-  "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = 0\<rbrakk> \<Longrightarrow> 
-     (fetch (mopup_a n @ shift mopup_b (2 * n)) s Bk) = (R, s - 1)"
-apply(subgoal_tac "\<exists> q. s = 2 * q", auto)
-apply(case_tac qa, simp, simp)
-apply(auto simp: fetch.simps nth_of.simps nth_append)
-apply(subgoal_tac "mopup_a n ! (4 * nat + 2) = 
-                     mopup_a (Suc nat) ! ((4 * nat) + 2)", 
-      simp add: mopup_a.simps nth_append)
-apply(rule mopup_a_nth, auto)
-done
-
-lemma fetch_jump_over1_o: 
- "fetch (mopup_a n @ shift mopup_b (2 * n)) (Suc (2 * n)) Oc
-  = (R, Suc (2 * n))"
-apply(subgoal_tac "length (mopup_a n) = 4 * n")
-apply(auto simp: fetch.simps nth_of.simps mopup_b_def nth_append 
-                 shift.simps)
-done
-
-lemma fetch_jump_over1_b: 
- "fetch (mopup_a n @ shift mopup_b (2 * n)) (Suc (2 * n)) Bk 
- = (R, Suc (Suc (2 * n)))"
-apply(subgoal_tac "length (mopup_a n) = 4 * n")
-apply(auto simp: fetch.simps nth_of.simps mopup_b_def 
-                 nth_append shift.simps)
-done
-
-lemma fetch_aft_erase_a_o: 
- "fetch (mopup_a n @ shift mopup_b (2 * n)) (Suc (Suc (2 * n))) Oc 
- = (W0, Suc (2 * n + 2))"
-apply(subgoal_tac "length (mopup_a n) = 4 * n")
-apply(auto simp: fetch.simps nth_of.simps mopup_b_def 
-                 nth_append shift.simps)
-done
-
-lemma fetch_aft_erase_a_b: 
- "fetch (mopup_a n @ shift mopup_b (2 * n)) (Suc (Suc (2 * n))) Bk
-  = (L, Suc (2 * n + 4))"
-apply(subgoal_tac "length (mopup_a n) = 4 * n")
-apply(auto simp: fetch.simps nth_of.simps mopup_b_def 
-                 nth_append shift.simps)
-done
-
-lemma fetch_aft_erase_b_b: 
- "fetch (mopup_a n @ shift mopup_b (2 * n)) (2*n + 3) Bk
-  = (R, Suc (2 * n + 3))"
-apply(subgoal_tac "length (mopup_a n) = 4 * n")
-apply(subgoal_tac "2*n + 3 = Suc (2*n + 2)", simp only: fetch.simps)
-apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps)
-done
-
-lemma fetch_aft_erase_c_o: 
- "fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 4) Oc 
- = (W0, Suc (2 * n + 2))"
-apply(subgoal_tac "length (mopup_a n) = 4 * n")
-apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps)
-apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps)
-done
-
-lemma fetch_aft_erase_c_b: 
- "fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 4) Bk 
- = (R, Suc (2 * n + 1))"
-apply(subgoal_tac "length (mopup_a n) = 4 * n")
-apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps)
-apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps)
-done
-
-lemma fetch_left_moving_o: 
- "(fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 5) Oc) 
- = (L, 2*n + 6)"
-apply(subgoal_tac "length (mopup_a n) = 4 * n")
-apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps)
-apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps)
-done
-
-lemma fetch_left_moving_b: 
- "(fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 5) Bk)
-  = (L, 2*n + 5)"
-apply(subgoal_tac "length (mopup_a n) = 4 * n")
-apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps)
-apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps)
-done
-
-lemma fetch_jump_over2_b:
-  "(fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 6) Bk) 
- = (R, 0)"
-apply(subgoal_tac "length (mopup_a n) = 4 * n")
-apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps)
-apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps)
-done
-
-lemma fetch_jump_over2_o: 
-"(fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 6) Oc) 
- = (L, 2*n + 6)"
-apply(subgoal_tac "length (mopup_a n) = 4 * n")
-apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps)
-apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps)
-done
-
-lemmas mopupfetchs = 
-fetch_bef_erase_a_o fetch_bef_erase_a_b fetch_bef_erase_b_b 
-fetch_jump_over1_o fetch_jump_over1_b fetch_aft_erase_a_o 
-fetch_aft_erase_a_b fetch_aft_erase_b_b fetch_aft_erase_c_o 
-fetch_aft_erase_c_b fetch_left_moving_o fetch_left_moving_b 
-fetch_jump_over2_b fetch_jump_over2_o
-
-declare 
-  mopup_jump_over2.simps[simp del] mopup_left_moving.simps[simp del]
-  mopup_aft_erase_c.simps[simp del] mopup_aft_erase_b.simps[simp del] 
-  mopup_aft_erase_a.simps[simp del] mopup_jump_over1.simps[simp del]
-  mopup_bef_erase_a.simps[simp del] mopup_bef_erase_b.simps[simp del]
-  mopup_stop.simps[simp del]
-
-lemma [simp]: 
-  "\<lbrakk>mopup_bef_erase_a (s, l, Oc # xs) lm n ires\<rbrakk> \<Longrightarrow> 
-  mopup_bef_erase_b (Suc s, l, Bk # xs) lm n ires"
-apply(auto simp: mopup_bef_erase_a.simps mopup_bef_erase_b.simps )
-apply(rule_tac x = "m - 1" in exI, rule_tac x = rn in exI)
-apply(case_tac m, simp, simp)
-done
-
-lemma mopup_false1:
-  "\<lbrakk>0 < s; s \<le> 2 * n; s mod 2 = Suc 0;  \<not> Suc s \<le> 2 * n\<rbrakk> 
-  \<Longrightarrow> RR"
-apply(arith)
-done
-
-lemma [simp]: 
- "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = Suc 0; 
-   mopup_bef_erase_a (s, l, Oc # xs) lm n ires; r = Oc # xs\<rbrakk>
- \<Longrightarrow> (Suc s \<le> 2 * n \<longrightarrow> mopup_bef_erase_b (Suc s, l, Bk # xs) lm n ires)  \<and>
-     (\<not> Suc s \<le> 2 * n \<longrightarrow> mopup_jump_over1 (Suc s, l, Bk # xs) lm n ires) "
-apply(auto elim: mopup_false1)
-done
-
-lemma drop_tape_of_cons: 
-  "\<lbrakk>Suc q < length lm; x = lm ! q\<rbrakk> \<Longrightarrow> <drop q lm> = Oc # Oc \<up> x @ Bk # <drop (Suc q) lm>"
-by (metis Suc_lessD append_Cons list.simps(2) nth_drop' replicate_Suc tape_of_nl_cons)
-
-lemma erase2jumpover1:
-  "\<lbrakk>q < length list; 
-             \<forall>rn. <drop q list> \<noteq> Oc # Oc \<up> abc_lm_v (a # list) (Suc q) @ Bk # <drop (Suc q) list> @ Bk \<up> rn\<rbrakk>
-       \<Longrightarrow> <drop q list> = Oc # Oc \<up> abc_lm_v (a # list) (Suc q)"
-apply(erule_tac x = 0 in allE, simp)
-apply(case_tac "Suc q < length list")
-apply(erule_tac notE)
-apply(rule_tac drop_tape_of_cons, simp_all add: abc_lm_v.simps)
-apply(subgoal_tac "length list = Suc q", auto)
-apply(subgoal_tac "drop q list = [list ! q]")
-apply(simp add: tape_of_nl_abv tape_of_nat_abv)
-by (metis append_Nil2 append_eq_conv_conj drop_Suc_conv_tl lessI)
-
-lemma erase2jumpover2:
-  "\<lbrakk>q < length list; \<forall>rn. <drop q list> @ Bk # Bk \<up> n \<noteq>
-  Oc # Oc \<up> abc_lm_v (a # list) (Suc q) @ Bk # <drop (Suc q) list> @ Bk \<up> rn\<rbrakk>
-  \<Longrightarrow> RR"
-apply(case_tac "Suc q < length list")
-apply(erule_tac x = "Suc n" in allE, simp)
-apply(erule_tac notE)
-apply(rule_tac drop_tape_of_cons, simp_all add: abc_lm_v.simps)
-apply(subgoal_tac "length list = Suc q", auto)
-apply(erule_tac x = "n" in allE, simp add: tape_of_nl_abv)
-by (metis append_Nil2 append_eq_conv_conj drop_Suc_conv_tl lessI replicate_Suc tape_of_nl_abv tape_of_nl_cons)
-
-lemma mopup_bef_erase_a_2_jump_over[simp]: 
- "\<lbrakk>n < length lm; 0 < s; s mod 2 = Suc 0;  s \<le> 2 * n;
-   mopup_bef_erase_a (s, l, Bk # xs) lm n ires; \<not> (Suc (Suc s) \<le> 2 * n)\<rbrakk> 
-\<Longrightarrow> mopup_jump_over1 (s', Bk # l, xs) lm n ires"
-apply(auto simp: mopup_bef_erase_a.simps mopup_jump_over1.simps)
-apply(case_tac m, auto simp: mod_ex1)
-apply(subgoal_tac "n = Suc q", auto)
-apply(rule_tac x = "Suc ln" in exI, rule_tac x = 0 in exI, auto)
-apply(case_tac [!] lm, simp_all)
-apply(case_tac [!] rn, auto elim: erase2jumpover1 erase2jumpover2)
-apply(erule_tac x = 0 in allE, simp)
-apply(rule_tac classical, simp)
-apply(erule_tac notE)
-apply(rule_tac drop_tape_of_cons, simp_all add: abc_lm_v.simps)
-done
-
-lemma Suc_Suc_div:  "\<lbrakk>0 < s; s mod 2 = Suc 0; Suc (Suc s) \<le> 2 * n\<rbrakk>
-           \<Longrightarrow> (Suc (Suc (s div 2))) \<le> n"
-apply(arith)
-done
-
-lemma mopup_bef_erase_a_2_a[simp]: 
- "\<lbrakk>n < length lm; 0 < s; s mod 2 = Suc 0; 
-   mopup_bef_erase_a (s, l, Bk # xs) lm n ires; 
-   Suc (Suc s) \<le> 2 * n\<rbrakk> \<Longrightarrow> 
-   mopup_bef_erase_a (Suc (Suc s), Bk # l, xs) lm n ires"
-apply(auto simp: mopup_bef_erase_a.simps)
-apply(subgoal_tac "drop (Suc (Suc (s div 2))) lm \<noteq> []")
-apply(case_tac m, simp_all)
-apply(rule_tac x = "Suc (abc_lm_v lm (Suc (s div 2)))" in exI, 
-      rule_tac x = rn in exI, auto simp: mod_ex1)
-apply(rule_tac drop_tape_of_cons)
-apply arith
-apply(simp add: abc_lm_v.simps)
-done
-
-lemma mopup_false2: 
- "\<lbrakk>0 < s; s \<le> 2 * n; 
-   s mod 2 = Suc 0; Suc s \<noteq> 2 * n;
-   \<not> Suc (Suc s) \<le> 2 * n\<rbrakk> \<Longrightarrow> RR"
-apply(arith)
-done
-
-lemma [simp]: "mopup_bef_erase_a (s, l, []) lm n ires \<Longrightarrow> 
-                        mopup_bef_erase_a (s, l, [Bk]) lm n ires"
-apply(auto simp: mopup_bef_erase_a.simps)
-done
-
-lemma [simp]:
-   "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = Suc 0; \<not> Suc (Suc s) \<le> 2 *n;
-     mopup_bef_erase_a (s, l, []) lm n ires\<rbrakk>
-    \<Longrightarrow>  mopup_jump_over1 (s', Bk # l, []) lm n ires"
-by auto
-
-lemma "mopup_bef_erase_b (s, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []"
-apply(auto simp: mopup_bef_erase_b.simps)
-done
-
-lemma [simp]: "mopup_bef_erase_b (s, l, Oc # xs) lm n ires = False"
-apply(auto simp: mopup_bef_erase_b.simps )
-done
- 
-lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow> 
-                                      (s - Suc 0) mod 2 = Suc 0"
-apply(arith)
-done
-
-lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow>
-                                       s - Suc 0 \<le> 2 * n"
-apply(simp)
-done
-
-lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow> \<not> s \<le> Suc 0"
-apply(arith)
-done
-
-lemma [simp]: "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; 
-               s mod 2 \<noteq> Suc 0; 
-               mopup_bef_erase_b (s, l, Bk # xs) lm n ires; r = Bk # xs\<rbrakk> 
-           \<Longrightarrow> mopup_bef_erase_a (s - Suc 0, Bk # l, xs) lm n ires"
-apply(auto simp: mopup_bef_erase_b.simps mopup_bef_erase_a.simps)
-done
-
-lemma [simp]: "\<lbrakk>mopup_bef_erase_b (s, l, []) lm n ires\<rbrakk> \<Longrightarrow> 
-                   mopup_bef_erase_a (s - Suc 0, Bk # l, []) lm n ires"
-apply(auto simp: mopup_bef_erase_b.simps mopup_bef_erase_a.simps)
-done
-
-lemma [simp]: 
-   "\<lbrakk>n < length lm;
-    mopup_jump_over1 (Suc (2 * n), l, Oc # xs) lm n ires;
-    r = Oc # xs\<rbrakk>
-  \<Longrightarrow> mopup_jump_over1 (Suc (2 * n), Oc # l, xs) lm n ires"
-apply(auto simp: mopup_jump_over1.simps)
-apply(rule_tac x = ln in exI, rule_tac x = "Suc m1" in exI,
-       rule_tac x = "m2 - 1" in exI, simp)
-apply(case_tac "m2", simp, simp)
-apply(rule_tac x = ln in exI, rule_tac x = "Suc m1" in exI, 
-      rule_tac x = "m2 - 1" in exI)
-apply(case_tac m2, simp, simp)
-done
-
-lemma mopup_jump_over1_2_aft_erase_a[simp]:  
- "\<lbrakk>n < length lm; mopup_jump_over1 (Suc (2 * n), l, Bk # xs) lm n ires\<rbrakk>
-  \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, xs) lm n ires"
-apply(simp only: mopup_jump_over1.simps mopup_aft_erase_a.simps)
-apply(erule_tac exE)+
-apply(rule_tac x = ln in exI, rule_tac x = "Suc 0" in exI)
-apply(case_tac m2, simp)
-apply(rule_tac x = rn in exI, rule_tac x = "drop (Suc n) lm" in exI, 
-      simp)
-apply(simp)
-done
-
-lemma [simp]: 
- "\<lbrakk>n < length lm; mopup_jump_over1 (Suc (2 * n), l, []) lm n ires\<rbrakk> \<Longrightarrow> 
-    mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, []) lm n ires"
-apply(rule mopup_jump_over1_2_aft_erase_a, simp)
-apply(auto simp: mopup_jump_over1.simps)
-apply(rule_tac x = ln in exI, rule_tac x = "Suc (abc_lm_v lm n)" in exI, 
-      rule_tac x = 0 in exI, simp add: )
-done
-
-
-lemma [simp]: 
- "\<lbrakk>n < length lm; 
-   mopup_aft_erase_a (Suc (Suc (2 * n)), l, Oc # xs) lm n ires\<rbrakk> 
- \<Longrightarrow> mopup_aft_erase_b (Suc (Suc (Suc (2 * n))), l, Bk # xs) lm n ires"
-apply(auto simp: mopup_aft_erase_a.simps mopup_aft_erase_b.simps )
-apply(case_tac ml)
-apply(simp_all add: tape_of_nl_cons split: if_splits)
-apply(case_tac a, simp_all)
-apply(rule_tac x = rn in exI, rule_tac x = "[]" in exI, simp)
-apply(rule_tac x = rn in exI, rule_tac x = "[nat]" in exI, simp)
-apply(case_tac a, simp_all)
-apply(rule_tac x = rn in exI, rule_tac x = "list" in exI, simp)
-apply(rule_tac x = rn in exI)
-apply(rule_tac x = "nat # list" in exI, simp add: tape_of_nl_cons)
-done
-
-lemma [simp]:
-  "mopup_aft_erase_a (Suc (Suc (2 * n)), l, Bk # xs) lm n ires \<Longrightarrow> l \<noteq> []"
-apply(auto simp: mopup_aft_erase_a.simps)
-done
-
-lemma [simp]:
-  "\<lbrakk>n < length lm;
-    mopup_aft_erase_a (Suc (Suc (2 * n)), l, Bk # xs) lm n ires\<rbrakk>
-  \<Longrightarrow> mopup_left_moving (5 + 2 * n, tl l, hd l # Bk # xs) lm n ires"
-apply(simp only: mopup_aft_erase_a.simps mopup_left_moving.simps)
-apply(erule exE)+
-apply(case_tac lnr, simp)
-apply(case_tac ml, simp, simp add: tape_of_nl_cons split: if_splits)
-apply(auto)
-apply(case_tac ml, simp_all add: tape_of_nl_cons split: if_splits)
-apply(rule_tac x = "Suc rn" in exI, simp)
-done
-
-lemma [simp]:
-  "mopup_aft_erase_a (Suc (Suc (2 * n)), l, []) lm n ires \<Longrightarrow> l \<noteq> []"
-apply(simp only: mopup_aft_erase_a.simps)
-apply(erule exE)+
-apply(auto)
-done
-
-lemma [simp]:
-  "\<lbrakk>n < length lm; mopup_aft_erase_a (Suc (Suc (2 * n)), l, []) lm n ires\<rbrakk>
-  \<Longrightarrow> mopup_left_moving (5 + 2 * n, tl l, [hd l]) lm n ires"
-apply(simp only: mopup_aft_erase_a.simps mopup_left_moving.simps)
-apply(erule exE)+
-apply(subgoal_tac "ml = [] \<and> rn = 0", erule conjE, erule conjE, simp)
-apply(case_tac lnr, simp)
-apply(rule_tac x = lnl in exI, simp)
-apply(rule_tac x = 1 in exI, simp)
-apply(case_tac ml, simp, simp)
-done
-
-
-lemma [simp]: "mopup_aft_erase_b (2 * n + 3, l, Oc # xs) lm n ires = False"
-apply(auto simp: mopup_aft_erase_b.simps )
-done
-
-lemma tape_of_ex1[intro]: 
-  "\<exists>rna ml. Oc \<up> a @ Bk \<up> rn = <ml::nat list> @ Bk \<up> rna \<or> Oc \<up> a @ Bk \<up> rn = Bk # <ml> @ Bk \<up> rna"
-apply(case_tac a, simp_all)
-apply(rule_tac x = rn in exI, rule_tac x = "[]" in exI, simp)
-apply(rule_tac x = rn in exI, rule_tac x = "[nat]" in exI, simp)
-done
-
-lemma [intro]: "\<exists>rna ml. Oc \<up> a @ Bk # <list::nat list> @ Bk \<up> rn = 
-  <ml> @ Bk \<up> rna \<or> Oc \<up> a @ Bk # <list> @ Bk \<up> rn = Bk # <ml::nat list> @ Bk \<up> rna"
-apply(case_tac "list = []", simp add: replicate_Suc[THEN sym] del: replicate_Suc)
-apply(rule_tac rn = "Suc rn" in tape_of_ex1)
-apply(case_tac a, simp)
-apply(rule_tac x = rn in exI, rule_tac x = list in exI, simp)
-apply(rule_tac x = rn in exI, rule_tac x = "nat # list" in exI)
-apply(simp add: tape_of_nl_cons)
-done
-
-lemma [simp]: 
- "\<lbrakk>n < length lm; 
-   mopup_aft_erase_c (2 * n + 4, l, Oc # xs) lm n ires\<rbrakk>
-  \<Longrightarrow> mopup_aft_erase_b (Suc (Suc (Suc (2 * n))), l, Bk # xs) lm n ires"
-apply(auto simp: mopup_aft_erase_c.simps mopup_aft_erase_b.simps )
-apply(case_tac ml, simp_all add: tape_of_nl_cons split: if_splits, auto)
-done
-
-lemma mopup_aft_erase_c_aft_erase_a[simp]: 
- "\<lbrakk>n < length lm; mopup_aft_erase_c (2 * n + 4, l, Bk # xs) lm n ires\<rbrakk> 
- \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, xs) lm n ires"
-apply(simp only: mopup_aft_erase_c.simps mopup_aft_erase_a.simps )
-apply(erule_tac exE)+
-apply(erule conjE, erule conjE, erule disjE)
-apply(subgoal_tac "ml = []", simp, case_tac rn, 
-      simp, simp, rule conjI)
-apply(rule_tac x = lnl in exI, rule_tac x = "Suc lnr" in exI, simp)
-apply(rule_tac x = nat in exI, rule_tac x = "[]" in exI, simp)
-apply(case_tac ml, simp, simp add: tape_of_nl_cons split: if_splits)
-apply(rule_tac x = lnl in exI, rule_tac x = "Suc lnr" in exI, simp)
-apply(rule_tac x = rn in exI, rule_tac x = "ml" in exI, simp)
-done
-
-lemma [simp]: 
- "\<lbrakk>n < length lm; mopup_aft_erase_c (2 * n + 4, l, []) lm n ires\<rbrakk> 
- \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, []) lm n ires"
-apply(rule mopup_aft_erase_c_aft_erase_a, simp)
-apply(simp only: mopup_aft_erase_c.simps)
-apply(erule exE)+
-apply(rule_tac x = lnl in exI, rule_tac x = lnr in exI, simp add: )
-apply(rule_tac x = 0 in exI, rule_tac x = "[]" in exI, simp)
-done
-
-lemma mopup_aft_erase_b_2_aft_erase_c[simp]:
-  "\<lbrakk>n < length lm; mopup_aft_erase_b (2 * n + 3, l, Bk # xs) lm n ires\<rbrakk>  
- \<Longrightarrow> mopup_aft_erase_c (4 + 2 * n, Bk # l, xs) lm n ires"
-apply(auto simp: mopup_aft_erase_b.simps mopup_aft_erase_c.simps)
-apply(rule_tac x = "lnl" in exI, rule_tac x = "Suc lnr" in exI, simp)
-apply(rule_tac x = "lnl" in exI, rule_tac x = "Suc lnr" in exI, simp)
-done
-
-lemma [simp]: 
- "\<lbrakk>n < length lm; mopup_aft_erase_b (2 * n + 3, l, []) lm n ires\<rbrakk> 
- \<Longrightarrow> mopup_aft_erase_c (4 + 2 * n, Bk # l, []) lm n ires"
-apply(rule_tac mopup_aft_erase_b_2_aft_erase_c, simp)
-apply(simp add: mopup_aft_erase_b.simps)
-done
-
-lemma [simp]: 
-    "mopup_left_moving (2 * n + 5, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []"
-apply(auto simp: mopup_left_moving.simps)
-done
-
-lemma [simp]:  
- "\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, Oc # xs) lm n ires\<rbrakk>
-  \<Longrightarrow> mopup_jump_over2 (2 * n + 6, tl l, hd l # Oc # xs) lm n ires"
-apply(simp only: mopup_left_moving.simps mopup_jump_over2.simps)
-apply(erule_tac exE)+
-apply(erule conjE, erule disjE, erule conjE)
-apply(case_tac rn, simp, simp add: )
-apply(case_tac "hd l", simp add:  )
-apply(case_tac "abc_lm_v lm n", simp)
-apply(rule_tac x = "lnl" in exI, rule_tac x = rn in exI, 
-      rule_tac x = "Suc 0" in exI, rule_tac x = 0 in exI)
-apply(case_tac lnl, simp, simp, simp add: exp_ind[THEN sym], simp)
-apply(case_tac "abc_lm_v lm n", simp)
-apply(case_tac lnl, simp, simp)
-apply(rule_tac x = lnl in exI, rule_tac x = rn in exI)
-apply(rule_tac x = nat in exI, rule_tac x = "Suc (Suc 0)" in exI, simp)
-done
-
-lemma [simp]: "mopup_left_moving (2 * n + 5, l, xs) lm n ires \<Longrightarrow> l \<noteq> []"
-apply(auto simp: mopup_left_moving.simps)
-done
-
-lemma [simp]:
-  "\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, Bk # xs) lm n ires\<rbrakk> 
- \<Longrightarrow> mopup_left_moving (2 * n + 5, tl l, hd l # Bk # xs) lm n ires"
-apply(simp only: mopup_left_moving.simps)
-apply(erule exE)+
-apply(case_tac lnr, simp)
-apply(rule_tac x = lnl in exI, rule_tac x = nat in exI, simp)
-apply(rule_tac x = "Suc rn" in exI, simp)
-done
-
-lemma [simp]: 
-"\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, []) lm n ires\<rbrakk>
-    \<Longrightarrow> mopup_left_moving (2 * n + 5, tl l, [hd l]) lm n ires"
-apply(simp only: mopup_left_moving.simps)
-apply(erule exE)+
-apply(case_tac lnr, auto)
-done
-
-
-lemma [simp]: 
- "mopup_jump_over2 (2 * n + 6, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []"
-apply(auto simp: mopup_jump_over2.simps )
-done
-
-lemma [simp]: 
-"\<lbrakk>n < length lm; mopup_jump_over2 (2 * n + 6, l, Oc # xs) lm n ires\<rbrakk>
- \<Longrightarrow>  mopup_jump_over2 (2 * n + 6, tl l, hd l # Oc # xs) lm n ires"
-apply(simp only: mopup_jump_over2.simps)
-apply(erule_tac exE)+
-apply(simp add:  , erule conjE, erule_tac conjE)
-apply(case_tac m1, simp)
-apply(rule_tac x = ln in exI, rule_tac x = rn in exI, 
-      rule_tac x = 0 in exI, simp)
-apply(case_tac ln, simp, simp, simp only: exp_ind[THEN sym], simp)
-apply(rule_tac x = ln in exI, rule_tac x = rn in exI, 
-      rule_tac x = nat in exI, rule_tac x = "Suc m2" in exI, simp)
-done
-
-lemma [simp]: 
- "\<lbrakk>n < length lm; mopup_jump_over2 (2 * n + 6, l, Bk # xs) lm n ires\<rbrakk> 
-  \<Longrightarrow> mopup_stop (0, Bk # l, xs) lm n ires"
-apply(auto simp: mopup_jump_over2.simps mopup_stop.simps)
-apply(simp_all add: tape_of_nat_abv exp_ind[THEN sym])
-done
-
-lemma [simp]: "mopup_jump_over2 (2 * n + 6, l, []) lm n ires = False"
-apply(simp only: mopup_jump_over2.simps, simp)
-done
-
-lemma mopup_inv_step:
-  "\<lbrakk>n < length lm; mopup_inv (s, l, r) lm n ires\<rbrakk>
-  \<Longrightarrow> mopup_inv (step (s, l, r) (mopup_a n @ shift mopup_b (2 * n), 0)) lm n ires"
-apply(case_tac r, case_tac [2] a)
-apply(auto split:if_splits simp add:step.simps)
-apply(simp_all add: mopupfetchs)
-done
-
-declare mopup_inv.simps[simp del]
-lemma mopup_inv_steps: 
-"\<lbrakk>n < length lm; mopup_inv (s, l, r) lm n ires\<rbrakk> \<Longrightarrow> 
-     mopup_inv (steps (s, l, r) (mopup_a n @ shift mopup_b (2 * n), 0)  stp) lm n ires"
-apply(induct_tac stp, simp add: steps.simps)
-apply(simp add: step_red)
-apply(case_tac "steps (s, l, r) 
-                (mopup_a n @ shift mopup_b (2 * n), 0) na", simp)
-apply(rule_tac mopup_inv_step, simp, simp)
-done
-
-fun abc_mopup_stage1 :: "config \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "abc_mopup_stage1 (s, l, r) n = 
-           (if s > 0 \<and> s \<le> 2*n then 6
-            else if s = 2*n + 1 then 4
-            else if s \<ge> 2*n + 2 \<and> s \<le> 2*n + 4 then 3
-            else if s = 2*n + 5 then 2
-            else if s = 2*n + 6 then 1
-            else 0)"
-
-fun abc_mopup_stage2 :: "config \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "abc_mopup_stage2 (s, l, r) n = 
-           (if s > 0 \<and> s \<le> 2*n then length r
-            else if s = 2*n + 1 then length r
-            else if s = 2*n + 5 then length l
-            else if s = 2*n + 6 then length l
-            else if s \<ge> 2*n + 2 \<and> s \<le> 2*n + 4 then length r
-            else 0)"
-
-fun abc_mopup_stage3 :: "config \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "abc_mopup_stage3 (s, l, r) n = 
-          (if s > 0 \<and> s \<le> 2*n then 
-              if hd r = Bk then 0
-              else 1
-           else if s = 2*n + 2 then 1 
-           else if s = 2*n + 3 then 0
-           else if s = 2*n + 4 then 2
-           else 0)"
-
-fun abc_mopup_measure :: "(config \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)"
-  where
-  "abc_mopup_measure (c, n) = 
-    (abc_mopup_stage1 c n, abc_mopup_stage2 c n, 
-                                       abc_mopup_stage3 c n)"
-
-definition abc_mopup_LE ::
-   "(((nat \<times> cell list \<times> cell list) \<times> nat) \<times> 
-    ((nat \<times> cell list \<times> cell list) \<times> nat)) set"
-  where
-  "abc_mopup_LE \<equiv> (inv_image lex_triple abc_mopup_measure)"
-
-lemma wf_abc_mopup_le[intro]: "wf abc_mopup_LE"
-by(auto intro:wf_inv_image simp:abc_mopup_LE_def lex_triple_def lex_pair_def)
-
-lemma [simp]: "mopup_bef_erase_a (a, aa, []) lm n ires = False"
-apply(auto simp: mopup_bef_erase_a.simps)
-done
-
-lemma [simp]: "mopup_bef_erase_b (a, aa, []) lm n ires = False"
-apply(auto simp: mopup_bef_erase_b.simps) 
-done
-
-lemma [simp]: "mopup_aft_erase_b (2 * n + 3, aa, []) lm n ires = False"
-apply(auto simp: mopup_aft_erase_b.simps)
-done
-
-declare mopup_inv.simps[simp del]
-term mopup_inv
-
-lemma [simp]: 
-  "\<lbrakk>0 < q; q \<le> n\<rbrakk> \<Longrightarrow> 
-     (fetch (mopup_a n @ shift mopup_b (2 * n)) (2*q) Bk) = (R, 2*q - 1)"
-apply(case_tac q, simp, simp)
-apply(auto simp: fetch.simps nth_of.simps nth_append)
-apply(subgoal_tac "mopup_a n ! (4 * nat + 2) = 
-                     mopup_a (Suc nat) ! ((4 * nat) + 2)", 
-      simp add: mopup_a.simps nth_append)
-apply(rule mopup_a_nth, auto)
-done
-
-(* FIXME: is also in uncomputable *)
-lemma halt_lemma: 
-  "\<lbrakk>wf LE; \<forall>n. (\<not> P (f n) \<longrightarrow> (f (Suc n), (f n)) \<in> LE)\<rbrakk> \<Longrightarrow> \<exists>n. P (f n)"
-by (metis wf_iff_no_infinite_down_chain)
-
-
-lemma mopup_halt:
-  assumes 
-  less: "n < length lm"
-  and inv: "mopup_inv (Suc 0, l, r) lm n ires"
-  and f: "f = (\<lambda> stp. (steps (Suc 0, l, r) (mopup_a n @ shift mopup_b (2 * n), 0) stp, n))"
-  and P: "P = (\<lambda> (c, n). is_final c)"
-  shows "\<exists> stp. P (f stp)"
-proof(rule_tac LE = abc_mopup_LE in halt_lemma)
-  show "wf abc_mopup_LE" by(auto)
-next
-  show "\<forall>n. \<not> P (f n) \<longrightarrow> (f (Suc n), f n) \<in> abc_mopup_LE"
-  proof(rule_tac allI, rule_tac impI)
-    fix na
-    assume h: "\<not> P (f na)"
-    show "(f (Suc na), f na) \<in> abc_mopup_LE"
-    proof(simp add: f)
-      obtain a b c where g:"steps (Suc 0, l, r) (mopup_a n @ shift mopup_b (2 * n), 0) na = (a, b, c)"
-        apply(case_tac "steps (Suc 0, l, r) (mopup_a n @ shift mopup_b (2 * n), 0) na", auto)
-        done
-      then have "mopup_inv (a, b, c) lm n ires"
-        thm mopup_inv_steps
-        using inv less mopup_inv_steps[of n lm "Suc 0" l r ires na]
-        apply(simp)
-        done
-      moreover have "a > 0"
-        using h g
-        apply(simp add: f P)
-        done
-      ultimately have "((step (a, b, c) (mopup_a n @ shift mopup_b (2 * n), 0), n), (a, b, c), n) \<in> abc_mopup_LE"
-        apply(case_tac c, case_tac [2] aa)
-        apply(auto split:if_splits simp add:step.simps mopup_inv.simps)
-        apply(simp_all add: mopupfetchs abc_mopup_LE_def lex_triple_def lex_pair_def )
-        done
-      thus "((step (steps (Suc 0, l, r) (mopup_a n @ shift mopup_b (2 * n), 0) na) 
-        (mopup_a n @ shift mopup_b (2 * n), 0), n),
-        steps (Suc 0, l, r) (mopup_a n @ shift mopup_b (2 * n), 0) na, n)
-        \<in> abc_mopup_LE"
-        using g by simp
-    qed
-  qed
-qed     
-
-lemma mopup_inv_start: 
-  "n < length am \<Longrightarrow> mopup_inv (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) am n ires"
-apply(auto simp: mopup_inv.simps mopup_bef_erase_a.simps mopup_jump_over1.simps)
-apply(case_tac [!] am, auto split: if_splits simp: tape_of_nl_cons)
-apply(rule_tac x = "Suc a" in exI, rule_tac x = k in exI, simp)
-apply(case_tac [!] n, simp_all add: abc_lm_v.simps)
-apply(case_tac k, simp, simp_all)
-done
-      
-lemma mopup_correct:
-  assumes less: "n < length (am::nat list)"
-  and rs: "abc_lm_v am n = rs"
-  shows "\<exists> stp i j. (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp)
-    = (0, Bk\<up>i @ Bk # Bk # ires, Oc # Oc\<up> rs @ Bk\<up>j)"
-using less
-proof -
-  have a: "mopup_inv (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) am n ires"
-    using less
-    apply(simp add: mopup_inv_start)
-    done    
-  then have "\<exists> stp. is_final (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp)"
-    using less mopup_halt[of n am  "Bk # Bk # ires" "<am> @ Bk \<up> k" ires
-      "(\<lambda>stp. (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp, n))"
-      "(\<lambda>(c, n). is_final c)"]
-    apply(simp)
-    done
-  from this obtain stp where b:
-    "is_final (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp)" ..
-  from a b have
-    "mopup_inv (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp)
-    am n ires"
-    apply(rule_tac mopup_inv_steps, simp_all add: less)
-    done    
-  from b and this show "?thesis"
-    apply(rule_tac x = stp in exI, simp)
-    apply(case_tac "steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) 
-      (mopup_a n @ shift mopup_b (2 * n), 0) stp")
-    apply(simp add: mopup_inv.simps mopup_stop.simps rs)
-    using rs
-    apply(simp add: tape_of_nat_abv)
-    done
-qed
-
-(*we can use Hoare_plus here*)
-
-lemma wf_mopup[intro]: "tm_wf (mopup n, 0)"
-apply(induct n, simp add: mopup.simps shift.simps mopup_b_def tm_wf.simps)
-apply(auto simp: mopup.simps shift.simps mopup_b_def tm_wf.simps)
-done
-
-lemma length_tp:
-  "\<lbrakk>ly = layout_of ap; tp = tm_of ap\<rbrakk> \<Longrightarrow> 
-  start_of ly (length ap) = Suc (length tp div 2)"
-apply(frule_tac length_tp', simp_all)
-apply(simp add: start_of.simps)
-done
-
-lemma compile_correct_halt: 
-  assumes layout: "ly = layout_of ap"
-  and compile: "tp = tm_of ap"
-  and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires"
-  and abc_halt: "abc_steps_l (0, lm) ap stp = (length ap, am)"
-  and rs_loc: "n < length am"
-  and rs: "abc_lm_v am n = rs"
-  and off: "off = length tp div 2"
-  shows "\<exists> stp i j. steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stp = (0, Bk\<up>i @ Bk # Bk # ires, Oc\<up>Suc rs @ Bk\<up>j)"
-proof -
-  have "\<exists> stp k. steps (Suc 0, l, r) (tp, 0) stp = (Suc off, Bk # Bk # ires, <am> @ Bk\<up>k)"
-    using assms tp_correct'[of ly ap tp lm l r ires stp am]
-    by(simp add: length_tp)
-  then obtain stp k where "steps (Suc 0, l, r) (tp, 0) stp = (Suc off, Bk # Bk # ires, <am> @ Bk\<up>k)"
-    by blast
-  then have a: "steps (Suc 0, l, r) (tp@shift (mopup n) off , 0) stp = (Suc off, Bk # Bk # ires, <am> @ Bk\<up>k)"
-    using assms
-    by(auto intro: tm_append_first_steps_eq)
-  have "\<exists> stp i j. (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp)
-    = (0, Bk\<up>i @ Bk # Bk # ires, Oc # Oc\<up> rs @ Bk\<up>j)"
-    using assms
-    by(auto intro: mopup_correct)
-  then obtain stpb i j where 
-    "steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stpb
-    = (0, Bk\<up>i @ Bk # Bk # ires, Oc # Oc\<up> rs @ Bk\<up>j)" by blast
-  then have b: "steps (Suc 0 + off, Bk # Bk # ires, <am> @ Bk \<up> k) (tp @ shift (mopup n) off, 0) stpb
-    = (0, Bk\<up>i @ Bk # Bk # ires, Oc # Oc\<up> rs @ Bk\<up>j)"
-    using assms wf_mopup
-   thm tm_append_second_halt_eq
-    apply(drule_tac tm_append_second_halt_eq, auto)
-    done
-  from a b show "?thesis"
-    by(rule_tac x = "stp + stpb" in exI, simp add: steps_add)
-qed
- 
-declare mopup.simps[simp del]
-lemma abc_step_red2:
-  "abc_steps_l (s, lm) p (Suc n) = (let (as', am') = abc_steps_l (s, lm) p n in
-                                    abc_step_l (as', am') (abc_fetch as' p))"
-apply(case_tac "abc_steps_l (s, lm) p n", simp)
-apply(drule_tac abc_step_red, simp)
-done
-
-lemma crsp_steps2:
-  assumes 
-  layout: "ly = layout_of ap"
-  and compile: "tp = tm_of ap"
-  and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires"
-  and nothalt: "as < length ap"
-  and aexec: "abc_steps_l (0, lm) ap stp = (as, am)"
-  shows "\<exists>stpa\<ge>stp. crsp ly (as, am) (steps (Suc 0, l, r) (tp, 0) stpa) ires"
-using nothalt aexec
-proof(induct stp arbitrary: as am)
-  case 0
-  thus "?case"
-    using crsp
-    by(rule_tac x = 0 in exI, auto simp: abc_steps_l.simps steps.simps crsp)
-next
-  case (Suc stp as am)
-  have ind: 
-    "\<And> as am.  \<lbrakk>as < length ap; abc_steps_l (0, lm) ap stp = (as, am)\<rbrakk> 
-    \<Longrightarrow> \<exists>stpa\<ge>stp. crsp ly (as, am) (steps (Suc 0, l, r) (tp, 0) stpa) ires" by fact
-  have a: "as < length ap" by fact
-  have b: "abc_steps_l (0, lm) ap (Suc stp) = (as, am)" by fact
-  obtain as' am' where c: "abc_steps_l (0, lm) ap stp = (as', am')" 
-    by(case_tac "abc_steps_l (0, lm) ap stp", auto)
-  then have d: "as' < length ap"
-    using a b
-    by(simp add: abc_step_red2, case_tac "as' < length ap", simp,
-      simp add: abc_fetch.simps abc_steps_l.simps abc_step_l.simps)
-  have "\<exists>stpa\<ge>stp. crsp ly (as', am') (steps (Suc 0, l, r) (tp, 0) stpa) ires"
-    using d c ind by simp
-  from this obtain stpa where e: 
-    "stpa \<ge> stp \<and>  crsp ly (as', am') (steps (Suc 0, l, r) (tp, 0) stpa) ires"
-    by blast
-  obtain s' l' r' where f: "steps (Suc 0, l, r) (tp, 0) stpa = (s', l', r')"
-    by(case_tac "steps (Suc 0, l, r) (tp, 0) stpa", auto)
-  obtain ins where g: "abc_fetch as' ap = Some ins" using d 
-    by(case_tac "abc_fetch as' ap",auto simp: abc_fetch.simps)
-  then have "\<exists>stp> (0::nat). crsp ly (abc_step_l (as', am') (Some ins)) 
-    (steps (s', l', r') (tp, 0) stp) ires "
-    using layout compile e f 
-    by(rule_tac crsp_step, simp_all)
-  then obtain stpb where "stpb > 0 \<and> crsp ly (abc_step_l (as', am') (Some ins)) 
-    (steps (s', l', r') (tp, 0) stpb) ires" ..
-  from this show "?case" using b e g f c
-    by(rule_tac x = "stpa + stpb" in exI, simp add: steps_add abc_step_red2)
-qed
-    
-lemma compile_correct_unhalt: 
-  assumes layout: "ly = layout_of ap"
-  and compile: "tp = tm_of ap"
-  and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires"
-  and off: "off = length tp div 2"
-  and abc_unhalt: "\<forall> stp. (\<lambda> (as, am). as < length ap) (abc_steps_l (0, lm) ap stp)"
-  shows "\<forall> stp.\<not> is_final (steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stp)"
-using assms
-proof(rule_tac allI, rule_tac notI)
-  fix stp
-  assume h: "is_final (steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stp)"
-  obtain as am where a: "abc_steps_l (0, lm) ap stp = (as, am)"
-    by(case_tac "abc_steps_l (0, lm) ap stp", auto)
-  then have b: "as < length ap"
-    using abc_unhalt
-    by(erule_tac x = stp in allE, simp)
-  have "\<exists> stpa\<ge>stp. crsp ly (as, am) (steps (Suc 0, l, r) (tp, 0) stpa) ires "
-    using assms b a
-    apply(rule_tac crsp_steps2, simp_all)
-    done
-  then obtain stpa where
-    "stpa\<ge>stp \<and> crsp ly (as, am) (steps (Suc 0, l, r) (tp, 0) stpa) ires" ..
-  then obtain s' l' r' where b: "(steps (Suc 0, l, r) (tp, 0) stpa) = (s', l', r') \<and> 
-       stpa\<ge>stp \<and> crsp ly (as, am) (s', l', r') ires"
-    by(case_tac "steps (Suc 0, l, r) (tp, 0) stpa", auto)
-  hence c:
-    "(steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stpa) = (s', l', r')"
-    by(rule_tac tm_append_first_steps_eq, simp_all add: crsp.simps)
-  from b have d: "s' > 0 \<and> stpa \<ge> stp"
-    by(simp add: crsp.simps)
-  then obtain diff where e: "stpa = stp + diff"   by (metis le_iff_add)
-  obtain s'' l'' r'' where f:
-    "steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stp = (s'', l'', r'') \<and> is_final (s'', l'', r'')"
-    using h
-    by(case_tac "steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stp", auto)
-
-  then have "is_final (steps (s'', l'', r'') (tp @ shift (mopup n) off, 0) diff)"
-    by(auto intro: after_is_final)
-  then have "is_final (steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stpa)"
-    using e
-    by(simp add: steps_add f)
-  from this and c d show "False" by simp
-qed
-
-end
-
--- a/thys/rec_def.thy	Thu Feb 07 06:39:06 2013 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,87 +0,0 @@
-theory rec_def
-imports Main
-begin
-
-section {*
-  Recursive functions
-*}
-
-text {*
-  Datatype of recursive operators.
-*}
-
-datatype recf = 
- -- {* The zero function, which always resturns @{text "0"} as result. *}
-  z | 
- -- {* The successor function, which increments its arguments. *}
-  s | 
- -- {*
-  The projection function, where @{text "id i j"} returns the @{text "j"}-th
-  argment out of the @{text "i"} arguments.
-  *}
-  id nat nat | 
- -- {*
-  The compostion operator, where "@{text "Cn n f [g1; g2; \<dots> ;gm]"} 
-  computes @{text "f (g1(x1, x2, \<dots>, xn), g2(x1, x2, \<dots>, xn), \<dots> , 
-  gm(x1, x2, \<dots> , xn))"} for input argments @{text "x1, \<dots>, xn"}.
-  *}
-  Cn nat recf "recf list" | 
--- {*
-  The primitive resursive operator, where @{text "Pr n f g"} computes:
-  @{text "Pr n f g (x1, x2, \<dots>, xn-1, 0) = f(x1, \<dots>, xn-1)"} 
-  and @{text "Pr n f g (x1, x2, \<dots>, xn-1, k') = g(x1, x2, \<dots>, xn-1, k, 
-                                                  Pr n f g (x1, \<dots>, xn-1, k))"}.
-  *}
-  Pr nat recf recf | 
--- {*
-  The minimization operator, where @{text "Mn n f (x1, x2, \<dots> , xn)"} 
-  computes the first i such that @{text "f (x1, \<dots>, xn, i) = 0"} and for all
-  @{text "j"}, @{text "f (x1, x2, \<dots>, xn, j) > 0"}.
-  *}
-  Mn nat recf 
-
-text {* 
-  The semantis of recursive operators is given by an inductively defined
-  relation as follows, where  
-  @{text "rec_calc_rel R [x1, x2, \<dots>, xn] r"} means the computation of 
-  @{text "R"} over input arguments @{text "[x1, x2, \<dots>, xn"} terminates
-  and gives rise to a result @{text "r"}
-*}
-
-inductive rec_calc_rel :: "recf \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> bool"
-where
-  calc_z: "rec_calc_rel z [n] 0" |
-  calc_s: "rec_calc_rel s [n] (Suc n)" |
-  calc_id: "\<lbrakk>length args = i; j < i; args!j = r\<rbrakk> \<Longrightarrow> rec_calc_rel (id i j) args r" |
-  calc_cn: "\<lbrakk>length args = n;
-             \<forall> k < length gs. rec_calc_rel (gs ! k) args (rs ! k);
-             length rs = length gs; 
-             rec_calc_rel f rs r\<rbrakk> 
-            \<Longrightarrow> rec_calc_rel (Cn n f gs) args r" |
-  calc_pr_zero: 
-           "\<lbrakk>length args = n;
-             rec_calc_rel f args r0 \<rbrakk> 
-            \<Longrightarrow> rec_calc_rel (Pr n f g) (args @ [0]) r0" |
-  calc_pr_ind: "
-           \<lbrakk> length args = n;
-             rec_calc_rel (Pr n f g) (args @ [k]) rk; 
-             rec_calc_rel g (args @ [k] @ [rk]) rk'\<rbrakk>
-            \<Longrightarrow> rec_calc_rel (Pr n f g) (args @ [Suc k]) rk'"  |
-  calc_mn: "\<lbrakk>length args = n; 
-             rec_calc_rel f (args@[r]) 0; 
-             \<forall> i < r. (\<exists> ri. rec_calc_rel f (args@[i]) ri \<and> ri \<noteq> 0)\<rbrakk> 
-            \<Longrightarrow> rec_calc_rel (Mn n f) args r" 
-
-inductive_cases calc_pr_reverse:
-              "rec_calc_rel (Pr n f g) (lm) rSucy"
-
-inductive_cases calc_z_reverse: "rec_calc_rel z lm x"
-
-inductive_cases calc_s_reverse: "rec_calc_rel s lm x"
-
-inductive_cases calc_id_reverse: "rec_calc_rel (id m n) lm x"
-
-inductive_cases calc_cn_reverse: "rec_calc_rel (Cn n f gs) lm x"
-
-inductive_cases calc_mn_reverse:"rec_calc_rel (Mn n f) lm x"
-end
\ No newline at end of file
--- a/thys/recursive.thy	Thu Feb 07 06:39:06 2013 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,5158 +0,0 @@
-theory recursive
-imports Main rec_def abacus
-begin
-
-section {* 
-  Compiling from recursive functions to Abacus machines
-  *}
-
-text {*
-  Some auxilliary Abacus machines used to construct the result Abacus machines.
-*}
-
-text {*
-  @{text "get_paras_num recf"} returns the arity of recursive function @{text "recf"}.
-*}
-fun get_paras_num :: "recf \<Rightarrow> nat"
-  where
-  "get_paras_num z = 1" |
-  "get_paras_num s = 1" |
-  "get_paras_num (id m n) = m" |
-  "get_paras_num (Cn n f gs) = n" |
-  "get_paras_num (Pr n f g) = Suc n"  |
-  "get_paras_num (Mn n f) = n"  
-
-fun addition :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
-  where
-  "addition m n p = [Dec m 4, Inc n, Inc p, Goto 0, Dec p 7, 
-                       Inc m, Goto 4]"
-
-fun mv_box :: "nat \<Rightarrow> nat \<Rightarrow> abc_prog"
-  where
-  "mv_box m n = [Dec m 3, Inc n, Goto 0]"
-
-fun abc_inst_shift :: "abc_inst \<Rightarrow> nat \<Rightarrow> abc_inst"
-  where
-  "abc_inst_shift (Inc m) n = Inc m" |
-  "abc_inst_shift (Dec m e) n = Dec m (e + n)" |
-  "abc_inst_shift (Goto m) n = Goto (m + n)"
-
-fun abc_shift :: "abc_inst list \<Rightarrow> nat \<Rightarrow> abc_inst list" 
-  where
-  "abc_shift xs n = map (\<lambda> x. abc_inst_shift x n) xs" 
-
-fun abc_append :: "abc_inst list \<Rightarrow> abc_inst list \<Rightarrow> 
-                           abc_inst list" (infixl "[+]" 60)
-  where
-  "abc_append al bl = (let al_len = length al in 
-                           al @ abc_shift bl al_len)"
-
-text {*
-  The compilation of @{text "z"}-operator.
-*}
-definition rec_ci_z :: "abc_inst list"
-  where
-  "rec_ci_z \<equiv> [Goto 1]"
-
-text {*
-  The compilation of @{text "s"}-operator.
-*}
-definition rec_ci_s :: "abc_inst list"
-  where
-  "rec_ci_s \<equiv> (addition 0 1 2 [+] [Inc 1])"
-
-
-text {*
-  The compilation of @{text "id i j"}-operator
-*}
-
-fun rec_ci_id :: "nat \<Rightarrow> nat \<Rightarrow> abc_inst list"
-  where
-  "rec_ci_id i j = addition j i (i + 1)"
-
-fun mv_boxes :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_inst list"
-  where
-  "mv_boxes ab bb 0 = []" |
-  "mv_boxes ab bb (Suc n) = mv_boxes ab bb n [+] mv_box (ab + n)
-  (bb + n)"
-
-fun empty_boxes :: "nat \<Rightarrow> abc_inst list"
-  where
-  "empty_boxes 0 = []" |
-  "empty_boxes (Suc n) = empty_boxes n [+] [Dec n 2, Goto 0]"
-
-fun cn_merge_gs ::
-  "(abc_inst list \<times> nat \<times> nat) list \<Rightarrow> nat \<Rightarrow> abc_inst list"
-  where
-  "cn_merge_gs [] p = []" |
-  "cn_merge_gs (g # gs) p = 
-      (let (gprog, gpara, gn) = g in 
-         gprog [+] mv_box gpara p [+] cn_merge_gs gs (Suc p))"
-
-
-text {*
-  The compiler of recursive functions, where @{text "rec_ci recf"} return 
-  @{text "(ap, arity, fp)"}, where @{text "ap"} is the Abacus program, @{text "arity"} is the 
-  arity of the recursive function @{text "recf"}, 
-@{text "fp"} is the amount of memory which is going to be
-  used by @{text "ap"} for its execution. 
-*}
-
-function rec_ci :: "recf \<Rightarrow> abc_inst list \<times> nat \<times> nat"
-  where
-  "rec_ci z = (rec_ci_z, 1, 2)" |
-  "rec_ci s = (rec_ci_s, 1, 3)" |
-  "rec_ci (id m n) = (rec_ci_id m n, m, m + 2)" |
-  "rec_ci (Cn n f gs) = 
-      (let cied_gs = map (\<lambda> g. rec_ci g) (f # gs) in
-       let (fprog, fpara, fn) = hd cied_gs in 
-       let pstr = 
-        Max (set (Suc n # fn # (map (\<lambda> (aprog, p, n). n) cied_gs))) in
-       let qstr = pstr + Suc (length gs) in 
-       (cn_merge_gs (tl cied_gs) pstr [+] mv_boxes 0 qstr n [+] 
-          mv_boxes pstr 0 (length gs) [+] fprog [+] 
-            mv_box fpara pstr [+] empty_boxes (length gs) [+] 
-             mv_box pstr n [+] mv_boxes qstr 0 n, n,  qstr + n))" |
-  "rec_ci (Pr n f g) = 
-         (let (fprog, fpara, fn) = rec_ci f in 
-          let (gprog, gpara, gn) = rec_ci g in 
-          let p = Max (set ([n + 3, fn, gn])) in 
-          let e = length gprog + 7 in 
-           (mv_box n p [+] fprog [+] mv_box n (Suc n) [+] 
-               (([Dec p e] [+] gprog [+] 
-                 [Inc n, Dec (Suc n) 3, Goto 1]) @
-                     [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gprog + 4)]),
-             Suc n, p + 1))" |
-  "rec_ci (Mn n f) =
-         (let (fprog, fpara, fn) = rec_ci f in 
-          let len = length (fprog) in 
-            (fprog @ [Dec (Suc n) (len + 5), Dec (Suc n) (len + 3),
-             Goto (len + 1), Inc n, Goto 0], n, max (Suc n) fn) )"
-  by pat_completeness auto
-termination 
-proof
-term size
-  show "wf (measure size)" by auto
-next
-  fix n f gs x
-  assume "(x::recf) \<in> set (f # gs)" 
-  thus "(x, Cn n f gs) \<in> measure size"
-    by(induct gs, auto)
-next
-  fix n f g
-  show "(f, Pr n f g) \<in> measure size" by auto
-next
-  fix n f g x xa y xb ya
-  show "(g, Pr n f g) \<in> measure size" by auto
-next
-  fix n f
-  show "(f, Mn n f) \<in> measure size" by auto
-qed
-
-declare rec_ci.simps [simp del] rec_ci_s_def[simp del] 
-        rec_ci_z_def[simp del] rec_ci_id.simps[simp del]
-        mv_boxes.simps[simp del] abc_append.simps[simp del]
-        mv_box.simps[simp del] addition.simps[simp del]
-  
-thm rec_calc_rel.induct
-
-declare abc_steps_l.simps[simp del] abc_fetch.simps[simp del] 
-        abc_step_l.simps[simp del] 
-
-lemma abc_steps_add: 
-  "abc_steps_l (as, lm) ap (m + n) = 
-         abc_steps_l (abc_steps_l (as, lm) ap m) ap n"
-apply(induct m arbitrary: n as lm, simp add: abc_steps_l.simps)
-proof -
-  fix m n as lm
-  assume ind: 
-    "\<And>n as lm. abc_steps_l (as, lm) ap (m + n) = 
-                   abc_steps_l (abc_steps_l (as, lm) ap m) ap n"
-  show "abc_steps_l (as, lm) ap (Suc m + n) = 
-             abc_steps_l (abc_steps_l (as, lm) ap (Suc m)) ap n"
-    apply(insert ind[of as lm "Suc n"], simp)
-    apply(insert ind[of as lm "Suc 0"], simp add: abc_steps_l.simps)
-    apply(case_tac "(abc_steps_l (as, lm) ap m)", simp)
-    apply(simp add: abc_steps_l.simps)
-    apply(case_tac "abc_step_l (a, b) (abc_fetch a ap)", 
-          simp add: abc_steps_l.simps)
-    done
-qed
-
-(*lemmas: rec_ci and rec_calc_rel*)
-
-lemma rec_calc_inj_case_z: 
-  "\<lbrakk>rec_calc_rel z l x; rec_calc_rel z l y\<rbrakk> \<Longrightarrow> x = y"
-apply(auto elim: calc_z_reverse)
-done
-
-lemma  rec_calc_inj_case_s: 
-  "\<lbrakk>rec_calc_rel s l x; rec_calc_rel s l y\<rbrakk> \<Longrightarrow> x = y"
-apply(auto elim: calc_s_reverse)
-done
-
-lemma rec_calc_inj_case_id:
-  "\<lbrakk>rec_calc_rel (recf.id nat1 nat2) l x;
-    rec_calc_rel (recf.id nat1 nat2) l y\<rbrakk> \<Longrightarrow> x = y"
-apply(auto elim: calc_id_reverse)
-done
-
-lemma rec_calc_inj_case_mn:
-  assumes ind: "\<And> l x y. \<lbrakk>rec_calc_rel f l x; rec_calc_rel f l y\<rbrakk> 
-           \<Longrightarrow> x = y" 
-  and h: "rec_calc_rel (Mn n f) l x" "rec_calc_rel (Mn n f) l y"
-  shows "x = y"
-  apply(insert h)
-  apply(elim  calc_mn_reverse)
-  apply(case_tac "x > y", simp)
-  apply(erule_tac x = "y" in allE, auto)
-proof -
-  fix v va
-  assume "rec_calc_rel f (l @ [y]) 0" 
-    "rec_calc_rel f (l @ [y]) v"  
-    "0 < v"
-  thus "False"
-    apply(insert ind[of "l @ [y]" 0 v], simp)
-    done
-next
-  fix v va
-  assume 
-    "rec_calc_rel f (l @ [x]) 0" 
-    "\<forall>x<y. \<exists>v. rec_calc_rel f (l @ [x]) v \<and> 0 < v" "\<not> y < x"
-  thus "x = y"
-    apply(erule_tac x = "x" in allE)
-    apply(case_tac "x = y", auto)
-    apply(drule_tac y = v in ind, simp, simp)
-    done
-qed 
-
-lemma rec_calc_inj_case_pr: 
-  assumes f_ind: 
-  "\<And>l x y. \<lbrakk>rec_calc_rel f l x; rec_calc_rel f l y\<rbrakk> \<Longrightarrow> x = y"
-  and g_ind:
-  "\<And>x xa y xb ya l xc yb. 
-  \<lbrakk>x = rec_ci f; (xa, y) = x; (xb, ya) = y; 
-  rec_calc_rel g l xc; rec_calc_rel g l yb\<rbrakk> \<Longrightarrow> xc = yb"
-  and h: "rec_calc_rel (Pr n f g) l x" "rec_calc_rel (Pr n f g) l y"  
-  shows "x = y"
-  apply(case_tac "rec_ci f")
-proof -
-  fix a b c
-  assume "rec_ci f = (a, b, c)"
-  hence ng_ind: 
-    "\<And> l xc yb. \<lbrakk>rec_calc_rel g l xc; rec_calc_rel g l yb\<rbrakk>
-    \<Longrightarrow> xc = yb"
-    apply(insert g_ind[of "(a, b, c)" "a" "(b, c)" b c], simp)
-    done
-  from h show "x = y"
-    apply(erule_tac calc_pr_reverse, erule_tac calc_pr_reverse)
-    apply(erule f_ind, simp, simp)
-    apply(erule_tac calc_pr_reverse, simp, simp)
-  proof -
-    fix la ya ry laa yaa rya
-    assume k1:  "rec_calc_rel g (la @ [ya, ry]) x" 
-      "rec_calc_rel g (la @ [ya, rya]) y"
-      and k2: "rec_calc_rel (Pr (length la) f g) (la @ [ya]) ry"
-              "rec_calc_rel (Pr (length la) f g) (la @ [ya]) rya"
-    from k2 have "ry = rya"
-      apply(induct ya arbitrary: ry rya)
-      apply(erule_tac calc_pr_reverse, 
-        erule_tac calc_pr_reverse, simp)
-      apply(erule f_ind, simp, simp, simp)
-      apply(erule_tac calc_pr_reverse, simp)
-      apply(erule_tac rSucy = rya in calc_pr_reverse, simp, simp)
-    proof -
-      fix ya ry rya l y ryb laa yb ryc
-      assume ind:
-        "\<And>ry rya. \<lbrakk>rec_calc_rel (Pr (length l) f g) (l @ [y]) ry; 
-                   rec_calc_rel (Pr (length l) f g) (l @ [y]) rya\<rbrakk> \<Longrightarrow> ry = rya"
-        and j: "rec_calc_rel (Pr (length l) f g) (l @ [y]) ryb"
-        "rec_calc_rel g (l @ [y, ryb]) ry" 
-        "rec_calc_rel (Pr (length l) f g) (l @ [y]) ryc" 
-        "rec_calc_rel g (l @ [y, ryc]) rya"
-      from j show "ry = rya"
-	apply(insert ind[of ryb ryc], simp)
-	apply(insert ng_ind[of "l @ [y, ryc]" ry rya], simp)
-	done
-    qed 
-    from k1 and this show "x = y"
-      apply(simp)
-      apply(insert ng_ind[of "la @ [ya, rya]" x y], simp)
-      done
-  qed  
-qed
-
-lemma Suc_nth_part_eq:
-  "\<forall>k<Suc (length list). (a # xs) ! k = (aa # list) ! k
-       \<Longrightarrow> \<forall>k<(length list). (xs) ! k = (list) ! k"
-apply(rule allI, rule impI)
-apply(erule_tac x = "Suc k" in allE, simp)
-done
-
-
-lemma list_eq_intro:  
-  "\<lbrakk>length xs = length ys; \<forall> k < length xs. xs ! k = ys ! k\<rbrakk> 
-  \<Longrightarrow> xs = ys"
-apply(induct xs arbitrary: ys, simp)
-apply(case_tac ys, simp, simp)
-proof -
-  fix a xs ys aa list
-  assume ind: 
-    "\<And>ys. \<lbrakk>length list = length ys; \<forall>k<length ys. xs ! k = ys ! k\<rbrakk>
-    \<Longrightarrow> xs = ys"
-    and h: "length xs = length list" 
-    "\<forall>k<Suc (length list). (a # xs) ! k = (aa # list) ! k"
-  from h show "a = aa \<and> xs = list"
-    apply(insert ind[of list], simp)
-    apply(frule Suc_nth_part_eq, simp)
-    apply(erule_tac x = "0" in allE, simp)
-    done
-qed
-
-lemma rec_calc_inj_case_cn: 
-  assumes ind: 
-  "\<And>x l xa y.
-  \<lbrakk>x = f \<or> x \<in> set gs; rec_calc_rel x l xa; rec_calc_rel x l y\<rbrakk>
-  \<Longrightarrow> xa = y"
-  and h: "rec_calc_rel (Cn n f gs) l x" 
-         "rec_calc_rel (Cn n f gs) l y"
-  shows "x = y"
-  apply(insert h, elim  calc_cn_reverse)
-  apply(subgoal_tac "rs = rsa")
-  apply(rule_tac x = f and l = rsa and xa = x and y = y in ind, 
-        simp, simp, simp)
-  apply(intro list_eq_intro, simp, rule allI, rule impI)
-  apply(erule_tac x = k in allE, rule_tac x = k in allE, simp, simp)
-  apply(rule_tac x = "gs ! k" in ind, simp, simp, simp)
-  done
-
-lemma rec_calc_inj:
-  "\<lbrakk>rec_calc_rel f l x; 
-    rec_calc_rel f l y\<rbrakk> \<Longrightarrow> x = y"
-apply(induct f arbitrary: l x y rule: rec_ci.induct)
-apply(simp add: rec_calc_inj_case_z)
-apply(simp add: rec_calc_inj_case_s)
-apply(simp add: rec_calc_inj_case_id, simp)
-apply(erule rec_calc_inj_case_cn,simp, simp)
-apply(erule rec_calc_inj_case_pr, auto)
-apply(erule rec_calc_inj_case_mn, auto)
-done
-
-
-lemma calc_rel_reverse_ind_step_ex: 
-  "\<lbrakk>rec_calc_rel (Pr n f g) (lm @ [Suc x]) rs\<rbrakk> 
-  \<Longrightarrow> \<exists> rs. rec_calc_rel (Pr n f g) (lm @ [x]) rs"
-apply(erule calc_pr_reverse, simp, simp)
-apply(rule_tac x = rk in exI, simp)
-done
-
-lemma [simp]: "Suc x \<le> y \<Longrightarrow> Suc (y - Suc x) = y - x"
-by arith
-
-lemma calc_pr_para_not_null: 
-  "rec_calc_rel (Pr n f g) lm rs \<Longrightarrow> lm \<noteq> []"
-apply(erule calc_pr_reverse, simp, simp)
-done
-
-lemma calc_pr_less_ex: 
- "\<lbrakk>rec_calc_rel (Pr n f g) lm rs; x \<le> last lm\<rbrakk> \<Longrightarrow> 
- \<exists>rs. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rs"
-apply(subgoal_tac "lm \<noteq> []")
-apply(induct x, rule_tac x = rs in exI, simp, simp, erule exE)
-apply(rule_tac rs = xa in calc_rel_reverse_ind_step_ex, simp)
-apply(simp add: calc_pr_para_not_null)
-done
-
-lemma calc_pr_zero_ex:
-  "rec_calc_rel (Pr n f g) lm rs \<Longrightarrow> 
-             \<exists>rs. rec_calc_rel f (butlast lm) rs"
-apply(drule_tac x = "last lm" in calc_pr_less_ex, simp,
-      erule_tac exE, simp)
-apply(erule_tac calc_pr_reverse, simp)
-apply(rule_tac x = rs in exI, simp, simp)
-done
-
-
-lemma abc_steps_ind: 
-  "abc_steps_l (as, am) ap (Suc stp) =
-          abc_steps_l (abc_steps_l (as, am) ap stp) ap (Suc 0)"
-apply(insert abc_steps_add[of as am ap stp "Suc 0"], simp)
-done
-
-lemma abc_steps_zero: "abc_steps_l asm ap 0 = asm"
-apply(case_tac asm, simp add: abc_steps_l.simps)
-done
-
-lemma abc_append_nth: 
-  "n < length ap + length bp \<Longrightarrow> 
-       (ap [+] bp) ! n =
-         (if n < length ap then ap ! n 
-          else abc_inst_shift (bp ! (n - length ap)) (length ap))"
-apply(simp add: abc_append.simps nth_append map_nth split: if_splits)
-done
-
-lemma abc_state_keep:  
-  "as \<ge> length bp \<Longrightarrow> abc_steps_l (as, lm) bp stp = (as, lm)"
-apply(induct stp, simp add: abc_steps_zero)
-apply(simp add: abc_steps_ind)
-apply(simp add: abc_steps_zero)
-apply(simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps)
-done
-
-lemma abc_halt_equal: 
-  "\<lbrakk>abc_steps_l (0, lm) bp stpa = (length bp, lm1); 
-    abc_steps_l (0, lm) bp stpb = (length bp, lm2)\<rbrakk> \<Longrightarrow> lm1 = lm2"
-apply(case_tac "stpa - stpb > 0")
-apply(insert abc_steps_add[of 0 lm bp stpb "stpa - stpb"], simp)
-apply(insert abc_state_keep[of bp "length bp" lm2 "stpa - stpb"], 
-      simp, simp add: abc_steps_zero)
-apply(insert abc_steps_add[of 0 lm bp stpa "stpb - stpa"], simp)
-apply(insert abc_state_keep[of bp "length bp" lm1 "stpb - stpa"], 
-      simp)
-done  
-
-lemma abc_halt_point_ex: 
-  "\<lbrakk>\<exists>stp. abc_steps_l (0, lm) bp stp = (bs, lm');
-    bs = length bp; bp \<noteq> []\<rbrakk> 
-  \<Longrightarrow> \<exists> stp. (\<lambda> (s, l). s < bs \<and> 
-              (abc_steps_l (s, l) bp (Suc 0)) = (bs, lm')) 
-      (abc_steps_l (0, lm) bp stp) "
-apply(erule_tac exE)
-proof -
-  fix stp
-  assume "bs = length bp" 
-         "abc_steps_l (0, lm) bp stp = (bs, lm')" 
-         "bp \<noteq> []"
-  thus 
-    "\<exists>stp. (\<lambda>(s, l). s < bs \<and> 
-      abc_steps_l (s, l) bp (Suc 0) = (bs, lm')) 
-                       (abc_steps_l (0, lm) bp stp)"
-    apply(induct stp, simp add: abc_steps_zero, simp)
-  proof -
-    fix stpa
-    assume ind: 
-     "abc_steps_l (0, lm) bp stpa = (length bp, lm')
-       \<Longrightarrow> \<exists>stp. (\<lambda>(s, l). s < length bp  \<and> abc_steps_l (s, l) bp 
-             (Suc 0) = (length bp, lm')) (abc_steps_l (0, lm) bp stp)"
-    and h: "abc_steps_l (0, lm) bp (Suc stpa) = (length bp, lm')" 
-           "abc_steps_l (0, lm) bp stp = (length bp, lm')" 
-           "bp \<noteq> []"
-    from h show 
-      "\<exists>stp. (\<lambda>(s, l). s < length bp \<and> abc_steps_l (s, l) bp (Suc 0)
-                    = (length bp, lm')) (abc_steps_l (0, lm) bp stp)"
-      apply(case_tac "abc_steps_l (0, lm) bp stpa", 
-            case_tac "a = length bp")
-      apply(insert ind, simp)
-      apply(subgoal_tac "b = lm'", simp)
-      apply(rule_tac abc_halt_equal, simp, simp)
-      apply(rule_tac x = stpa in exI, simp add: abc_steps_ind)
-      apply(simp add: abc_steps_zero)
-      apply(rule classical, simp add: abc_steps_l.simps 
-                             abc_fetch.simps abc_step_l.simps)
-      done
-  qed
-qed  
-
-
-lemma abc_append_empty_r[simp]: "[] [+] ab = ab"
-apply(simp add: abc_append.simps abc_inst_shift.simps)
-apply(induct ab, simp, simp)
-apply(case_tac a, simp_all add: abc_inst_shift.simps)
-done
-
-lemma abc_append_empty_l[simp]:  "ab [+] [] = ab"
-apply(simp add: abc_append.simps abc_inst_shift.simps)
-done
-
-
-lemma abc_append_length[simp]:  
-  "length (ap [+] bp) = length ap + length bp"
-apply(simp add: abc_append.simps)
-done
-
-declare Let_def[simp]
-
-lemma abc_append_commute: "as [+] bs [+] cs = as [+] (bs [+] cs)"
-apply(simp add: abc_append.simps abc_shift.simps abc_inst_shift.simps)
-apply(induct cs, simp, simp)
-apply(case_tac a, auto simp: abc_inst_shift.simps Let_def)
-done
-
-lemma abc_halt_point_step[simp]: 
-  "\<lbrakk>a < length bp; abc_steps_l (a, b) bp (Suc 0) = (length bp, lm')\<rbrakk>
-  \<Longrightarrow> abc_steps_l (length ap + a, b) (ap [+] bp [+] cp) (Suc 0) = 
-                                        (length ap + length bp, lm')"
-apply(simp add: abc_steps_l.simps abc_fetch.simps abc_append_nth)
-apply(case_tac "bp ! a", 
-                      auto simp: abc_steps_l.simps abc_step_l.simps)
-done
-
-lemma abc_step_state_in:
-  "\<lbrakk>bs < length bp;  abc_steps_l (a, b) bp (Suc 0) = (bs, l)\<rbrakk>
-  \<Longrightarrow> a < length bp"
-apply(simp add: abc_steps_l.simps abc_fetch.simps)
-apply(rule_tac classical, 
-      simp add: abc_step_l.simps abc_steps_l.simps)
-done
-
-
-lemma abc_append_state_in_exc: 
-  "\<lbrakk>bs < length bp; abc_steps_l (0, lm) bp stpa = (bs, l)\<rbrakk>
- \<Longrightarrow> abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa = 
-                                             (length ap + bs, l)"
-apply(induct stpa arbitrary: bs l, simp add: abc_steps_zero)
-proof -
-  fix stpa bs l
-  assume ind: 
-    "\<And>bs l. \<lbrakk>bs < length bp; abc_steps_l (0, lm) bp stpa = (bs, l)\<rbrakk>
-    \<Longrightarrow> abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa = 
-                                                (length ap + bs, l)"
-    and h: "bs < length bp" 
-           "abc_steps_l (0, lm) bp (Suc stpa) = (bs, l)"
-  from h show 
-    "abc_steps_l (length ap, lm) (ap [+] bp [+] cp) (Suc stpa) = 
-                                                (length ap + bs, l)"
-    apply(simp add: abc_steps_ind)
-    apply(case_tac "(abc_steps_l (0, lm) bp stpa)", simp)
-  proof -
-    fix a b
-    assume g: "abc_steps_l (0, lm) bp stpa = (a, b)" 
-              "abc_steps_l (a, b) bp (Suc 0) = (bs, l)"
-    from h and g have k1: "a < length bp"
-      apply(simp add: abc_step_state_in)
-      done
-    from h and g and k1 show 
-   "abc_steps_l (abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa) 
-              (ap [+] bp [+] cp) (Suc 0) = (length ap + bs, l)"
-      apply(insert ind[of a b], simp)
-      apply(simp add: abc_steps_l.simps abc_fetch.simps 
-                      abc_append_nth)
-      apply(case_tac "bp ! a", auto simp: 
-                                 abc_steps_l.simps abc_step_l.simps)
-      done
-  qed
-qed
-
-lemma [simp]: "abc_steps_l (0, am) [] stp = (0, am)"
-apply(induct stp, simp add: abc_steps_zero)
-apply(simp add: abc_steps_ind)
-apply(simp add: abc_steps_zero abc_steps_l.simps 
-                abc_fetch.simps abc_step_l.simps)
-done
-
-lemma abc_append_exc1:
-  "\<lbrakk>\<exists> stp. abc_steps_l (0, lm) bp stp = (bs, lm');
-    bs = length bp; 
-    as = length ap\<rbrakk>
-    \<Longrightarrow> \<exists> stp. abc_steps_l (as, lm) (ap [+] bp [+] cp) stp 
-                                                 = (as + bs, lm')"
-apply(case_tac "bp = []", erule_tac exE, simp,
-      rule_tac x = 0 in exI, simp add: abc_steps_zero)
-apply(frule_tac abc_halt_point_ex, simp, simp,
-      erule_tac exE, erule_tac exE) 
-apply(rule_tac x = "stpa + Suc 0" in exI)
-apply(case_tac "(abc_steps_l (0, lm) bp stpa)", 
-      simp add: abc_steps_ind)
-apply(subgoal_tac 
-  "abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa 
-                                   = (length ap + a, b)", simp)
-apply(simp add: abc_steps_zero)
-apply(rule_tac abc_append_state_in_exc, simp, simp)
-done
-
-lemma abc_append_exc3: 
-  "\<lbrakk>\<exists> stp. abc_steps_l (0, am) bp stp = (bs, bm); ss = length ap\<rbrakk>
-   \<Longrightarrow>  \<exists> stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
-apply(erule_tac exE)
-proof -
-  fix stp
-  assume h: "abc_steps_l (0, am) bp stp = (bs, bm)" "ss = length ap"
-  thus " \<exists>stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
-  proof(induct stp arbitrary: bs bm)
-    fix bs bm
-    assume "abc_steps_l (0, am) bp 0 = (bs, bm)"
-    thus "\<exists>stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
-      apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
-      done
-  next
-    fix stp bs bm
-    assume ind: 
-      "\<And>bs bm. \<lbrakk>abc_steps_l (0, am) bp stp = (bs, bm);
-                 ss = length ap\<rbrakk> \<Longrightarrow> 
-          \<exists>stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
-    and g: "abc_steps_l (0, am) bp (Suc stp) = (bs, bm)"
-    from g show 
-      "\<exists>stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
-      apply(insert abc_steps_add[of 0 am bp stp "Suc 0"], simp)
-      apply(case_tac "(abc_steps_l (0, am) bp stp)", simp)
-    proof -
-      fix a b
-      assume "(bs, bm) = abc_steps_l (a, b) bp (Suc 0)" 
-             "abc_steps_l (0, am) bp (Suc stp) = 
-                       abc_steps_l (a, b) bp (Suc 0)" 
-              "abc_steps_l (0, am) bp stp = (a, b)"
-      thus "?thesis"
-	apply(insert ind[of a b], simp add: h, erule_tac exE)
-	apply(rule_tac x = "Suc stp" in exI)
-	apply(simp only: abc_steps_ind, simp add: abc_steps_zero)
-      proof -
-	fix stp
-	assume "(bs, bm) = abc_steps_l (a, b) bp (Suc 0)"
-	thus "abc_steps_l (a + length ap, b) (ap [+] bp) (Suc 0)
-                                              = (bs + length ap, bm)"
-	  apply(simp add: abc_steps_l.simps abc_steps_zero
-                          abc_fetch.simps split: if_splits)
-	  apply(case_tac "bp ! a", 
-                simp_all add: abc_inst_shift.simps abc_append_nth
-                   abc_steps_l.simps abc_steps_zero abc_step_l.simps)
-	  apply(auto)
-	  done
-      qed
-    qed
-  qed
-qed
-
-lemma abc_add_equal:
-  "\<lbrakk>ap \<noteq> []; 
-    abc_steps_l (0, am) ap astp = (a, b);
-    a < length ap\<rbrakk>
-     \<Longrightarrow> (abc_steps_l (0, am) (ap @ bp) astp) = (a, b)"
-apply(induct astp arbitrary: a b, simp add: abc_steps_l.simps, simp)
-apply(simp add: abc_steps_ind)
-apply(case_tac "(abc_steps_l (0, am) ap astp)")
-proof -
-  fix astp a b aa ba
-  assume ind: 
-    "\<And>a b. \<lbrakk>abc_steps_l (0, am) ap astp = (a, b); 
-             a < length ap\<rbrakk> \<Longrightarrow> 
-                  abc_steps_l (0, am) (ap @ bp) astp = (a, b)"
-  and h: "abc_steps_l (abc_steps_l (0, am) ap astp) ap (Suc 0)
-                                                            = (a, b)"
-        "a < length ap" 
-        "abc_steps_l (0, am) ap astp = (aa, ba)"
-  from h show "abc_steps_l (abc_steps_l (0, am) (ap @ bp) astp)
-                                        (ap @ bp) (Suc 0) = (a, b)"
-    apply(insert ind[of aa ba], simp)
-    apply(subgoal_tac "aa < length ap", simp)
-    apply(simp add: abc_steps_l.simps abc_fetch.simps
-                     nth_append abc_steps_zero)
-    apply(rule abc_step_state_in, auto)
-    done
-qed
-
-
-lemma abc_add_exc1: 
-  "\<lbrakk>\<exists> astp. abc_steps_l (0, am) ap astp = (as, bm); as = length ap\<rbrakk>
-  \<Longrightarrow> \<exists> stp. abc_steps_l (0, am) (ap @ bp) stp = (as, bm)"
-apply(case_tac "ap = []", simp, 
-      rule_tac x = 0 in exI, simp add: abc_steps_zero)
-apply(drule_tac abc_halt_point_ex, simp, simp)
-apply(erule_tac exE, case_tac "(abc_steps_l (0, am) ap astp)", simp)
-apply(rule_tac x = "Suc astp" in exI, simp add: abc_steps_ind, auto)
-apply(frule_tac bp = bp in abc_add_equal, simp, simp, simp)
-apply(simp add: abc_steps_l.simps abc_steps_zero 
-                abc_fetch.simps nth_append)
-done
-
-declare abc_shift.simps[simp del] 
-
-lemma abc_append_exc2: 
-  "\<lbrakk>\<exists> astp. abc_steps_l (0, am) ap astp = (as, bm); as = length ap; 
-    \<exists> bstp. abc_steps_l (0, bm) bp bstp = (bs, bm'); bs = length bp;
-    cs = as + bs; bp \<noteq> []\<rbrakk>
-  \<Longrightarrow> \<exists> stp. abc_steps_l (0, am) (ap [+] bp) stp = (cs, bm')"
-apply(insert abc_append_exc1[of bm bp bs bm' as ap "[]"], simp)
-apply(drule_tac bp = "abc_shift bp (length ap)" in abc_add_exc1, simp)
-apply(subgoal_tac "ap @ abc_shift bp (length ap) = ap [+] bp", 
-      simp, auto)
-apply(rule_tac x = "stpa + stp" in exI, simp add: abc_steps_add)
-apply(simp add: abc_append.simps)
-done
-lemma exponent_add_iff: "a\<up>b @ a\<up>c@ xs = a\<up>(b+c) @ xs"
-apply(auto simp: replicate_add)
-done
-
-lemma exponent_cons_iff: "a # a\<up>c @ xs = a\<up>(Suc c) @ xs"
-apply(auto simp: replicate_add)
-done
-
-lemma  [simp]: "length lm = n \<Longrightarrow>  
-  abc_steps_l (Suc 0, lm @ Suc x # 0 # suf_lm) 
-       [Inc n, Dec (Suc n) 3, Goto (Suc 0)] (Suc (Suc 0))
-                                  = (3, lm @ Suc x # 0 # suf_lm)"
-apply(simp add: abc_steps_l.simps abc_fetch.simps 
-                abc_step_l.simps abc_lm_v.simps abc_lm_s.simps 
-                nth_append list_update_append)
-done
-
-lemma [simp]: 
-  "length lm = n \<Longrightarrow> 
-  abc_steps_l (Suc 0, lm @ Suc x # Suc y # suf_lm) 
-     [Inc n, Dec (Suc n) 3, Goto (Suc 0)] (Suc (Suc 0))
-  = (Suc 0, lm @ Suc x # y # suf_lm)"
-apply(simp add: abc_steps_l.simps abc_fetch.simps 
-                abc_step_l.simps abc_lm_v.simps abc_lm_s.simps 
-                nth_append list_update_append)
-done
-
-lemma pr_cycle_part_middle_inv: 
-  "\<lbrakk>length lm = n\<rbrakk> \<Longrightarrow> 
-  \<exists> stp. abc_steps_l (0, lm @ x # y # suf_lm) 
-                         [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp 
-  = (3, lm @ Suc x # 0 # suf_lm)"
-proof -
-  assume h: "length lm = n"
-  hence k1: "\<exists> stp. abc_steps_l (0, lm @ x # y # suf_lm) 
-                           [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp 
-    = (Suc 0, lm @ Suc x # y # suf_lm)"
-    apply(rule_tac x = "Suc 0" in exI)
-    apply(simp add: abc_steps_l.simps abc_step_l.simps 
-                    abc_lm_v.simps abc_lm_s.simps nth_append 
-                    list_update_append abc_fetch.simps)
-    done
-  from h have k2: 
-    "\<exists> stp. abc_steps_l (Suc 0, lm @ Suc x # y # suf_lm)
-                      [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp 
-    = (3, lm @ Suc x # 0 # suf_lm)"
-    apply(induct y)
-    apply(rule_tac x = "Suc (Suc 0)" in exI, simp, simp, 
-          erule_tac exE)
-    apply(rule_tac x = "Suc (Suc 0) + stp" in exI, 
-          simp only: abc_steps_add, simp)
-    done      
-  from k1 and k2 show 
-    "\<exists> stp. abc_steps_l (0, lm @ x # y # suf_lm) 
-                       [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp 
-    = (3, lm @ Suc x # 0 # suf_lm)"
-    apply(erule_tac exE, erule_tac exE)
-    apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
-    done
-qed
-
-lemma [simp]: 
-  "length lm = Suc n \<Longrightarrow> 
-  (abc_steps_l (length ap, lm @ x # Suc y # suf_lm) 
-           (ap @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length ap)]) 
-                    (Suc (Suc (Suc 0))))
-  = (length ap, lm @ Suc x # y # suf_lm)"
-apply(simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps 
-         abc_lm_v.simps list_update_append nth_append abc_lm_s.simps)
-done
-
-lemma switch_para_inv:
-  assumes bp_def:"bp =  ap @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto ss]"
-  and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" 
-         "ss = length ap" 
-         "length lm = Suc n"
-  shows " \<exists>stp. abc_steps_l (ss, lm @ x # y # suf_lm) bp stp =
-                               (0, lm @ (x + y) # 0 # suf_lm)"
-apply(induct y arbitrary: x)
-apply(rule_tac x = "Suc 0" in exI,
-  simp add: bp_def mv_box.simps abc_steps_l.simps 
-            abc_fetch.simps h abc_step_l.simps 
-            abc_lm_v.simps list_update_append nth_append
-            abc_lm_s.simps)
-proof -
-  fix y x
-  assume ind: 
-    "\<And>x. \<exists>stp. abc_steps_l (ss, lm @ x # y # suf_lm) bp stp = 
-                                     (0, lm @ (x + y) # 0 # suf_lm)"
-  show "\<exists>stp. abc_steps_l (ss, lm @ x # Suc y # suf_lm) bp stp = 
-                                  (0, lm @ (x + Suc y) # 0 # suf_lm)"
-    apply(insert ind[of "Suc x"], erule_tac exE)
-    apply(rule_tac x = "Suc (Suc (Suc 0)) + stp" in exI, 
-          simp only: abc_steps_add bp_def h)
-    apply(simp add: h)
-    done
-qed
-
-lemma [simp]:
-  "length lm = rs_pos \<and> Suc (Suc rs_pos) < a_md \<and> 0 < rs_pos \<Longrightarrow> 
-      a_md - Suc 0 < Suc (Suc (Suc (a_md + length suf_lm - 
-                                         Suc (Suc (Suc 0)))))"
-apply(arith)
-done
-
-lemma [simp]: 
-  "Suc (Suc rs_pos) < a_md \<and> 0 < rs_pos \<Longrightarrow> 
-                           \<not> a_md - Suc 0 < rs_pos - Suc 0"
-apply(arith)
-done
-
-lemma [simp]: 
-  "Suc (Suc rs_pos) < a_md \<and> 0 < rs_pos \<Longrightarrow> 
-           \<not> a_md - rs_pos < Suc (Suc (a_md - Suc (Suc rs_pos)))"
-apply(arith)
-done
-
-lemma butlast_append_last: "lm \<noteq> [] \<Longrightarrow> lm = butlast lm @ [last lm]"
-apply(auto)
-done
-
-lemma [simp]: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)
-           \<Longrightarrow> (Suc (Suc rs_pos)) < a_md"
-apply(simp add: rec_ci.simps)
-apply(case_tac "rec_ci f", simp)
-apply(case_tac "rec_ci g", simp)
-apply(arith)
-done
-
-(*
-lemma pr_para_ge_suc0: "rec_calc_rel (Pr n f g) lm xs \<Longrightarrow> 0 < n"
-apply(erule calc_pr_reverse, simp, simp)
-done
-*)
-
-lemma ci_pr_para_eq: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)
-                  \<Longrightarrow> rs_pos = Suc n"
-apply(simp add: rec_ci.simps)
-apply(case_tac "rec_ci g",  case_tac "rec_ci f", simp)
-done
-
-lemma [intro]:  
-  "\<lbrakk>rec_ci z = (aprog, rs_pos, a_md); rec_calc_rel z lm xs\<rbrakk>
-  \<Longrightarrow> length lm = rs_pos"
-apply(simp add: rec_ci.simps rec_ci_z_def)
-apply(erule_tac calc_z_reverse, simp)
-done
-
-lemma [intro]: 
-  "\<lbrakk>rec_ci s = (aprog, rs_pos, a_md); rec_calc_rel s lm xs\<rbrakk>
-  \<Longrightarrow> length lm = rs_pos"
-apply(simp add: rec_ci.simps rec_ci_s_def)
-apply(erule_tac calc_s_reverse, simp)
-done
-
-lemma [intro]: 
-  "\<lbrakk>rec_ci (recf.id nat1 nat2) = (aprog, rs_pos, a_md); 
-    rec_calc_rel (recf.id nat1 nat2) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
-apply(simp add: rec_ci.simps rec_ci_id.simps)
-apply(erule_tac calc_id_reverse, simp)
-done
-
-lemma [intro]: 
-  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
-    rec_calc_rel (Cn n f gs) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
-apply(erule_tac calc_cn_reverse, simp)
-apply(simp add: rec_ci.simps)
-apply(case_tac "rec_ci f",  simp)
-done
-
-lemma [intro]:
-  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
-    rec_calc_rel (Pr n f g) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
-apply(erule_tac  calc_pr_reverse, simp)
-apply(drule_tac ci_pr_para_eq, simp, simp)
-apply(drule_tac ci_pr_para_eq, simp)
-done
-
-lemma [intro]: 
-  "\<lbrakk>rec_ci (Mn n f) = (aprog, rs_pos, a_md);
-    rec_calc_rel (Mn n f) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
-apply(erule_tac calc_mn_reverse)
-apply(simp add: rec_ci.simps)
-apply(case_tac "rec_ci f",  simp)
-done
-
-lemma para_pattern: 
-  "\<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm xs\<rbrakk>
-  \<Longrightarrow> length lm = rs_pos"
-apply(case_tac f, auto)
-done
-
-lemma ci_pr_g_paras:
-  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
-    rec_ci g = (a, aa, ba);
-    rec_calc_rel (Pr n f g) (lm @ [x]) rs; x > 0\<rbrakk> \<Longrightarrow> 
-    aa = Suc rs_pos "
-apply(erule calc_pr_reverse, simp)
-apply(subgoal_tac "length (args @ [k, rk]) = aa", simp)
-apply(subgoal_tac "rs_pos = Suc n", simp)
-apply(simp add: ci_pr_para_eq)
-apply(erule para_pattern, simp)
-done
-
-lemma ci_pr_g_md_less: 
-  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
-    rec_ci g = (a, aa, ba)\<rbrakk> \<Longrightarrow> ba < a_md"
-apply(simp add: rec_ci.simps)
-apply(case_tac "rec_ci f",  auto)
-done
-
-lemma [intro]: "rec_ci z = (ap, rp, ad) \<Longrightarrow> rp < ad"
-  by(simp add: rec_ci.simps)
-
-lemma [intro]: "rec_ci s = (ap, rp, ad) \<Longrightarrow> rp < ad"
-  by(simp add: rec_ci.simps)
-
-lemma [intro]: "rec_ci (recf.id nat1 nat2) = (ap, rp, ad) \<Longrightarrow> rp < ad"
-  by(simp add: rec_ci.simps)
-
-lemma [intro]: "rec_ci (Cn n f gs) = (ap, rp, ad) \<Longrightarrow> rp < ad"
-apply(simp add: rec_ci.simps)
-apply(case_tac "rec_ci f",  simp)
-done
-
-lemma [intro]: "rec_ci (Pr n f g) = (ap, rp, ad) \<Longrightarrow> rp < ad"
-apply(simp add: rec_ci.simps)
-by(case_tac "rec_ci f", case_tac "rec_ci g",  auto)
-
-lemma [intro]: "rec_ci (Mn n f) = (ap, rp, ad) \<Longrightarrow> rp < ad"
-apply(simp add: rec_ci.simps)
-apply(case_tac "rec_ci f", simp)
-apply(arith)
-done
-
-lemma ci_ad_ge_paras: "rec_ci f = (ap, rp, ad) \<Longrightarrow> ad > rp"
-apply(case_tac f, auto)
-done
-
-lemma [elim]: "\<lbrakk>a [+] b = []; a \<noteq> [] \<or> b \<noteq> []\<rbrakk> \<Longrightarrow> RR"
-apply(auto simp: abc_append.simps abc_shift.simps)
-done
-
-lemma [intro]: "rec_ci z = ([], aa, ba) \<Longrightarrow> False"
-by(simp add: rec_ci.simps rec_ci_z_def)
-
-lemma [intro]: "rec_ci s = ([], aa, ba) \<Longrightarrow> False"
-by(auto simp: rec_ci.simps rec_ci_s_def addition.simps)
-
-lemma [intro]: "rec_ci (id m n) = ([], aa, ba) \<Longrightarrow> False"
-by(auto simp: rec_ci.simps rec_ci_id.simps addition.simps)
-
-lemma [intro]: "rec_ci (Cn n f gs) = ([], aa, ba) \<Longrightarrow> False"
-apply(case_tac "rec_ci f", auto simp: rec_ci.simps abc_append.simps)
-apply(simp add: abc_shift.simps mv_box.simps)
-done
-
-lemma [intro]: "rec_ci (Pr n f g) = ([], aa, ba) \<Longrightarrow> False"
-apply(simp add: rec_ci.simps)
-apply(case_tac "rec_ci f", case_tac "rec_ci g")
-by(auto)
-
-lemma [intro]: "rec_ci (Mn n f) = ([], aa, ba) \<Longrightarrow> False"
-apply(case_tac "rec_ci f", auto simp: rec_ci.simps)
-done
-
-lemma rec_ci_not_null:  "rec_ci g = (a, aa, ba) \<Longrightarrow> a \<noteq> []"
-by(case_tac g, auto)
-
-lemma calc_pr_g_def:
- "\<lbrakk>rec_calc_rel (Pr rs_pos f g) (lm @ [Suc x]) rsa;
-   rec_calc_rel (Pr rs_pos f g) (lm @ [x]) rsxa\<rbrakk>
- \<Longrightarrow> rec_calc_rel g (lm @ [x, rsxa]) rsa"
-apply(erule_tac calc_pr_reverse, simp, simp)
-apply(subgoal_tac "rsxa = rk", simp)
-apply(erule_tac rec_calc_inj, auto)
-done
-
-lemma ci_pr_md_def: 
-  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
-    rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\<rbrakk>
-  \<Longrightarrow> a_md = Suc (max (n + 3) (max bc ba))"
-by(simp add: rec_ci.simps)
-
-lemma  ci_pr_f_paras: 
-  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
-    rec_calc_rel (Pr n f g) lm rs;
-    rec_ci f = (ab, ac, bc)\<rbrakk>  \<Longrightarrow> ac = rs_pos - Suc 0"
-apply(subgoal_tac "\<exists>rs. rec_calc_rel f (butlast lm) rs", 
-      erule_tac exE)
-apply(drule_tac f = f and lm = "butlast lm" in para_pattern, 
-      simp, simp)
-apply(drule_tac para_pattern, simp)
-apply(subgoal_tac "lm \<noteq> []", simp)
-apply(erule_tac calc_pr_reverse, simp, simp)
-apply(erule calc_pr_zero_ex)
-done
-
-lemma ci_pr_md_ge_f:  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
-        rec_ci f = (ab, ac, bc)\<rbrakk> \<Longrightarrow> Suc bc \<le> a_md"
-apply(case_tac "rec_ci g")
-apply(simp add: rec_ci.simps, auto)
-done
-
-lemma ci_pr_md_ge_g:  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
-        rec_ci g = (ab, ac, bc)\<rbrakk> \<Longrightarrow> bc < a_md"
-apply(case_tac "rec_ci f")
-apply(simp add: rec_ci.simps, auto)
-done 
-
-lemma rec_calc_rel_def0: 
-  "\<lbrakk>rec_calc_rel (Pr n f g) lm rs; rec_calc_rel f (butlast lm) rsa\<rbrakk>
-  \<Longrightarrow> rec_calc_rel (Pr n f g) (butlast lm @ [0]) rsa"
-  apply(rule_tac calc_pr_zero, simp)
-apply(erule_tac calc_pr_reverse, simp, simp, simp)
-done
-
-lemma [simp]:  "length (mv_box m n) = 3"
-by (auto simp: mv_box.simps)
-(*
-lemma
-  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
-  rec_calc_rel (Pr n f g) lm rs;
-  rec_ci g = (a, aa, ba);
-  rec_ci f = (ab, ac, bc)\<rbrakk>
-\<Longrightarrow> \<exists>ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = 3 + length ab \<and> bp = recursive.mv_box (n - Suc 0) n 3"
-apply(simp add: rec_ci.simps)
-apply(rule_tac x = "recursive.mv_box (n - Suc 0) (max (Suc (Suc n)) (max bc ba)) 3 [+] ab" in exI, simp)
-apply(rule_tac x = "([Dec (max (Suc (Suc n)) (max bc ba)) (length a + 7)] [+] a [+] 
-  [Inc (n - Suc 0), Dec n 3, Goto (Suc 0)]) @ [Dec (Suc n) 0, Inc n, Goto (length a + 4)]" in exI, simp)
-apply(auto simp: abc_append_commute)
-done
-
-lemma  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
-        rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\<rbrakk>
-    \<Longrightarrow> \<exists>ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = 3 \<and> bp = ab"
-apply(simp add: rec_ci.simps)
-apply(rule_tac x = "recursive.mv_box (n - Suc 0) (max (Suc (Suc n)) (max bc ba)) 3" in exI, simp)
-apply(rule_tac x = "recursive.mv_box (n - Suc 0) n 3 [+]
-     ([Dec (max (Suc (Suc n)) (max bc ba)) (length a + 7)] [+] a 
-  [+] [Inc (n - Suc 0), Dec n 3, Goto (Suc 0)]) @ [Dec (Suc n) 0, Inc n, Goto (length a + 4)]" in exI, auto)
-apply(simp add: abc_append_commute)
-done
-*)
-
-lemma [simp]: "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); rec_calc_rel (Pr n f g) lm rs\<rbrakk>
-    \<Longrightarrow> rs_pos = Suc n"
-apply(simp add: ci_pr_para_eq)
-done
-
-
-lemma [simp]: "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); rec_calc_rel (Pr n f g) lm rs\<rbrakk>
-    \<Longrightarrow> length lm = Suc n"
-apply(subgoal_tac "rs_pos = Suc n", rule_tac para_pattern, simp, simp)
-apply(case_tac "rec_ci f", case_tac "rec_ci g", simp add: rec_ci.simps)
-done
-
-lemma [simp]: "rec_ci (Pr n f g) = (a, rs_pos, a_md) \<Longrightarrow> Suc (Suc n) < a_md"
-apply(case_tac "rec_ci f", case_tac "rec_ci g", simp add: rec_ci.simps)
-apply arith
-done
-
-lemma [simp]: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md) \<Longrightarrow> 0 < rs_pos"
-apply(case_tac "rec_ci f", case_tac "rec_ci g")
-apply(simp add: rec_ci.simps)
-done
-
-lemma [simp]: "Suc (Suc rs_pos) < a_md \<Longrightarrow> 
-       butlast lm @ (last lm - xa) # (rsa::nat) # 0 # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm =
-       butlast lm @ (last lm - xa) # rsa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm"
-apply(simp add: replicate_Suc[THEN sym])
-done
-
-lemma pr_cycle_part_ind: 
-  assumes g_ind: 
-  "\<And>lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow> 
-  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(ba - aa) @ suf_lm) a stp = 
-                    (length a, lm @ rs # 0\<up>(ba - Suc aa) @ suf_lm)"
-  and ap_def: 
-  "ap = ([Dec (a_md - Suc 0) (length a + 7)] [+]
-        (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)])) @
-         [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]"
-  and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" 
-         "rec_calc_rel (Pr n f g) 
-                   (butlast lm @ [last lm - Suc xa]) rsxa" 
-         "Suc xa \<le> last lm" 
-         "rec_ci g = (a, aa, ba)"
-         "rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rsa"
-         "lm \<noteq> []"
-  shows 
-  "\<exists>stp. abc_steps_l 
-     (0, butlast lm @ (last lm - Suc xa) # rsxa # 
-               0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) ap stp =
-     (0, butlast lm @ (last lm - xa) # rsa
-                 # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)"
-proof -
-  have k1: "\<exists>stp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) #
-    rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) ap stp =
-         (length a + 4, butlast lm @ (last lm - xa) # 0 # rsa #
-                           0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)"
-    apply(simp add: ap_def, rule_tac abc_add_exc1)
-    apply(rule_tac as = "Suc 0" and 
-      bm = "butlast lm @ (last lm - Suc xa) # 
-      rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm" in abc_append_exc2,
-      auto)
-  proof -
-    show 
-      "\<exists>astp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # rsxa 
-                   # 0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) 
-              [Dec (a_md - Suc 0)(length a + 7)] astp =
-      (Suc 0, butlast lm @ (last lm - Suc xa) # 
-             rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)"
-      apply(rule_tac x = "Suc 0" in exI, 
-          simp add: abc_steps_l.simps abc_step_l.simps
-                     abc_fetch.simps)
-      apply(subgoal_tac "length lm = Suc n \<and> rs_pos = Suc n \<and>
-                              a_md > Suc (Suc rs_pos)")
-      apply(simp add: abc_lm_v.simps nth_append abc_lm_s.simps)
-      apply(insert nth_append[of 
-                 "(last lm - Suc xa) # rsxa # 0\<up>(a_md - Suc (Suc rs_pos))" 
-                 "Suc xa # suf_lm" "(a_md - rs_pos)"], simp)
-      apply(simp add: list_update_append del: list_update.simps)
-      apply(insert list_update_append[of "(last lm - Suc xa) # rsxa # 
-                                          0\<up>(a_md - Suc (Suc rs_pos))" 
-                    "Suc xa # suf_lm" "a_md - rs_pos" "xa"], simp)
-      apply(case_tac a_md, simp, simp)
-      apply(insert h, simp)
-      apply(insert para_pattern[of "Pr n f g" aprog rs_pos a_md 
-                    "(butlast lm @ [last lm - Suc xa])" rsxa], simp)
-      done
-  next
-    show "\<exists>bstp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # 
-           rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm) (a [+] 
-            [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)]) bstp =
-         (3 + length a, butlast lm @ (last lm - xa) # 0 # rsa #
-                          0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)"
-      apply(rule_tac as = "length a" and
-               bm = "butlast lm @ (last lm - Suc xa) # rsxa # rsa #
-                     0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm" 
-        in abc_append_exc2, simp_all)
-    proof -
-      from h have j1: "aa = Suc rs_pos \<and> a_md > ba \<and> ba > Suc rs_pos"
-	apply(insert h)
-	apply(insert ci_pr_g_paras[of n f g aprog rs_pos
-                 a_md a aa ba "butlast lm" "last lm - xa" rsa], simp)
-	apply(drule_tac ci_pr_md_ge_g, auto)
-	apply(erule_tac ci_ad_ge_paras)
-	done
-      from h have j2: "rec_calc_rel g (butlast lm @ 
-                                  [last lm - Suc xa, rsxa]) rsa"
-	apply(rule_tac  calc_pr_g_def, simp, simp)
-	done
-      from j1 and j2 show 
-        "\<exists>astp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) #
-                rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm) a astp =
-        (length a, butlast lm @ (last lm - Suc xa) # rsxa # rsa 
-                         # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)"
-	apply(insert g_ind[of
-          "butlast lm @ (last lm - Suc xa) # [rsxa]" rsa 
-          "0\<up>(a_md - ba - Suc 0) @ xa # suf_lm"], simp, auto)
-	apply(simp add: exponent_add_iff)
-	apply(rule_tac x = stp in exI, simp add: numeral_3_eq_3)
-	done
-    next
-      from h have j3: "length lm = rs_pos \<and> rs_pos > 0"
-	apply(rule_tac conjI)
-	apply(drule_tac lm = "(butlast lm @ [last lm - Suc xa])"
-                          and xs = rsxa in para_pattern, simp, simp, simp)
-        done
-      from h have j4: "Suc (last lm - Suc xa) = last lm - xa"
-	apply(case_tac "last lm", simp, simp)
-	done
-      from j3 and j4 show
-      "\<exists>bstp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # rsxa #
-                     rsa # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)
-            [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)] bstp =
-        (3, butlast lm @ (last lm - xa) # 0 # rsa #
-                       0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)"
-	apply(insert pr_cycle_part_middle_inv[of "butlast lm" 
-          "rs_pos - Suc 0" "(last lm - Suc xa)" rsxa 
-          "rsa # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm"], simp)
-	done
-    qed
-  qed
-  from h have k2: 
-    "\<exists>stp. abc_steps_l (length a + 4, butlast lm @ (last lm - xa) # 0 
-           # rsa # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm) ap stp =
-    (0, butlast lm @ (last lm - xa) # rsa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)"
-    apply(insert switch_para_inv[of ap 
-      "([Dec (a_md - Suc 0) (length a + 7)] [+] 
-      (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)]))"
-      n "length a + 4" f g aprog rs_pos a_md 
-      "butlast lm @ [last lm - xa]" 0 rsa 
-      "0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm"])
-    apply(simp add: h ap_def)
-    apply(subgoal_tac "length lm = Suc n \<and> Suc (Suc rs_pos) < a_md", 
-          simp)
-    apply(insert h, simp)
-    apply(frule_tac lm = "(butlast lm @ [last lm - Suc xa])" 
-      and xs = rsxa in para_pattern, simp, simp)
-    done   
-  from k1 and k2 show "?thesis"
-    apply(auto)
-    apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
-    done
-qed
-
-lemma ci_pr_ex1: 
-  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
-    rec_ci g = (a, aa, ba);
-    rec_ci f = (ab, ac, bc)\<rbrakk>
-\<Longrightarrow> \<exists>ap bp. length ap = 6 + length ab \<and>
-    aprog = ap [+] bp \<and>
-    bp = ([Dec (a_md - Suc 0) (length a + 7)] [+] (a [+]
-         [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)])) @ 
-         [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]"
-apply(simp add: rec_ci.simps)
-apply(rule_tac x = "recursive.mv_box n (max (Suc (Suc (Suc n)))
-    (max bc ba)) [+] ab [+] recursive.mv_box n (Suc n)" in exI,
-     simp)
-apply(auto simp add: abc_append_commute numeral_3_eq_3)
-done
-
-lemma pr_cycle_part:
-  "\<lbrakk>\<And>lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow>
-     \<exists>stp. abc_steps_l (0, lm @ 0\<up>(ba - aa) @ suf_lm) a stp = 
-                        (length a, lm @ rs # 0\<up>(ba - Suc aa) @ suf_lm);
-  rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
-  rec_calc_rel (Pr n f g) lm rs;
-  rec_ci g = (a, aa, ba);
-  rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rsx;
-  rec_ci f = (ab, ac, bc);
-  lm \<noteq> [];
-  x \<le> last lm\<rbrakk> \<Longrightarrow> 
-  \<exists>stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - x) #
-              rsx # 0\<up>(a_md - Suc (Suc rs_pos)) @ x # suf_lm) aprog stp =
-  (6 + length ab, butlast lm @ last lm # rs #
-                                0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)"
-proof -
-  assume g_ind:
-    "\<And>lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow> 
-    \<exists>stp. abc_steps_l (0, lm @ 0\<up>(ba - aa) @ suf_lm) a stp =
-                      (length a, lm @ rs # 0\<up>(ba - Suc aa) @ suf_lm)"
-    and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" 
-           "rec_calc_rel (Pr n f g) lm rs" 
-           "rec_ci g = (a, aa, ba)"
-           "rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rsx" 
-           "lm \<noteq> []"
-           "x \<le> last lm" 
-           "rec_ci f = (ab, ac, bc)" 
-  from h show 
-    "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - x) # 
-            rsx # 0\<up>(a_md - Suc (Suc rs_pos)) @ x # suf_lm) aprog stp =
-    (6 + length ab, butlast lm @ last lm # rs #
-                               0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)" 
-  proof(induct x arbitrary: rsx, simp_all)
-    fix rsxa
-    assume "rec_calc_rel (Pr n f g) lm rsxa" 
-           "rec_calc_rel (Pr n f g) lm rs"
-    from h and this have "rs = rsxa"
-      apply(subgoal_tac "lm \<noteq> [] \<and> rs_pos = Suc n", simp)
-      apply(rule_tac rec_calc_inj, simp, simp)
-      apply(simp)
-      done
-    thus "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @  last lm # 
-             rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm) aprog stp =
-      (6 + length ab, butlast lm @ last lm # rs #
-                               0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)"
-      by(rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
-  next
-    fix xa rsxa
-    assume ind:
-   "\<And>rsx. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rsx 
-  \<Longrightarrow> \<exists>stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - xa) #
-             rsx # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm) aprog stp =
-      (6 + length ab, butlast lm @ last lm # rs # 
-                               0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)"
-      and g: "rec_calc_rel (Pr n f g) 
-                      (butlast lm @ [last lm - Suc xa]) rsxa"
-      "Suc xa \<le> last lm"
-      "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" 
-      "rec_calc_rel (Pr n f g) lm rs"
-      "rec_ci g = (a, aa, ba)" 
-      "rec_ci f = (ab, ac, bc)" "lm \<noteq> []"
-    from g have k1: 
-      "\<exists> rs. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rs"
-      apply(rule_tac rs = rs in  calc_pr_less_ex, simp, simp)
-      done
-    from g and this show 
-      "\<exists>stp. abc_steps_l (6 + length ab, 
-           butlast lm @ (last lm - Suc xa) # rsxa # 
-              0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) aprog stp =
-              (6 + length ab, butlast lm @ last lm # rs # 
-                                0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)"
-    proof(erule_tac exE)
-      fix rsa
-      assume k2: "rec_calc_rel (Pr n f g) 
-                           (butlast lm @ [last lm - xa]) rsa"
-      from g and k2 have
-      "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @ 
-       (last lm - Suc xa) # rsxa # 
-               0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) aprog stp
-        = (6 + length ab, butlast lm @ (last lm - xa) # rsa # 
-                               0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)"
-	proof -
-	  from g have k2_1: 
-            "\<exists> ap bp. length ap = 6 + length ab \<and>
-                   aprog = ap [+] bp \<and> 
-                   bp = ([Dec (a_md - Suc 0) (length a + 7)] [+]
-                  (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, 
-                  Goto (Suc 0)])) @
-                  [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]"
-            apply(rule_tac ci_pr_ex1, auto)
-	    done
-	  from k2_1 and k2 and g show "?thesis"
-	    proof(erule_tac exE, erule_tac exE)
-	      fix ap bp
-	      assume 
-                "length ap = 6 + length ab \<and> 
-                 aprog = ap [+] bp \<and> bp =
-                ([Dec (a_md - Suc 0) (length a + 7)] [+] 
-                (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3,
-                Goto (Suc 0)])) @ 
-                [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" 
-	      from g and this and k2 and g_ind show "?thesis"
-		apply(insert abc_append_exc3[of 
-                  "butlast lm @ (last lm - Suc xa) # rsxa #
-                  0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm" bp 0
-                  "butlast lm @ (last lm - xa) # rsa #
-                0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm" "length ap" ap],
-                 simp)
-		apply(subgoal_tac 
-                "\<exists>stp. abc_steps_l (0, butlast lm @ (last lm - Suc xa)
-                           # rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # 
-                              suf_lm) bp stp =
-	          (0, butlast lm @ (last lm - xa) # rsa #
-                           0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)",
-                      simp, erule_tac conjE, erule conjE)
-		apply(erule pr_cycle_part_ind, auto)
-		done
-	    qed
-	  qed  
-      from g and k2 and this show "?thesis"
-	apply(erule_tac exE)
-	apply(insert ind[of rsa], simp)
-	apply(erule_tac exE)
-	apply(rule_tac x = "stp + stpa" in exI, 
-              simp add: abc_steps_add)
-	done
-    qed
-  qed
-qed
-
-lemma ci_pr_length: 
-  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
-    rec_ci g = (a, aa, ba);  
-    rec_ci f = (ab, ac, bc)\<rbrakk>
-    \<Longrightarrow>  length aprog = 13 + length ab + length a"
-apply(auto simp: rec_ci.simps)
-done
-
-fun mv_box_inv :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> bool"
-  where
-  "mv_box_inv (as, lm) m n initlm = 
-         (let plus = initlm ! m + initlm ! n in
-           length initlm > max m n \<and> m \<noteq> n \<and> 
-              (if as = 0 then \<exists> k l. lm = initlm[m := k, n := l] \<and> 
-                    k + l = plus \<and> k \<le> initlm ! m 
-              else if as = 1 then \<exists> k l. lm = initlm[m := k, n := l]
-                             \<and> k + l + 1 = plus \<and> k < initlm ! m 
-              else if as = 2 then \<exists> k l. lm = initlm[m := k, n := l] 
-                              \<and> k + l = plus \<and> k \<le> initlm ! m
-              else if as = 3 then lm = initlm[m := 0, n := plus]
-              else False))"
-
-fun mv_box_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "mv_box_stage1 (as, lm) m  = 
-            (if as = 3 then 0 
-             else 1)"
-
-fun mv_box_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "mv_box_stage2 (as, lm) m = (lm ! m)"
-
-fun mv_box_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "mv_box_stage3 (as, lm) m = (if as = 1 then 3 
-                                else if as = 2 then 2
-                                else if as = 0 then 1 
-                                else 0)"
- 
-fun mv_box_measure :: "((nat \<times> nat list) \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)"
-  where
-  "mv_box_measure ((as, lm), m) = 
-     (mv_box_stage1 (as, lm) m, mv_box_stage2 (as, lm) m,
-      mv_box_stage3 (as, lm) m)"
-
-definition lex_pair :: "((nat \<times> nat) \<times> nat \<times> nat) set"
-  where
-  "lex_pair = less_than <*lex*> less_than"
-
-definition lex_triple :: 
- "((nat \<times> (nat \<times> nat)) \<times> (nat \<times> (nat \<times> nat))) set"
-  where
-  "lex_triple \<equiv> less_than <*lex*> lex_pair"
-
-definition mv_box_LE :: 
- "(((nat \<times> nat list) \<times> nat) \<times> ((nat \<times> nat list) \<times> nat)) set"
-  where 
-  "mv_box_LE \<equiv> (inv_image lex_triple mv_box_measure)"
-
-lemma wf_lex_triple: "wf lex_triple"
-  by (auto intro:wf_lex_prod simp:lex_triple_def lex_pair_def)
-
-lemma wf_mv_box_le[intro]: "wf mv_box_LE"
-by(auto intro:wf_inv_image wf_lex_triple simp: mv_box_LE_def)
-
-declare mv_box_inv.simps[simp del]
-
-lemma mv_box_inv_init:  
-"\<lbrakk>m < length initlm; n < length initlm; m \<noteq> n\<rbrakk> \<Longrightarrow> 
-  mv_box_inv (0, initlm) m n initlm"
-apply(simp add: abc_steps_l.simps mv_box_inv.simps)
-apply(rule_tac x = "initlm ! m" in exI, 
-      rule_tac x = "initlm ! n" in exI, simp)
-done
-
-lemma [simp]: "abc_fetch 0 (recursive.mv_box m n) = Some (Dec m 3)"
-apply(simp add: mv_box.simps abc_fetch.simps)
-done
-
-lemma [simp]: "abc_fetch (Suc 0) (recursive.mv_box m n) =
-               Some (Inc n)"
-apply(simp add: mv_box.simps abc_fetch.simps)
-done
-
-lemma [simp]: "abc_fetch 2 (recursive.mv_box m n) = Some (Goto 0)"
-apply(simp add: mv_box.simps abc_fetch.simps)
-done
-
-lemma [simp]: "abc_fetch 3 (recursive.mv_box m n) = None"
-apply(simp add: mv_box.simps abc_fetch.simps)
-done
-
-lemma [simp]: 
-  "\<lbrakk>m \<noteq> n; m < length initlm; n < length initlm;
-    k + l = initlm ! m + initlm ! n; k \<le> initlm ! m; 0 < k\<rbrakk>
- \<Longrightarrow> \<exists>ka la. initlm[m := k, n := l, m := k - Suc 0] = 
-     initlm[m := ka, n := la] \<and>
-     Suc (ka + la) = initlm ! m + initlm ! n \<and> 
-     ka < initlm ! m"
-apply(rule_tac x = "k - Suc 0" in exI, rule_tac x = l in exI, 
-      simp, auto)
-apply(subgoal_tac 
-      "initlm[m := k, n := l, m := k - Suc 0] = 
-       initlm[n := l, m := k, m := k - Suc 0]")
-apply(simp add: list_update_overwrite )
-apply(simp add: list_update_swap)
-apply(simp add: list_update_swap)
-done
-
-lemma [simp]:
-  "\<lbrakk>m \<noteq> n; m < length initlm; n < length initlm; 
-    Suc (k + l) = initlm ! m + initlm ! n;
-    k < initlm ! m\<rbrakk>
-    \<Longrightarrow> \<exists>ka la. initlm[m := k, n := l, n := Suc l] = 
-                initlm[m := ka, n := la] \<and> 
-                ka + la = initlm ! m + initlm ! n \<and> 
-                ka \<le> initlm ! m"
-apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, auto)
-done
-
-lemma [simp]: 
-  "\<lbrakk>length initlm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
-   \<forall>na. \<not> (\<lambda>(as, lm) m. as = 3) 
-    (abc_steps_l (0, initlm) (recursive.mv_box m n) na) m \<and> 
-  mv_box_inv (abc_steps_l (0, initlm) 
-           (recursive.mv_box m n) na) m n initlm \<longrightarrow>
-  mv_box_inv (abc_steps_l (0, initlm) 
-           (recursive.mv_box m n) (Suc na)) m n initlm \<and>
-  ((abc_steps_l (0, initlm) (recursive.mv_box m n) (Suc na), m),
-   abc_steps_l (0, initlm) (recursive.mv_box m n) na, m) \<in> mv_box_LE"
-apply(rule allI, rule impI, simp add: abc_steps_ind)
-apply(case_tac "(abc_steps_l (0, initlm) (recursive.mv_box m n) na)",
-      simp)
-apply(auto split:if_splits simp add:abc_steps_l.simps mv_box_inv.simps)
-apply(auto simp add: mv_box_LE_def lex_triple_def lex_pair_def 
-                     abc_step_l.simps abc_steps_l.simps
-                     mv_box_inv.simps abc_lm_v.simps abc_lm_s.simps
-                split: if_splits )
-apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, simp)
-done
-
-lemma mv_box_inv_halt: 
-  "\<lbrakk>length initlm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
-  \<exists> stp. (\<lambda> (as, lm). as = 3 \<and> 
-  mv_box_inv (as, lm) m n initlm) 
-             (abc_steps_l (0::nat, initlm) (mv_box m n) stp)"
-thm halt_lemma2
-apply(insert halt_lemma2[of mv_box_LE
-    "\<lambda> ((as, lm), m). mv_box_inv (as, lm) m n initlm"
-    "\<lambda> stp. (abc_steps_l (0, initlm) (recursive.mv_box m n) stp, m)"
-    "\<lambda> ((as, lm), m). as = (3::nat)"
-    ])
-apply(insert wf_mv_box_le)
-apply(simp add: mv_box_inv_init abc_steps_zero)
-apply(erule_tac exE)
-apply(rule_tac x = na in exI)
-apply(case_tac "(abc_steps_l (0, initlm) (recursive.mv_box m n) na)",
-      simp, auto)
-done
-
-lemma mv_box_halt_cond:
-  "\<lbrakk>m \<noteq> n; mv_box_inv (a, b) m n lm; a = 3\<rbrakk> \<Longrightarrow> 
-  b = lm[n := lm ! m + lm ! n, m := 0]"
-apply(simp add: mv_box_inv.simps, auto)
-apply(simp add: list_update_swap)
-done
-
-lemma mv_box_ex:
-  "\<lbrakk>length lm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
-  \<exists> stp. abc_steps_l (0::nat, lm) (mv_box m n) stp
-  = (3, (lm[n := (lm ! m + lm ! n)])[m := 0::nat])"
-apply(drule mv_box_inv_halt, simp, erule_tac exE)
-apply(rule_tac x = stp in exI)
-apply(case_tac "abc_steps_l (0, lm) (recursive.mv_box m n) stp",
-      simp)
-apply(erule_tac mv_box_halt_cond, auto)
-done
-
-lemma [simp]: 
-  "\<lbrakk>a_md = Suc (max (Suc (Suc n)) (max bc ba)); 
-   length lm = rs_pos \<and> rs_pos = n \<and> n > 0\<rbrakk>
-  \<Longrightarrow> n - Suc 0 < length lm + 
-  (Suc (max (Suc (Suc n)) (max bc ba)) - rs_pos + length suf_lm) \<and>
-   Suc (Suc n) < length lm + (Suc (max (Suc (Suc n)) (max bc ba)) -
-  rs_pos + length suf_lm) \<and> bc < length lm + (Suc (max (Suc (Suc n)) 
- (max bc ba)) - rs_pos + length suf_lm) \<and> ba < length lm + 
-  (Suc (max (Suc (Suc n)) (max bc ba)) - rs_pos + length suf_lm)"
-apply(arith)
-done
-
-lemma [simp]:
-  "\<lbrakk>a_md = Suc (max (Suc (Suc n)) (max bc ba)); 
-   length lm = rs_pos \<and> rs_pos = n \<and> n > 0\<rbrakk>
- \<Longrightarrow> n - Suc 0 < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba)) \<and>
-     Suc n < length suf_lm + max (Suc (Suc n)) (max bc ba) \<and> 
-     bc < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba)) \<and> 
-     ba < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba))"
-apply(arith)
-done
-
-lemma [simp]: "n - Suc 0 \<noteq> max (Suc (Suc n)) (max bc ba)"
-apply(arith)
-done
-
-lemma [simp]: 
-  "a_md \<ge> Suc bc \<and> rs_pos > 0 \<and> bc \<ge> rs_pos \<Longrightarrow> 
- bc - (rs_pos - Suc 0) + a_md - Suc bc = Suc (a_md - rs_pos - Suc 0)"
-apply(arith)
-done
-
-lemma [simp]: "length lm = n \<and> rs_pos = n \<and> 0 < rs_pos \<and> 
-                                                  Suc rs_pos < a_md 
-       \<Longrightarrow> n - Suc 0 < Suc (Suc (a_md + length suf_lm - Suc (Suc 0))) 
-        \<and> n < Suc (Suc (a_md + length suf_lm - Suc (Suc 0)))"
-apply(arith)
-done
-     
-lemma [simp]: "length lm = n \<and> rs_pos = n \<and> 0 < rs_pos \<and> 
-               Suc rs_pos < a_md \<Longrightarrow> n - Suc 0 \<noteq> n"
-by arith
-
-lemma ci_pr_ex2: 
-  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
-    rec_calc_rel (Pr n f g) lm rs; 
-    rec_ci g = (a, aa, ba); 
-    rec_ci f = (ab, ac, bc)\<rbrakk>
-  \<Longrightarrow> \<exists>ap bp. aprog = ap [+] bp \<and> 
-         ap = mv_box n (max (Suc (Suc (Suc n))) (max bc ba))"
-apply(simp add: rec_ci.simps)
-apply(rule_tac x = "(ab [+] (recursive.mv_box n (Suc n) [+]
-              ([Dec (max (n + 3) (max bc ba)) (length a + 7)] 
-      [+] (a [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @ 
-      [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]))" in exI, auto)
-apply(simp add: abc_append_commute numeral_3_eq_3)
-done
-
-lemma [simp]: 
-  "max (Suc (Suc (Suc n))) (max bc ba) - n < 
-     Suc (max (Suc (Suc (Suc n))) (max bc ba)) - n"
-apply(arith)
-done
-
-thm nth_replicate
-(*
-lemma exp_nth[simp]: "n < m \<Longrightarrow> a\<up>m ! n = a"
-apply(sim)
-done
-*)
-lemma [simp]: "length lm = n \<and> rs_pos = n \<and> 0 < n \<Longrightarrow> 
-                      lm[n - Suc 0 := 0::nat] = butlast lm @ [0]"
-apply(auto)
-apply(insert list_update_append[of "butlast lm" "[last lm]" 
-                                   "length lm - Suc 0" "0"], simp)
-done
-
-lemma [simp]: "\<lbrakk>length lm = n; 0 < n\<rbrakk>  \<Longrightarrow> lm ! (n - Suc 0) = last lm"
-apply(insert nth_append[of "butlast lm" "[last lm]" "n - Suc 0"],
-      simp)
-apply(insert butlast_append_last[of lm], auto)
-done
-lemma exp_suc_iff: "a\<up>b @ [a] = a\<up>(b + Suc 0)"
-apply(simp add: exp_ind del: replicate.simps)
-done
-
-lemma less_not_less[simp]: "n > 0 \<Longrightarrow> \<not> n < n - Suc 0"
-by auto
-
-lemma [simp]:
-  "Suc n < length suf_lm + max (Suc (Suc n)) (max bc ba) \<and> 
-  bc < Suc (length suf_lm + max (Suc (Suc n)) 
-  (max bc ba)) \<and> 
-  ba < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba))"
-  by arith
-
-lemma [simp]: "length lm = n \<and> rs_pos = n \<and> n > 0 \<Longrightarrow> 
-(lm @ 0\<up>(Suc (max (Suc (Suc n)) (max bc ba)) - n) @ suf_lm) 
-  [max (Suc (Suc n)) (max bc ba) :=
-   (lm @ 0\<up>(Suc (max (Suc (Suc n)) (max bc ba)) - n) @ suf_lm) ! (n - Suc 0) + 
-       (lm @ 0\<up>(Suc (max (Suc (Suc n)) (max bc ba)) - n) @ suf_lm) ! 
-                   max (Suc (Suc n)) (max bc ba), n - Suc 0 := 0::nat]
- = butlast lm @ 0 # 0\<up>(max (Suc (Suc n)) (max bc ba) - n) @ last lm # suf_lm"
-apply(simp add: nth_append nth_replicate list_update_append)
-apply(insert list_update_append[of "0\<up>((max (Suc (Suc n)) (max bc ba)) - n)"
-         "[0]" "max (Suc (Suc n)) (max bc ba) - n" "last lm"], simp)
-apply(simp add: exp_suc_iff Suc_diff_le del: list_update.simps)
-done
-
-lemma exp_eq: "(a = b) = (c\<up>a = c\<up>b)"
-apply(auto)
-done
-
-lemma [simp]:
-  "\<lbrakk>length lm = n; 0 < n;  Suc n < a_md\<rbrakk> \<Longrightarrow> 
-   (butlast lm @ rsa # 0\<up>(a_md - Suc n) @ last lm # suf_lm)
-    [n := (butlast lm @ rsa # 0\<up>(a_md - Suc n) @ last lm # suf_lm) ! 
-        (n - Suc 0) + (butlast lm @ rsa # (0::nat)\<up>(a_md - Suc n) @ 
-                                last lm # suf_lm) ! n, n - Suc 0 := 0]
- = butlast lm @ 0 # rsa # 0\<up>(a_md - Suc (Suc n)) @ last lm # suf_lm"
-apply(simp add: nth_append list_update_append)
-apply(case_tac "a_md - Suc n", auto)
-done
-
-lemma [simp]: 
-  "Suc (Suc rs_pos) \<le> a_md \<and> length lm = rs_pos \<and> 0 < rs_pos
-  \<Longrightarrow> a_md - Suc 0 < 
-          Suc (Suc (Suc (a_md + length suf_lm - Suc (Suc (Suc 0)))))"
-by arith
-
-lemma [simp]: 
-  "Suc (Suc rs_pos) \<le> a_md \<and> length lm = rs_pos \<and> 0 < rs_pos \<Longrightarrow> 
-                                   \<not> a_md - Suc 0 < rs_pos - Suc 0"
-by arith
-
-lemma [simp]: "Suc (Suc rs_pos) \<le> a_md \<Longrightarrow> 
-                                \<not> a_md - Suc 0 < rs_pos - Suc 0"
-by arith
-
-lemma [simp]: "\<lbrakk>Suc (Suc rs_pos) \<le> a_md\<rbrakk> \<Longrightarrow> 
-               \<not> a_md - rs_pos < Suc (Suc (a_md - Suc (Suc rs_pos)))"
-by arith 
-
-lemma [simp]: 
-  "Suc (Suc rs_pos) \<le> a_md \<and> length lm = rs_pos \<and> 0 < rs_pos
- \<Longrightarrow> (abc_lm_v (butlast lm @ last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos)) @
-        0 # suf_lm) (a_md - Suc 0) = 0 \<longrightarrow>
-      abc_lm_s (butlast lm @ last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos)) @ 
-        0 # suf_lm) (a_md - Suc 0) 0 = 
-         lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) \<and>
-     abc_lm_v (butlast lm @ last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos)) @ 
-               0 # suf_lm) (a_md - Suc 0) = 0"
-apply(simp add: abc_lm_v.simps nth_append abc_lm_s.simps)
-apply(insert nth_append[of "last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos))" 
-               "0 # suf_lm" "(a_md - rs_pos)"], auto)
-apply(simp only: exp_suc_iff)
-apply(subgoal_tac "a_md - Suc 0 < a_md + length suf_lm", simp)
-apply(case_tac "lm = []", auto)
-done
-
-lemma pr_prog_ex[simp]: "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
-      rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\<rbrakk>
-    \<Longrightarrow> \<exists>cp. aprog = recursive.mv_box n (max (n + 3) 
-                    (max bc ba)) [+] cp"
-apply(simp add: rec_ci.simps)
-apply(rule_tac x = "(ab [+] (recursive.mv_box n (Suc n) [+]
-              ([Dec (max (n + 3) (max bc ba)) (length a + 7)] 
-             [+] (a [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]))
-             @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]))" in exI)
-apply(auto simp: abc_append_commute)
-done
-
-lemma [simp]: "mv_box m n \<noteq> []"
-by (simp add: mv_box.simps)
-(*
-lemma [simp]: "\<lbrakk>rs_pos = n; 0 < rs_pos ; Suc rs_pos < a_md\<rbrakk> \<Longrightarrow> 
-                        n - Suc 0 < a_md + length suf_lm"
-by arith
-*)
-lemma [intro]: 
-  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
-    rec_ci f = (ab, ac, bc)\<rbrakk> \<Longrightarrow> 
-   \<exists>ap. (\<exists>cp. aprog = ap [+] ab [+] cp) \<and> length ap = 3"
-apply(case_tac "rec_ci g", simp add: rec_ci.simps)
-apply(rule_tac x = "mv_box n 
-              (max (n + 3) (max bc c))" in exI, simp)
-apply(rule_tac x = "recursive.mv_box n (Suc n) [+]
-                 ([Dec (max (n + 3) (max bc c)) (length a + 7)]
-                 [+] a [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])
-               @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" in exI, 
-      auto)
-apply(simp add: abc_append_commute)
-done
-
-lemma [intro]: 
-  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
-    rec_ci g = (a, aa, ba); 
-    rec_ci f = (ab, ac, bc)\<rbrakk> \<Longrightarrow> 
-    \<exists>ap. (\<exists>cp. aprog = ap [+] recursive.mv_box n (Suc n) [+] cp)
-      \<and> length ap = 3 + length ab"
-apply(simp add: rec_ci.simps)
-apply(rule_tac x = "recursive.mv_box n (max (n + 3)
-                                (max bc ba)) [+] ab" in exI, simp)
-apply(rule_tac x = "([Dec (max (n + 3) (max bc ba))
-  (length a + 7)] [+] a [+] 
-  [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ 
-  [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" in exI)
-apply(auto simp: abc_append_commute)
-done
-
-lemma [intro]: 
-  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
-    rec_ci g = (a, aa, ba); 
-    rec_ci f = (ab, ac, bc)\<rbrakk>
-    \<Longrightarrow> \<exists>ap. (\<exists>cp. aprog = ap [+] ([Dec (a_md - Suc 0) (length a + 7)]
-             [+] (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, 
-             Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n),
-             Goto (length a + 4)] [+] cp) \<and>
-             length ap = 6 + length ab"
-apply(simp add: rec_ci.simps)
-apply(rule_tac x = "recursive.mv_box n
-    (max (n + 3) (max bc ba)) [+] ab [+] 
-     recursive.mv_box n (Suc n)" in exI, simp)
-apply(rule_tac x = "[]" in exI, auto)
-apply(simp add: abc_append_commute)
-done
-
-lemma [simp]: 
-  "n < Suc (max (n + 3) (max bc ba) + length suf_lm) \<and> 
-   Suc (Suc n) < max (n + 3) (max bc ba) + length suf_lm \<and> 
-   bc < Suc (max (n + 3) (max bc ba) + length suf_lm) \<and> 
-   ba < Suc (max (n + 3) (max bc ba) + length suf_lm)"
-by arith
-
-lemma [simp]: "n \<noteq> max (n + (3::nat)) (max bc ba)"
-by arith
-
-lemma [simp]:"length lm = Suc n \<Longrightarrow> lm[n := (0::nat)] = butlast lm @ [0]"
-apply(subgoal_tac "\<exists> xs x. lm = xs @ [x]", auto simp: list_update_append)
-apply(rule_tac x = "butlast lm" in exI, rule_tac x = "last lm" in exI)
-apply(case_tac lm, auto)
-done
-
-lemma [simp]:  "length lm = Suc n \<Longrightarrow> lm ! n =last lm"
-apply(subgoal_tac "lm \<noteq> []")
-apply(simp add: last_conv_nth, case_tac lm, simp_all)
-done
-
-lemma [simp]: "length lm = Suc n \<Longrightarrow> 
-      (lm @ (0::nat)\<up>(max (n + 3) (max bc ba) - n) @ suf_lm)
-           [max (n + 3) (max bc ba) := (lm @ 0\<up>(max (n + 3) (max bc ba) - n) @ suf_lm) ! n + 
-                  (lm @ 0\<up>(max (n + 3) (max bc ba) - n) @ suf_lm) ! max (n + 3) (max bc ba), n := 0]
-       = butlast lm @ 0 # 0\<up>(max (n + 3) (max bc ba) - Suc n) @ last lm # suf_lm"
-apply(auto simp: list_update_append nth_append)
-apply(subgoal_tac "(0\<up>(max (n + 3) (max bc ba) - n)) = 0\<up>(max (n + 3) (max bc ba) - Suc n) @ [0::nat]")
-apply(simp add: list_update_append)
-apply(simp add: exp_suc_iff)
-done
-
-lemma [simp]: "Suc (Suc n) < a_md \<Longrightarrow>  
-      n < Suc (Suc (a_md + length suf_lm - 2)) \<and>
-        n < Suc (a_md + length suf_lm - 2)"
-by(arith)
-
-lemma [simp]: "\<lbrakk>length lm = Suc n; Suc (Suc n) < a_md\<rbrakk>
-        \<Longrightarrow>(butlast lm @ (rsa::nat) # 0\<up>(a_md - Suc (Suc n)) @ last lm # suf_lm)
-          [Suc n := (butlast lm @ rsa # 0\<up>(a_md - Suc (Suc n)) @ last lm # suf_lm) ! n +
-                  (butlast lm @ rsa # 0\<up>(a_md - Suc (Suc n)) @ last lm # suf_lm) ! Suc n, n := 0]
-    = butlast lm @ 0 # rsa # 0\<up>(a_md - Suc (Suc (Suc n))) @ last lm # suf_lm"
-apply(auto simp: list_update_append)
-apply(subgoal_tac "(0\<up>(a_md - Suc (Suc n))) = (0::nat) # (0\<up>(a_md - Suc (Suc (Suc n))))", simp add: nth_append)
-apply(simp add: replicate_Suc[THEN sym])
-done
-
-lemma pr_case:
-  assumes nf_ind:
-  "\<And> lm rs suf_lm. rec_calc_rel f lm rs \<Longrightarrow> 
-  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(bc - ac) @ suf_lm) ab stp = 
-                (length ab, lm @ rs # 0\<up>(bc - Suc ac) @ suf_lm)"
-  and ng_ind: "\<And> lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow> 
-        \<exists>stp. abc_steps_l (0, lm @ 0\<up>(ba - aa) @ suf_lm) a stp = 
-                       (length a, lm @ rs # 0\<up>(ba - Suc aa) @ suf_lm)"
-    and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)"  "rec_calc_rel (Pr n f g) lm rs" 
-           "rec_ci g = (a, aa, ba)" "rec_ci f = (ab, ac, bc)" 
-  shows "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-proof -
-  from h have k1: "\<exists> stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp
-    = (3, butlast lm @ 0 # 0\<up>(a_md - rs_pos - 1) @ last lm # suf_lm)"
-  proof -
-    have "\<exists>bp cp. aprog = bp [+] cp \<and> bp = mv_box n 
-                 (max (n + 3) (max bc ba))"
-      apply(insert h, simp)
-      apply(erule pr_prog_ex, auto)
-      done
-    thus "?thesis"
-      apply(erule_tac exE, erule_tac exE, simp)
-      apply(subgoal_tac 
-           "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm)
-              ([] [+] recursive.mv_box n
-                  (max (n + 3) (max bc ba)) [+] cp) stp =
-             (0 + 3, butlast lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ 
-                                        last lm # suf_lm)", simp)
-      apply(rule_tac abc_append_exc1, simp_all)
-      apply(insert mv_box_ex[of "n" "(max (n + 3) 
-                 (max bc ba))" "lm @ 0\<up>(a_md - rs_pos) @ suf_lm"], simp)
-      apply(subgoal_tac "a_md = Suc (max (n + 3) (max bc ba))",
-            simp)
-      apply(subgoal_tac "length lm = Suc n \<and> rs_pos = Suc n", simp)
-      apply(insert h)
-      apply(simp add: para_pattern ci_pr_para_eq)
-      apply(rule ci_pr_md_def, auto)
-      done
-  qed
-  from h have k2: 
-  "\<exists> stp. abc_steps_l (3,  butlast lm @ 0 # 0\<up>(a_md - rs_pos - 1) @ 
-             last lm # suf_lm) aprog stp 
-    = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-  proof -
-    from h have k2_1: "\<exists> rs. rec_calc_rel f (butlast lm) rs"
-      apply(erule_tac calc_pr_zero_ex)
-      done
-    thus "?thesis"
-    proof(erule_tac exE)
-      fix rsa
-      assume k2_2: "rec_calc_rel f (butlast lm) rsa"
-      from h and k2_2 have k2_2_1: 
-       "\<exists> stp. abc_steps_l (3, butlast lm @ 0 # 0\<up>(a_md - rs_pos - 1) 
-                 @ last lm # suf_lm) aprog stp
-        = (3 + length ab, butlast lm @ rsa # 0\<up>(a_md - rs_pos - 1) @ 
-                                             last lm # suf_lm)"
-      proof -
-	from h have j1: "
-          \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = 3 \<and> 
-              bp = ab"
-	  apply(auto)
-	  done
-	from h have j2: "ac = rs_pos - 1"
-	  apply(drule_tac ci_pr_f_paras, simp, auto)
-	  done
-	from h and j2 have j3: "a_md \<ge> Suc bc \<and> rs_pos > 0 \<and> bc \<ge> rs_pos"
-	  apply(rule_tac conjI)
-	  apply(erule_tac ab = ab and ac = ac in ci_pr_md_ge_f, simp)
-	  apply(rule_tac context_conjI)
-          apply(simp_all add: rec_ci.simps)
-	  apply(drule_tac ci_ad_ge_paras, drule_tac ci_ad_ge_paras)
-	  apply(arith)
-	  done	  
-	from j1 and j2 show "?thesis"
-	  apply(auto simp del: abc_append_commute)
-	  apply(rule_tac abc_append_exc1, simp_all)
-	  apply(insert nf_ind[of "butlast lm" "rsa" 
-                "0\<up>(a_md - bc - Suc 0) @ last lm # suf_lm"], 
-               simp add: k2_2 j2, erule_tac exE)
-	  apply(simp add: exponent_add_iff j3)
-	  apply(rule_tac x = "stp" in exI, simp)
-	  done
-      qed
-      from h have k2_2_2: 
-      "\<exists> stp. abc_steps_l (3 + length ab, butlast lm @ rsa # 
-                  0\<up>(a_md - rs_pos - 1) @ last lm # suf_lm) aprog stp
-        = (6 + length ab, butlast lm @ 0 # rsa # 
-                       0\<up>(a_md - rs_pos - 2) @ last lm # suf_lm)"
-      proof -	     
-	from h have "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
-          length ap = 3 + length ab \<and> bp = recursive.mv_box n (Suc n)"
-	  by auto
-	thus "?thesis"
-	proof(erule_tac exE, erule_tac exE, erule_tac exE, 
-              erule_tac exE)
-	  fix ap cp bp apa
-	  assume "aprog = ap [+] bp [+] cp \<and> length ap = 3 + 
-                    length ab \<and> bp = recursive.mv_box n (Suc n)"
-	  thus "?thesis"
-	    apply(simp del: abc_append_commute)
-	    apply(subgoal_tac 
-              "\<exists>stp. abc_steps_l (3 + length ab, 
-               butlast lm @ rsa # 0\<up>(a_md - Suc rs_pos) @
-                 last lm # suf_lm) (ap [+] 
-                   recursive.mv_box n (Suc n) [+] cp) stp =
-              ((3 + length ab) + 3, butlast lm @ 0 # rsa # 
-                  0\<up>(a_md - Suc (Suc rs_pos)) @ last lm # suf_lm)", simp)
-	    apply(rule_tac abc_append_exc1, simp_all)
-	    apply(insert mv_box_ex[of n "Suc n" 
-                    "butlast lm @ rsa # 0\<up>(a_md - Suc rs_pos) @ 
-                          last lm # suf_lm"], simp)
-	    apply(subgoal_tac "length lm = Suc n \<and> rs_pos = Suc n \<and> a_md > Suc (Suc n)", simp)
-	    apply(insert h, simp)
-            done
-	qed
-      qed
-      from h have k2_3: "lm \<noteq> []"
-	apply(rule_tac calc_pr_para_not_null, simp)
-	done
-      from h and k2_2 and k2_3 have k2_2_3: 
-      "\<exists> stp. abc_steps_l (6 + length ab, butlast lm @ 
-          (last lm - last lm) # rsa # 
-            0\<up>(a_md - (Suc (Suc rs_pos))) @ last lm # suf_lm) aprog stp
-        = (6 + length ab, butlast lm @ last lm # rs # 
-                        0\<up>(a_md - Suc (Suc (rs_pos))) @ 0 # suf_lm)"
-	apply(rule_tac x = "last lm" and g = g in pr_cycle_part, auto)
-	apply(rule_tac ng_ind, simp)
-	apply(rule_tac rec_calc_rel_def0, simp, simp)
-	done
-      from h  have k2_2_4: 
-       "\<exists> stp. abc_steps_l (6 + length ab,
-             butlast lm @ last lm # rs # 0\<up>(a_md - rs_pos - 2) @
-                  0 # suf_lm) aprog stp
-        = (13 + length ab + length a,
-                   lm @ rs # 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-      proof -
-	from h have 
-        "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
-                     length ap = 6 + length ab \<and> 
-                    bp = ([Dec (a_md - Suc 0) (length a + 7)] [+] 
-                         (a [+] [Inc (rs_pos - Suc 0), 
-                         Dec rs_pos 3, Goto (Suc 0)])) @ 
-                        [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]"
-	  by auto
-	thus "?thesis"
-	  apply(auto)
-	  apply(subgoal_tac  
-            "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @ 
-                last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)
-                (ap [+] ([Dec (a_md - Suc 0) (length a + 7)] [+] 
-                (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, 
-                Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), 
-                Goto (length a + 4)] [+] cp) stp =
-            (6 + length ab + (length a + 7) , 
-                 lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)", simp)
-	  apply(subgoal_tac "13 + (length ab + length a) = 
-                              13 + length ab + length a", simp)
-	  apply(arith)
-	  apply(rule abc_append_exc1, simp_all)
-	  apply(rule_tac x = "Suc 0" in exI, 
-                simp add: abc_steps_l.simps abc_fetch.simps
-                         nth_append abc_append_nth abc_step_l.simps)
-	  apply(subgoal_tac "a_md > Suc (Suc rs_pos) \<and> 
-                            length lm = rs_pos \<and> rs_pos > 0", simp)
-	  apply(insert h, simp)
-	  apply(subgoal_tac "rs_pos = Suc n", simp, simp)
-          done
-      qed
-      from h have k2_2_5: "length aprog = 13 + length ab + length a"
-	apply(rule_tac ci_pr_length, simp_all)
-	done
-      from k2_2_1 and k2_2_2 and k2_2_3 and k2_2_4 and k2_2_5 
-      show "?thesis"
-	apply(auto)
-	apply(rule_tac x = "stp + stpa + stpb + stpc" in exI, 
-              simp add: abc_steps_add)
-	done
-    qed
-  qed	
-  from k1 and k2 show 
-    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp 
-               = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-    apply(erule_tac exE)
-    apply(erule_tac exE)
-    apply(rule_tac x = "stp + stpa" in exI)
-    apply(simp add: abc_steps_add)
-    done
-qed
-
-thm rec_calc_rel.induct
-
-lemma eq_switch: "x = y \<Longrightarrow> y = x"
-by simp
-
-lemma [simp]: 
-  "\<lbrakk>rec_ci f = (a, aa, ba); 
-    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk> \<Longrightarrow> \<exists>bp. aprog = a @ bp"
-apply(simp add: rec_ci.simps)
-apply(rule_tac x = "[Dec (Suc n) (length a + 5), 
-      Dec (Suc n) (length a + 3), Goto (Suc (length a)), 
-      Inc n, Goto 0]" in exI, auto)
-done
-
-lemma ci_mn_para_eq[simp]: 
-  "rec_ci (Mn n f) = (aprog, rs_pos, a_md) \<Longrightarrow> rs_pos = n"
-apply(case_tac "rec_ci f", simp add: rec_ci.simps)
-done
-(*
-lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md); rec_calc_rel (Mn n f) lm rs\<rbrakk> \<Longrightarrow> aa = Suc rs_pos"
-apply(rule_tac calc_mn_reverse, simp)
-apply(insert para_pattern [of f a aa ba "lm @ [rs]" 0], simp)
-apply(subgoal_tac "rs_pos = length lm", simp)
-apply(drule_tac ci_mn_para_eq, simp)
-done
-*)
-lemma [simp]: "rec_ci f = (a, aa, ba) \<Longrightarrow> aa < ba"
-apply(simp add: ci_ad_ge_paras)
-done
-
-lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); 
-                rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
-    \<Longrightarrow> ba \<le> a_md"
-apply(simp add: rec_ci.simps)
-by arith
-
-lemma mn_calc_f: 
-  assumes ind: 
-  "\<And>aprog a_md rs_pos rs suf_lm lm.
-  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\<rbrakk>  
-  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp    
-           = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-  and h: "rec_ci f = (a, aa, ba)" 
-         "rec_ci (Mn n f) = (aprog, rs_pos, a_md)"  
-         "rec_calc_rel f (lm @ [x]) rsx" 
-         "aa = Suc n"
-  shows "\<exists>stp. abc_steps_l (0, lm @ x # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
-                  aprog stp = (length a, 
-                   lm @ x # rsx # 0\<up>(a_md - Suc (Suc rs_pos)) @ suf_lm)"
-proof -
-  from h have k1: "\<exists> ap bp. aprog = ap @ bp \<and> ap = a"
-    by simp
-  from h have k2: "rs_pos = n"
-    apply(erule_tac ci_mn_para_eq)
-    done
-  from h and k1 and k2 show "?thesis"
-  
-  proof(erule_tac exE, erule_tac exE, simp, 
-        rule_tac abc_add_exc1, auto)
-    fix bp
-    show 
-      "\<exists>astp. abc_steps_l (0, lm @ x # 0\<up>(a_md - Suc n) @ suf_lm) a astp
-      = (length a, lm @ x # rsx # 0\<up>(a_md - Suc (Suc n)) @ suf_lm)"
-      apply(insert ind[of a "Suc n" ba  "lm @ [x]" rsx 
-             "0\<up>(a_md - ba) @ suf_lm"], simp add: exponent_add_iff h k2)
-      apply(subgoal_tac "ba > aa \<and> a_md \<ge> ba \<and> aa = Suc n", 
-            insert h, auto)
-      done
-  qed
-qed
-
-fun mn_ind_inv ::
-  "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> nat list \<Rightarrow> bool"
-  where
-  "mn_ind_inv (as, lm') ss x rsx suf_lm lm = 
-           (if as = ss then lm' = lm @ x # rsx # suf_lm
-            else if as = ss + 1 then 
-                 \<exists>y. (lm' = lm @ x # y # suf_lm) \<and> y \<le> rsx
-            else if as = ss + 2 then 
-                 \<exists>y. (lm' = lm @ x # y # suf_lm) \<and> y \<le> rsx
-            else if as = ss + 3 then lm' = lm @ x # 0 # suf_lm
-            else if as = ss + 4 then lm' = lm @ Suc x # 0 # suf_lm
-            else if as = 0 then lm' = lm @ Suc x # 0 # suf_lm
-            else False
-)"
-
-fun mn_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "mn_stage1 (as, lm) ss n = 
-            (if as = 0 then 0 
-             else if as = ss + 4 then 1
-             else if as = ss + 3 then 2
-             else if as = ss + 2 \<or> as = ss + 1 then 3
-             else if as = ss then 4
-             else 0
-)"
-
-fun mn_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "mn_stage2 (as, lm) ss n = 
-            (if as = ss + 1 \<or> as = ss + 2 then (lm ! (Suc n))
-             else 0)"
-
-fun mn_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "mn_stage3 (as, lm) ss n = (if as = ss + 2 then 1 else 0)"
-
- 
-fun mn_measure :: "((nat \<times> nat list) \<times> nat \<times> nat) \<Rightarrow>
-                                                (nat \<times> nat \<times> nat)"
-  where
-  "mn_measure ((as, lm), ss, n) = 
-     (mn_stage1 (as, lm) ss n, mn_stage2 (as, lm) ss n,
-                                       mn_stage3 (as, lm) ss n)"
-
-definition mn_LE :: "(((nat \<times> nat list) \<times> nat \<times> nat) \<times>
-                     ((nat \<times> nat list) \<times> nat \<times> nat)) set"
-  where "mn_LE \<equiv> (inv_image lex_triple mn_measure)"
-
-thm halt_lemma2
-lemma wf_mn_le[intro]: "wf mn_LE"
-by(auto intro:wf_inv_image wf_lex_triple simp: mn_LE_def)
-
-declare mn_ind_inv.simps[simp del]
-
-lemma mn_inv_init: 
-  "mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 0)
-                                         (length a) x rsx suf_lm lm"
-apply(simp add: mn_ind_inv.simps abc_steps_zero)
-done
-
-lemma mn_halt_init: 
-  "rec_ci f = (a, aa, ba) \<Longrightarrow> 
-  \<not> (\<lambda>(as, lm') (ss, n). as = 0) 
-    (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 0) 
-                                                       (length a, n)"
-apply(simp add: abc_steps_zero)
-apply(erule_tac rec_ci_not_null)
-done
-
-thm rec_ci.simps
-lemma [simp]: 
-  "\<lbrakk>rec_ci f = (a, aa, ba); 
-    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
-    \<Longrightarrow> abc_fetch (length a) aprog =
-                      Some (Dec (Suc n) (length a + 5))"
-apply(simp add: rec_ci.simps abc_fetch.simps, 
-                erule_tac conjE, erule_tac conjE, simp)
-apply(drule_tac eq_switch, drule_tac eq_switch, simp)
-done
-
-lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
-    \<Longrightarrow> abc_fetch (Suc (length a)) aprog = Some (Dec (Suc n) (length a + 3))"
-apply(simp add: rec_ci.simps abc_fetch.simps, erule_tac conjE, erule_tac conjE, simp)
-apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append)
-done
-
-lemma [simp]:
-  "\<lbrakk>rec_ci f = (a, aa, ba);
-    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
-    \<Longrightarrow> abc_fetch (Suc (Suc (length a))) aprog = 
-                                     Some (Goto (length a + 1))"
-apply(simp add: rec_ci.simps abc_fetch.simps,
-      erule_tac conjE, erule_tac conjE, simp)
-apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append)
-done
-
-lemma [simp]: 
-  "\<lbrakk>rec_ci f = (a, aa, ba);
-    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
-    \<Longrightarrow> abc_fetch (length a + 3) aprog = Some (Inc n)"
-apply(simp add: rec_ci.simps abc_fetch.simps, 
-      erule_tac conjE, erule_tac conjE, simp)
-apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append)
-done
-
-lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
-    \<Longrightarrow> abc_fetch (length a + 4) aprog = Some (Goto 0)"
-apply(simp add: rec_ci.simps abc_fetch.simps, erule_tac conjE, erule_tac conjE, simp)
-apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append)
-done
-
-lemma [simp]: 
-  "0 < rsx
-   \<Longrightarrow> \<exists>y. (lm @ x # rsx # suf_lm)[Suc (length lm) := rsx - Suc 0]   
-    = lm @ x # y # suf_lm \<and> y \<le> rsx"
-apply(case_tac rsx, simp, simp)
-apply(rule_tac x = nat in exI, simp add: list_update_append)
-done
-
-lemma [simp]: 
-  "\<lbrakk>y \<le> rsx; 0 < y\<rbrakk>
-   \<Longrightarrow> \<exists>ya. (lm @ x # y # suf_lm)[Suc (length lm) := y - Suc 0] 
-          = lm @ x # ya # suf_lm \<and> ya \<le> rsx"
-apply(case_tac y, simp, simp)
-apply(rule_tac x = nat in exI, simp add: list_update_append)
-done
-
-lemma mn_halt_lemma: 
-  "\<lbrakk>rec_ci f = (a, aa, ba);
-    rec_ci (Mn n f) = (aprog, rs_pos, a_md);
-     0 < rsx; length lm = n\<rbrakk>
-    \<Longrightarrow>
-  \<forall>na. \<not> (\<lambda>(as, lm') (ss, n). as = 0)
-  (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog na) 
-                                                       (length a, n)
- \<and> mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm)
-                       aprog na) (length a) x rsx suf_lm lm 
-\<longrightarrow> mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 
-                         (Suc na)) (length a) x rsx suf_lm lm
- \<and> ((abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog (Suc na), 
-                                                    length a, n), 
-    abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog na,
-                              length a, n) \<in> mn_LE"
-apply(rule allI, rule impI, simp add: abc_steps_ind)
-apply(case_tac "(abc_steps_l (length a, lm @ x # rsx # suf_lm) 
-                                                   aprog na)", simp)
-apply(auto split:if_splits simp add:abc_steps_l.simps 
-                           mn_ind_inv.simps abc_steps_zero)
-apply(auto simp add: mn_LE_def lex_triple_def lex_pair_def 
-            abc_step_l.simps abc_steps_l.simps mn_ind_inv.simps
-            abc_lm_v.simps abc_lm_s.simps nth_append
-           split: if_splits)
-apply(drule_tac  rec_ci_not_null, simp)
-done
-
-lemma mn_halt:
-  "\<lbrakk>rec_ci f = (a, aa, ba);
-    rec_ci (Mn n f) = (aprog, rs_pos, a_md);
-    0 < rsx; length lm = n\<rbrakk>
-    \<Longrightarrow> \<exists> stp. (\<lambda> (as, lm'). (as = 0 \<and> 
-           mn_ind_inv (as, lm')  (length a) x rsx suf_lm lm))
-            (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog stp)"
-apply(insert wf_mn_le)	  
-apply(insert halt_lemma2[of mn_LE
-  "\<lambda> ((as, lm'), ss, n). mn_ind_inv (as, lm') ss x rsx suf_lm lm"
-  "\<lambda> stp. (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog stp, 
-  length a, n)"
-  "\<lambda> ((as, lm'), ss, n). as = 0"], 
-   simp)
-apply(simp add: mn_halt_init mn_inv_init)
-apply(drule_tac x = x and suf_lm = suf_lm in mn_halt_lemma, auto)
-apply(rule_tac x = n in exI, 
-      case_tac "(abc_steps_l (length a, lm @ x # rsx # suf_lm)
-                              aprog n)", simp)
-done
-
-lemma [simp]: "Suc rs_pos < a_md \<Longrightarrow> 
-                Suc (a_md - Suc (Suc rs_pos)) = a_md - Suc rs_pos"
-by arith
-
-term rec_ci
-(*
-lemma [simp]: "\<lbrakk>rec_ci (Mn n f) = (aprog, rs_pos, a_md); rec_calc_rel (Mn n f) lm rs\<rbrakk>  \<Longrightarrow> Suc rs_pos < a_md"
-apply(case_tac "rec_ci f")
-apply(subgoal_tac "c > b \<and> b = Suc rs_pos \<and> a_md \<ge> c")
-apply(arith, auto)
-done
-*)
-lemma mn_ind_step: 
-  assumes ind:  
-  "\<And>aprog a_md rs_pos rs suf_lm lm.
-  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md);
-   rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
-  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp
-            = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-  and h: "rec_ci f = (a, aa, ba)" 
-         "rec_ci (Mn n f) = (aprog, rs_pos, a_md)"  
-         "rec_calc_rel f (lm @ [x]) rsx" 
-         "rsx > 0" 
-         "Suc rs_pos < a_md" 
-         "aa = Suc rs_pos"
-  shows "\<exists>stp. abc_steps_l (0, lm @ x # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
-             aprog stp = (0, lm @ Suc x # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-thm abc_add_exc1
-proof -
-  have k1: 
-    "\<exists> stp. abc_steps_l (0, lm @ x #  0\<up>(a_md - Suc (rs_pos)) @ suf_lm)
-         aprog stp = 
-       (length a, lm @ x # rsx # 0\<up>(a_md  - Suc (Suc rs_pos)) @ suf_lm)"
-    apply(insert h)
-    apply(auto intro: mn_calc_f ind)
-    done
-  from h have k2: "length lm = n"
-    apply(subgoal_tac "rs_pos = n")
-    apply(drule_tac  para_pattern, simp, simp, simp)
-    done
-  from h have k3: "a_md > (Suc rs_pos)"
-    apply(simp)
-    done
-  from k2 and h and k3 have k4: 
-    "\<exists> stp. abc_steps_l (length a,
-       lm @ x # rsx # 0\<up>(a_md  - Suc (Suc rs_pos)) @ suf_lm) aprog stp = 
-        (0, lm @ Suc x # 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-    apply(frule_tac x = x and 
-       suf_lm = "0\<up>(a_md - Suc (Suc rs_pos)) @ suf_lm" in mn_halt, auto)
-    apply(rule_tac x = "stp" in exI, 
-          simp add: mn_ind_inv.simps rec_ci_not_null)
-    apply(simp only: replicate.simps[THEN sym], simp)
-    done
-  from k1 and k4 show "?thesis"
-    apply(auto)
-    apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
-    done
-qed
-
-lemma [simp]: 
-  "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md);
-    rec_calc_rel (Mn n f) lm rs\<rbrakk> \<Longrightarrow> aa = Suc rs_pos"
-apply(rule_tac calc_mn_reverse, simp)
-apply(insert para_pattern [of f a aa ba "lm @ [rs]" 0], simp)
-apply(subgoal_tac "rs_pos = length lm", simp)
-apply(drule_tac ci_mn_para_eq, simp)
-done
-
-lemma [simp]: "\<lbrakk>rec_ci (Mn n f) = (aprog, rs_pos, a_md);      
-                rec_calc_rel (Mn n f) lm rs\<rbrakk>  \<Longrightarrow> Suc rs_pos < a_md"
-apply(case_tac "rec_ci f")
-apply(subgoal_tac "c > b \<and> b = Suc rs_pos \<and> a_md \<ge> c")
-apply(arith, auto)
-done
-
-lemma mn_ind_steps:  
-  assumes ind:
-  "\<And>aprog a_md rs_pos rs suf_lm lm. 
-  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
-  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
-              (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-  and h: "rec_ci f = (a, aa, ba)" 
-  "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
-  "rec_calc_rel (Mn n f) lm rs"
-  "rec_calc_rel f (lm @ [rs]) 0" 
-  "\<forall>x<rs. (\<exists> v. rec_calc_rel f (lm @ [x]) v \<and> 0 < v)"
-  "n = length lm" 
-  "x \<le> rs"
-  shows "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
-                 aprog stp = (0, lm @ x # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-apply(insert h, induct x, 
-      rule_tac x = 0 in exI, simp add: abc_steps_zero, simp)
-proof -
-  fix x
-  assume k1: 
-    "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
-                aprog stp = (0, lm @ x # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-  and k2: "rec_ci (Mn (length lm) f) = (aprog, rs_pos, a_md)" 
-          "rec_calc_rel (Mn (length lm) f) lm rs" 
-          "rec_calc_rel f (lm @ [rs]) 0" 
-          "\<forall>x<rs.(\<exists> v. rec_calc_rel f (lm @ [x]) v \<and> v > 0)" 
-          "n = length lm" 
-          "Suc x \<le> rs" 
-          "rec_ci f = (a, aa, ba)"
-  hence k2:
-    "\<exists>stp. abc_steps_l (0, lm @ x # 0\<up>(a_md - rs_pos - 1) @ suf_lm) aprog
-               stp = (0, lm @ Suc x # 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-    apply(erule_tac x = x in allE)
-    apply(auto)
-    apply(rule_tac  x = x in mn_ind_step)
-    apply(rule_tac ind, auto)      
-    done
-  from k1 and k2 show 
-    "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
-          aprog stp = (0, lm @ Suc x # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-    apply(auto)
-    apply(rule_tac x = "stp + stpa" in exI, simp only: abc_steps_add)
-    done
-qed
-    
-lemma [simp]: 
-"\<lbrakk>rec_ci f = (a, aa, ba); 
-  rec_ci (Mn n f) = (aprog, rs_pos, a_md); 
-  rec_calc_rel (Mn n f) lm rs;
-  length lm = n\<rbrakk>
- \<Longrightarrow> abc_lm_v (lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) (Suc n) = 0"
-apply(auto simp: abc_lm_v.simps nth_append)
-done
-
-lemma [simp]: 
-  "\<lbrakk>rec_ci f = (a, aa, ba); 
-    rec_ci (Mn n f) = (aprog, rs_pos, a_md); 
-    rec_calc_rel (Mn n f) lm rs;
-     length lm = n\<rbrakk>
-    \<Longrightarrow> abc_lm_s (lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) (Suc n) 0 =
-                           lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm"
-apply(auto simp: abc_lm_s.simps list_update_append)
-done
-
-lemma mn_length: 
-  "\<lbrakk>rec_ci f = (a, aa, ba);
-    rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
-  \<Longrightarrow> length aprog = length a + 5"
-apply(simp add: rec_ci.simps, erule_tac conjE)
-apply(drule_tac eq_switch, drule_tac eq_switch, simp)
-done
-
-lemma mn_final_step:
-  assumes ind:
-  "\<And>aprog a_md rs_pos rs suf_lm lm.
-  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); 
-  rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
-  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
-              (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-  and h: "rec_ci f = (a, aa, ba)" 
-         "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
-         "rec_calc_rel (Mn n f) lm rs" 
-         "rec_calc_rel f (lm @ [rs]) 0" 
-  shows "\<exists>stp. abc_steps_l (0, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
-     aprog stp = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-proof -
-  from h and ind have k1:
-    "\<exists>stp.  abc_steps_l (0, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
-        aprog stp = (length a,  lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-    thm mn_calc_f
-    apply(insert mn_calc_f[of f a aa ba n aprog 
-                               rs_pos a_md lm rs 0 suf_lm], simp)
-    apply(subgoal_tac "aa = Suc n", simp add: exponent_cons_iff)
-    apply(subgoal_tac "rs_pos = n", simp, simp)
-    done
-  from h have k2: "length lm = n"
-    apply(subgoal_tac "rs_pos = n")
-    apply(drule_tac f = "Mn n f" in para_pattern, simp, simp, simp)
-    done
-  from h and k2 have k3: 
-  "\<exists>stp. abc_steps_l (length a, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
-    aprog stp = (length a + 5, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-    apply(rule_tac x = "Suc 0" in exI, 
-          simp add: abc_step_l.simps abc_steps_l.simps)
-    done
-  from h have k4: "length aprog = length a + 5"
-    apply(simp add: mn_length)
-    done
-  from k1 and k3 and k4 show "?thesis"
-    apply(auto)
-    apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
-    done
-qed
-
-lemma mn_case: 
-  assumes ind: 
-  "\<And>aprog a_md rs_pos rs suf_lm lm.
-  \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
-  \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
-               (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-  and h: "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
-         "rec_calc_rel (Mn n f) lm rs"
-  shows "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp 
-  = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-apply(case_tac "rec_ci f", simp)
-apply(insert h, rule_tac calc_mn_reverse, simp)
-proof -
-  fix a b c v
-  assume h: "rec_ci f = (a, b, c)" 
-            "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
-            "rec_calc_rel (Mn n f) lm rs" 
-            "rec_calc_rel f (lm @ [rs]) 0" 
-            "\<forall>x<rs. \<exists>v. rec_calc_rel f (lm @ [x]) v \<and> 0 < v"
-            "n = length lm"
-  hence k1:
-    "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm) aprog
-                  stp = (0, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-    thm mn_ind_steps
-    apply(auto intro: mn_ind_steps ind)
-    done
-  from h have k2: 
-    "\<exists>stp. abc_steps_l (0, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) aprog
-         stp = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-    apply(auto intro: mn_final_step ind)
-    done
-  from k1 and k2 show 
-    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
-  (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-    apply(auto, insert h)
-    apply(subgoal_tac "Suc rs_pos < a_md")
-    apply(rule_tac x = "stp + stpa" in exI, 
-      simp only: abc_steps_add exponent_cons_iff, simp, simp)
-    done
-qed
-
-lemma z_rs: "rec_calc_rel z lm rs \<Longrightarrow> rs = 0"
-apply(rule_tac calc_z_reverse, auto)
-done
-
-lemma z_case:
-  "\<lbrakk>rec_ci z = (aprog, rs_pos, a_md); rec_calc_rel z lm rs\<rbrakk>
-  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
-           (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-apply(simp add: rec_ci.simps rec_ci_z_def, auto)
-apply(rule_tac x = "Suc 0" in exI, simp add: abc_steps_l.simps 
-                               abc_fetch.simps abc_step_l.simps z_rs)
-done
-
-fun addition_inv :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow>     
-                     nat list \<Rightarrow> bool"
-  where
-  "addition_inv (as, lm') m n p lm = 
-        (let sn = lm ! n in
-         let sm = lm ! m in
-         lm ! p = 0 \<and>
-             (if as = 0 then \<exists> x. x \<le> lm ! m \<and> lm' = lm[m := x,
-                                    n := (sn + sm - x), p := (sm - x)]
-             else if as = 1 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x,
-                            n := (sn + sm - x - 1), p := (sm - x - 1)]
-             else if as = 2 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x, 
-                               n := (sn + sm - x), p := (sm - x - 1)]
-             else if as = 3 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x,
-                                   n := (sn + sm - x), p := (sm - x)]
-             else if as = 4 then \<exists> x. x \<le> lm ! m \<and> lm' = lm[m := x,
-                                       n := (sn + sm), p := (sm - x)] 
-             else if as = 5 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x, 
-                                  n := (sn + sm), p := (sm - x - 1)] 
-             else if as = 6 then \<exists> x. x < lm ! m \<and> lm' =
-                     lm[m := Suc x, n := (sn + sm), p := (sm - x - 1)]
-             else if as = 7 then lm' = lm[m := sm, n := (sn + sm)]
-             else False))"
-
-fun addition_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "addition_stage1 (as, lm) m p = 
-          (if as = 0 \<or> as = 1 \<or> as = 2 \<or> as = 3 then 2 
-           else if as = 4 \<or> as = 5 \<or> as = 6 then 1
-           else 0)"
-
-fun addition_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow>  nat \<Rightarrow> nat"
-  where
-  "addition_stage2 (as, lm) m p = 
-              (if 0 \<le> as \<and> as \<le> 3 then lm ! m
-               else if 4 \<le> as \<and> as \<le> 6 then lm ! p
-               else 0)"
-
-fun addition_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
-  where
-  "addition_stage3 (as, lm) m p = 
-             (if as = 1 then 4  
-              else if as = 2 then 3 
-              else if as = 3 then 2
-              else if as = 0 then 1 
-              else if as = 5 then 2
-              else if as = 6 then 1 
-              else if as = 4 then 0 
-              else 0)"
-
-fun addition_measure :: "((nat \<times> nat list) \<times> nat \<times> nat) \<Rightarrow> 
-                                                 (nat \<times> nat \<times> nat)"
-  where
-  "addition_measure ((as, lm), m, p) =
-     (addition_stage1 (as, lm) m p, 
-      addition_stage2 (as, lm) m p,
-      addition_stage3 (as, lm) m p)"
-
-definition addition_LE :: "(((nat \<times> nat list) \<times> nat \<times> nat) \<times> 
-                          ((nat \<times> nat list) \<times> nat \<times> nat)) set"
-  where "addition_LE \<equiv> (inv_image lex_triple addition_measure)"
-
-lemma [simp]: "wf addition_LE"
-by(simp add: wf_inv_image wf_lex_triple addition_LE_def)
-
-declare addition_inv.simps[simp del]
-
-lemma addition_inv_init: 
-  "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow>
-                                   addition_inv (0, lm) m n p lm"
-apply(simp add: addition_inv.simps)
-apply(rule_tac x = "lm ! m" in exI, simp)
-done
-
-thm addition.simps
-
-lemma [simp]: "abc_fetch 0 (addition m n p) = Some (Dec m 4)"
-by(simp add: abc_fetch.simps addition.simps)
-
-lemma [simp]: "abc_fetch (Suc 0) (addition m n p) = Some (Inc n)"
-by(simp add: abc_fetch.simps addition.simps)
-
-lemma [simp]: "abc_fetch 2 (addition m n p) = Some (Inc p)"
-by(simp add: abc_fetch.simps addition.simps)
-
-lemma [simp]: "abc_fetch 3 (addition m n p) = Some (Goto 0)"
-by(simp add: abc_fetch.simps addition.simps)
-
-lemma [simp]: "abc_fetch 4 (addition m n p) = Some (Dec p 7)"
-by(simp add: abc_fetch.simps addition.simps)
-
-lemma [simp]: "abc_fetch 5 (addition m n p) = Some (Inc m)"
-by(simp add: abc_fetch.simps addition.simps)
-
-lemma [simp]: "abc_fetch 6 (addition m n p) = Some (Goto 4)"
-by(simp add: abc_fetch.simps addition.simps)
-
-lemma [simp]:
-  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x \<le> lm ! m; 0 < x\<rbrakk>
- \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m - x, 
-                    p := lm ! m - x, m := x - Suc 0] =
-                 lm[m := xa, n := lm ! n + lm ! m - Suc xa,
-                    p := lm ! m - Suc xa]"
-apply(case_tac x, simp, simp)
-apply(rule_tac x = nat in exI, simp add: list_update_swap 
-                                         list_update_overwrite)
-done
-
-lemma [simp]:
-  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
-   \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m - Suc x,
-                      p := lm ! m - Suc x, n := lm ! n + lm ! m - x]
-                 = lm[m := xa, n := lm ! n + lm ! m - xa, 
-                      p := lm ! m - Suc xa]"
-apply(rule_tac x = x in exI, 
-      simp add: list_update_swap list_update_overwrite)
-done
-
-lemma [simp]: 
-  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
-   \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m - x, 
-                          p := lm ! m - Suc x, p := lm ! m - x]
-                 = lm[m := xa, n := lm ! n + lm ! m - xa, 
-                          p := lm ! m - xa]"
-apply(rule_tac x = x in exI, simp add: list_update_overwrite)
-done
-
-lemma [simp]: 
-  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = (0::nat); m < p; n < p; x < lm ! m\<rbrakk>
-  \<Longrightarrow> \<exists>xa\<le>lm ! m. lm[m := x, n := lm ! n + lm ! m - x,
-                                   p := lm ! m - x] = 
-                  lm[m := xa, n := lm ! n + lm ! m - xa, 
-                                   p := lm ! m - xa]"
-apply(rule_tac x = x in exI, simp)
-done
-
-lemma [simp]: 
-  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p;
-    x \<le> lm ! m; lm ! m \<noteq> x\<rbrakk>
-  \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m, 
-                       p := lm ! m - x, p := lm ! m - Suc x] 
-               = lm[m := xa, n := lm ! n + lm ! m, 
-                       p := lm ! m - Suc xa]"
-apply(rule_tac x = x in exI, simp add: list_update_overwrite)
-done
-
-lemma [simp]:
-  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
-  \<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m,
-                             p := lm ! m - Suc x, m := Suc x]
-                = lm[m := Suc xa, n := lm ! n + lm ! m, 
-                             p := lm ! m - Suc xa]"
-apply(rule_tac x = x in exI, 
-     simp add: list_update_swap list_update_overwrite)
-done
-
-lemma [simp]: 
-  "\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
-  \<Longrightarrow> \<exists>xa\<le>lm ! m. lm[m := Suc x, n := lm ! n + lm ! m, 
-                             p := lm ! m - Suc x] 
-               = lm[m := xa, n := lm ! n + lm ! m, p := lm ! m - xa]"
-apply(rule_tac x = "Suc x" in exI, simp)
-done
-
-lemma addition_halt_lemma: 
-  "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow>
-  \<forall>na. \<not> (\<lambda>(as, lm') (m, p). as = 7) 
-        (abc_steps_l (0, lm) (addition m n p) na) (m, p) \<and> 
-  addition_inv (abc_steps_l (0, lm) (addition m n p) na) m n p lm 
-\<longrightarrow> addition_inv (abc_steps_l (0, lm) (addition m n p) 
-                                 (Suc na)) m n p lm 
-  \<and> ((abc_steps_l (0, lm) (addition m n p) (Suc na), m, p), 
-     abc_steps_l (0, lm) (addition m n p) na, m, p) \<in> addition_LE"
-apply(rule allI, rule impI, simp add: abc_steps_ind)
-apply(case_tac "(abc_steps_l (0, lm) (addition m n p) na)", simp)
-apply(auto split:if_splits simp add: addition_inv.simps
-                                 abc_steps_zero)
-apply(simp_all add: abc_steps_l.simps abc_steps_zero)
-apply(auto simp add: addition_LE_def lex_triple_def lex_pair_def 
-                     abc_step_l.simps addition_inv.simps 
-                     abc_lm_v.simps abc_lm_s.simps nth_append
-                split: if_splits)
-apply(rule_tac x = x in exI, simp)
-done
-
-lemma  addition_ex: 
-  "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow> 
-  \<exists> stp. (\<lambda> (as, lm'). as = 7 \<and> addition_inv (as, lm') m n p lm) 
-                        (abc_steps_l (0, lm) (addition m n p) stp)"
-apply(insert halt_lemma2[of addition_LE
-  "\<lambda> ((as, lm'), m, p). addition_inv (as, lm') m n p lm"
-  "\<lambda> stp. (abc_steps_l (0, lm) (addition m n p) stp, m, p)"
-  "\<lambda> ((as, lm'), m, p). as = 7"], 
-  simp add: abc_steps_zero addition_inv_init)
-apply(drule_tac addition_halt_lemma, simp, simp, simp,
-      simp, erule_tac exE)
-apply(rule_tac x = na in exI, 
-      case_tac "(abc_steps_l (0, lm) (addition m n p) na)", auto)
-done
-
-lemma [simp]: "length (addition m n p) = 7"
-by (simp add: addition.simps)
-
-lemma [elim]: "addition 0 (Suc 0) 2 = [] \<Longrightarrow> RR"
-by(simp add: addition.simps)
-
-lemma [simp]: "(0\<up>2)[0 := n] = [n, 0::nat]"
-apply(subgoal_tac "2 = Suc 1", 
-      simp only: replicate.simps)
-apply(auto)
-done
-
-lemma [simp]: 
-  "\<exists>stp. abc_steps_l (0, n # 0\<up>2 @ suf_lm) 
-     (addition 0 (Suc 0) 2 [+] [Inc (Suc 0)]) stp = 
-                                      (8, n # Suc n # 0 # suf_lm)"
-apply(rule_tac bm = "n # n # 0 # suf_lm" in abc_append_exc2, auto)
-apply(insert addition_ex[of 0 "Suc 0" 2 "n # 0\<up>2 @ suf_lm"], 
-      simp add: nth_append numeral_2_eq_2, erule_tac exE)
-apply(rule_tac x = stp in exI,
-      case_tac "(abc_steps_l (0, n # 0\<up>2 @ suf_lm)
-                      (addition 0 (Suc 0) 2) stp)", 
-      simp add: addition_inv.simps nth_append list_update_append numeral_2_eq_2)
-apply(simp add: nth_append numeral_2_eq_2, erule_tac exE)
-apply(rule_tac x = "Suc 0" in exI,
-      simp add: abc_steps_l.simps abc_fetch.simps 
-      abc_steps_zero abc_step_l.simps abc_lm_s.simps abc_lm_v.simps)
-done
-
-lemma s_case:
-  "\<lbrakk>rec_ci s = (aprog, rs_pos, a_md); rec_calc_rel s lm rs\<rbrakk>
-  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
-               (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-apply(simp add: rec_ci.simps rec_ci_s_def, auto)
-apply(rule_tac calc_s_reverse, auto)
-done
-
-lemma [simp]: 
-  "\<lbrakk>n < length lm; lm ! n = rs\<rbrakk>
-    \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0 # 0 #suf_lm)
-                     (addition n (length lm) (Suc (length lm))) stp 
-             = (7, lm @ rs # 0 # suf_lm)"
-apply(insert addition_ex[of n "length lm"
-                           "Suc (length lm)" "lm @ 0 # 0 # suf_lm"])
-apply(simp add: nth_append, erule_tac exE)
-apply(rule_tac x = stp in exI)
-apply(case_tac "abc_steps_l (0, lm @ 0 # 0 # suf_lm) (addition n (length lm)
-                 (Suc (length lm))) stp", simp)
-apply(simp add: addition_inv.simps)
-apply(insert nth_append[of lm "0 # 0 # suf_lm" "n"], simp)
-done
-
-lemma [simp]: "0\<up>2 = [0, 0::nat]"
-apply(auto simp:numeral_2_eq_2)
-done
-
-lemma id_case: 
-  "\<lbrakk>rec_ci (id m n) = (aprog, rs_pos, a_md); 
-    rec_calc_rel (id m n) lm rs\<rbrakk>
-  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
-               (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-apply(simp add: rec_ci.simps rec_ci_id.simps, auto)
-apply(rule_tac calc_id_reverse, simp, simp)
-done   
-
-lemma list_tl_induct:
-  "\<lbrakk>P []; \<And>a list. P list \<Longrightarrow> P (list @ [a::'a])\<rbrakk> \<Longrightarrow> 
-                                            P ((list::'a list))"
-apply(case_tac "length list", simp)
-proof -
-  fix nat
-  assume ind: "\<And>a list. P list \<Longrightarrow> P (list @ [a])"
-  and h: "length list = Suc nat" "P []"
-  from h show "P list"
-  proof(induct nat arbitrary: list, case_tac lista, simp, simp)
-    fix lista a listaa
-    from h show "P [a]"
-      by(insert ind[of "[]"], simp add: h)
-  next
-    fix nat list
-    assume nind: "\<And>list. \<lbrakk>length list = Suc nat; P []\<rbrakk> \<Longrightarrow> P list" 
-    and g: "length (list:: 'a list) = Suc (Suc nat)"
-    from g show "P (list::'a list)"
-      apply(insert nind[of "butlast list"], simp add: h)
-      apply(insert ind[of "butlast list" "last list"], simp)
-      apply(subgoal_tac "butlast list @ [last list] = list", simp)
-      apply(case_tac "list::'a list", simp, simp)
-      done
-  qed
-qed      
-  
-lemma nth_eq_butlast_nth: "\<lbrakk>length ys > Suc k\<rbrakk> \<Longrightarrow> 
-                                        ys ! k = butlast ys ! k"
-apply(subgoal_tac "\<exists> xs y. ys = xs @ [y]", auto simp: nth_append)
-apply(rule_tac x = "butlast ys" in exI, rule_tac x = "last ys" in exI)
-apply(case_tac "ys = []", simp, simp)
-done
-
-lemma [simp]: 
-"\<lbrakk>\<forall>k<Suc (length list). rec_calc_rel ((list @ [a]) ! k) lm (ys ! k);
-  length ys = Suc (length list)\<rbrakk>
-   \<Longrightarrow> \<forall>k<length list. rec_calc_rel (list ! k) lm (butlast ys ! k)"
-apply(rule allI, rule impI)
-apply(erule_tac  x = k in allE, simp add: nth_append)
-apply(subgoal_tac "ys ! k = butlast ys ! k", simp)
-apply(rule_tac nth_eq_butlast_nth, arith)
-done
-
-lemma cn_merge_gs_tl_app: 
-  "cn_merge_gs (gs @ [g]) pstr = 
-        cn_merge_gs gs pstr [+] cn_merge_gs [g] (pstr + length gs)"
-apply(induct gs arbitrary: pstr, simp add: cn_merge_gs.simps, simp)
-apply(case_tac a, simp add: abc_append_commute)
-done
-
-lemma cn_merge_gs_length: 
-  "length (cn_merge_gs (map rec_ci list) pstr) = 
-      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci list. length ap) + 3 * length list "
-apply(induct list arbitrary: pstr, simp, simp)
-apply(case_tac "rec_ci a", simp)
-done
-
-lemma [simp]: "Suc n \<le> pstr \<Longrightarrow> pstr + x - n > 0"
-by arith
-
-lemma [simp]:
-  "\<lbrakk>Suc (pstr + length list) \<le> a_md; 
-    length ys = Suc (length list);
-    length lm = n;
-     Suc n \<le> pstr\<rbrakk>
-   \<Longrightarrow>  (ys ! length list # 0\<up>(pstr - Suc n) @ butlast ys @
-             0\<up>(a_md - (pstr + length list)) @ suf_lm) ! 
-                      (pstr + length list - n) = (0 :: nat)"
-apply(insert nth_append[of "ys ! length list # 0\<up>(pstr - Suc n) @
-     butlast ys" "0\<up>(a_md - (pstr + length list)) @ suf_lm"
-      "(pstr + length list - n)"], simp add: nth_append)
-done
-
-lemma [simp]:
-  "\<lbrakk>Suc (pstr + length list) \<le> a_md; 
-    length ys = Suc (length list);
-    length lm = n;
-     Suc n \<le> pstr\<rbrakk>
-    \<Longrightarrow> (lm @ last ys # 0\<up>(pstr - Suc n) @ butlast ys @
-         0\<up>(a_md - (pstr + length list)) @ suf_lm)[pstr + length list := 
-                                        last ys, n := 0] =
-        lm @ (0::nat)\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm"
-apply(insert list_update_length[of 
-   "lm @ last ys # 0\<up>(pstr - Suc n) @ butlast ys" 0 
-   "0\<up>(a_md - Suc (pstr + length list)) @ suf_lm" "last ys"], simp)
-apply(simp add: exponent_cons_iff)
-apply(insert list_update_length[of "lm" 
-        "last ys" "0\<up>(pstr - Suc n) @ butlast ys @ 
-      last ys # 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm" 0], simp)
-apply(simp add: exponent_cons_iff)
-apply(case_tac "ys = []", simp_all add: append_butlast_last_id)
-done
-
-lemma cn_merge_gs_ex: 
-  "\<lbrakk>\<And>x aprog a_md rs_pos rs suf_lm lm.
-    \<lbrakk>x \<in> set gs; rec_ci x = (aprog, rs_pos, a_md);
-     rec_calc_rel x lm rs\<rbrakk>
-     \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp 
-           = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm); 
-   pstr + length gs\<le> a_md;
-   \<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k);
-   length ys = length gs; length lm = n;
-   pstr \<ge> Max (set (Suc n # map (\<lambda>(aprog, p, n). n) (map rec_ci gs)))\<rbrakk>
-  \<Longrightarrow> \<exists> stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suf_lm)
-                   (cn_merge_gs (map rec_ci gs) pstr) stp 
-   = (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) gs) +
-  3 * length gs, lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - (pstr + length gs)) @ suf_lm)"
-apply(induct gs arbitrary: ys rule: list_tl_induct)
-apply(simp add: exponent_add_iff, simp)
-proof -
-  fix a list ys
-  assume ind: "\<And>x aprog a_md rs_pos rs suf_lm lm.
-    \<lbrakk>x = a \<or> x \<in> set list; rec_ci x = (aprog, rs_pos, a_md); 
-     rec_calc_rel x lm rs\<rbrakk>
-    \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
-                (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-  and ind2: 
-    "\<And>ys. \<lbrakk>\<And>x aprog a_md rs_pos rs suf_lm lm.
-    \<lbrakk>x \<in> set list; rec_ci x = (aprog, rs_pos, a_md);
-     rec_calc_rel x lm rs\<rbrakk>
-    \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp
-        = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm);
-    \<forall>k<length list. rec_calc_rel (list ! k) lm (ys ! k); 
-    length ys = length list\<rbrakk>
-    \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suf_lm) 
-                   (cn_merge_gs (map rec_ci list) pstr) stp =
-    (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) +
-     3 * length list,
-                lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm)"
-    and h: "Suc (pstr + length list) \<le> a_md" 
-            "\<forall>k<Suc (length list). 
-                   rec_calc_rel ((list @ [a]) ! k) lm (ys ! k)" 
-            "length ys = Suc (length list)" 
-            "length lm = n"
-            "Suc n \<le> pstr \<and> (\<lambda>(aprog, p, n). n) (rec_ci a) \<le> pstr \<and> 
-            (\<forall>a\<in>set list. (\<lambda>(aprog, p, n). n) (rec_ci a) \<le> pstr)"
-  from h have k1: 
-    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suf_lm)
-                     (cn_merge_gs (map rec_ci list) pstr) stp =
-    (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) +
-     3 * length list, lm @ 0\<up>(pstr - n) @ butlast ys @
-                               0\<up>(a_md - (pstr + length list)) @ suf_lm) "
-    apply(rule_tac ind2)
-    apply(rule_tac ind, auto)
-    done
-  from k1 and h show 
-    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suf_lm) 
-          (cn_merge_gs (map rec_ci list @ [rec_ci a]) pstr) stp =
-        (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) + 
-        (\<lambda>(ap, pos, n). length ap) (rec_ci a) + (3 + 3 * length list),
-             lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm)"
-    apply(simp add: cn_merge_gs_tl_app)
-    thm abc_append_exc2
-    apply(rule_tac as = 
-  "(\<Sum>(ap, pos, n)\<leftarrow>map rec_ci list. length ap) + 3 * length list"    
-      and bm = "lm @ 0\<up>(pstr - n) @ butlast ys @ 
-                              0\<up>(a_md - (pstr + length list)) @ suf_lm" 
-      and bs = "(\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3" 
-      and bm' = "lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ 
-                                  suf_lm" in abc_append_exc2, simp)
-    apply(simp add: cn_merge_gs_length)
-  proof -
-    from h show 
-      "\<exists>bstp. abc_steps_l (0, lm @ 0\<up>(pstr - n) @ butlast ys @ 
-                                  0\<up>(a_md - (pstr + length list)) @ suf_lm) 
-              ((\<lambda>(gprog, gpara, gn). gprog [+] recursive.mv_box gpara 
-              (pstr + length list)) (rec_ci a)) bstp =
-              ((\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3, 
-             lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm)"
-      apply(case_tac "rec_ci a", simp)
-      apply(rule_tac as = "length aa" and 
-                     bm = "lm @ (ys ! (length list)) # 
-          0\<up>(pstr - Suc n) @ butlast ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm" 
-        and bs = "3" and bm' = "lm @ 0\<up>(pstr - n) @ ys @
-             0\<up>(a_md - Suc (pstr + length list)) @ suf_lm" in abc_append_exc2)
-    proof -
-      fix aa b c
-      assume g: "rec_ci a = (aa, b, c)"
-      from h and g have k2: "b = n"
-	apply(erule_tac x = "length list" in allE, simp)
-	apply(subgoal_tac "length lm = b", simp)
-	apply(rule para_pattern, simp, simp)
-	done
-      from h and g and this show 
-        "\<exists>astp. abc_steps_l (0, lm @ 0\<up>(pstr - n) @ butlast ys @ 
-                         0\<up>(a_md - (pstr + length list)) @ suf_lm) aa astp =
-        (length aa, lm @ ys ! length list # 0\<up>(pstr - Suc n) @ 
-                       butlast ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm)"
-	apply(subgoal_tac "c \<ge> Suc n")
-	apply(insert ind[of a aa b c lm "ys ! length list" 
-          "0\<up>(pstr - c) @ butlast ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm"], simp)
-	apply(erule_tac x = "length list" in allE, 
-              simp add: exponent_add_iff)
-	apply(rule_tac Suc_leI, rule_tac ci_ad_ge_paras, simp)
-	done
-    next
-      fix aa b c
-      show "length aa = length aa" by simp 
-    next
-      fix aa b c
-      assume "rec_ci a = (aa, b, c)"
-      from h and this show     
-      "\<exists>bstp. abc_steps_l (0, lm @ ys ! length list #
-          0\<up>(pstr - Suc n) @ butlast ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm)
-                 (recursive.mv_box b (pstr + length list)) bstp =
-       (3, lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm)"
-	apply(insert mv_box_ex [of b "pstr + length list" 
-         "lm @ ys ! length list # 0\<up>(pstr - Suc n) @ butlast ys @ 
-         0\<up>(a_md - (pstr + length list)) @ suf_lm"], simp)
-        apply(subgoal_tac "b = n")
-	apply(simp add: nth_append split: if_splits)
-	apply(erule_tac x = "length list" in allE, simp)
-        apply(drule para_pattern, simp, simp)
-	done
-    next
-      fix aa b c
-      show "3 = length (recursive.mv_box b (pstr + length list))" 
-        by simp
-    next
-      fix aa b aaa ba
-      show "length aa + 3 = length aa + 3" by simp
-    next
-      fix aa b c
-      show "mv_box b (pstr + length list) \<noteq> []" 
-        by(simp add: mv_box.simps)
-    qed
-  next
-    show "(\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3 = 
-        length ((\<lambda>(gprog, gpara, gn). gprog [+]
-           recursive.mv_box gpara (pstr + length list)) (rec_ci a))"
-      by(case_tac "rec_ci a", simp)
-  next
-    show "listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) +
-      (\<lambda>(ap, pos, n). length ap) (rec_ci a) + (3 + 3 * length list)=
-      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci list. length ap) + 3 * length list + 
-                ((\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3)" by simp
-  next
-    show "(\<lambda>(gprog, gpara, gn). gprog [+] 
-      recursive.mv_box gpara (pstr + length list)) (rec_ci a) \<noteq> []"
-      by(case_tac "rec_ci a", 
-         simp add: abc_append.simps abc_shift.simps)
-  qed
-qed
- 
-lemma [simp]: "length (mv_boxes aa ba n) = 3*n"
-by(induct n, auto simp: mv_boxes.simps)
-
-lemma exp_suc: "a\<up>Suc b = a\<up>b @ [a]"
-by(simp add: exp_ind del: replicate.simps)
-
-lemma [simp]: 
-  "\<lbrakk>Suc n \<le> ba - aa;  length lm2 = Suc n;
-    length lm3 = ba - Suc (aa + n)\<rbrakk>
-  \<Longrightarrow> (last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba - aa) = (0::nat)"
-proof -
-  assume h: "Suc n \<le> ba - aa"
-  and g: "length lm2 = Suc n" "length lm3 = ba - Suc (aa + n)"
-  from h and g have k: "ba - aa = Suc (length lm3 + n)"
-    by arith
-  from  k show 
-    "(last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba - aa) = 0"
-    apply(simp, insert g)
-    apply(simp add: nth_append)
-    done
-qed
-
-lemma [simp]: "length lm1 = aa \<Longrightarrow>
-      (lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (aa + n) = last lm2"
-apply(simp add: nth_append)
-done
-
-lemma [simp]: "\<lbrakk>Suc n \<le> ba - aa; aa < ba\<rbrakk> \<Longrightarrow> 
-                    (ba < Suc (aa + (ba - Suc (aa + n) + n))) = False"
-apply arith
-done
-
-lemma [simp]: "\<lbrakk>Suc n \<le> ba - aa; aa < ba; length lm1 = aa; 
-       length lm2 = Suc n; length lm3 = ba - Suc (aa + n)\<rbrakk>
-     \<Longrightarrow> (lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba + n) = 0"
-using nth_append[of "lm1 @ (0\<Colon>'a)\<up>n @ last lm2 # lm3 @ butlast lm2" 
-                     "(0\<Colon>'a) # lm4" "ba + n"]
-apply(simp)
-done
-
-lemma [simp]: 
- "\<lbrakk>Suc n \<le> ba - aa; aa < ba; length lm1 = aa; length lm2 = Suc n;
-                 length lm3 = ba - Suc (aa + n)\<rbrakk>
-  \<Longrightarrow> (lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ (0::nat) # lm4)
-  [ba + n := last lm2, aa + n := 0] = 
-  lm1 @ 0 # 0\<up>n @ lm3 @ lm2 @ lm4"
-using list_update_append[of "lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2" "0 # lm4" 
-                            "ba + n" "last lm2"]
-apply(simp)
-apply(simp add: list_update_append)
-apply(simp only: exponent_cons_iff exp_suc, simp)
-apply(case_tac lm2, simp, simp)
-done
-
-lemma mv_boxes_ex:
-  "\<lbrakk>n \<le> ba - aa; ba > aa; length lm1 = aa; 
-    length (lm2::nat list) = n; length lm3 = ba - aa - n\<rbrakk>
-     \<Longrightarrow> \<exists> stp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<up>n @ lm4)
-       (mv_boxes aa ba n) stp = (3 * n, lm1 @ 0\<up>n @ lm3 @ lm2 @ lm4)"
-apply(induct n arbitrary: lm2 lm3 lm4, simp)
-apply(rule_tac x = 0 in exI, simp add: abc_steps_zero, 
-              simp add: mv_boxes.simps del: exp_suc_iff)
-apply(rule_tac as = "3 *n" and bm = "lm1 @ 0\<up>n @ last lm2 # lm3 @
-               butlast lm2 @ 0 # lm4" in abc_append_exc2, simp_all)
-apply(simp only: exponent_cons_iff, simp only: exp_suc, simp)
-proof -
-  fix n lm2 lm3 lm4
-  assume ind:
-    "\<And>lm2 lm3 lm4. \<lbrakk>length lm2 = n; length lm3 = ba - (aa + n)\<rbrakk> \<Longrightarrow>
-    \<exists>stp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<up>n @ lm4) 
-       (mv_boxes aa ba n) stp = (3 * n, lm1 @ 0\<up>n @ lm3 @ lm2 @ lm4)"
-  and h: "Suc n \<le> ba - aa" "aa < ba" "length (lm1::nat list) = aa" 
-         "length (lm2::nat list) = Suc n" 
-         "length (lm3::nat list) = ba - Suc (aa + n)"
-  from h show 
-    "\<exists>astp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<up>n @ 0 # lm4) 
-                       (mv_boxes aa ba n) astp = 
-        (3 * n, lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4)"
-    apply(insert ind[of "butlast lm2" "last lm2 # lm3" "0 # lm4"], 
-          simp)
-    apply(subgoal_tac "lm1 @ butlast lm2 @ last lm2 # lm3 @ 0\<up>n @ 
-              0 # lm4 = lm1 @ lm2 @ lm3 @ 0\<up>n @ 0 # lm4", simp, simp)
-    apply(case_tac "lm2 = []", simp, simp)
-    done
-next
-  fix n lm2 lm3 lm4
-  assume h: "Suc n \<le> ba - aa"
-            "aa < ba" 
-            "length (lm1::nat list) = aa" 
-            "length (lm2::nat list) = Suc n" 
-            "length (lm3::nat list) = ba - Suc (aa + n)"
-  thus " \<exists>bstp. abc_steps_l (0, lm1 @ 0\<up>n @ last lm2 # lm3 @
-                       butlast lm2 @ 0 # lm4) 
-                         (recursive.mv_box (aa + n) (ba + n)) bstp
-               = (3, lm1 @ 0 # 0\<up>n @ lm3 @ lm2 @ lm4)"
-    apply(insert mv_box_ex[of "aa + n" "ba + n" 
-       "lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4"], simp)
-    done
-qed
-(*    
-lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; 
-                ba < aa; 
-               length lm2 = aa - Suc (ba + n)\<rbrakk>
-      \<Longrightarrow> ((0::nat) # lm2 @ 0\<up>n @ last lm3 # lm4) ! (aa - ba)
-         = last lm3"
-proof -
-  assume h: "Suc n \<le> aa - ba"
-    and g: " ba < aa" "length lm2 = aa - Suc (ba + n)"
-  from h and g have k: "aa - ba = Suc (length lm2 + n)"
-    by arith
-  thus "((0::nat) # lm2 @ 0\<up>n @ last lm3 # lm4) ! (aa - ba) = last lm3"
-    apply(simp,  simp add: nth_append)
-    done
-qed
-*)
-
-lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; ba < aa; length lm1 = ba; 
-        length lm2 = aa - Suc (ba + n); length lm3 = Suc n\<rbrakk>
-   \<Longrightarrow> (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ last lm3 # lm4) ! (aa + n) = last lm3"
-using nth_append[of "lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n" "last lm3 # lm4" "aa + n"]
-apply(simp)
-done
-
-lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; ba < aa; length lm1 = ba; 
-        length lm2 = aa - Suc (ba + n); length lm3 = Suc n\<rbrakk>
-     \<Longrightarrow> (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ last lm3 # lm4) ! (ba + n) = 0"
-apply(simp add: nth_append)
-done
-
-lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; ba < aa; length lm1 = ba; 
-        length lm2 = aa - Suc (ba + n); length lm3 = Suc n\<rbrakk> 
-     \<Longrightarrow> (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ last lm3 # lm4)[ba + n := last lm3, aa + n := 0]
-      = lm1 @ lm3 @ lm2 @ 0 # 0\<up>n @ lm4"
-using list_update_append[of "lm1 @ butlast lm3" "(0\<Colon>'a) # lm2 @ (0\<Colon>'a)\<up>n @ last lm3 # lm4"]
-apply(simp)
-using list_update_append[of "lm1 @ butlast lm3 @ last lm3 # lm2 @ (0\<Colon>'a)\<up>n"
-                            "last lm3 # lm4" "aa + n" "0"]
-apply(simp)
-apply(simp only: replicate_Suc[THEN sym] exp_suc, simp)
-apply(case_tac lm3, simp, simp)
-done
-
-lemma mv_boxes_ex2:
-  "\<lbrakk>n \<le> aa - ba; 
-    ba < aa; 
-    length (lm1::nat list) = ba;
-    length (lm2::nat list) = aa - ba - n; 
-    length (lm3::nat list) = n\<rbrakk>
-     \<Longrightarrow> \<exists> stp. abc_steps_l (0, lm1 @ 0\<up>n @ lm2 @ lm3 @ lm4) 
-                (mv_boxes aa ba n) stp =
-                    (3 * n, lm1 @ lm3 @ lm2 @ 0\<up>n @ lm4)"
-apply(induct n arbitrary: lm2 lm3 lm4, simp)
-apply(rule_tac x = 0 in exI, simp add: abc_steps_zero, 
-                   simp add: mv_boxes.simps del: exp_suc_iff)
-apply(rule_tac as = "3 *n" and bm = "lm1 @ butlast lm3 @ 0 # lm2 @
-                  0\<up>n @ last lm3 # lm4" in abc_append_exc2, simp_all)
-apply(simp only: exponent_cons_iff, simp only: exp_suc, simp)
-proof -
-  fix n lm2 lm3 lm4
-  assume ind: 
-"\<And>lm2 lm3 lm4. \<lbrakk>length lm2 = aa - (ba + n); length lm3 = n\<rbrakk> \<Longrightarrow> 
-  \<exists>stp. abc_steps_l (0, lm1 @ 0\<up>n @ lm2 @ lm3 @ lm4) 
-                 (mv_boxes aa ba n) stp = 
-                            (3 * n, lm1 @ lm3 @ lm2 @ 0\<up>n @ lm4)"
-  and h: "Suc n \<le> aa - ba" 
-         "ba < aa"  
-         "length (lm1::nat list) = ba" 
-         "length (lm2::nat list) = aa - Suc (ba + n)" 
-         "length (lm3::nat list) = Suc n"
-  from h show
-    "\<exists>astp. abc_steps_l (0, lm1 @ 0\<up>n @ 0 # lm2 @ lm3 @ lm4)
-        (mv_boxes aa ba n) astp = 
-          (3 * n, lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ last lm3 # lm4)"
-    apply(insert ind[of "0 # lm2" "butlast lm3" "last lm3 # lm4"],
-          simp)
-    apply(subgoal_tac
-      "lm1 @ 0\<up>n @ 0 # lm2 @ butlast lm3 @ last lm3 # lm4 =
-           lm1 @ 0\<up>n @ 0 # lm2 @ lm3 @ lm4", simp, simp)
-    apply(case_tac "lm3 = []", simp, simp)
-    done
-next
-  fix n lm2 lm3 lm4
-  assume h:
-    "Suc n \<le> aa - ba" 
-    "ba < aa"
-    "length lm1 = ba"
-    "length (lm2::nat list) = aa - Suc (ba + n)" 
-    "length (lm3::nat list) = Suc n"
-  thus
-    "\<exists>bstp. abc_steps_l (0, lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ 
-                               last lm3 # lm4) 
-           (recursive.mv_box (aa + n) (ba + n)) bstp =
-                 (3, lm1 @ lm3 @ lm2 @ 0 # 0\<up>n @ lm4)"
-    apply(insert mv_box_ex[of "aa + n" "ba + n" "lm1 @ butlast lm3 @ 
-                          0 # lm2 @ 0\<up>n @ last lm3 # lm4"], simp)
-    done
-qed
-
-lemma cn_merge_gs_len: 
-  "length (cn_merge_gs (map rec_ci gs) pstr) = 
-      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs"
-apply(induct gs arbitrary: pstr, simp, simp)
-apply(case_tac "rec_ci a", simp)
-done
-
-lemma [simp]: "n < pstr \<Longrightarrow>
-     Suc (pstr + length ys - n) = Suc (pstr + length ys) - n"
-by arith
-
-lemma save_paras':  
-  "\<lbrakk>length lm = n; pstr > n; a_md > pstr + length ys + n\<rbrakk>
-  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(pstr - n) @ ys @
-               0\<up>(a_md - pstr - length ys) @ suf_lm) 
-                 (mv_boxes 0 (pstr + Suc (length ys)) n) stp
-        = (3 * n, 0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
-thm mv_boxes_ex
-apply(insert mv_boxes_ex[of n "pstr + Suc (length ys)" 0 "[]" "lm" 
-         "0\<up>(pstr - n) @ ys @ [0]" "0\<up>(a_md - pstr - length ys - n - Suc 0) @ suf_lm"], simp)
-apply(erule_tac exE, rule_tac x = stp in exI,
-                            simp add: exponent_add_iff)
-apply(simp only: exponent_cons_iff, simp)
-done
-
-lemma [simp]:
- "(max ba (Max (insert ba (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs))))
- = (Max (insert ba (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs)))"
-apply(rule min_max.sup_absorb2, auto)
-done
-
-lemma [simp]:
-  "((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs) = 
-                  (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs)"
-apply(induct gs)
-apply(simp, simp)
-done
-
-lemma ci_cn_md_def:  
-  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
-  rec_ci f = (a, aa, ba)\<rbrakk>
-    \<Longrightarrow> a_md = max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) o 
-  rec_ci) ` set gs))) + Suc (length gs) + n"
-apply(simp add: rec_ci.simps, auto)
-done
-
-lemma save_paras_prog_ex:
-  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
-    rec_ci f = (a, aa, ba); 
-    pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
-                                    (map rec_ci (f # gs))))\<rbrakk>
-    \<Longrightarrow> \<exists>ap bp cp. 
-      aprog = ap [+] bp [+] cp \<and>
-      length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
-              3 * length gs \<and> bp = mv_boxes 0 (pstr + Suc (length gs)) n"
-apply(simp add: rec_ci.simps)
-apply(rule_tac x = 
-  "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba 
-      (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs))))" in exI,
-      simp add: cn_merge_gs_len)
-apply(rule_tac x = 
-  "mv_boxes (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
-   0 (length gs) [+] a [+]recursive.mv_box aa (max (Suc n) 
-   (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
-   empty_boxes (length gs) [+] recursive.mv_box (max (Suc n) 
-  (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
-   mv_boxes (Suc (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) 
-   ` set gs))) + length gs)) 0 n" in exI, auto)
-apply(simp add: abc_append_commute)
-done
-
-lemma save_paras: 
-  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
-    rs_pos = n;
-    \<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k);
-    length ys = length gs;
-    length lm = n;
-    rec_ci f = (a, aa, ba);
-    pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
-                                          (map rec_ci (f # gs))))\<rbrakk>
-  \<Longrightarrow> \<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
-          3 * length gs, lm @ 0\<up>(pstr - n) @ ys @
-                 0\<up>(a_md - pstr - length ys) @ suf_lm) aprog stp = 
-           ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
-                      3 * length gs + 3 * n, 
-             0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
-proof -
-  assume h:
-    "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
-    "rs_pos = n" 
-    "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)" 
-    "length ys = length gs"  
-    "length lm = n"    
-    "rec_ci f = (a, aa, ba)"
-    and g: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
-                                        (map rec_ci (f # gs))))"
-  from h and g have k1: 
-    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
-    length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
-                3 *length gs \<and> bp = mv_boxes 0 (pstr + Suc (length ys)) n"
-    apply(drule_tac save_paras_prog_ex, auto)
-    done
-  from h have k2: 
-    "\<exists> stp. abc_steps_l (0, lm @ 0\<up>(pstr - n) @ ys @ 
-                         0\<up>(a_md - pstr - length ys) @ suf_lm)
-         (mv_boxes 0 (pstr + Suc (length ys)) n) stp = 
-        (3 * n, 0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
-    apply(rule_tac save_paras', simp, simp_all add: g)
-    apply(drule_tac a = a and aa = aa and ba = ba in 
-                                        ci_cn_md_def, simp, simp)
-    done
-  from k1 show 
-    "\<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
-         3 * length gs, lm @ 0\<up>(pstr - n) @ ys @ 
-                 0\<up>(a_md - pstr - length ys) @ suf_lm) aprog stp =
-             ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
-               3 * length gs + 3 * n, 
-                0\<up> pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
-  proof(erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
-    fix ap bp apa cp
-    assume "aprog = ap [+] bp [+] cp \<and> length ap = 
-            (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs
-            \<and> bp = mv_boxes 0 (pstr + Suc (length ys)) n"
-    from this and k2 show "?thesis"
-      apply(simp)
-      apply(rule_tac abc_append_exc1, simp, simp, simp)
-      done
-  qed
-qed
- 
-lemma ci_cn_para_eq:
-  "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md) \<Longrightarrow> rs_pos = n"
-apply(simp add: rec_ci.simps, case_tac "rec_ci f", simp)
-done
-
-lemma calc_gs_prog_ex: 
-  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
-    rec_ci f = (a, aa, ba);
-    Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
-                         (map rec_ci (f # gs)))) = pstr\<rbrakk>
-   \<Longrightarrow> \<exists>ap bp. aprog = ap [+] bp \<and> 
-                 ap = cn_merge_gs (map rec_ci gs) pstr"
-apply(simp add: rec_ci.simps)
-apply(rule_tac x = "mv_boxes 0 (Suc (max (Suc n)  
-   (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n [+]
-   mv_boxes (max (Suc n) (Max (insert ba 
-  (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
-   a [+] recursive.mv_box aa (max (Suc n)
-    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
-   empty_boxes (length gs) [+] recursive.mv_box (max (Suc n)
-    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
-    mv_boxes (Suc (max (Suc n) (Max 
-    (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n"
-   in exI)
-apply(auto simp: abc_append_commute)
-done
-
-lemma cn_calc_gs: 
-  assumes ind: 
-  "\<And>x aprog a_md rs_pos rs suf_lm lm.
-  \<lbrakk>x \<in> set gs; 
-   rec_ci x = (aprog, rs_pos, a_md); 
-   rec_calc_rel x lm rs\<rbrakk>
-  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
-     (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-  and h:  "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"  
-          "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
-          "length ys = length gs" 
-          "length lm = n" 
-          "rec_ci f = (a, aa, ba)" 
-          "Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
-                               (map rec_ci (f # gs)))) = pstr"
-  shows  
-  "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
-  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs, 
-   lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md -pstr - length ys) @ suf_lm) "
-proof -
-  from h have k1:
-    "\<exists> ap bp. aprog = ap [+] bp \<and> ap = 
-                        cn_merge_gs (map rec_ci gs) pstr"
-    by(erule_tac calc_gs_prog_ex, auto)
-  from h have j1: "rs_pos = n"
-    by(simp add: ci_cn_para_eq)
-  from h have j2: "a_md \<ge> pstr"
-    by(drule_tac a = a and aa = aa and ba = ba in 
-                                ci_cn_md_def, simp, simp)
-  from h have j3: "pstr > n"
-    by(auto)    
-  from j1 and j2 and j3 and h have k2:
-    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) 
-                         (cn_merge_gs (map rec_ci gs) pstr) stp 
-    = ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs, 
-                  lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - pstr - length ys) @ suf_lm)"
-    apply(simp)
-    apply(rule_tac cn_merge_gs_ex, rule_tac ind, simp, simp, auto)
-    apply(drule_tac a = a and aa = aa and ba = ba in 
-                                 ci_cn_md_def, simp, simp)
-    apply(rule min_max.le_supI2, auto)
-    done
-  from k1 show "?thesis"
-  proof(erule_tac exE, erule_tac exE, simp)
-    fix ap bp
-    from k2 show 
-      "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm)
-           (cn_merge_gs (map rec_ci gs) pstr [+] bp) stp =
-      (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) gs) +
-         3 * length gs, 
-         lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - (pstr + length ys)) @ suf_lm)"
-      apply(insert abc_append_exc1[of 
-        "lm @ 0\<up>(a_md - rs_pos) @ suf_lm" 
-        "(cn_merge_gs (map rec_ci gs) pstr)" 
-        "length (cn_merge_gs (map rec_ci gs) pstr)" 
-        "lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - pstr - length ys) @ suf_lm" 0 
-        "[]" bp], simp add: cn_merge_gs_len)
-      done      
-  qed
-qed
-
-lemma reset_new_paras': 
-  "\<lbrakk>length lm = n; 
-    pstr > 0; 
-    a_md \<ge> pstr + length ys + n;
-     pstr > length ys\<rbrakk> \<Longrightarrow>
-   \<exists>stp. abc_steps_l (0, 0\<up>pstr @ ys @ 0 # lm @  0\<up>(a_md - Suc (pstr + length ys + n)) @
-          suf_lm) (mv_boxes pstr 0 (length ys)) stp =
-  (3 * length ys, ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
-thm mv_boxes_ex2
-apply(insert mv_boxes_ex2[of "length ys" "pstr" 0 "[]"
-     "0\<up>(pstr - length ys)" "ys" 
-     "0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm"], 
-     simp add: exponent_add_iff)
-done
-
-lemma [simp]:  
-  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
-  rec_calc_rel f ys rs; rec_ci f = (a, aa, ba);
-  pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
-               (map rec_ci (f # gs))))\<rbrakk>
-  \<Longrightarrow> length ys < pstr"
-apply(subgoal_tac "length ys = aa", simp)
-apply(subgoal_tac "aa < ba \<and> ba \<le> pstr", 
-      rule basic_trans_rules(22), auto)
-apply(rule min_max.le_supI2)
-apply(auto)
-apply(erule_tac para_pattern, simp)
-done
-
-lemma reset_new_paras_prog_ex: 
-  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
-   rec_ci f = (a, aa, ba);
-   Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
-  (map rec_ci (f # gs)))) = pstr\<rbrakk>
-  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
-  length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
-           3 *length gs + 3 * n \<and> bp = mv_boxes pstr 0 (length gs)"
-apply(simp add: rec_ci.simps)
-apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) 
-          (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
-          mv_boxes 0 (Suc (max (Suc n) (Max (insert ba 
-           (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n" in exI, 
-       simp add: cn_merge_gs_len)
-apply(rule_tac x = "a [+]
-     recursive.mv_box aa (max (Suc n) (Max (insert ba 
-     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
-     empty_boxes (length gs) [+] recursive.mv_box 
-     (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n
-      [+] mv_boxes (Suc (max (Suc n) (Max (insert ba 
-     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI,
-       auto simp: abc_append_commute)
-done
-
-lemma reset_new_paras:
-       "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
-        rs_pos = n;
-        \<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k);
-        length ys = length gs;
-        length lm = n;
-        length ys = aa;
-        rec_ci f = (a, aa, ba);
-        pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
-                                    (map rec_ci (f # gs))))\<rbrakk>
-\<Longrightarrow> \<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
-                                               3 * length gs + 3 * n,
-        0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
-  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n,
-           ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
-proof -
-  assume h:
-    "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
-    "rs_pos = n" 
-    "length ys = aa"
-    "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
-    "length ys = length gs"  "length lm = n"    
-    "rec_ci f = (a, aa, ba)"
-    and g: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
-                                         (map rec_ci (f # gs))))"
-  thm rec_ci.simps
-  from h and g have k1:
-    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = 
-    (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
-          3 *length gs + 3 * n \<and> bp = mv_boxes pstr 0 (length ys)"
-    by(drule_tac reset_new_paras_prog_ex, auto)
-  from h have k2:
-    "\<exists> stp. abc_steps_l (0, 0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @
-              suf_lm) (mv_boxes pstr 0 (length ys)) stp = 
-    (3 * (length ys), 
-     ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
-    apply(rule_tac reset_new_paras', simp)
-    apply(simp add: g)
-    apply(drule_tac a = a and aa = aa and ba = ba in ci_cn_md_def,
-      simp, simp add: g, simp)
-    apply(subgoal_tac "length gs = aa \<and> aa < ba \<and> ba \<le> pstr", arith,
-          simp add: para_pattern)
-    apply(insert g, auto intro: min_max.le_supI2)
-    done
-  from k1 show 
-    "\<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3
-    * length gs + 3 * n, 0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ 
-     suf_lm) aprog stp =
-    ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs +
-      3 * n, ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
-  proof(erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
-    fix ap bp apa cp
-    assume "aprog = ap [+] bp [+] cp \<and> length ap = 
-      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs +
-                  3 * n \<and> bp = mv_boxes pstr 0 (length ys)"
-    from this and k2 show "?thesis"
-      apply(simp)
-      apply(drule_tac as = 
-        "(\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs +
-        3 * n" and ap = ap and cp = cp in abc_append_exc1, auto)
-      apply(rule_tac x = stp in exI, simp add: h)
-      using h
-      apply(simp)
-      done
-  qed
-qed
-
-thm rec_ci.simps 
-
-lemma calc_f_prog_ex: 
-  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
-    rec_ci f = (a, aa, ba);
-    Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
-                   (map rec_ci (f # gs)))) = pstr\<rbrakk>
-   \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
-  length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
-                                6 *length gs + 3 * n \<and> bp = a"
-apply(simp add: rec_ci.simps)
-apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba
-     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
-     mv_boxes 0 (Suc (max (Suc n) (Max (insert ba 
-            (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n [+]
-     mv_boxes (max (Suc n) (Max (insert ba 
-      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs)" in exI,
-     simp add: cn_merge_gs_len)
-apply(rule_tac x = "recursive.mv_box aa (max (Suc n) (Max (insert ba 
-     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
-     empty_boxes (length gs) [+] recursive.mv_box (max (Suc n) (
-     Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
-     mv_boxes (Suc (max (Suc n) (Max (insert ba 
-     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI,
-  auto simp: abc_append_commute)
-done
-
-lemma calc_cn_f:
-  assumes ind:
-  "\<And>x aprog a_md rs_pos rs suf_lm lm.
-  \<lbrakk>x \<in> set (f # gs);
-  rec_ci x = (aprog, rs_pos, a_md); 
-  rec_calc_rel x lm rs\<rbrakk>
-  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
-  (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-  and h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
-  "rec_calc_rel (Cn n f gs) lm rs"
-  "length ys = length gs"
-  "rec_calc_rel f ys rs"
-  "length lm = n"
-  "rec_ci f = (a, aa, ba)" 
-  and p: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
-                                (map rec_ci (f # gs))))"
-  shows "\<exists>stp. abc_steps_l   
-  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n,
-  ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
-  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 
-                3 * n + length a,
-  ys @ rs # 0\<up>pstr @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
-proof -
-  from h have k1: 
-    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
-    length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
-    6 *length gs + 3 * n \<and> bp = a"
-    by(drule_tac calc_f_prog_ex, auto)
-  from h and k1 show "?thesis"
-  proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
-    fix ap bp apa cp
-    assume
-      "aprog = ap [+] bp [+] cp \<and> 
-      length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
-      6 * length gs + 3 * n \<and> bp = a"
-    from h and this show "?thesis"
-      apply(simp, rule_tac abc_append_exc1, simp_all)
-      apply(insert ind[of f "a" aa ba ys rs 
-        "0\<up>(pstr - ba + length gs) @ 0 # lm @ 
-        0\<up>(a_md - Suc (pstr + length gs + n)) @ suf_lm"], simp)
-      apply(subgoal_tac "ba > aa \<and> aa = length gs\<and> pstr \<ge> ba", simp)
-      apply(simp add: exponent_add_iff)
-      apply(case_tac pstr, simp add: p)
-      apply(simp only: exp_suc, simp)
-      apply(rule conjI, rule ci_ad_ge_paras, simp, rule conjI)
-      apply(subgoal_tac "length ys = aa", simp,
-        rule para_pattern, simp, simp)
-      apply(insert p, simp)
-      apply(auto intro: min_max.le_supI2)
-      done
-  qed
-qed
-(*
-lemma [simp]: 
-  "\<lbrakk>pstr + length ys + n \<le> a_md; ys \<noteq> []\<rbrakk> \<Longrightarrow> 
-                          pstr < a_md + length suf_lm"
-apply(case_tac "length ys", simp)
-apply(arith)
-done
-*)
-
-lemma [simp]: 
-  "pstr > length ys 
-  \<Longrightarrow> (ys @ rs # 0\<up>pstr @ lm @
-  0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) ! pstr = (0::nat)"
-apply(simp add: nth_append)
-done
-
-(*
-lemma [simp]: "\<lbrakk>length ys < pstr; pstr - length ys = Suc x\<rbrakk>
-  \<Longrightarrow> pstr - Suc (length ys) = x"
-by arith
-*)
-
-lemma [simp]: "pstr > length ys \<Longrightarrow> 
-      (ys @ rs # 0\<up>pstr @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)
-                                         [pstr := rs, length ys := 0] =
-       ys @ 0\<up>(pstr - length ys) @ (rs::nat) # 0\<up>length ys @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm"
-apply(auto simp: list_update_append)
-apply(case_tac "pstr - length ys",simp_all)
-using list_update_length[of 
-  "0\<up>(pstr - Suc (length ys))" "0" "0\<up>length ys @ lm @ 
-  0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm" rs]
-apply(simp only: exponent_cons_iff exponent_add_iff, simp)
-apply(subgoal_tac "pstr - Suc (length ys) = nat", simp, simp)
-done
-
-lemma save_rs': 
-  "\<lbrakk>pstr > length ys\<rbrakk>
-  \<Longrightarrow> \<exists>stp. abc_steps_l (0, ys @ rs # 0\<up>pstr @ lm @ 
-  0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) 
-  (recursive.mv_box (length ys) pstr) stp =
-  (3, ys @ 0\<up>(pstr - (length ys)) @ rs # 
-  0\<up>length ys  @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
-apply(insert mv_box_ex[of "length ys" pstr 
-  "ys @ rs # 0\<up>pstr @ lm @ 0\<up>(a_md - Suc(pstr + length ys + n)) @ suf_lm"], 
-  simp)
-done
-
-
-lemma save_rs_prog_ex:
-  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
-  rec_ci f = (a, aa, ba);
-  Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
-                        (map rec_ci (f # gs)))) = pstr\<rbrakk>
-  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
-  length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
-              6 *length gs + 3 * n + length a
-  \<and> bp = mv_box aa pstr"
-apply(simp add: rec_ci.simps)
-apply(rule_tac x =
-  "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba 
-   (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
-   [+] mv_boxes 0 (Suc (max (Suc n) (Max (insert ba 
-   (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n [+]
-   mv_boxes (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
-    0 (length gs) [+] a" 
-  in exI, simp add: cn_merge_gs_len)
-apply(rule_tac x = 
-  "empty_boxes (length gs) [+]
-   recursive.mv_box (max (Suc n) (Max (insert ba 
-    (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
-   mv_boxes (Suc (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))
-    + length gs)) 0 n" in exI, 
-  auto simp: abc_append_commute)
-done
-
-lemma save_rs:  
-  assumes h: 
-  "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
-  "rec_calc_rel (Cn n f gs) lm rs"
-  "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
-  "length ys = length gs" 
-  "rec_calc_rel f ys rs" 
-  "rec_ci f = (a, aa, ba)"  
-  "length lm = n"
-  and pdef: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
-                                            (map rec_ci (f # gs))))"
-  shows "\<exists>stp. abc_steps_l
-           ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs
-          + 3 * n + length a, ys @ rs # 0\<up>pstr @ lm @
-             0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
-  ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs 
-  + 3 * n + length a + 3,
-  ys @ 0\<up>(pstr - length ys) @ rs # 0\<up>length ys @ lm @ 
-                               0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
-proof -
-  thm rec_ci.simps
-  from h and pdef have k1: 
-    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
-    length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
-    6 *length gs + 3 * n + length a \<and> bp = mv_box (length ys) pstr "
-    apply(subgoal_tac "length ys = aa")
-    apply(drule_tac a = a and aa = aa and ba = ba in save_rs_prog_ex, 
-      simp, simp, simp)
-    by(rule_tac para_pattern, simp, simp)
-  from k1 show 
-    "\<exists>stp. abc_steps_l
-    ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n
-    + length a, ys @ rs # 0\<up>pstr @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) 
-    @ suf_lm) aprog stp =
-    ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n
-    + length a + 3, ys @ 0\<up>(pstr - length ys) @ rs # 
-    0\<up>length ys @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
-  proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
-    fix ap bp apa cp
-    assume "aprog = ap [+] bp [+] cp \<and> length ap = 
-      (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 
-      3 * n + length a \<and> bp = recursive.mv_box (length ys) pstr"
-    thus"?thesis"
-      apply(simp, rule_tac abc_append_exc1, simp_all)
-      apply(rule_tac save_rs', insert h)
-      apply(subgoal_tac "length gs = aa \<and> pstr \<ge> ba \<and> ba > aa",
-            arith)
-      apply(simp add: para_pattern, insert pdef, auto)
-      apply(rule_tac min_max.le_supI2, simp)
-      done
-  qed
-qed
-
-lemma [simp]: "length (empty_boxes n) = 2*n"
-apply(induct n, simp, simp)
-done
-
-lemma mv_box_step_ex: "length lm = n \<Longrightarrow> 
-      \<exists>stp. abc_steps_l (0, lm @ Suc x # suf_lm) [Dec n 2, Goto 0] stp
-  = (0, lm @ x # suf_lm)"
-apply(rule_tac x = "Suc (Suc 0)" in exI, 
-  simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps 
-         abc_lm_v.simps abc_lm_s.simps nth_append list_update_append)
-done
-
-lemma mv_box_ex': 
-  "\<lbrakk>length lm = n\<rbrakk> \<Longrightarrow> 
-  \<exists> stp. abc_steps_l (0, lm @ x # suf_lm) [Dec n 2, Goto 0] stp =
-  (Suc (Suc 0), lm @ 0 # suf_lm)"
-apply(induct x)
-apply(rule_tac x = "Suc 0" in exI, 
-  simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps
-            abc_lm_v.simps nth_append abc_lm_s.simps, simp)
-apply(drule_tac x = x and suf_lm = suf_lm in mv_box_step_ex, 
-      erule_tac exE, erule_tac exE)
-apply(rule_tac x = "stpa + stp" in exI, simp add: abc_steps_add)
-done
-
-lemma [simp]: "drop n lm = a # list \<Longrightarrow> list = drop (Suc n) lm"
-apply(induct n arbitrary: lm a list, simp)
-apply(case_tac "lm", simp, simp)
-done
-
-lemma empty_boxes_ex: "\<lbrakk>length lm \<ge> n\<rbrakk>
-     \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm) (empty_boxes n) stp = 
-                                          (2*n, 0\<up>n @ drop n lm)"
-apply(induct n, simp, simp)
-apply(rule_tac abc_append_exc2, auto)
-apply(case_tac "drop n lm", simp, simp)
-proof -
-  fix n stp a list
-  assume h: "Suc n \<le> length lm"  "drop n lm = a # list"
-  thus "\<exists>bstp. abc_steps_l (0, 0\<up>n @ a # list) [Dec n 2, Goto 0] bstp =
-                       (Suc (Suc 0), 0 # 0\<up>n @ drop (Suc n) lm)"
-    apply(insert mv_box_ex'[of "0\<up>n" n a list], simp, erule_tac exE)
-    apply(rule_tac x = stp in exI, simp, simp only: exponent_cons_iff)
-    apply(simp add:exp_ind del: replicate.simps)
-    done
-qed
-
-
-lemma mv_box_paras_prog_ex:
-  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
-  rec_ci f = (a, aa, ba); 
-  Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
-                    (map rec_ci (f # gs)))) = pstr\<rbrakk>
-  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
-  length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
-  6 *length gs + 3 * n + length a + 3 \<and> bp = empty_boxes (length gs)"
-apply(simp add: rec_ci.simps)
-apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) 
-    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
-    mv_boxes 0 (Suc (max (Suc n) (Max 
-     (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n
-    [+] mv_boxes (max (Suc n) (Max (insert ba 
-    (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
-     a [+] recursive.mv_box aa (max (Suc n) 
-   (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))" 
-    in exI, simp add: cn_merge_gs_len)
-apply(rule_tac x = " recursive.mv_box (max (Suc n) (Max (insert ba
-     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
-     mv_boxes (Suc (max (Suc n) (Max (insert ba 
-     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI, 
-  auto simp: abc_append_commute)
-done
-
-lemma mv_box_paras: 
- assumes h: 
-  "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
-  "rec_calc_rel (Cn n f gs) lm rs" 
-  "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
-  "length ys = length gs" 
-  "rec_calc_rel f ys rs" 
-  "rec_ci f = (a, aa, ba)" 
-  and pdef: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
-                                             (map rec_ci (f # gs))))"
-  and starts: "ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
-                              6 * length gs + 3 * n + length a + 3"
-  shows "\<exists>stp. abc_steps_l
-           (ss, ys @ 0\<up>(pstr - length ys) @ rs # 0\<up>length ys 
-               @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
-   (ss + 2 * length gs, 0\<up>pstr @ rs # 0\<up>length ys  @ lm @ 
-                                0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
-proof -
-  from h and pdef and starts have k1: 
-    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
-    length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
-                               6 *length gs + 3 * n + length a + 3
-    \<and> bp = empty_boxes (length ys)"
-    by(drule_tac mv_box_paras_prog_ex, auto)
-  from h have j1: "aa < ba"
-    by(simp add: ci_ad_ge_paras)
-  from h have j2: "length gs = aa"
-    by(drule_tac f = f in para_pattern, simp, simp)
-  from h and pdef have j3: "ba \<le> pstr"
-    apply simp 
-    apply(rule_tac min_max.le_supI2, simp)
-    done
-  from k1 show "?thesis"
-  proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
-    fix ap bp apa cp
-    assume "aprog = ap [+] bp [+] cp \<and> 
-      length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
-      6 * length gs + 3 * n + length a + 3 \<and> 
-      bp = empty_boxes (length ys)"
-    thus"?thesis"
-      apply(simp, rule_tac abc_append_exc1, simp_all add: starts h)
-      apply(insert empty_boxes_ex[of 
-        "length gs" "ys @ 0\<up>(pstr - (length gs)) @ rs #
-        0\<up>length gs @ lm @ 0\<up>(a_md - Suc (pstr + length gs + n)) @ suf_lm"], 
-        simp add: h)
-      apply(erule_tac exE, rule_tac x = stp in exI, 
-        simp add: replicate.simps[THEN sym]
-        replicate_add[THEN sym] del: replicate.simps)
-      apply(subgoal_tac "pstr >(length gs)", simp)
-      apply(subgoal_tac "ba > aa \<and> length gs = aa \<and> pstr \<ge> ba", simp)
-      apply(simp add: j1 j2 j3)
-      done     
-  qed
-qed
-
-lemma restore_rs_prog_ex:
-  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
-  rec_ci f = (a, aa, ba);
-  Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
-  (map rec_ci (f # gs)))) = pstr;
-  ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
-  8 * length gs + 3 * n + length a + 3\<rbrakk>
-  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
-                                           bp = mv_box pstr n"
-apply(simp add: rec_ci.simps)
-apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) 
-      (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
-      mv_boxes 0 (Suc (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n)
-        \<circ> rec_ci) ` set gs))) + length gs)) n [+]
-     mv_boxes (max (Suc n) (Max (insert ba 
-      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
-     a [+] recursive.mv_box aa (max (Suc n)
-       (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
-     empty_boxes (length gs)" in exI, simp add: cn_merge_gs_len)
-apply(rule_tac x = "mv_boxes (Suc (max (Suc n) 
-       (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) 
-        + length gs)) 0 n" 
-  in exI, auto simp: abc_append_commute)
-done
-
-lemma exp_add: "a\<up>(b+c) = a\<up>b @ a\<up>c"
-apply(simp add:replicate_add)
-done
-
-lemma [simp]: "n < pstr \<Longrightarrow> (0\<up>pstr)[n := rs] @ [0::nat] = 0\<up>n @ rs # 0\<up>(pstr - n)"
-using list_update_length[of "0\<up>n" "0::nat" "0\<up>(pstr - Suc n)" rs]
-apply(simp add: replicate_Suc[THEN sym] exp_add[THEN sym] exp_suc[THEN sym])
-done
-
-lemma restore_rs:
-  assumes h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
-  "rec_calc_rel (Cn n f gs) lm rs" 
-  "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
-  "length ys = length gs"
-  "rec_calc_rel f ys rs" 
-  "rec_ci f = (a, aa, ba)" 
-  and pdef: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
-                                        (map rec_ci (f # gs))))"
-  and starts: "ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
-                              8 * length gs + 3 * n + length a + 3" 
-  shows "\<exists>stp. abc_steps_l
-           (ss, 0\<up>pstr @ rs # 0\<up>length ys  @ lm @
-                    0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
-  (ss + 3, 0\<up>n @ rs # 0\<up>(pstr - n) @ 0\<up>length ys  @ lm @ 
-                                   0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
-proof -
- from h and pdef and starts have k1:
-   "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
-                                            bp = mv_box pstr n"
-   by(drule_tac restore_rs_prog_ex, auto)
- from k1 show "?thesis"
- proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
-   fix ap bp apa cp
-   assume "aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
-                                 bp = recursive.mv_box pstr n"
-   thus"?thesis"
-     apply(simp, rule_tac abc_append_exc1, simp_all add: starts h)
-     apply(insert mv_box_ex[of pstr n "0\<up>pstr @ rs # 0\<up>length gs @
-                     lm @ 0\<up>(a_md - Suc (pstr + length gs + n)) @ suf_lm"], simp)
-     apply(subgoal_tac "pstr > n", simp)
-     apply(erule_tac exE, rule_tac x = stp in exI, 
-                         simp add: nth_append list_update_append)
-     apply(simp add: pdef)
-     done
-  qed
-qed
-
-lemma [simp]:"xs \<noteq> [] \<Longrightarrow> length xs \<ge> Suc 0"
-by(case_tac xs, auto)
-
-lemma  [simp]: "n < max (Suc n) (max ba (Max (((\<lambda>(aprog, p, n). n) o 
-                                                  rec_ci) ` set gs)))"
-by(simp)
-
-lemma restore_paras_prog_ex: 
-  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
-  rec_ci f = (a, aa, ba);
-  Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
-                          (map rec_ci (f # gs)))) = pstr;
-  ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
-                         8 * length gs + 3 * n + length a + 6\<rbrakk>
-  \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
-                      bp = mv_boxes (pstr + Suc (length gs)) (0::nat) n"
-apply(simp add: rec_ci.simps)
-apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) 
-      (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
-      [+] mv_boxes 0 (Suc (max (Suc n) 
-       (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) 
-     + length gs)) n [+] mv_boxes (max (Suc n) 
-    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
-     a [+] recursive.mv_box aa (max (Suc n) 
-      (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
-     empty_boxes (length gs) [+]
-     recursive.mv_box (max (Suc n) (Max (insert ba 
-     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n" in exI, simp add: cn_merge_gs_len)
-apply(rule_tac x = "[]" in exI, auto simp: abc_append_commute)
-done
-
-lemma restore_paras: 
-  assumes h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
-  "rec_calc_rel (Cn n f gs) lm rs" 
-  "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
-  "length ys = length gs"
-  "rec_calc_rel f ys rs" 
-  "rec_ci f = (a, aa, ba)"
-  and pdef: 
-  "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
-                         (map rec_ci (f # gs))))"
-  and starts: "ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
-                              8 * length gs + 3 * n + length a + 6" 
-  shows "\<exists>stp. abc_steps_l (ss, 0\<up>n @ rs # 0\<up>(pstr - n+ length ys) @
-                         lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)
-  aprog stp = (ss + 3 * n, lm @ rs # 0\<up>(a_md - Suc n) @ suf_lm)"
-proof -
-  thm rec_ci.simps
-  from h and pdef and starts have k1:
-    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and>
-                     bp = mv_boxes (pstr + Suc (length gs)) (0::nat) n"
-    by(drule_tac restore_paras_prog_ex, auto)
-  from k1 show "?thesis"
-  proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
-    fix ap bp apa cp
-    assume "aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
-                              bp = mv_boxes (pstr + Suc (length gs)) 0 n"
-    thus"?thesis"
-      apply(simp, rule_tac abc_append_exc1, simp_all add: starts h)
-      apply(insert mv_boxes_ex2[of n "pstr + Suc (length gs)" 0 "[]" 
-        "rs # 0\<up>(pstr - n + length gs)" "lm" 
-        "0\<up>(a_md - Suc (pstr + length gs + n)) @ suf_lm"], simp)
-      apply(subgoal_tac "pstr > n \<and> 
-        a_md > pstr + length gs + n \<and> length lm = n" , simp add: exponent_add_iff h)
-      using h pdef
-      apply(simp)     
-      apply(frule_tac a = a and 
-        aa = aa and ba = ba in ci_cn_md_def, simp, simp)
-      apply(subgoal_tac "length lm = rs_pos",
-        simp add: ci_cn_para_eq, erule_tac para_pattern, simp)
-      done
-  qed
-qed
-
-lemma ci_cn_length:
-  "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
-  rec_calc_rel (Cn n f gs) lm rs;
-  rec_ci f = (a, aa, ba)\<rbrakk>
-  \<Longrightarrow> length aprog = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
-                             8 * length gs + 6 * n + length a + 6"
-apply(simp add: rec_ci.simps, auto simp: cn_merge_gs_len)
-done
-
-lemma  cn_case: 
-  assumes ind:
-  "\<And>x aprog a_md rs_pos rs suf_lm lm.
-  \<lbrakk>x \<in> set (f # gs);
-  rec_ci x = (aprog, rs_pos, a_md);
-  rec_calc_rel x lm rs\<rbrakk>
-  \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
-               (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-  and h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
-         "rec_calc_rel (Cn n f gs) lm rs"
-  shows "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp 
-  = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-apply(insert h, case_tac "rec_ci f",  rule_tac calc_cn_reverse, simp)
-proof -
-  fix a b c ys
-  let ?pstr = "Max (set (Suc n # c # (map (\<lambda>(aprog, p, n). n) 
-                                         (map rec_ci (f # gs)))))"  
-  let ?gs_len = "listsum (map (\<lambda> (ap, pos, n). length ap) 
-                                                (map rec_ci (gs)))"
-  assume g: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
-    "rec_calc_rel (Cn n f gs) lm rs"
-    "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)" 
-    "length ys = length gs" 
-    "rec_calc_rel f ys rs"
-    "n = length lm"
-    "rec_ci f = (a, b, c)"  
-  hence k1:
-    "\<exists> stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
-    (?gs_len + 3 * length gs, lm @ 0\<up>(?pstr - n) @ ys @
-                               0\<up>(a_md - ?pstr - length ys) @ suf_lm)"	
-    apply(rule_tac a = a and aa = b and ba = c in cn_calc_gs)
-    apply(rule_tac ind, auto)
-    done  
-  thm rec_ci.simps
-  from g have k2: 
-    "\<exists> stp. abc_steps_l (?gs_len + 3 * length gs,  lm @ 
-        0\<up>(?pstr - n) @ ys @ 0\<up>(a_md - ?pstr - length ys) @ suf_lm) aprog stp = 
-    (?gs_len + 3 * length gs + 3 * n, 0\<up>?pstr @ ys @ 0 # lm @ 
-                              0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm)"
-    thm save_paras
-    apply(erule_tac ba = c in save_paras, auto intro: ci_cn_para_eq)
-    done
-  from g have k3: 
-    "\<exists> stp. abc_steps_l (?gs_len + 3 * length gs + 3 * n,
-    0\<up>?pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm) aprog stp =
-    (?gs_len + 6 * length gs + 3 * n,  
-           ys @ 0\<up>?pstr @ 0 # lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm)"
-    apply(erule_tac ba = c in reset_new_paras, 
-          auto intro: ci_cn_para_eq)
-    using para_pattern[of f a b c ys rs]
-    apply(simp)
-    done
-  from g have k4: 
-    "\<exists>stp. abc_steps_l  (?gs_len + 6 * length gs + 3 * n,  
-    ys @ 0\<up>?pstr @ 0 # lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm) aprog stp =
-    (?gs_len + 6 * length gs + 3 * n + length a, 
-   ys @ rs # 0\<up>?pstr  @ lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm)"
-    apply(rule_tac ba = c in calc_cn_f, rule_tac ind, auto)
-    done
-thm rec_ci.simps
-  from g h have k5:
-    "\<exists> stp. abc_steps_l (?gs_len + 6 * length gs + 3 * n + length a,
-    ys @ rs # 0\<up>?pstr @ lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm)
-    aprog stp =
-    (?gs_len + 6 * length gs + 3 * n + length a + 3,
-    ys @ 0\<up>(?pstr - length ys) @ rs # 0\<up>length ys @ lm @ 
-    0\<up>(a_md  - Suc (?pstr + length ys + n)) @ suf_lm)"
-    apply(rule_tac save_rs, auto simp: h)
-    done
-  from g have k6: 
-    "\<exists> stp. abc_steps_l (?gs_len + 6 * length gs + 3 * n + 
-    length a + 3, ys @ 0\<up>(?pstr - length ys) @ rs # 0\<up>length ys @ lm @ 
-    0\<up>(a_md  - Suc (?pstr + length ys + n)) @ suf_lm) 
-    aprog stp =
-    (?gs_len + 8 * length gs + 3 *n + length a + 3,
-    0\<up>?pstr @ rs # 0\<up>length ys @ lm @ 
-                        0\<up>(a_md -Suc (?pstr + length ys + n)) @ suf_lm)"
-    apply(drule_tac suf_lm = suf_lm in mv_box_paras, auto)
-    apply(rule_tac x = stp in exI, simp)
-    done
-  from g have k7: 
-    "\<exists> stp. abc_steps_l (?gs_len + 8 * length gs + 3 *n + 
-    length a + 3, 0\<up>?pstr  @ rs # 0\<up>length ys @ lm @ 
-    0\<up>(a_md -Suc (?pstr + length ys + n)) @ suf_lm) aprog stp =
-    (?gs_len + 8 * length gs + 3 * n + length a + 6, 
-    0\<up>n @ rs # 0\<up>(?pstr  - n) @ 0\<up>length ys @ lm @
-                        0\<up>(a_md -Suc (?pstr + length ys + n)) @ suf_lm)"
-    apply(drule_tac suf_lm = suf_lm in restore_rs, auto)
-    apply(rule_tac x = stp in exI, simp)
-    done
-  from g have k8: "\<exists> stp. abc_steps_l (?gs_len + 8 * length gs + 
-    3 * n + length a + 6,
-    0\<up>n @ rs # 0\<up>(?pstr  - n) @ 0\<up>length ys @ lm @
-                      0\<up>(a_md -Suc (?pstr + length ys + n)) @ suf_lm) aprog stp =
-    (?gs_len + 8 * length gs + 6 * n + length a + 6,
-                           lm @ rs # 0\<up>(a_md - Suc n) @ suf_lm)"
-    apply(drule_tac suf_lm = suf_lm in restore_paras, auto)
-    apply(simp add: exponent_add_iff)
-    apply(rule_tac x = stp in exI, simp)
-    done
-  from g have j1: 
-    "length aprog = ?gs_len + 8 * length gs + 6 * n + length a + 6"
-    by(drule_tac a = a and aa = b and ba = c in ci_cn_length,
-      simp, simp, simp)
-  from g have j2: "rs_pos = n"
-    by(simp add: ci_cn_para_eq)
-  from k1 and k2 and k3 and k4 and k5 and k6 and k7 and k8
-    and j1 and j2 show 
-    "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
-    (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
-    apply(auto)
-    apply(rule_tac x = "stp + stpa + stpb + stpc +
-      stpd + stpe + stpf + stpg" in exI, simp add: abc_steps_add)
-    done
-qed
-
-text {*
-  Correctness of the complier (terminate case), which says if the execution of 
-  a recursive function @{text "recf"} terminates and gives result, then 
-  the Abacus program compiled from @{text "recf"} termintes and gives the same result.
-  Additionally, to facilitate induction proof, we append @{text "anything"} to the
-  end of Abacus memory.
-*}
-
-lemma recursive_compile_correct:
-  "\<lbrakk>rec_ci recf = (ap, arity, fp);
-    rec_calc_rel recf args r\<rbrakk>
-  \<Longrightarrow> (\<exists> stp. (abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp) = 
-              (length ap, args@[r]@0\<up>(fp - arity - 1) @ anything))"
-apply(induct arbitrary: ap fp arity r anything args
-  rule: rec_ci.induct)
-prefer 5
-proof(case_tac "rec_ci g", case_tac "rec_ci f", simp)
-  fix n f g ap fp arity r anything args  a b c aa ba ca
-  assume f_ind:
-    "\<And>ap fp arity r anything args.
-    \<lbrakk>aa = ap \<and> ba = arity \<and> ca = fp; rec_calc_rel f args r\<rbrakk> \<Longrightarrow> 
-    \<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
-    (length ap, args @ r # 0\<up>(fp - Suc arity) @ anything)"
-    and g_ind:
-    "\<And>x xa y xb ya ap fp arity r anything args.
-    \<lbrakk>x = (aa, ba, ca); xa = aa \<and> y = (ba, ca); xb = ba \<and> ya = ca; 
-    a = ap \<and> b = arity \<and> c = fp; rec_calc_rel g args r\<rbrakk>
-    \<Longrightarrow> \<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
-    (length ap, args @ r # 0\<up>(fp - Suc arity) @ anything)"
-    and h: "rec_ci (Pr n f g) = (ap, arity, fp)" 
-    "rec_calc_rel (Pr n f g) args r" 
-    "rec_ci g = (a, b, c)" 
-    "rec_ci f = (aa, ba, ca)"
-  from h have nf_ind: 
-    "\<And> args r anything. rec_calc_rel f args r \<Longrightarrow> 
-    \<exists>stp. abc_steps_l (0, args @ 0\<up>(ca - ba) @ anything) aa stp = 
-    (length aa, args @ r # 0\<up>(ca - Suc ba) @ anything)"
-    and ng_ind: 
-    "\<And> args r anything. rec_calc_rel g args r \<Longrightarrow> 
-    \<exists>stp. abc_steps_l (0, args @ 0\<up>(c - b) @ anything) a stp = 
-         (length a, args @ r # 0\<up>(c - Suc b)  @ anything)"
-    apply(insert f_ind[of aa ba ca], simp)
-    apply(insert g_ind[of "(aa, ba, ca)" aa "(ba, ca)" ba ca a b c],
-      simp)
-    done
-  from nf_ind and ng_ind and h show 
-    "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp = 
-    (length ap, args @ r # 0\<up>(fp - Suc arity) @ anything)"
-    apply(auto intro: nf_ind ng_ind pr_case)
-    done
-next
-  fix ap fp arity r anything args
-  assume h:
-    "rec_ci z = (ap, arity, fp)" "rec_calc_rel z args r"
-  thus "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
-    (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
-    by (rule_tac z_case)    
-next
-  fix ap fp arity r anything args
-  assume h: 
-    "rec_ci s = (ap, arity, fp)" 
-    "rec_calc_rel s args r"
-  thus 
-    "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
-    (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
-    by(erule_tac s_case, simp)
-next
-  fix m n ap fp arity r anything args
-  assume h: "rec_ci (id m n) = (ap, arity, fp)" 
-    "rec_calc_rel (id m n) args r"
-  thus "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp 
-    = (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
-    by(erule_tac id_case)
-next
-  fix n f gs ap fp arity r anything args
-  assume ind: "\<And>x ap fp arity r anything args.
-    \<lbrakk>x \<in> set (f # gs); 
-    rec_ci x = (ap, arity, fp); 
-    rec_calc_rel x args r\<rbrakk>
-    \<Longrightarrow> \<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
-    (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
-  and h: "rec_ci (Cn n f gs) = (ap, arity, fp)" 
-    "rec_calc_rel (Cn n f gs) args r"
-  from h show
-    "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) 
-       ap stp = (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
-    apply(rule_tac cn_case, rule_tac ind, auto)
-    done
-next
-  fix n f ap fp arity r anything args
-  assume ind:
-    "\<And>ap fp arity r anything args.
-    \<lbrakk>rec_ci f = (ap, arity, fp); rec_calc_rel f args r\<rbrakk> \<Longrightarrow> 
-    \<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
-    (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
-  and h: "rec_ci (Mn n f) = (ap, arity, fp)" 
-    "rec_calc_rel (Mn n f) args r"
-  from h show 
-    "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp = 
-              (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
-    apply(rule_tac mn_case, rule_tac ind, auto)
-    done    
-qed
-
-lemma abc_append_uhalt1:
-  "\<lbrakk>\<forall> stp. (\<lambda> (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp);
-    p = ap [+] bp [+] cp\<rbrakk>
-  \<Longrightarrow> \<forall> stp. (\<lambda> (ss, e). ss < length p) 
-                     (abc_steps_l (length ap, lm) p stp)"
-apply(auto)
-apply(erule_tac x = stp in allE, auto)
-apply(frule_tac ap = ap and cp = cp in abc_append_state_in_exc, auto)
-done
-
-
-lemma abc_append_unhalt2:
-  "\<lbrakk>abc_steps_l (0, am) ap stp = (length ap, lm); bp \<noteq> [];
-  \<forall> stp. (\<lambda> (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp);
-  p = ap [+] bp [+] cp\<rbrakk>
-  \<Longrightarrow> \<forall> stp. (\<lambda> (ss, e). ss < length p) (abc_steps_l (0, am) p stp)"
-proof -
-  assume h: 
-    "abc_steps_l (0, am) ap stp = (length ap, lm)" 
-    "bp \<noteq> []"
-    "\<forall> stp. (\<lambda> (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp)"
-    "p = ap [+] bp [+] cp"
-  have "\<exists> stp. (abc_steps_l (0, am) p stp) = (length ap, lm)"
-    using h
-    thm abc_add_exc1
-    apply(simp add: abc_append.simps)
-    apply(rule_tac abc_add_exc1, auto)
-    done
-  from this obtain stpa where g1: 
-    "(abc_steps_l (0, am) p stpa) = (length ap, lm)" ..
-  moreover have g2: "\<forall> stp. (\<lambda> (ss, e). ss < length p) 
-                          (abc_steps_l (length ap, lm) p stp)"
-    using h
-    apply(erule_tac abc_append_uhalt1, simp)
-    done
-  moreover from g1 and g2 have
-    "\<forall> stp. (\<lambda> (ss, e). ss < length p) 
-                    (abc_steps_l (0, am) p (stpa + stp))"
-    apply(simp add: abc_steps_add)
-    done
-  thus "\<forall> stp. (\<lambda> (ss, e). ss < length p) 
-                           (abc_steps_l (0, am) p stp)"
-    apply(rule_tac allI, auto)
-    apply(case_tac "stp \<ge>  stpa")
-    apply(erule_tac x = "stp - stpa" in allE, simp)
-  proof - 	
-    fix stp a b
-    assume g3:  "abc_steps_l (0, am) p stp = (a, b)" 
-                "\<not> stpa \<le> stp"
-    thus "a < length p"
-      using g1 h
-      apply(case_tac "a < length p", simp, simp)
-      apply(subgoal_tac "\<exists> d. stpa = stp + d")
-      using  abc_state_keep[of p a b "stpa - stp"]
-      apply(erule_tac exE, simp add: abc_steps_add)
-      apply(rule_tac x = "stpa - stp" in exI, simp)
-      done
-  qed
-qed
-
-text {*
-  Correctness of the complier (non-terminating case for Mn). There are many cases when a 
-  recursive function does not terminate. For the purpose of Uiversal Turing Machine, we only 
-  need to prove the case for @{text "Mn"} and @{text "Cn"}.
-  This lemma is for @{text "Mn"}. For @{text "Mn n f"}, this lemma describes what 
-  happens when @{text "f"} always terminates but always does not return zero, so that
-  @{text "Mn"} has to loop forever.
-  *}
-
-lemma Mn_unhalt:
-  assumes mn_rf: "rf = Mn n f"
-  and compiled_mnrf: "rec_ci rf = (aprog, rs_pos, a_md)"
-  and compiled_f: "rec_ci f = (aprog', rs_pos', a_md')"
-  and args: "length lm = n"
-  and unhalt_condition: "\<forall> y. (\<exists> rs. rec_calc_rel f (lm @ [y]) rs \<and> rs \<noteq> 0)"
-  shows "\<forall> stp. case abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm)
-               aprog stp of (ss, e) \<Rightarrow> ss < length aprog"
-  using mn_rf compiled_mnrf compiled_f args unhalt_condition
-proof(rule_tac allI)
-  fix stp
-  assume h: "rf = Mn n f" 
-            "rec_ci rf = (aprog, rs_pos, a_md)"
-            "rec_ci f = (aprog', rs_pos', a_md')" 
-            "\<forall>y. \<exists>rs. rec_calc_rel f (lm @ [y]) rs \<and> rs \<noteq> 0" "length lm = n"
-  thm mn_ind_step
-  have "\<exists>stpa \<ge> stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm) aprog stpa 
-         = (0, lm @ stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-  proof(induct stp, auto)
-    show "\<exists>stpa. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
-          aprog stpa = (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-      apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
-      done
-  next
-    fix stp stpa
-    assume g1: "stp \<le> stpa"
-      and g2: "abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
-                            aprog stpa
-               = (0, lm @ stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-    have "\<exists>rs. rec_calc_rel f (lm @ [stp]) rs \<and> rs \<noteq> 0"
-      using h
-      apply(erule_tac x = stp in allE, simp)
-      done
-    from this obtain rs where g3:
-      "rec_calc_rel f (lm @ [stp]) rs \<and> rs \<noteq> 0" ..
-    hence "\<exists> stpb. abc_steps_l (0, lm @ stp # 0\<up>(a_md - Suc rs_pos) @
-                     suf_lm) aprog stpb 
-      = (0, lm @ Suc stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-      using h
-      apply(rule_tac mn_ind_step)
-      apply(rule_tac recursive_compile_correct, simp, simp)
-    proof -
-      show "rec_ci f = ((aprog', rs_pos', a_md'))" using h by simp
-    next
-      show "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" using h by simp
-    next
-      show "rec_calc_rel f (lm @ [stp]) rs" using g3 by simp
-    next
-      show "0 < rs" using g3 by simp
-    next
-      show "Suc rs_pos < a_md"
-        using g3 h
-        apply(auto)
-        apply(frule_tac f = f in para_pattern, simp, simp)
-        apply(simp add: rec_ci.simps, auto)
-        apply(subgoal_tac "Suc (length lm) < a_md'")
-        apply(arith)
-        apply(simp add: ci_ad_ge_paras)
-        done
-    next
-      show "rs_pos' = Suc rs_pos"
-        using g3 h
-        apply(auto)
-        apply(frule_tac f = f in para_pattern, simp, simp)
-        apply(simp add: rec_ci.simps)
-        done
-    qed
-    thus "\<exists>stpa\<ge>Suc stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @
-                 suf_lm) aprog stpa 
-      = (0, lm @ Suc stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
-      using g2
-      apply(erule_tac exE)
-      apply(case_tac "stpb = 0", simp add: abc_steps_l.simps)
-      apply(rule_tac x = "stpa + stpb" in exI, simp add:
-        abc_steps_add)
-      using g1
-      apply(arith)
-      done
-  qed
-  from this obtain stpa where 
-    "stp \<le> stpa \<and> abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
-         aprog stpa = (0, lm @ stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)" ..
-  thus "case abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp
-    of (ss, e) \<Rightarrow> ss < length aprog"
-    apply(case_tac "abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog
-      stp", simp, case_tac "a \<ge> length aprog", 
-        simp, simp)
-    apply(subgoal_tac "\<exists> d. stpa = stp + d", erule_tac exE)
-    apply(subgoal_tac "lm @ 0\<up>(a_md - rs_pos) @ suf_lm = lm @ 0 # 
-             0\<up>(a_md - Suc rs_pos) @ suf_lm", simp add: abc_steps_add)
-    apply(frule_tac as = a and lm = b and stp = d in abc_state_keep, 
-          simp)
-    using h  
-    apply(simp add: rec_ci.simps, simp, 
-              simp only: replicate_Suc[THEN sym])
-    apply(case_tac rs_pos, simp, simp)
-    apply(rule_tac x = "stpa - stp" in exI, simp, simp)
-    done
-qed   
-
-lemma abc_append_cons_eq[intro!]: 
-  "\<lbrakk>ap = bp; cp = dp\<rbrakk> \<Longrightarrow> ap [+] cp = bp [+] dp"
-by simp 
-
-lemma cn_merge_gs_split: 
-  "\<lbrakk>i < length gs; rec_ci (gs!i) = (ga, gb, gc)\<rbrakk> \<Longrightarrow> 
-     cn_merge_gs (map rec_ci gs) p = 
-        cn_merge_gs (map rec_ci (take i gs)) p [+] ga [+] 
-       mv_box gb (p + i) [+] 
-      cn_merge_gs (map rec_ci (drop (Suc i) gs)) (p + Suc i)"
-apply(induct i arbitrary: gs p, case_tac gs, simp, simp)
-apply(case_tac gs, simp, case_tac "rec_ci a", 
-       simp add: abc_append_commute[THEN sym])
-done
-
-text {*
-  Correctness of the complier (non-terminating case for Mn). There are many cases when a 
-  recursive function does not terminate. For the purpose of Uiversal Turing Machine, we only 
-  need to prove the case for @{text "Mn"} and @{text "Cn"}.
-  This lemma is for @{text "Cn"}. For @{text "Cn f g1 g2 \<dots>gi, gi+1, \<dots> gn"}, this lemma describes what 
-  happens when every one of @{text "g1, g2, \<dots> gi"} terminates, but 
-  @{text "gi+1"} does not terminate, so that whole function @{text "Cn f g1 g2 \<dots>gi, gi+1, \<dots> gn"}
-  does not terminate.
-  *}
-
-lemma cn_gi_uhalt: 
-  assumes cn_recf: "rf = Cn n f gs"
-  and compiled_cn_recf: "rec_ci rf = (aprog, rs_pos, a_md)"
-  and args_length: "length lm = n"
-  and exist_unhalt_recf: "i < length gs" "gi = gs ! i"
-  and complied_unhalt_recf: "rec_ci gi = (ga, gb, gc)"  "gb = n"
-  and all_halt_before_gi: "\<forall> j < i. (\<exists> rs. rec_calc_rel (gs!j) lm rs)" 
-  and unhalt_condition: "\<And> slm. \<forall> stp. case abc_steps_l (0, lm @ 0\<up>(gc - gb) @ slm) 
-     ga stp of (se, e) \<Rightarrow> se < length ga"
-  shows " \<forall> stp. case abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suflm) aprog
-  stp of (ss, e) \<Rightarrow> ss < length aprog"
-  using cn_recf compiled_cn_recf args_length exist_unhalt_recf complied_unhalt_recf
-        all_halt_before_gi unhalt_condition
-proof(case_tac "rec_ci f", simp)
-  fix a b c
-  assume h1: "rf = Cn n f gs" 
-    "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
-    "length lm = n" 
-    "gi = gs ! i" 
-    "rec_ci (gs!i) = (ga, n, gc)" 
-    "gb = n" "rec_ci f = (a, b, c)"
-    and h2: "\<forall>j<i. \<exists>rs. rec_calc_rel (gs ! j) lm rs"
-    "i < length gs"
-  and ind:
-    "\<And> slm. \<forall> stp. case abc_steps_l (0, lm @ 0\<up>(gc - n) @ slm) ga stp of (se, e) \<Rightarrow> se < length ga"
-  have h3: "rs_pos = n"
-    using h1
-    by(rule_tac ci_cn_para_eq, simp)
-  let ?ggs = "take i gs"
-  have "\<exists> ys. (length ys = i \<and> 
-    (\<forall> k < i. rec_calc_rel (?ggs ! k) lm (ys ! k)))"
-    using h2
-    apply(induct i, simp, simp)
-    apply(erule_tac exE)
-    apply(erule_tac x = ia in allE, simp)
-    apply(erule_tac exE)
-    apply(rule_tac x = "ys @ [x]" in exI, simp add: nth_append, auto)
-    apply(subgoal_tac "k = length ys", simp, simp)
-    done
-  from this obtain ys where g1:
-    "(length ys = i \<and> (\<forall> k < i. rec_calc_rel (?ggs ! k)
-                        lm (ys ! k)))" ..
-  let ?pstr = "Max (set (Suc n # c # map (\<lambda>(aprog, p, n). n)
-    (map rec_ci (f # gs))))"
-  have "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suflm) 
-    (cn_merge_gs (map rec_ci ?ggs) ?pstr) stp =
-    (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) ?ggs) +
-    3 * length ?ggs, lm @ 0\<up>(?pstr - n) @ ys @ 0\<up>(a_md -(?pstr + length ?ggs)) @
-    suflm) "
-    apply(rule_tac  cn_merge_gs_ex)
-    apply(rule_tac  recursive_compile_correct, simp, simp)
-    using h1
-    apply(simp add: rec_ci.simps, auto)
-    using g1
-    apply(simp)
-    using h2 g1
-    apply(simp)
-    apply(rule_tac min_max.le_supI2)
-    apply(rule_tac Max_ge, simp, simp, rule_tac disjI2)
-    apply(subgoal_tac "aa \<in> set gs", simp)
-    using h2
-    apply(rule_tac A = "set (take i gs)" in subsetD, 
-      simp add: set_take_subset, simp)
-    done
-  thm cn_merge_gs.simps
-  from this obtain stpa where g2: 
-    "abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suflm) 
-    (cn_merge_gs (map rec_ci ?ggs) ?pstr) stpa =
-    (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) ?ggs) +
-    3 * length ?ggs, lm @ 0\<up>(?pstr - n) @ ys @ 0\<up>(a_md -(?pstr + length ?ggs)) @
-    suflm)" ..
-  moreover have 
-    "\<exists> cp. aprog = (cn_merge_gs
-    (map rec_ci ?ggs) ?pstr) [+] ga [+] cp"
-    using h1
-    apply(simp add: rec_ci.simps)
-    apply(rule_tac x = "mv_box n (?pstr + i) [+] 
-      (cn_merge_gs (map rec_ci (drop (Suc i) gs)) (?pstr + Suc i))
-      [+]mv_boxes 0 (Suc (max (Suc n) (Max (insert c 
-     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) +
-      length gs)) n [+] mv_boxes (max (Suc n) (Max (insert c 
-      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
-      a [+] recursive.mv_box b (max (Suc n) 
-      (Max (insert c (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
-     empty_boxes (length gs) [+] recursive.mv_box (max (Suc n) 
-      (Max (insert c (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
-      mv_boxes (Suc (max (Suc n) (Max (insert c 
-    (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI)
-    apply(simp add: abc_append_commute [THEN sym])
-    apply(auto)
-    using cn_merge_gs_split[of i gs ga "length lm" gc 
-      "(max (Suc (length lm))
-       (Max (insert c (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))"] 
-      h2
-    apply(simp)
-    done
-  from this obtain cp where g3: 
-    "aprog = (cn_merge_gs (map rec_ci ?ggs) ?pstr) [+] ga [+] cp" ..
-  show "\<forall> stp. case abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suflm) 
-    aprog stp of (ss, e) \<Rightarrow> ss < length aprog"
-  proof(rule_tac abc_append_unhalt2)
-    show "abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suflm) (
-      cn_merge_gs (map rec_ci ?ggs) ?pstr) stpa =
-         (length ((cn_merge_gs (map rec_ci ?ggs) ?pstr)),  
-          lm @ 0\<up>(?pstr - n) @ ys @ 0\<up>(a_md -(?pstr + length ?ggs)) @ suflm)"
-      using h3 g2
-      apply(simp add: cn_merge_gs_length)
-      done
-  next
-    show "ga \<noteq> []"
-      using h1
-      apply(simp add: rec_ci_not_null)
-      done
-  next
-    show "\<forall>stp. case abc_steps_l (0, lm @ 0\<up>(?pstr - n) @ ys
-      @ 0\<up>(a_md - (?pstr + length (take i gs))) @ suflm) ga  stp of
-          (ss, e) \<Rightarrow> ss < length ga"
-      using ind[of "0\<up>(?pstr - gc) @ ys @ 0\<up>(a_md - (?pstr + length (take i gs)))
-        @ suflm"]
-      apply(subgoal_tac "lm @ 0\<up>(?pstr - n) @ ys
-        @ 0\<up>(a_md - (?pstr + length (take i gs))) @ suflm
-                       = lm @ 0\<up>(gc - n) @ 
-        0\<up>(?pstr - gc) @ ys @ 0\<up>(a_md - (?pstr + length (take i gs))) @ suflm", simp)
-      apply(simp add: replicate_add[THEN sym])
-      apply(subgoal_tac "gc > n \<and> ?pstr \<ge> gc")
-      apply(erule_tac conjE)
-      apply(simp add: h1)
-      using h1
-      apply(auto)
-      apply(rule_tac min_max.le_supI2)
-      apply(rule_tac Max_ge, simp, simp)
-      apply(rule_tac disjI2)
-      using h2
-      thm rev_image_eqI
-      apply(rule_tac x = "gs!i" in rev_image_eqI, simp, simp)
-      done
-  next
-    show "aprog = cn_merge_gs (map rec_ci (take i gs)) 
-              ?pstr [+] ga [+] cp"
-      using g3 by simp
-  qed
-qed
-
-lemma recursive_compile_correct_spec: 
-  "\<lbrakk>rec_ci re = (ap, ary, fp); 
-    rec_calc_rel re args r\<rbrakk>
-  \<Longrightarrow> (\<exists> stp. (abc_steps_l (0, args @ 0\<up>(fp - ary)) ap stp) = 
-                     (length ap, args@[r]@0\<up>(fp - ary - 1)))"
-using recursive_compile_correct[of re ap ary fp args r "[]"]
-by simp
-
-definition dummy_abc :: "nat \<Rightarrow> abc_inst list"
-where
-"dummy_abc k = [Inc k, Dec k 0, Goto 3]"
-
-definition abc_list_crsp:: "nat list \<Rightarrow> nat list \<Rightarrow> bool"
-  where
-  "abc_list_crsp xs ys = (\<exists> n. xs = ys @ 0\<up>n \<or> ys = xs @ 0\<up>n)"
-
-lemma [intro]: "abc_list_crsp (lm @ 0\<up>m) lm"
-apply(auto simp: abc_list_crsp_def)
-done
-
-lemma abc_list_crsp_lm_v: 
-  "abc_list_crsp lma lmb \<Longrightarrow> abc_lm_v lma n = abc_lm_v lmb n"
-apply(auto simp: abc_list_crsp_def abc_lm_v.simps 
-                 nth_append)
-done
-
-lemma  rep_app_cons_iff: 
-  "k < n \<Longrightarrow> replicate n a[k:=b] = 
-          replicate k a @ b # replicate (n - k - 1) a"
-apply(induct n arbitrary: k, simp)
-apply(simp split:nat.splits)
-done
-
-lemma abc_list_crsp_lm_s: 
-  "abc_list_crsp lma lmb \<Longrightarrow> 
-      abc_list_crsp (abc_lm_s lma m n) (abc_lm_s lmb m n)"
-apply(auto simp: abc_list_crsp_def abc_lm_v.simps abc_lm_s.simps)
-apply(simp_all add: list_update_append, auto)
-proof -
-  fix na
-  assume h: "m < length lmb + na" " \<not> m < length lmb"
-  hence "m - length lmb < na" by simp
-  hence "replicate na 0[(m- length lmb):= n] = 
-           replicate (m - length lmb) 0 @ n # 
-              replicate (na - (m - length lmb) - 1) 0"
-    apply(erule_tac rep_app_cons_iff)
-    done
-  thus "\<exists>nb. replicate na 0[m - length lmb := n] =
-                 replicate (m - length lmb) 0 @ n # replicate nb 0 \<or>
-                 replicate (m - length lmb) 0 @ [n] =
-                 replicate na 0[m - length lmb := n] @ replicate nb 0"
-    apply(auto)
-    done
-next
-  fix na
-  assume h: "\<not> m < length lmb + na"
-  show 
-    "\<exists>nb. replicate na 0 @ replicate (m - (length lmb + na)) 0 @ [n] =
-           replicate (m - length lmb) 0 @ n # replicate nb 0 \<or>
-          replicate (m - length lmb) 0 @ [n] =
-            replicate na 0 @
-            replicate (m - (length lmb + na)) 0 @ n # replicate nb 0"
-    apply(rule_tac x = 0 in exI, simp, auto)
-    using h
-    apply(simp add: replicate_add[THEN sym])
-    done
-next
-  fix na
-  assume h: "\<not> m < length lma" "m < length lma + na"
-  hence "m - length lma < na" by simp
-  hence 
-    "replicate na 0[(m- length lma):= n] = replicate (m - length lma) 
-                  0 @ n # replicate (na - (m - length lma) - 1) 0"
-    apply(erule_tac rep_app_cons_iff)
-    done
-  thus "\<exists>nb. replicate (m - length lma) 0 @ [n] =
-                 replicate na 0[m - length lma := n] @ replicate nb 0 
-           \<or> replicate na 0[m - length lma := n] =
-                 replicate (m - length lma) 0 @ n # replicate nb 0"
-    apply(auto)
-    done
-next
-  fix na
-  assume "\<not> m < length lma + na"
-  thus " \<exists>nb. replicate (m - length lma) 0 @ [n] =
-            replicate na 0 @
-            replicate (m - (length lma + na)) 0 @ n # replicate nb 0 
-        \<or>   replicate na 0 @ 
-               replicate (m - (length lma + na)) 0 @ [n] =
-            replicate (m - length lma) 0 @ n # replicate nb 0"
-    apply(rule_tac x = 0 in exI, simp, auto)
-    apply(simp add: replicate_add[THEN sym])
-    done
-qed
-
-lemma abc_list_crsp_step: 
-  "\<lbrakk>abc_list_crsp lma lmb; abc_step_l (aa, lma) i = (a, lma'); 
-    abc_step_l (aa, lmb) i = (a', lmb')\<rbrakk>
-    \<Longrightarrow> a' = a \<and> abc_list_crsp lma' lmb'"
-apply(case_tac i, auto simp: abc_step_l.simps 
-       abc_list_crsp_lm_s abc_list_crsp_lm_v Let_def 
-                       split: abc_inst.splits if_splits)
-done
-
-lemma abc_list_crsp_steps: 
-  "\<lbrakk>abc_steps_l (0, lm @ 0\<up>m) aprog stp = (a, lm'); aprog \<noteq> []\<rbrakk> 
-      \<Longrightarrow> \<exists> lma. abc_steps_l (0, lm) aprog stp = (a, lma) \<and> 
-                                          abc_list_crsp lm' lma"
-apply(induct stp arbitrary: a lm', simp add: abc_steps_l.simps, auto)
-apply(case_tac "abc_steps_l (0, lm @ 0\<up>m) aprog stp", 
-      simp add: abc_step_red)
-proof -
-  fix stp a lm' aa b
-  assume ind:
-    "\<And>a lm'. aa = a \<and> b = lm' \<Longrightarrow> 
-     \<exists>lma. abc_steps_l (0, lm) aprog stp = (a, lma) \<and>
-                                          abc_list_crsp lm' lma"
-    and h: "abc_steps_l (0, lm @ 0\<up>m) aprog (Suc stp) = (a, lm')" 
-           "abc_steps_l (0, lm @ 0\<up>m) aprog stp = (aa, b)" 
-           "aprog \<noteq> []"
-  hence g1: "abc_steps_l (0, lm @ 0\<up>m) aprog (Suc stp)
-          = abc_step_l (aa, b) (abc_fetch aa aprog)"
-    apply(rule_tac abc_step_red, simp)
-    done
-  have "\<exists>lma. abc_steps_l (0, lm) aprog stp = (aa, lma) \<and> 
-              abc_list_crsp b lma"
-    apply(rule_tac ind, simp)
-    done
-  from this obtain lma where g2: 
-    "abc_steps_l (0, lm) aprog stp = (aa, lma) \<and> 
-     abc_list_crsp b lma"   ..
-  hence g3: "abc_steps_l (0, lm) aprog (Suc stp)
-          = abc_step_l (aa, lma) (abc_fetch aa aprog)"
-    apply(rule_tac abc_step_red, simp)
-    done
-  show "\<exists>lma. abc_steps_l (0, lm) aprog (Suc stp) = (a, lma) \<and> abc_list_crsp lm' lma"
-    using g1 g2 g3 h
-    apply(auto)
-    apply(case_tac "abc_step_l (aa, b) (abc_fetch aa aprog)",
-          case_tac "abc_step_l (aa, lma) (abc_fetch aa aprog)", simp)
-    apply(rule_tac abc_list_crsp_step, auto)
-    done
-qed
-
-lemma recursive_compile_correct_norm: 
-  "\<lbrakk>rec_ci re = (aprog, rs_pos, a_md);  
-   rec_calc_rel re lm rs\<rbrakk>
-  \<Longrightarrow> (\<exists> stp lm' m. (abc_steps_l (0, lm) aprog stp) = 
-  (length aprog, lm') \<and> abc_list_crsp lm' (lm @ rs # 0\<up>m))"
-apply(frule_tac recursive_compile_correct_spec, auto)
-apply(drule_tac abc_list_crsp_steps)
-apply(rule_tac rec_ci_not_null, simp)
-apply(erule_tac exE, rule_tac x = stp in exI, 
-  auto simp: abc_list_crsp_def)
-done
-
-lemma [simp]: "length (dummy_abc (length lm)) = 3"
-apply(simp add: dummy_abc_def)
-done
-
-lemma [simp]: "dummy_abc (length lm) \<noteq> []"
-apply(simp add: dummy_abc_def)
-done
-
-lemma dummy_abc_steps_ex: 
-  "\<exists>bstp. abc_steps_l (0, lm') (dummy_abc (length lm)) bstp = 
-  ((Suc (Suc (Suc 0))), abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)))"
-apply(rule_tac x = "Suc (Suc (Suc 0))" in exI)
-apply(auto simp: abc_steps_l.simps abc_step_l.simps 
-  dummy_abc_def abc_fetch.simps)
-apply(auto simp: abc_lm_s.simps abc_lm_v.simps nth_append)
-apply(simp add: butlast_append)
-done
-
-lemma [simp]: 
-  "\<lbrakk>Suc (length lm) - length lm' \<le> n; \<not> length lm < length lm'; lm @ rs # 0 \<up> m = lm' @ 0 \<up> n\<rbrakk> 
-  \<Longrightarrow> lm' @ 0 \<up> Suc (length lm - length lm') = lm @ [rs]"
-apply(subgoal_tac "n > m")
-apply(subgoal_tac "\<exists> d. n = d + m", erule_tac exE)
-apply(simp add: replicate_add)
-apply(drule_tac length_equal, simp)
-apply(simp add: replicate_Suc[THEN sym] del: replicate_Suc)
-apply(rule_tac x = "n - m" in exI, simp)
-apply(drule_tac length_equal, simp)
-done
-
-lemma [elim]: 
-  "lm @ rs # 0\<up>m = lm' @ 0\<up>n \<Longrightarrow> 
-  \<exists>m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) = 
-                            lm @ rs # 0\<up>m"
-proof(cases "length lm' > length lm")
-  case True 
-  assume h: "lm @ rs # 0\<up>m = lm' @ 0\<up>n" "length lm < length lm'"
-  hence "m \<ge> n"
-    apply(drule_tac length_equal)
-    apply(simp)
-    done
-  hence "\<exists> d. m = d + n"
-    apply(rule_tac x = "m - n" in exI, simp)
-    done
-  from this obtain d where "m = d + n" ..
-  from h and this show "?thesis"
-    apply(auto simp: abc_lm_s.simps abc_lm_v.simps 
-                     replicate_add)
-    done
-next
-  case False
-  assume h:"lm @ rs # 0\<up>m = lm' @ 0\<up>n" 
-    and    g: "\<not> length lm < length lm'"
-  have "take (Suc (length lm)) (lm @ rs # 0\<up>m) = 
-                        take (Suc (length lm)) (lm' @ 0\<up>n)"
-    using h by simp
-  moreover have "n \<ge> (Suc (length lm) - length lm')"
-    using h g
-    apply(drule_tac length_equal)
-    apply(simp)
-    done
-  ultimately show 
-    "\<exists>m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) =
-                                                       lm @ rs # 0\<up>m"
-    using g h
-    apply(simp add: abc_lm_s.simps abc_lm_v.simps min_def)
-    apply(rule_tac x = 0 in exI, 
-      simp add:replicate_append_same replicate_Suc[THEN sym]
-                                      del:replicate_Suc)
-    done    
-qed
-
-lemma [elim]: 
-  "abc_list_crsp lm' (lm @ rs # 0\<up>m)
-  \<Longrightarrow> \<exists>m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) 
-             = lm @ rs # 0\<up>m"
-apply(auto simp: abc_list_crsp_def)
-apply(simp add: abc_lm_v.simps abc_lm_s.simps)
-apply(rule_tac x =  "m + n" in exI, 
-      simp add: replicate_add)
-done
-
-lemma abc_append_dummy_complie:
-  "\<lbrakk>rec_ci recf = (ap, ary, fp);  
-    rec_calc_rel recf args r; 
-    length args = k\<rbrakk>
-  \<Longrightarrow> (\<exists> stp m. (abc_steps_l (0, args) (ap [+] dummy_abc k) stp) = 
-                  (length ap + 3, args @ r # 0\<up>m))"
-apply(drule_tac recursive_compile_correct_norm, auto simp: numeral_3_eq_3)
-proof -
-  fix stp lm' m
-  assume h: "rec_calc_rel recf args r"  
-    "abc_steps_l (0, args) ap stp = (length ap, lm')" 
-    "abc_list_crsp lm' (args @ r # 0\<up>m)"
-  thm abc_append_exc2
-  thm abc_lm_s.simps
-  have "\<exists>stp. abc_steps_l (0, args) (ap [+] 
-    (dummy_abc (length args))) stp = (length ap + 3, 
-    abc_lm_s lm' (length args) (abc_lm_v lm' (length args)))"
-    using h
-    apply(rule_tac bm = lm' in abc_append_exc2,
-          auto intro: dummy_abc_steps_ex simp: numeral_3_eq_3)
-    done
-  thus "\<exists>stp m. abc_steps_l (0, args) (ap [+] 
-    dummy_abc (length args)) stp = (Suc (Suc (Suc (length ap))), args @ r # 0\<up>m)"
-    using h
-    apply(erule_tac exE)
-    apply(rule_tac x = stpa in exI, auto)
-    done
-qed
-
-lemma [simp]: "length (dummy_abc k) = 3"
-apply(simp add: dummy_abc_def)
-done
-
-lemma [simp]: "length args = k \<Longrightarrow> abc_lm_v (args @ r # 0\<up>m) k = r "
-apply(simp add: abc_lm_v.simps nth_append)
-done
-
-lemma [simp]: "crsp (layout_of (ap [+] dummy_abc k)) (0, args)
-  (Suc 0, Bk # Bk # ires, <args> @ Bk \<up> rn) ires"
-apply(auto simp: crsp.simps start_of.simps)
-done
-
-(* cccc *)
-
-fun tm_of_rec :: "recf \<Rightarrow> instr list"
-where "tm_of_rec recf = (let (ap, k, fp) = rec_ci recf in
-                         let tp = tm_of (ap [+] dummy_abc k) in 
-                             tp @ (shift (mopup k) (length tp div 2)))"
-
-lemma recursive_compile_to_tm_correct: 
-  "\<lbrakk>rec_ci recf = (ap, ary, fp); 
-    rec_calc_rel recf args r;
-    length args = k;
-    ly = layout_of (ap [+] dummy_abc k);
-    tp = tm_of (ap [+] dummy_abc k)\<rbrakk>
-  \<Longrightarrow> \<exists> stp m l. steps0 (Suc 0, Bk # Bk # ires, <args> @ Bk\<up>rn)
-  (tp @ shift (mopup k) (length tp div 2)) stp
-  = (0, Bk\<up>m @ Bk # Bk # ires, Oc\<up>Suc r @ Bk\<up>l)"
-  using abc_append_dummy_complie[of recf ap ary fp args r k]
-apply(simp)
-apply(erule_tac exE)+
-apply(frule_tac tp = tp and n = k 
-               and ires = ires in compile_correct_halt, simp_all add: length_append)
-apply(simp_all add: length_append)
-done
-
-lemma recursive_compile_to_tm_correct2: 
-  assumes "rec_ci recf = (ap, ary, fp)" 
-  and     "rec_calc_rel recf args r"
-  and     "length args = k"
-  and     "tp = tm_of (ap [+] dummy_abc k)"
-  shows "\<exists> m n. {\<lambda>tp. tp = ([Bk, Bk], <args>)}
-             (tp @ (shift (mopup k) (length tp div 2)))
-             {\<lambda>tp. tp = (Bk \<up> m, Oc \<up> (Suc r) @ Bk \<up> n)}"
-using recursive_compile_to_tm_correct[where ires="[]" and rn="0", OF assms(1-3) _ assms(4)]
-apply(simp add: Hoare_halt_def)
-apply(drule_tac x="layout_of (ap [+] dummy_abc k)" in meta_spec)
-apply(auto)
-apply(rule_tac x="m + 2" in exI)
-apply(rule_tac x="l" in exI)
-apply(rule_tac x="stp" in exI)
-apply(auto)
-by (metis append_Nil2 replicate_app_Cons_same)
-
-lemma recursive_compile_to_tm_correct3: 
-  assumes "rec_calc_rel recf args r"
-  shows "{\<lambda>tp. tp = ([Bk, Bk], <args>)} tm_of_rec recf {\<lambda>tp. \<exists>k l. tp = (Bk \<up> k, <r> @ Bk \<up> l)}"
-using recursive_compile_to_tm_correct2[OF _ assms] 
-apply(auto)
-apply(case_tac "rec_ci recf")
-apply(auto)
-apply(drule_tac x="a" in meta_spec)
-apply(drule_tac x="b" in meta_spec)
-apply(drule_tac x="c" in meta_spec)
-apply(drule_tac x="length args" in meta_spec)
-apply(drule_tac x="tm_of (a [+] dummy_abc (length args))" in meta_spec)
-apply(auto)
-apply(simp add: tape_of_nat_abv)
-apply(subgoal_tac "b = length args")
-apply(simp add: Hoare_halt_def)
-apply(auto)[1]
-apply(rule_tac x="na" in exI)
-apply(auto)[1]
-apply(case_tac "steps0 (Suc 0, [Bk, Bk], <args>)
-                                   (tm_of (a [+] dummy_abc (length args)) @
-                                    shift (mopup (length args))
-                                     (listsum
- (layout_of (a [+] dummy_abc (length args)))))
-                                   na")
-apply(simp)
-by (metis assms para_pattern)
-
-
-lemma [simp]:
-  "list_all (\<lambda>(acn, s). s \<le> Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))) xs \<Longrightarrow>
-  list_all (\<lambda>(acn, s). s \<le> Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))))) xs"
-apply(induct xs, simp, simp)
-apply(case_tac a, simp)
-done
-
-lemma shift_append: "shift (xs @ ys) n = shift xs n @ shift ys n"
-apply(simp add: shift.simps)
-done
-
-lemma [simp]: "length (shift (mopup n) ss) = 4 * n + 12"
-apply(auto simp: mopup.simps shift_append mopup_b_def)
-done
-
-lemma length_tm_even[intro]: "length (tm_of ap) mod 2 = 0"
-apply(simp add: tm_of.simps)
-done
-
-lemma [simp]: "k < length ap \<Longrightarrow> tms_of ap ! k  = 
- ci (layout_of ap) (start_of (layout_of ap) k) (ap ! k)"
-apply(simp add: tms_of.simps tpairs_of.simps)
-done
-
-lemma start_of_suc_inc:
-  "\<lbrakk>k < length ap; ap ! k = Inc n\<rbrakk> \<Longrightarrow> start_of (layout_of ap) (Suc k) =
-                        start_of (layout_of ap) k + 2 * n + 9"
-apply(rule_tac start_of_Suc1, auto simp: abc_fetch.simps)
-done
-
-lemma start_of_suc_dec:
-  "\<lbrakk>k < length ap; ap ! k = (Dec n e)\<rbrakk> \<Longrightarrow> start_of (layout_of ap) (Suc k) =
-                        start_of (layout_of ap) k + 2 * n + 16"
-apply(rule_tac start_of_Suc2, auto simp: abc_fetch.simps)
-done
-
-lemma inc_state_all_le:
-  "\<lbrakk>k < length ap; ap ! k = Inc n; 
-       (a, b) \<in> set (shift (shift tinc_b (2 * n)) 
-                            (start_of (layout_of ap) k - Suc 0))\<rbrakk>
-       \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
-apply(subgoal_tac "b \<le> start_of (layout_of ap) (Suc k)")
-apply(subgoal_tac "start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap) ")
-apply(arith)
-apply(case_tac "Suc k = length ap", simp)
-apply(rule_tac start_of_less, simp)
-apply(auto simp: tinc_b_def shift.simps start_of_suc_inc length_of.simps startof_not0)
-done
-
-lemma findnth_le[elim]: 
-  "(a, b) \<in> set (shift (findnth n) (start_of (layout_of ap) k - Suc 0))
-  \<Longrightarrow> b \<le> Suc (start_of (layout_of ap) k + 2 * n)"
-apply(induct n, simp add: findnth.simps shift.simps)
-apply(simp add: findnth.simps shift_append, auto)
-apply(auto simp: shift.simps)
-done
-
-lemma findnth_state_all_le1:
-  "\<lbrakk>k < length ap; ap ! k = Inc n;
-  (a, b) \<in> 
-  set (shift (findnth n) (start_of (layout_of ap) k - Suc 0))\<rbrakk> 
-  \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
-apply(subgoal_tac "b \<le> start_of (layout_of ap) (Suc k)")
-apply(subgoal_tac "start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap) ")
-apply(arith)
-apply(case_tac "Suc k = length ap", simp)
-apply(rule_tac start_of_less, simp)
-apply(subgoal_tac "b \<le> start_of (layout_of ap) k + 2*n + 1 \<and> 
-     start_of (layout_of ap) k + 2*n + 1 \<le>  start_of (layout_of ap) (Suc k)", auto)
-apply(auto simp: tinc_b_def shift.simps length_of.simps startof_not0 start_of_suc_inc)
-done
-
-lemma start_of_eq: "length ap < as \<Longrightarrow> start_of (layout_of ap) as = start_of (layout_of ap) (length ap)"
-apply(induct as, simp)
-apply(case_tac "length ap < as", simp add: start_of.simps)
-apply(subgoal_tac "as = length ap")
-apply(simp add: start_of.simps)
-apply arith
-done
-
-lemma start_of_all_le: "start_of (layout_of ap) as \<le> start_of (layout_of ap) (length ap)"
-apply(subgoal_tac "as > length ap \<or> as = length ap \<or> as < length ap", 
-      auto simp: start_of_eq start_of_less)
-done
-
-lemma findnth_state_all_le2: 
-  "\<lbrakk>k < length ap; 
-  ap ! k = Dec n e;
-  (a, b) \<in> set (shift (findnth n) (start_of (layout_of ap) k - Suc 0))\<rbrakk>
-  \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
-apply(subgoal_tac "b \<le> start_of (layout_of ap) k + 2*n + 1 \<and> 
-     start_of (layout_of ap) k + 2*n + 1 \<le>  start_of (layout_of ap) (Suc k) \<and>
-      start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap)", auto)
-apply(subgoal_tac "start_of (layout_of ap) (Suc k) = 
-  start_of  (layout_of ap)  k + 2*n + 16", simp)
-apply(simp add: start_of_suc_dec)
-apply(rule_tac start_of_all_le)
-done
-
-lemma dec_state_all_le[simp]:
-  "\<lbrakk>k < length ap; ap ! k = Dec n e; 
-  (a, b) \<in> set (shift (shift tdec_b (2 * n))
-  (start_of (layout_of ap) k - Suc 0))\<rbrakk>
-       \<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
-apply(subgoal_tac "2*n + start_of (layout_of ap) k + 16 \<le> start_of (layout_of ap) (length ap) \<and> start_of (layout_of ap) k > 0")
-prefer 2
-apply(subgoal_tac "start_of (layout_of ap) (Suc k) = start_of (layout_of ap) k + 2*n + 16
-                 \<and> start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap)")
-apply(simp add: startof_not0, rule_tac conjI)
-apply(simp add: start_of_suc_dec)
-apply(rule_tac start_of_all_le)
-apply(auto simp: tdec_b_def shift.simps)
-done
-
-lemma tms_any_less: 
-  "\<lbrakk>k < length ap; (a, b) \<in> set (tms_of ap ! k)\<rbrakk> \<Longrightarrow> 
-  b \<le> start_of (layout_of ap) (length ap)"
-apply(case_tac "ap!k", auto simp: tms_of.simps tpairs_of.simps ci.simps shift_append sete.simps)
-apply(erule_tac findnth_state_all_le1, simp_all)
-apply(erule_tac inc_state_all_le, simp_all)
-apply(erule_tac findnth_state_all_le2, simp_all)
-apply(rule_tac start_of_all_le)
-apply(rule_tac dec_state_all_le, simp_all)
-apply(rule_tac start_of_all_le)
-done
-
-lemma concat_in: "i < length (concat xs) \<Longrightarrow> \<exists>k < length xs. concat xs ! i \<in> set (xs ! k)"
-apply(induct xs rule: list_tl_induct, simp, simp)
-apply(case_tac "i < length (concat list)", simp)
-apply(erule_tac exE, rule_tac x = k in exI)
-apply(simp add: nth_append)
-apply(rule_tac x = "length list" in exI, simp)
-apply(simp add: nth_append)
-done 
-
-lemma [simp]: "length (tms_of ap) = length ap"
-apply(simp add: tms_of.simps tpairs_of.simps)
-done
-
-declare length_concat[simp]
-
-lemma in_tms: "i < length (tm_of ap) \<Longrightarrow> \<exists> k < length ap. (tm_of ap ! i) \<in> set (tms_of ap ! k)"
-apply(simp only: tm_of.simps)
-using concat_in[of i "tms_of ap"]
-apply(auto)
-done
-
-lemma all_le_start_of: "list_all (\<lambda>(acn, s). 
-  s \<le> start_of (layout_of ap) (length ap)) (tm_of ap)"
-apply(simp only: list_all_length)
-apply(rule_tac allI, rule_tac impI)
-apply(drule_tac in_tms, auto elim: tms_any_less)
-done
-
-lemma length_ci: 
-"\<lbrakk>k < length ap; length (ci ly y (ap ! k)) = 2 * qa\<rbrakk>
-      \<Longrightarrow> layout_of ap ! k = qa"
-apply(case_tac "ap ! k")
-apply(auto simp: layout_of.simps ci.simps 
-  length_of.simps tinc_b_def tdec_b_def length_findnth sete.simps)
-done
-
-lemma [intro]: "length (ci ly y i) mod 2 = 0"
-apply(case_tac i, auto simp: ci.simps length_findnth
-  tinc_b_def sete.simps tdec_b_def)
-done
-
-lemma [intro]: "listsum (map (length \<circ> (\<lambda>(x, y). ci ly x y)) zs) mod 2 = 0"
-apply(induct zs rule: list_tl_induct, simp)
-apply(case_tac a, simp)
-apply(subgoal_tac "length (ci ly aa b) mod 2 = 0")
-apply(auto)
-done
-
-lemma zip_pre:
-  "(length ys) \<le> length ap \<Longrightarrow>
-  zip ys ap = zip ys (take (length ys) (ap::'a list))"
-proof(induct ys arbitrary: ap, simp, case_tac ap, simp)
-  fix a ys ap aa list
-  assume ind: "\<And>(ap::'a list). length ys \<le> length ap \<Longrightarrow> 
-    zip ys ap = zip ys (take (length ys) ap)"
-  and h: "length (a # ys) \<le> length ap" "(ap::'a list) = aa # (list::'a list)"
-  from h show "zip (a # ys) ap = zip (a # ys) (take (length (a # ys)) ap)"
-    using ind[of list]
-    apply(simp)
-    done
-qed
-
-lemma length_start_of_tm: "start_of (layout_of ap) (length ap) = Suc (length (tm_of ap)  div 2)"
-using tpa_states[of "tm_of ap"  "length ap" ap]
-apply(simp add: tm_of.simps)
-done
-
-lemma [elim]: "list_all (\<lambda>(acn, s). s \<le> Suc q) xs
-        \<Longrightarrow> list_all (\<lambda>(acn, s). s \<le> q + (2 * n + 6)) xs"
-apply(simp add: list_all_length)
-apply(auto)
-done
-
-lemma [simp]: "length mopup_b = 12"
-apply(simp add: mopup_b_def)
-done
-(*
-lemma [elim]: "\<lbrakk>na < 4 * n; tshift (mop_bef n) q ! na = (a, b)\<rbrakk> \<Longrightarrow> 
-  b \<le> q + (2 * n + 6)"
-apply(induct n, simp, simp add: mop_bef.simps nth_append tshift_append shift_length)
-apply(case_tac "na < 4*n", simp, simp)
-apply(subgoal_tac "na = 4*n \<or> na = 1 + 4*n \<or> na = 2 + 4*n \<or> na = 3 + 4*n",
-  auto simp: shift_length)
-apply(simp_all add: tshift.simps)
-done
-*)
-
-lemma mp_up_all_le: "list_all  (\<lambda>(acn, s). s \<le> q + (2 * n + 6)) 
-  [(R, Suc (Suc (2 * n + q))), (R, Suc (2 * n + q)), 
-  (L, 5 + 2 * n + q), (W0, Suc (Suc (Suc (2 * n + q)))), (R, 4 + 2 * n + q),
-  (W0, Suc (Suc (Suc (2 * n + q)))), (R, Suc (Suc (2 * n + q))),
-  (W0, Suc (Suc (Suc (2 * n + q)))), (L, 5 + 2 * n + q),
-  (L, 6 + 2 * n + q), (R, 0),  (L, 6 + 2 * n + q)]"
-apply(auto)
-done
-
-lemma [simp]: "(a, b) \<in> set (mopup_a n) \<Longrightarrow> b \<le> 2 * n + 6"
-apply(induct n, auto simp: mopup_a.simps)
-done
-
-lemma [simp]: "(a, b) \<in> set (shift (mopup n) (listsum (layout_of ap)))
-  \<Longrightarrow> b \<le> (2 * listsum (layout_of ap) + length (mopup n)) div 2"
-apply(auto simp: mopup.simps shift_append shift.simps)
-apply(auto simp: mopup_a.simps mopup_b_def)
-done
-
-lemma [intro]: " 2 \<le> 2 * listsum (layout_of ap) + length (mopup n)"
-apply(simp add: mopup.simps)
-done
-
-lemma [intro]: " (2 * listsum (layout_of ap) + length (mopup n)) mod 2 = 0"
-apply(auto simp: mopup.simps)
-apply arith
-done
-
-lemma [simp]: "b \<le> Suc x
-          \<Longrightarrow> b \<le> (2 * x + length (mopup n)) div 2"
-apply(auto simp: mopup.simps)
-done
-
-lemma t_compiled_correct: 
-  "\<lbrakk>tp = tm_of ap; ly = layout_of ap; mop_ss = start_of ly (length ap)\<rbrakk> \<Longrightarrow> 
-    tm_wf (tp @ shift( mopup n) (length tp div 2), 0)"
-  using length_start_of_tm[of ap] all_le_start_of[of ap]
-apply(auto simp: tm_wf.simps List.list_all_iff)
-done
-
-end
-
-    
-  
-
-
-  
-
--- a/thys/turing_basic.thy	Thu Feb 07 06:39:06 2013 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,428 +0,0 @@
-(* Title: Turing machines
-   Author: Xu Jian <xujian817@hotmail.com>
-   Maintainer: Xu Jian
-*)
-
-theory turing_basic
-imports Main
-begin
-
-section {* Basic definitions of Turing machine *}
-
-datatype action = W0 | W1 | L | R | Nop
-
-datatype cell = Bk | Oc
-
-type_synonym tape = "cell list \<times> cell list"
-
-type_synonym state = nat
-
-type_synonym instr = "action \<times> state"
-
-type_synonym tprog = "instr list \<times> nat"
-
-type_synonym tprog0 = "instr list"
-
-type_synonym config = "state \<times> tape"
-
-fun nth_of where
-  "nth_of xs i = (if i \<ge> length xs then None
-                  else Some (xs ! i))"
-
-lemma nth_of_map [simp]:
-  shows "nth_of (map f p) n = (case (nth_of p n) of None \<Rightarrow> None | Some x \<Rightarrow> Some (f x))"
-apply(induct p arbitrary: n)
-apply(auto)
-apply(case_tac n)
-apply(auto)
-done
-
-fun 
-  fetch :: "instr list \<Rightarrow> state \<Rightarrow> cell \<Rightarrow> instr"
-where
-  "fetch p 0 b = (Nop, 0)"
-| "fetch p (Suc s) Bk = 
-     (case nth_of p (2 * s) of
-        Some i \<Rightarrow> i
-      | None \<Rightarrow> (Nop, 0))"
-|"fetch p (Suc s) Oc = 
-     (case nth_of p ((2 * s) + 1) of
-         Some i \<Rightarrow> i
-       | None \<Rightarrow> (Nop, 0))"
-
-lemma fetch_Nil [simp]:
-  shows "fetch [] s b = (Nop, 0)"
-apply(case_tac s)
-apply(auto)
-apply(case_tac b)
-apply(auto)
-done
-
-fun 
-  update :: "action \<Rightarrow> tape \<Rightarrow> tape"
-where 
-  "update W0 (l, r) = (l, Bk # (tl r))" 
-| "update W1 (l, r) = (l, Oc # (tl r))"
-| "update L (l, r) = (if l = [] then ([], Bk # r) else (tl l, (hd l) # r))" 
-| "update R (l, r) = (if r = [] then (Bk # l, []) else ((hd r) # l, tl r))" 
-| "update Nop (l, r) = (l, r)"
-
-abbreviation 
-  "read r == if (r = []) then Bk else hd r"
-
-fun step :: "config \<Rightarrow> tprog \<Rightarrow> config"
-  where 
-  "step (s, l, r) (p, off) = 
-     (let (a, s') = fetch p (s - off) (read r) in (s', update a (l, r)))"
-
-fun steps :: "config \<Rightarrow> tprog \<Rightarrow> nat \<Rightarrow> config"
-  where
-  "steps c p 0 = c" |
-  "steps c p (Suc n) = steps (step c p) p n"
-
-
-abbreviation
-  "step0 c p \<equiv> step c (p, 0)"
-
-abbreviation
-  "steps0 c p n \<equiv> steps c (p, 0) n"
-
-lemma step_red [simp]: 
-  shows "steps c p (Suc n) = step (steps c p n) p"
-by (induct n arbitrary: c) (auto)
-
-lemma steps_add [simp]: 
-  shows "steps c p (m + n) = steps (steps c p m) p n"
-by (induct m arbitrary: c) (auto)
-
-lemma step_0 [simp]: 
-  shows "step (0, (l, r)) p = (0, (l, r))"
-by (case_tac p, simp)
-
-lemma steps_0 [simp]: 
-  shows "steps (0, (l, r)) p n = (0, (l, r))"
-by (induct n) (simp_all)
-
-
-
-fun
-  is_final :: "config \<Rightarrow> bool"
-where
-  "is_final (s, l, r) = (s = 0)"
-
-lemma is_final_eq: 
-  shows "is_final (s, tp) = (s = 0)"
-by (case_tac tp) (auto)
-
-lemma after_is_final:
-  assumes "is_final c"
-  shows "is_final (steps c p n)"
-using assms 
-apply(induct n) 
-apply(case_tac [!] c)
-apply(auto)
-done
-
-lemma not_is_final:
-  assumes a: "\<not> is_final (steps c p n1)"
-  and b: "n2 \<le> n1"
-  shows "\<not> is_final (steps c p n2)"
-proof (rule notI)
-  obtain n3 where eq: "n1 = n2 + n3" using b by (metis le_iff_add)
-  assume "is_final (steps c p n2)"
-  then have "is_final (steps c p n1)" unfolding eq
-    by (simp add: after_is_final)
-  with a show "False" by simp
-qed
-
-(* if the machine is in the halting state, there must have 
-   been a state just before the halting state *)
-lemma before_final: 
-  assumes "steps0 (1, tp) A n = (0, tp')"
-  shows "\<exists> n'. \<not> is_final (steps0 (1, tp) A n') \<and> steps0 (1, tp) A (Suc n') = (0, tp')"
-using assms
-proof(induct n arbitrary: tp')
-  case (0 tp')
-  have asm: "steps0 (1, tp) A 0 = (0, tp')" by fact
-  then show "\<exists>n'. \<not> is_final (steps0 (1, tp) A n') \<and> steps0 (1, tp) A (Suc n') = (0, tp')"
-    by simp
-next
-  case (Suc n tp')
-  have ih: "\<And>tp'. steps0 (1, tp) A n = (0, tp') \<Longrightarrow>
-    \<exists>n'. \<not> is_final (steps0 (1, tp) A n') \<and> steps0 (1, tp) A (Suc n') = (0, tp')" by fact
-  have asm: "steps0 (1, tp) A (Suc n) = (0, tp')" by fact
-  obtain s l r where cases: "steps0 (1, tp) A n = (s, l, r)"
-    by (auto intro: is_final.cases)
-  then show "\<exists>n'. \<not> is_final (steps0 (1, tp) A n') \<and> steps0 (1, tp) A (Suc n') = (0, tp')"
-  proof (cases "s = 0")
-    case True (* in halting state *)
-    then have "steps0 (1, tp) A n = (0, tp')"
-      using asm cases by (simp del: steps.simps)
-    then show ?thesis using ih by simp
-  next
-    case False (* not in halting state *)
-    then have "\<not> is_final (steps0 (1, tp) A n) \<and> steps0 (1, tp) A (Suc n) = (0, tp')"
-      using asm cases by simp
-    then show ?thesis by auto
-  qed
-qed
-
-(* well-formedness of Turing machine programs *)
-abbreviation "is_even n \<equiv> (n::nat) mod 2 = 0"
-
-fun 
-  tm_wf :: "tprog \<Rightarrow> bool"
-where
-  "tm_wf (p, off) = (length p \<ge> 2 \<and> is_even (length p) \<and> 
-                    (\<forall>(a, s) \<in> set p. s \<le> length p div 2 + off \<and> s \<ge> off))"
-
-abbreviation
-  "tm_wf0 p \<equiv> tm_wf (p, 0)"
-
-abbreviation exponent :: "'a \<Rightarrow> nat \<Rightarrow> 'a list" ("_ \<up> _" [100, 99] 100)
-  where "x \<up> n == replicate n x"
-
-consts tape_of :: "'a \<Rightarrow> cell list" ("<_>" 100)
-
-defs (overloaded)
-  tape_of_nat_abv: "<(n::nat)> \<equiv> Oc \<up> (Suc n)"
-
-fun tape_of_nat_list :: "'a list \<Rightarrow> cell list" 
-  where 
-  "tape_of_nat_list [] = []" |
-  "tape_of_nat_list [n] = <n>" |
-  "tape_of_nat_list (n#ns) = <n> @ Bk # (tape_of_nat_list ns)"
-
-fun tape_of_nat_pair :: "'a \<times> 'b \<Rightarrow> cell list" 
-  where 
-  "tape_of_nat_pair (n, m) = <n> @ [Bk] @ <m>" 
-
-
-defs (overloaded)
-  tape_of_nl_abv: "<(ns::'a list)> \<equiv> tape_of_nat_list ns"
-  tape_of_nat_pair: "<(np::'a\<times>'b)> \<equiv> tape_of_nat_pair np"
-
-fun 
-  shift :: "instr list \<Rightarrow> nat \<Rightarrow> instr list"
-where
-  "shift p n = (map (\<lambda> (a, s). (a, (if s = 0 then 0 else s + n))) p)"
-
-fun 
-  adjust :: "instr list \<Rightarrow> instr list"
-where
-  "adjust p = map (\<lambda> (a, s). (a, if s = 0 then (Suc (length p div 2)) else s)) p"
-
-lemma length_shift [simp]: 
-  shows "length (shift p n) = length p"
-by simp
-
-lemma length_adjust [simp]: 
-  shows "length (adjust p) = length p"
-by (induct p) (auto)
-
-
-(* composition of two Turing machines *)
-fun
-  tm_comp :: "instr list \<Rightarrow> instr list \<Rightarrow> instr list" ("_ |+| _" [0, 0] 100)
-where
-  "tm_comp p1 p2 = ((adjust p1) @ (shift p2 (length p1 div 2)))"
-
-lemma tm_comp_length:
-  shows "length (A |+| B) = length A + length B"
-by auto
-
-lemma tm_comp_wf[intro]: 
-  "\<lbrakk>tm_wf (A, 0); tm_wf (B, 0)\<rbrakk> \<Longrightarrow> tm_wf (A |+| B, 0)"
-by (auto simp: tm_wf.simps shift.simps adjust.simps tm_comp_length tm_comp.simps)
-
-
-lemma tm_comp_step: 
-  assumes unfinal: "\<not> is_final (step0 c A)"
-  shows "step0 c (A |+| B) = step0 c A"
-proof -
-  obtain s l r where eq: "c = (s, l, r)" by (metis is_final.cases) 
-  have "\<not> is_final (step0 (s, l, r) A)" using unfinal eq by simp
-  then have "case (fetch A s (read r)) of (a, s) \<Rightarrow> s \<noteq> 0"
-    by (auto simp add: is_final_eq)
-  then  have "fetch (A |+| B) s (read r) = fetch A s (read r)"
-    apply(case_tac [!] "read r")
-    apply(case_tac [!] s)
-    apply(auto simp: tm_comp_length nth_append)
-    done
-  then show "step0 c (A |+| B) = step0 c A" by (simp add: eq) 
-qed
-
-lemma tm_comp_steps:  
-  assumes "\<not> is_final (steps0 c A n)" 
-  shows "steps0 c (A |+| B) n = steps0 c A n"
-using assms
-proof(induct n)
-  case 0
-  then show "steps0 c (A |+| B) 0 = steps0 c A 0" by auto
-next 
-  case (Suc n)
-  have ih: "\<not> is_final (steps0 c A n) \<Longrightarrow> steps0 c (A |+| B) n = steps0 c A n" by fact
-  have fin: "\<not> is_final (steps0 c A (Suc n))" by fact
-  then have fin1: "\<not> is_final (step0 (steps0 c A n) A)" 
-    by (auto simp only: step_red)
-  then have fin2: "\<not> is_final (steps0 c A n)"
-    by (metis is_final_eq step_0 surj_pair) 
- 
-  have "steps0 c (A |+| B) (Suc n) = step0 (steps0 c (A |+| B) n) (A |+| B)" 
-    by (simp only: step_red)
-  also have "... = step0 (steps0 c A n) (A |+| B)" by (simp only: ih[OF fin2])
-  also have "... = step0 (steps0 c A n) A" by (simp only: tm_comp_step[OF fin1])
-  finally show "steps0 c (A |+| B) (Suc n) = steps0 c A (Suc n)"
-    by (simp only: step_red)
-qed
-
-lemma tm_comp_fetch_in_A:
-  assumes h1: "fetch A s x = (a, 0)"
-  and h2: "s \<le> length A div 2" 
-  and h3: "s \<noteq> 0"
-  shows "fetch (A |+| B) s x = (a, Suc (length A div 2))"
-using h1 h2 h3
-apply(case_tac s)
-apply(case_tac [!] x)
-apply(auto simp: tm_comp_length nth_append)
-done
-
-lemma tm_comp_exec_after_first:
-  assumes h1: "\<not> is_final c" 
-  and h2: "step0 c A = (0, tp)"
-  and h3: "fst c \<le> length A div 2"
-  shows "step0 c (A |+| B) = (Suc (length A div 2), tp)"
-using h1 h2 h3
-apply(case_tac c)
-apply(auto simp del: tm_comp.simps)
-apply(case_tac "fetch A a Bk")
-apply(simp del: tm_comp.simps)
-apply(subst tm_comp_fetch_in_A)
-apply(auto)[4]
-apply(case_tac "fetch A a (hd c)")
-apply(simp del: tm_comp.simps)
-apply(subst tm_comp_fetch_in_A)
-apply(auto)[4]
-done
-
-lemma step_in_range: 
-  assumes h1: "\<not> is_final (step0 c A)"
-  and h2: "tm_wf (A, 0)"
-  shows "fst (step0 c A) \<le> length A div 2"
-using h1 h2
-apply(case_tac c)
-apply(case_tac a)
-apply(auto simp add: prod_case_unfold Let_def)
-apply(case_tac "hd c")
-apply(auto simp add: prod_case_unfold)
-done
-
-lemma steps_in_range: 
-  assumes h1: "\<not> is_final (steps0 (1, tp) A stp)"
-  and h2: "tm_wf (A, 0)"
-  shows "fst (steps0 (1, tp) A stp) \<le> length A div 2"
-using h1
-proof(induct stp)
-  case 0
-  then show "fst (steps0 (1, tp) A 0) \<le> length A div 2" using h2
-    by (auto simp add: steps.simps tm_wf.simps)
-next
-  case (Suc stp)
-  have ih: "\<not> is_final (steps0 (1, tp) A stp) \<Longrightarrow> fst (steps0 (1, tp) A stp) \<le> length A div 2" by fact
-  have h: "\<not> is_final (steps0 (1, tp) A (Suc stp))" by fact
-  from ih h h2 show "fst (steps0 (1, tp) A (Suc stp)) \<le> length A div 2"
-    by (metis step_in_range step_red)
-qed
-
-lemma tm_comp_pre_halt_same: 
-  assumes a_ht: "steps0 (1, tp) A n = (0, tp')"
-  and a_wf: "tm_wf (A, 0)"
-  obtains n' where "steps0 (1, tp) (A |+| B) n' = (Suc (length A div 2), tp')"
-proof -
-  assume a: "\<And>n. steps (1, tp) (A |+| B, 0) n = (Suc (length A div 2), tp') \<Longrightarrow> thesis"
-  obtain stp' where fin: "\<not> is_final (steps0 (1, tp) A stp')" and h: "steps0 (1, tp) A (Suc stp') = (0, tp')"
-  using before_final[OF a_ht] by blast
-  from fin have h1:"steps0 (1, tp) (A |+| B) stp' = steps0 (1, tp) A stp'"
-    by (rule tm_comp_steps)
-  from h have h2: "step0 (steps0 (1, tp) A stp') A = (0, tp')"
-    by (simp only: step_red)
-
-  have "steps0 (1, tp) (A |+| B) (Suc stp') = step0 (steps0 (1, tp) (A |+| B) stp') (A |+| B)" 
-    by (simp only: step_red)
-  also have "... = step0 (steps0 (1, tp) A stp') (A |+| B)" using h1 by simp
-  also have "... = (Suc (length A div 2), tp')" 
-    by (rule tm_comp_exec_after_first[OF fin h2 steps_in_range[OF fin a_wf]])
-  finally show thesis using a by blast
-qed
-
-lemma tm_comp_fetch_second_zero:
-  assumes h1: "fetch B s x = (a, 0)"
-  and hs: "tm_wf (A, 0)" "s \<noteq> 0"
-  shows "fetch (A |+| B) (s + (length A div 2)) x = (a, 0)"
-using h1 hs
-apply(case_tac x)
-apply(case_tac [!] s)
-apply(auto simp: tm_comp_length nth_append)
-done 
-
-lemma tm_comp_fetch_second_inst:
-  assumes h1: "fetch B sa x = (a, s)"
-  and hs: "tm_wf (A, 0)" "sa \<noteq> 0" "s \<noteq> 0"
-  shows "fetch (A |+| B) (sa + length A div 2) x = (a, s + length A div 2)"
-using h1 hs
-apply(case_tac x)
-apply(case_tac [!] sa)
-apply(auto simp: tm_comp_length nth_append)
-done 
-
-
-lemma tm_comp_second_same:
-  assumes a_wf: "tm_wf (A, 0)"
-  and steps: "steps0 (1, l, r) B stp = (s', l', r')"
-  shows "steps0 (Suc (length A div 2), l, r)  (A |+| B) stp 
-    = (if s' = 0 then 0 else s' + length A div 2, l', r')"
-using steps
-proof(induct stp arbitrary: s' l' r')
-  case 0
-  then show ?case by (simp add: steps.simps)
-next
-  case (Suc stp s' l' r')
-  obtain s'' l'' r'' where a: "steps0 (1, l, r) B stp = (s'', l'', r'')"
-    by (metis is_final.cases)
-  then have ih1: "s'' = 0 \<Longrightarrow> steps0 (Suc (length A div 2), l, r) (A |+| B) stp = (0, l'', r'')"
-  and ih2: "s'' \<noteq> 0 \<Longrightarrow> steps0 (Suc (length A div 2), l, r) (A |+| B) stp = (s'' + length A div 2, l'', r'')"
-    using Suc by (auto)
-  have h: "steps0 (1, l, r) B (Suc stp) = (s', l', r')" by fact
-
-  { assume "s'' = 0"
-    then have ?case using a h ih1 by (simp del: steps.simps) 
-  } moreover
-  { assume as: "s'' \<noteq> 0" "s' = 0"
-    from as a h 
-    have "step0 (s'', l'', r'') B = (0, l', r')" by (simp del: steps.simps)
-    with as have ?case
-    apply(simp add: ih2[OF as(1)] step.simps del: tm_comp.simps steps.simps)
-    apply(case_tac "fetch B s'' (read r'')")
-    apply(auto simp add: tm_comp_fetch_second_zero[OF _ a_wf] simp del: tm_comp.simps)
-    done
-  } moreover
-  { assume as: "s'' \<noteq> 0" "s' \<noteq> 0"
-    from as a h
-    have "step0 (s'', l'', r'') B = (s', l', r')" by (simp del: steps.simps)
-    with as have ?case
-    apply(simp add: ih2[OF as(1)] step.simps del: tm_comp.simps steps.simps)
-    apply(case_tac "fetch B s'' (read r'')")
-    apply(auto simp add: tm_comp_fetch_second_inst[OF _ a_wf as] simp del: tm_comp.simps)
-    done
-  }
-  ultimately show ?case by blast
-qed
-
-lemma tm_comp_second_halt_same:
-  assumes "tm_wf (A, 0)"  
-  and "steps0 (1, l, r) B stp = (0, l', r')"
-  shows "steps0 (Suc (length A div 2), l, r)  (A |+| B) stp = (0, l', r')"
-using tm_comp_second_same[OF assms] by (simp)
-
-end
-
--- a/thys/turing_hoare.thy	Thu Feb 07 06:39:06 2013 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,159 +0,0 @@
-theory turing_hoare
-imports turing_basic
-begin
-
-
-type_synonym assert = "tape \<Rightarrow> bool"
-
-definition 
-  assert_imp :: "assert \<Rightarrow> assert \<Rightarrow> bool" ("_ \<mapsto> _" [0, 0] 100)
-where
-  "P \<mapsto> Q \<equiv> \<forall>l r. P (l, r) \<longrightarrow> Q (l, r)"
-
-lemma [intro, simp]:
-  "P \<mapsto> P"
-unfolding assert_imp_def by simp
-
-fun 
-  holds_for :: "(tape \<Rightarrow> bool) \<Rightarrow> config \<Rightarrow> bool" ("_ holds'_for _" [100, 99] 100)
-where
-  "P holds_for (s, l, r) = P (l, r)"  
-
-lemma is_final_holds[simp]:
-  assumes "is_final c"
-  shows "Q holds_for (steps c p n) = Q holds_for c"
-using assms 
-apply(induct n)
-apply(auto)
-apply(case_tac [!] c)
-apply(auto)
-done
-
-(* Hoare Rules *)
-
-(* halting case *)
-definition
-  Hoare_halt :: "assert \<Rightarrow> tprog0 \<Rightarrow> assert \<Rightarrow> bool" ("({(1_)}/ (_)/ {(1_)})" 50)
-where
-  "{P} p {Q} \<equiv> \<forall>tp. P tp \<longrightarrow> (\<exists>n. is_final (steps0 (1, tp) p n) \<and> Q holds_for (steps0 (1, tp) p n))"
-
-
-(* not halting case *)
-definition
-  Hoare_unhalt :: "assert \<Rightarrow> tprog0 \<Rightarrow> bool" ("({(1_)}/ (_)) \<up>" 50)
-where
-  "{P} p \<up> \<equiv> \<forall>tp. P tp \<longrightarrow> (\<forall> n . \<not> (is_final (steps0 (1, tp) p n)))"
-
-
-lemma Hoare_haltI:
-  assumes "\<And>l r. P (l, r) \<Longrightarrow> \<exists>n. is_final (steps0 (1, (l, r)) p n) \<and> Q holds_for (steps0 (1, (l, r)) p n)"
-  shows "{P} p {Q}"
-unfolding Hoare_halt_def 
-using assms by auto
-
-lemma Hoare_unhaltI:
-  assumes "\<And>l r n. P (l, r) \<Longrightarrow> \<not> is_final (steps0 (1, (l, r)) p n)"
-  shows "{P} p \<up>"
-unfolding Hoare_unhalt_def 
-using assms by auto
-
-
-
-
-text {*
-  {P} A {Q}   {Q} B {S}  A well-formed
-  -----------------------------------------
-  {P} A |+| B {S}
-*}
-
-
-lemma Hoare_plus_halt [case_names A_halt B_halt A_wf]: 
-  assumes A_halt : "{P} A {Q}"
-  and B_halt : "{Q} B {S}"
-  and A_wf : "tm_wf (A, 0)"
-  shows "{P} A |+| B {S}"
-proof(rule Hoare_haltI)
-  fix l r
-  assume h: "P (l, r)"
-  then obtain n1 l' r' 
-    where "is_final (steps0 (1, l, r) A n1)"  
-      and a1: "Q holds_for (steps0 (1, l, r) A n1)"
-      and a2: "steps0 (1, l, r) A n1 = (0, l', r')"
-    using A_halt unfolding Hoare_halt_def
-    by (metis is_final_eq surj_pair) 
-  then obtain n2 
-    where "steps0 (1, l, r) (A |+| B) n2 = (Suc (length A div 2), l', r')"
-    using A_wf by (rule_tac tm_comp_pre_halt_same) 
-  moreover
-  from a1 a2 have "Q (l', r')" by (simp)
-  then obtain n3 l'' r''
-    where "is_final (steps0 (1, l', r') B n3)" 
-    and b1: "S holds_for (steps0 (1, l', r') B n3)"
-    and b2: "steps0 (1, l', r') B n3 = (0, l'', r'')"
-    using B_halt unfolding Hoare_halt_def 
-    by (metis is_final_eq surj_pair) 
-  then have "steps0 (Suc (length A div 2), l', r')  (A |+| B) n3 = (0, l'', r'')"
-    using A_wf by (rule_tac tm_comp_second_halt_same) 
-  ultimately show 
-    "\<exists>n. is_final (steps0 (1, l, r) (A |+| B) n) \<and> S holds_for (steps0 (1, l, r) (A |+| B) n)"
-    using b1 b2 by (rule_tac x = "n2 + n3" in exI) (simp)
-qed
-
-text {*
-  {P} A {Q}   {Q} B loops   A well-formed
-  ------------------------------------------
-          {P} A |+| B  loops
-*}
-
-lemma Hoare_plus_unhalt [case_names A_halt B_unhalt A_wf]:
-  assumes A_halt: "{P} A {Q}"
-  and B_uhalt: "{Q} B \<up>"
-  and A_wf : "tm_wf (A, 0)"
-  shows "{P} (A |+| B) \<up>"
-proof(rule_tac Hoare_unhaltI)
-  fix n l r 
-  assume h: "P (l, r)"
-  then obtain n1 l' r'
-    where a: "is_final (steps0 (1, l, r) A n1)" 
-    and b: "Q holds_for (steps0 (1, l, r) A n1)"
-    and c: "steps0 (1, l, r) A n1 = (0, l', r')"
-    using A_halt unfolding Hoare_halt_def 
-    by (metis is_final_eq surj_pair) 
-  then obtain n2 where eq: "steps0 (1, l, r) (A |+| B) n2 = (Suc (length A div 2), l', r')"
-    using A_wf by (rule_tac tm_comp_pre_halt_same)
-  then show "\<not> is_final (steps0 (1, l, r) (A |+| B) n)"
-  proof(cases "n2 \<le> n")
-    case True
-    from b c have "Q (l', r')" by simp
-    then have "\<forall> n. \<not> is_final (steps0 (1, l', r') B n)  "
-      using B_uhalt unfolding Hoare_unhalt_def by simp
-    then have "\<not> is_final (steps0 (1, l', r') B (n - n2))" by auto
-    then obtain s'' l'' r'' 
-      where "steps0 (1, l', r') B (n - n2) = (s'', l'', r'')" 
-      and "\<not> is_final (s'', l'', r'')" by (metis surj_pair)
-    then have "steps0 (Suc (length A div 2), l', r') (A |+| B) (n - n2) = (s''+ length A div 2, l'', r'')"
-      using A_wf by (auto dest: tm_comp_second_same simp del: tm_wf.simps)
-    then have "\<not> is_final (steps0 (1, l, r) (A |+| B) (n2 + (n  - n2)))"
-      using A_wf by (simp only: steps_add eq) (simp add: tm_wf.simps)
-    then show "\<not> is_final (steps0 (1, l, r) (A |+| B) n)" 
-      using `n2 \<le> n` by simp
-  next 
-    case False
-    then obtain n3 where "n = n2 - n3"
-      by (metis diff_le_self le_imp_diff_is_add nat_add_commute nat_le_linear)
-    moreover
-    with eq show "\<not> is_final (steps0 (1, l, r) (A |+| B) n)"
-      by (simp add: not_is_final[where ?n1.0="n2"])
-  qed
-qed
-
-lemma Hoare_consequence:
-  assumes "P' \<mapsto> P" "{P} p {Q}" "Q \<mapsto> Q'"
-  shows "{P'} p {Q'}"
-using assms
-unfolding Hoare_halt_def assert_imp_def
-by (metis holds_for.simps surj_pair)
-
-
-
-end
\ No newline at end of file
--- a/thys/uncomputable.thy	Thu Feb 07 06:39:06 2013 +0000
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1178 +0,0 @@
-(* Title: Turing machine's definition and its charater
-   Author: XuJian <xujian817@hotmail.com>
-   Maintainer: Xujian
-*)
-
-header {* Undeciablity of the {\em Halting problem} *}
-
-theory uncomputable
-imports Main turing_hoare
-begin
-
-lemma numeral:
-  shows "1 = Suc 0"
-  and "2 = Suc 1"
-  and "3 = Suc 2"
-  and "4 = Suc 3" 
-  and "5 = Suc 4" 
-  and "6 = Suc 5" 
-  and "7 = Suc 6"
-  and "8 = Suc 7" 
-  and "9 = Suc 8" 
-  and "10 = Suc 9" 
-  by arith+
-
-text {*
-  The {\em Copying} TM, which duplicates its input. 
-*}
-
-definition 
-  tcopy_begin :: "instr list"
-where
-  "tcopy_begin \<equiv> [(W0, 0), (R, 2), (R, 3), (R, 2),
-                 (W1, 3), (L, 4), (L, 4), (L, 0)]"
-
-definition 
-  tcopy_loop :: "instr list"
-where
-  "tcopy_loop \<equiv> [(R, 0), (R, 2),  (R, 3), (W0, 2),
-                 (R, 3), (R, 4), (W1, 5), (R, 4),
-                 (L, 6), (L, 5), (L, 6), (L, 1)]"
-
-definition 
-  tcopy_end :: "instr list"
-where
-  "tcopy_end \<equiv> [(L, 0), (R, 2), (W1, 3), (L, 4),
-                (R, 2), (R, 2), (L, 5), (W0, 4),
-                (R, 0), (L, 5)]"
-
-definition 
-  tcopy :: "instr list"
-where
-  "tcopy \<equiv> (tcopy_begin |+| tcopy_loop) |+| tcopy_end"
-
-
-(* tcopy_begin *)
-
-fun 
-  inv_begin0 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_begin1 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_begin2 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_begin3 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_begin4 :: "nat \<Rightarrow> tape \<Rightarrow> bool"
-where
-  "inv_begin0 n (l, r) = ((n > 1 \<and> (l, r) = (Oc \<up> (n - 2), [Oc, Oc, Bk, Oc])) \<or>   
-                          (n = 1 \<and> (l, r) = ([], [Bk, Oc, Bk, Oc])))"
-| "inv_begin1 n (l, r) = ((l, r) = ([], Oc \<up> n))"
-| "inv_begin2 n (l, r) = (\<exists> i j. i > 0 \<and> i + j = n \<and> (l, r) = (Oc \<up> i, Oc \<up> j))"
-| "inv_begin3 n (l, r) = (n > 0 \<and> (l, tl r) = (Bk # Oc \<up> n, []))"
-| "inv_begin4 n (l, r) = (n > 0 \<and> (l, r) = (Oc \<up> n, [Bk, Oc]) \<or> (l, r) = (Oc \<up> (n - 1), [Oc, Bk, Oc]))"
-
-fun inv_begin :: "nat \<Rightarrow> config \<Rightarrow> bool"
-  where
-  "inv_begin n (s, tp) = 
-        (if s = 0 then inv_begin0 n tp else
-         if s = 1 then inv_begin1 n tp else
-         if s = 2 then inv_begin2 n tp else
-         if s = 3 then inv_begin3 n tp else
-         if s = 4 then inv_begin4 n tp 
-         else False)"
-
-lemma [elim]: "\<lbrakk>0 < i; 0 < j\<rbrakk> \<Longrightarrow> 
-  \<exists>ia>0. ia + j - Suc 0 = i + j \<and> Oc # Oc \<up> i = Oc \<up> ia"
-by (rule_tac x = "Suc i" in exI, simp)
-
-lemma inv_begin_step: 
-  assumes "inv_begin n cf"
-  and "n > 0"
-  shows "inv_begin n (step0 cf tcopy_begin)"
-using assms
-unfolding tcopy_begin_def
-apply(case_tac cf)
-apply(auto simp: numeral split: if_splits)
-apply(case_tac "hd c")
-apply(auto)
-apply(case_tac c)
-apply(simp_all)
-done
-
-lemma inv_begin_steps: 
-  assumes "inv_begin n cf"
-  and "n > 0"
-  shows "inv_begin n (steps0 cf tcopy_begin stp)"
-apply(induct stp)
-apply(simp add: assms)
-apply(auto simp del: steps.simps)
-apply(rule_tac inv_begin_step)
-apply(simp_all add: assms)
-done
-
-lemma begin_partial_correctness:
-  assumes "is_final (steps0 (1, [], Oc \<up> n) tcopy_begin stp)"
-  shows "0 < n \<Longrightarrow> {inv_begin1 n} tcopy_begin {inv_begin0 n}"
-proof(rule_tac Hoare_haltI)
-  fix l r
-  assume h: "0 < n" "inv_begin1 n (l, r)"
-  have "inv_begin n (steps0 (1, [], Oc \<up> n) tcopy_begin stp)"
-    using h by (rule_tac inv_begin_steps) (simp_all add: inv_begin.simps)
-  then show
-    "\<exists>stp. is_final (steps0 (1, l, r) tcopy_begin stp) \<and> 
-    inv_begin0 n holds_for steps (1, l, r) (tcopy_begin, 0) stp"
-    using h assms
-    apply(rule_tac x = stp in exI)
-    apply(case_tac "(steps0 (1, [], Oc \<up> n) tcopy_begin stp)", simp add: inv_begin.simps)
-    done
-qed
-
-fun measure_begin_state :: "config \<Rightarrow> nat"
-  where
-  "measure_begin_state (s, l, r) = (if s = 0 then 0 else 5 - s)"
-
-fun measure_begin_step :: "config \<Rightarrow> nat"
-  where
-  "measure_begin_step (s, l, r) = 
-        (if s = 2 then length r else
-         if s = 3 then (if r = [] \<or> r = [Bk] then 1 else 0) else
-         if s = 4 then length l 
-         else 0)"
-
-definition
-  "measure_begin = measures [measure_begin_state, measure_begin_step]"
-
-lemma wf_measure_begin:
-  shows "wf measure_begin" 
-unfolding measure_begin_def 
-by auto
-
-lemma measure_begin_induct [case_names Step]: 
-  "\<lbrakk>\<And>n. \<not> P (f n) \<Longrightarrow> (f (Suc n), (f n)) \<in> measure_begin\<rbrakk> \<Longrightarrow> \<exists>n. P (f n)"
-using wf_measure_begin
-by (metis wf_iff_no_infinite_down_chain)
-
-lemma begin_halts: 
-  assumes h: "x > 0"
-  shows "\<exists> stp. is_final (steps0 (1, [], Oc \<up> x) tcopy_begin stp)"
-proof (induct rule: measure_begin_induct) 
-  case (Step n)
-  have "\<not> is_final (steps0 (1, [], Oc \<up> x) tcopy_begin n)" by fact
-  moreover
-  have "inv_begin x (steps0 (1, [], Oc \<up> x) tcopy_begin n)"
-    by (rule_tac inv_begin_steps) (simp_all add: inv_begin.simps h)
-  moreover
-  obtain s l r where eq: "(steps0 (1, [], Oc \<up> x) tcopy_begin n) = (s, l, r)"
-    by (metis measure_begin_state.cases)
-  ultimately 
-  have "(step0 (s, l, r) tcopy_begin, s, l, r) \<in> measure_begin"
-    apply(auto simp: measure_begin_def tcopy_begin_def numeral split: if_splits)
-    apply(subgoal_tac "r = [Oc]")
-    apply(auto)
-    by (metis cell.exhaust list.exhaust tl.simps(2))
-  then 
-  show "(steps0 (1, [], Oc \<up> x) tcopy_begin (Suc n), steps0 (1, [], Oc \<up> x) tcopy_begin n) \<in> measure_begin"
-    using eq by (simp only: step_red)
-qed
-
-lemma begin_correct: 
-  shows "0 < n \<Longrightarrow> {inv_begin1 n} tcopy_begin {inv_begin0 n}"
-using begin_partial_correctness begin_halts by blast
-
-declare tm_comp.simps [simp del] 
-declare adjust.simps[simp del] 
-declare shift.simps[simp del]
-declare tm_wf.simps[simp del]
-declare step.simps[simp del]
-declare steps.simps[simp del]
-
-(* tcopy_loop *)
-
-fun 
-  inv_loop1_loop :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_loop1_exit :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_loop5_loop :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_loop5_exit :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_loop6_loop :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_loop6_exit :: "nat \<Rightarrow> tape \<Rightarrow> bool"
-where
-  "inv_loop1_loop n (l, r) = (\<exists> i j. i + j + 1 = n \<and> (l, r) = (Oc\<up>i, Oc#Oc#Bk\<up>j @ Oc\<up>j) \<and> j > 0)"
-| "inv_loop1_exit n (l, r) = (0 < n \<and> (l, r) = ([], Bk#Oc#Bk\<up>n @ Oc\<up>n))"
-| "inv_loop5_loop x (l, r) = 
-     (\<exists> i j k t. i + j = Suc x \<and> i > 0 \<and> j > 0 \<and> k + t = j \<and> t > 0 \<and> (l, r) = (Oc\<up>k@Bk\<up>j@Oc\<up>i, Oc\<up>t))"
-| "inv_loop5_exit x (l, r) = 
-     (\<exists> i j. i + j = Suc x \<and> i > 0 \<and> j > 0 \<and> (l, r) = (Bk\<up>(j - 1)@Oc\<up>i, Bk # Oc\<up>j))"
-| "inv_loop6_loop x (l, r) = 
-     (\<exists> i j k t. i + j = Suc x \<and> i > 0 \<and> k + t + 1 = j \<and> (l, r) = (Bk\<up>k @ Oc\<up>i, Bk\<up>(Suc t) @ Oc\<up>j))"
-| "inv_loop6_exit x (l, r) = 
-     (\<exists> i j. i + j = x \<and> j > 0 \<and> (l, r) = (Oc\<up>i, Oc#Bk\<up>j @ Oc\<up>j))"
-
-fun 
-  inv_loop0 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_loop1 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_loop2 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_loop3 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_loop4 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_loop5 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_loop6 :: "nat \<Rightarrow> tape \<Rightarrow> bool"
-where
-  "inv_loop0 n (l, r) =  (0 < n \<and> (l, r) = ([Bk], Oc # Bk\<up>n @ Oc\<up>n))"
-| "inv_loop1 n (l, r) = (inv_loop1_loop n (l, r) \<or> inv_loop1_exit n (l, r))"
-| "inv_loop2 n (l, r) = (\<exists> i j any. i + j = n \<and> n > 0 \<and> i > 0 \<and> j > 0 \<and> (l, r) = (Oc\<up>i, any#Bk\<up>j@Oc\<up>j))"
-| "inv_loop3 n (l, r) = 
-     (\<exists> i j k t. i + j = n \<and> i > 0 \<and> j > 0 \<and>  k + t = Suc j \<and> (l, r) = (Bk\<up>k@Oc\<up>i, Bk\<up>t@Oc\<up>j))"
-| "inv_loop4 n (l, r) = 
-     (\<exists> i j k t. i + j = n \<and> i > 0 \<and> j > 0 \<and>  k + t = j \<and> (l, r) = (Oc\<up>k @ Bk\<up>(Suc j)@Oc\<up>i, Oc\<up>t))"
-| "inv_loop5 n (l, r) = (inv_loop5_loop n (l, r) \<or> inv_loop5_exit n (l, r))"
-| "inv_loop6 n (l, r) = (inv_loop6_loop n (l, r) \<or> inv_loop6_exit n (l, r))"
-
-fun inv_loop :: "nat \<Rightarrow> config \<Rightarrow> bool"
-  where
-  "inv_loop x (s, l, r) = 
-         (if s = 0 then inv_loop0 x (l, r)
-          else if s = 1 then inv_loop1 x (l, r)
-          else if s = 2 then inv_loop2 x (l, r)
-          else if s = 3 then inv_loop3 x (l, r)
-          else if s = 4 then inv_loop4 x (l, r)
-          else if s = 5 then inv_loop5 x (l, r)
-          else if s = 6 then inv_loop6 x (l, r)
-          else False)"
-       
-declare inv_loop.simps[simp del] inv_loop1.simps[simp del]
-        inv_loop2.simps[simp del] inv_loop3.simps[simp del] 
-        inv_loop4.simps[simp del] inv_loop5.simps[simp del] 
-        inv_loop6.simps[simp del]
-
-lemma [elim]: "Bk # list = Oc \<up> t \<Longrightarrow> RR"
-by (case_tac t, auto)
-
-lemma [simp]: "inv_loop1 x (b, []) = False"
-by (simp add: inv_loop1.simps)
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop2 x (b, [])\<rbrakk> \<Longrightarrow> inv_loop3 x (Bk # b, [])"
-by (auto simp: inv_loop2.simps inv_loop3.simps)
-
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop3 x (b, [])\<rbrakk> \<Longrightarrow> inv_loop3 x (Bk # b, [])"
-by (auto simp: inv_loop3.simps)
-
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop4 x (b, [])\<rbrakk> \<Longrightarrow> inv_loop5 x (b, [Oc])"
-apply(auto simp: inv_loop4.simps inv_loop5.simps)
-apply(rule_tac [!] x = i in exI, 
-      rule_tac [!] x  = "Suc j" in exI, simp_all)
-done
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop5 x ([], [])\<rbrakk> \<Longrightarrow> RR"
-by (auto simp: inv_loop4.simps inv_loop5.simps)
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop5 x (b, []); b \<noteq> []\<rbrakk> \<Longrightarrow> RR"
-by (auto simp: inv_loop4.simps inv_loop5.simps)
-
-lemma [elim]: "inv_loop6 x ([], []) \<Longrightarrow> RR"
-by (auto simp: inv_loop6.simps)
-
-lemma [elim]: "inv_loop6 x (b, []) \<Longrightarrow> RR" 
-by (auto simp: inv_loop6.simps)
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop1 x (b, Bk # list)\<rbrakk> \<Longrightarrow> b = []"
-by (auto simp: inv_loop1.simps)
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop1 x (b, Bk # list)\<rbrakk> \<Longrightarrow> list = Oc # Bk \<up> x @ Oc \<up> x"
-by (auto simp: inv_loop1.simps)
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop2 x (b, Bk # list)\<rbrakk> \<Longrightarrow> inv_loop3 x (Bk # b, list)"
-apply(auto simp: inv_loop2.simps inv_loop3.simps)
-apply(rule_tac [!] x = i  in exI, rule_tac [!] x = j in exI, simp_all)
-done
-
-lemma [elim]: "Bk # list = Oc \<up> j \<Longrightarrow> RR"
-by (case_tac j, simp_all)
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop3 x (b, Bk # list)\<rbrakk> \<Longrightarrow> inv_loop3 x (Bk # b, list)"
-apply(auto simp: inv_loop3.simps)
-apply(rule_tac [!] x = i in exI, 
-      rule_tac [!] x = j in exI, simp_all)
-apply(case_tac [!] t, auto)
-done
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop4 x (b, Bk # list)\<rbrakk> \<Longrightarrow> inv_loop5 x (b, Oc # list)"
-by (auto simp: inv_loop4.simps inv_loop5.simps)
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop5 x ([], Bk # list)\<rbrakk> \<Longrightarrow> inv_loop6 x ([], Bk # Bk # list)"
-by (auto simp: inv_loop6.simps inv_loop5.simps)
-
-lemma [simp]: "inv_loop5_loop x (b, Bk # list) = False"
-by (auto simp: inv_loop5.simps)
-
-lemma [simp]: "inv_loop6_exit x (b, Bk # list) = False"
-by (auto simp: inv_loop6.simps)
-
-declare inv_loop5_loop.simps[simp del]  inv_loop5_exit.simps[simp del]
-       inv_loop6_loop.simps[simp del]  inv_loop6_exit.simps[simp del]
-
-lemma [elim]:"\<lbrakk>0 < x; inv_loop5_exit x (b, Bk # list); b \<noteq> []; hd b = Bk\<rbrakk> 
-          \<Longrightarrow> inv_loop6_loop x (tl b, Bk # Bk # list)"
-apply(simp only: inv_loop5_exit.simps inv_loop6_loop.simps )
-apply(erule_tac exE)+
-apply(rule_tac x = i in exI, 
-      rule_tac x = j in exI,
-      rule_tac x = "j - Suc (Suc 0)" in exI, 
-      rule_tac x = "Suc 0" in exI, auto)
-apply(case_tac [!] j, simp_all)
-apply(case_tac [!] nat, simp_all)
-done
-
-lemma [simp]: "inv_loop6_loop x (b, Oc # Bk # list) = False"
-by (auto simp: inv_loop6_loop.simps)
-
-lemma [elim]: "\<lbrakk>x > 0; inv_loop5_exit x (b, Bk # list); b \<noteq> []; hd b = Oc\<rbrakk> \<Longrightarrow> 
-  inv_loop6_exit x (tl b, Oc # Bk # list)"
-apply(simp only: inv_loop5_exit.simps inv_loop6_exit.simps)
-apply(erule_tac exE)+
-apply(rule_tac x = "x - 1" in exI, rule_tac x = 1 in exI, simp)
-apply(case_tac j, auto)
-apply(case_tac [!] nat, auto)
-done
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop5 x (b, Bk # list); b \<noteq> []\<rbrakk> \<Longrightarrow> inv_loop6 x (tl b, hd b # Bk # list)"
-apply(simp add: inv_loop5.simps inv_loop6.simps)
-apply(case_tac "hd b", simp_all, auto)
-done
-
-lemma  [simp]: "inv_loop6 x ([], Bk # xs) = False"
-apply(simp add: inv_loop6.simps inv_loop6_loop.simps
-                inv_loop6_exit.simps)
-apply(auto)
-done
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop6 x ([], Bk # list)\<rbrakk> \<Longrightarrow> inv_loop6 x ([], Bk # Bk # list)"
-by (simp)
-
-lemma [simp]: "inv_loop6_exit x (b, Bk # list) = False"
-by (simp add: inv_loop6_exit.simps)
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop6_loop x (b, Bk # list); b \<noteq> []; hd b = Bk\<rbrakk>
-              \<Longrightarrow> inv_loop6_loop x (tl b, Bk # Bk # list)"
-apply(simp only: inv_loop6_loop.simps)
-apply(erule_tac exE)+
-apply(rule_tac x = i in exI, rule_tac x = j in exI, 
-      rule_tac x = "k - 1" in exI, rule_tac x = "Suc t" in exI, auto)
-apply(case_tac [!] k, auto)
-done
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop6_loop x (b, Bk # list); b \<noteq> []; hd b = Oc\<rbrakk> 
-        \<Longrightarrow> inv_loop6_exit x (tl b, Oc # Bk # list)"
-apply(simp only: inv_loop6_loop.simps inv_loop6_exit.simps)
-apply(erule_tac exE)+
-apply(rule_tac x = "i - 1" in exI, rule_tac x = j in exI, auto)
-apply(case_tac [!] k, auto)
-done
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop6 x (b, Bk # list); b \<noteq> []\<rbrakk> \<Longrightarrow> inv_loop6 x (tl b, hd b # Bk # list)"
-apply(simp add: inv_loop6.simps)
-apply(case_tac "hd b", simp_all, auto)
-done
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop1 x (b, Oc # list)\<rbrakk> \<Longrightarrow> inv_loop2 x (Oc # b, list)"
-apply(auto simp: inv_loop1.simps inv_loop2.simps)
-apply(rule_tac x = "Suc i" in exI, auto)
-done
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop2 x (b, Oc # list)\<rbrakk> \<Longrightarrow> inv_loop2 x (b, Bk # list)"
-by (auto simp: inv_loop2.simps)
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop3 x (b, Oc # list)\<rbrakk> \<Longrightarrow> inv_loop4 x (Oc # b, list)"
-apply(auto simp: inv_loop3.simps inv_loop4.simps)
-apply(rule_tac [!] x = i in exI, auto)
-apply(rule_tac [!] x = "Suc 0" in exI, rule_tac [!] x = "j - 1" in exI, auto)
-apply(case_tac [!] t, auto)
-apply(case_tac [!] j, auto)
-done
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop4 x (b, Oc # list)\<rbrakk> \<Longrightarrow> inv_loop4 x (Oc # b, list)"
-apply(auto simp: inv_loop4.simps)
-apply(rule_tac [!] x = "i" in exI, auto)
-apply(rule_tac [!] x = "Suc k" in exI, rule_tac [!] x = "t - 1" in exI, auto)
-apply(case_tac [!] t, simp_all)
-done
-
-lemma [simp]: "inv_loop5 x ([], list) = False"
-by (auto simp: inv_loop5.simps inv_loop5_exit.simps inv_loop5_loop.simps)
-
-lemma [simp]: "inv_loop5_exit x (b, Oc # list) = False"
-by (auto simp: inv_loop5_exit.simps)
-
-lemma [elim]: " \<lbrakk>inv_loop5_loop x (b, Oc # list); b \<noteq> []; hd b = Bk\<rbrakk>
-  \<Longrightarrow> inv_loop5_exit x (tl b, Bk # Oc # list)"
-apply(simp only: inv_loop5_loop.simps inv_loop5_exit.simps)
-apply(erule_tac exE)+
-apply(rule_tac x = i in exI, auto)
-apply(case_tac [!] k, auto)
-done
-
-lemma [elim]: "\<lbrakk>inv_loop5_loop x (b, Oc # list); b \<noteq> []; hd b = Oc\<rbrakk> 
-           \<Longrightarrow> inv_loop5_loop x (tl b, Oc # Oc # list)"
-apply(simp only:  inv_loop5_loop.simps)
-apply(erule_tac exE)+
-apply(rule_tac x = i in exI, rule_tac x = j in exI)
-apply(rule_tac x = "k - 1" in exI, rule_tac x = "Suc t" in exI, auto)
-apply(case_tac [!] k, auto)
-done
-
-lemma [elim]: "\<lbrakk>inv_loop5 x (b, Oc # list); b \<noteq> []\<rbrakk> \<Longrightarrow> inv_loop5 x (tl b, hd b # Oc # list)"
-apply(simp add: inv_loop5.simps)
-apply(case_tac "hd b", simp_all, auto)
-done
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop6 x ([], Oc # list)\<rbrakk> \<Longrightarrow> inv_loop1 x ([], Bk # Oc # list)"
-apply(auto simp: inv_loop6.simps inv_loop1.simps 
-  inv_loop6_loop.simps inv_loop6_exit.simps)
-done
-
-lemma [elim]: "\<lbrakk>0 < x; inv_loop6 x (b, Oc # list); b \<noteq> []\<rbrakk> 
-           \<Longrightarrow> inv_loop1 x (tl b, hd b # Oc # list)"
-apply(auto simp: inv_loop6.simps inv_loop1.simps 
-  inv_loop6_loop.simps inv_loop6_exit.simps)
-done
-
-lemma inv_loop_step: 
-  "\<lbrakk>inv_loop x cf; x > 0\<rbrakk> \<Longrightarrow> inv_loop x (step cf (tcopy_loop, 0))"
-apply(case_tac cf, case_tac c, case_tac [2] aa)
-apply(auto simp: inv_loop.simps step.simps tcopy_loop_def numeral split: if_splits)
-done
-
-lemma inv_loop_steps:
-  "\<lbrakk>inv_loop x cf; x > 0\<rbrakk> \<Longrightarrow> inv_loop x (steps cf (tcopy_loop, 0) stp)"
-apply(induct stp, simp add: steps.simps, simp)
-apply(erule_tac inv_loop_step, simp)
-done
-
-fun loop_stage :: "config \<Rightarrow> nat"
-  where
-  "loop_stage (s, l, r) = (if s = 0 then 0
-                           else (Suc (length (takeWhile (\<lambda>a. a = Oc) (rev l @ r)))))"
-
-fun loop_state :: "config \<Rightarrow> nat"
-  where
-  "loop_state (s, l, r) = (if s = 2 \<and> hd r = Oc then 0
-                           else if s = 1 then 1
-                           else 10 - s)"
-
-fun loop_step :: "config \<Rightarrow> nat"
-  where
-  "loop_step (s, l, r) = (if s = 3 then length r
-                          else if s = 4 then length r
-                          else if s = 5 then length l 
-                          else if s = 6 then length l
-                          else 0)"
-
-definition measure_loop :: "(config \<times> config) set"
-  where
-   "measure_loop = measures [loop_stage, loop_state, loop_step]"
-
-lemma wf_measure_loop: "wf measure_loop"
-unfolding measure_loop_def
-by (auto)
-
-lemma measure_loop_induct [case_names Step]: 
-  "\<lbrakk>\<And>n. \<not> P (f n) \<Longrightarrow> (f (Suc n), (f n)) \<in> measure_loop\<rbrakk> \<Longrightarrow> \<exists>n. P (f n)"
-using wf_measure_loop
-by (metis wf_iff_no_infinite_down_chain)
-
-
-
-lemma [simp]: "inv_loop2 x ([], b) = False"
-by (auto simp: inv_loop2.simps)
-
-lemma  [simp]: "inv_loop2 x (l', []) = False"
-by (auto simp: inv_loop2.simps)
-
-lemma [simp]: "inv_loop3 x (b, []) = False"
-by (auto simp: inv_loop3.simps)
-
-lemma [simp]: "inv_loop4 x ([], b) = False"
-by (auto simp: inv_loop4.simps)
-
-
-lemma [elim]: 
-  "\<lbrakk>inv_loop4 x (l', []); l' \<noteq> []; x > 0;
-  length (takeWhile (\<lambda>a. a = Oc) (rev l' @ [Oc])) \<noteq> 
-  length (takeWhile (\<lambda>a. a = Oc) (rev l'))\<rbrakk>
-  \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (rev l' @ [Oc])) < length (takeWhile (\<lambda>a. a = Oc) (rev l'))"
-apply(auto simp: inv_loop4.simps)
-apply(case_tac [!] j, simp_all add: List.takeWhile_tail)
-done
-
-
-lemma [elim]: 
-  "\<lbrakk>inv_loop4 x (l', Bk # list); l' \<noteq> []; 0 < x;
-   length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Oc # list)) \<noteq> 
-    length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Bk # list))\<rbrakk>
-    \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Oc # list)) < 
-    length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Bk # list))"
-by (auto simp: inv_loop4.simps)
-
-lemma takeWhile_replicate_append: 
-  "P a \<Longrightarrow> takeWhile P (a\<up>x @ ys) = a\<up>x @ takeWhile P ys"
-by (induct x, auto)
-
-lemma takeWhile_replicate: 
-  "P a \<Longrightarrow> takeWhile P (a\<up>x) = a\<up>x"
-by (induct x, auto)
-
-lemma [elim]: 
-   "\<lbrakk>inv_loop5 x (l', Bk # list); l' \<noteq> []; 0 < x;
-   length (takeWhile (\<lambda>a. a = Oc) (rev (tl l') @ hd l' # Bk # list)) \<noteq>
-   length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Bk # list))\<rbrakk>
-   \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (rev (tl l') @ hd l' # Bk # list)) < 
-   length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Bk # list))"
-apply(auto simp: inv_loop5.simps inv_loop5_exit.simps)
-apply(case_tac [!] j, simp_all)
-apply(case_tac [!] "nat", simp_all)
-apply(case_tac  nata, simp_all add: List.takeWhile_tail)
-apply(simp add: takeWhile_replicate_append takeWhile_replicate)
-apply(case_tac  nata, simp_all add: List.takeWhile_tail)
-done
-
-lemma [elim]: "\<lbrakk>inv_loop1 x (l', Oc # list)\<rbrakk> \<Longrightarrow> hd list = Oc"
-by (auto simp: inv_loop1.simps)
-
-lemma [elim]: 
-  "\<lbrakk>inv_loop6 x (l', Bk # list); l' \<noteq> []; 0 < x;
-  length (takeWhile (\<lambda>a. a = Oc) (rev (tl l') @ hd l' # Bk # list)) \<noteq> 
-  length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Bk # list))\<rbrakk>
-  \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (rev (tl l') @ hd l' # Bk # list)) < 
-  length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Bk # list))"
-apply(auto simp: inv_loop6.simps)
-apply(case_tac l', simp_all)
-done
-
-lemma [elim]:
-  "\<lbrakk>inv_loop2 x (l', Oc # list); l' \<noteq> []; 0 < x\<rbrakk> \<Longrightarrow> 
-  length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Bk # list)) <
-  length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Oc # list))"
-apply(auto simp: inv_loop2.simps)
-apply(simp_all add: takeWhile_tail takeWhile_replicate_append
-                takeWhile_replicate)
-done
-
-lemma [elim]: 
-  "\<lbrakk>inv_loop5 x (l', Oc # list); l' \<noteq> []; 0 < x;
-  length (takeWhile (\<lambda>a. a = Oc) (rev (tl l') @ hd l' # Oc # list)) \<noteq> 
-  length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Oc # list))\<rbrakk>
-  \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (rev (tl l') @ hd l' # Oc # list)) < 
-  length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Oc # list))"
-apply(auto simp: inv_loop5.simps)
-apply(case_tac l', auto)
-done
-
-lemma[elim]: 
-  "\<lbrakk>inv_loop6 x (l', Oc # list); l' \<noteq> []; 0 < x;
-  length (takeWhile (\<lambda>a. a = Oc) (rev (tl l') @ hd l' # Oc # list))
-  \<noteq> length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Oc # list))\<rbrakk>
-  \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (rev (tl l') @ hd l' # Oc # list)) < 
-      length (takeWhile (\<lambda>a. a = Oc) (rev l' @ Oc # list))"
-apply(case_tac l')
-apply(auto simp: inv_loop6.simps)
-done
-
-lemma loop_halts: 
-  assumes h: "n > 0" "inv_loop n (1, l, r)"
-  shows "\<exists> stp. is_final (steps0 (1, l, r) tcopy_loop stp)"
-proof (induct rule: measure_loop_induct) 
-  case (Step stp)
-  have "\<not> is_final (steps0 (1, l, r) tcopy_loop stp)" by fact
-  moreover
-  have "inv_loop n (steps0 (1, l, r) tcopy_loop stp)"
-    by (rule_tac inv_loop_steps) (simp_all only: h)
-  moreover
-  obtain s l' r' where eq: "(steps0 (1, l, r) tcopy_loop stp) = (s, l', r')"
-    by (metis measure_begin_state.cases)
-  ultimately 
-  have "(step0 (s, l', r') tcopy_loop, s, l', r') \<in> measure_loop"
-    using h(1)
-    apply(case_tac r')
-    apply(case_tac [2] a)
-    apply(auto simp: inv_loop.simps step.simps tcopy_loop_def numeral measure_loop_def split: if_splits)
-    done
-  then 
-  show "(steps0 (1, l, r) tcopy_loop (Suc stp), steps0 (1, l, r) tcopy_loop stp) \<in> measure_loop"
-    using eq by (simp only: step_red)
-qed
-
-lemma loop_correct:
-  shows "0 < n \<Longrightarrow> {inv_loop1 n} tcopy_loop {inv_loop0 n}"
-  using assms
-proof(rule_tac Hoare_haltI)
-  fix l r
-  assume h: "0 < n" "inv_loop1 n (l, r)"
-  then obtain stp where k: "is_final (steps0 (1, l, r) tcopy_loop stp)" 
-    using loop_halts
-    apply(simp add: inv_loop.simps)
-    apply(blast)
-    done
-  moreover
-  have "inv_loop n (steps0 (1, l, r) tcopy_loop stp)"
-    using h 
-    by (rule_tac inv_loop_steps) (simp_all add: inv_loop.simps)
-  ultimately show
-    "\<exists>stp. is_final (steps0 (1, l, r) tcopy_loop stp) \<and> 
-    inv_loop0 n holds_for steps0 (1, l, r) tcopy_loop stp"
-    using h(1) 
-    apply(rule_tac x = stp in exI)
-    apply(case_tac "(steps0 (1, l, r) tcopy_loop stp)")
-    apply(simp add: inv_loop.simps)
-    done
-qed
-
-
-
-
-(* tcopy_end *)
-
-fun
-  inv_end5_loop :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_end5_exit :: "nat \<Rightarrow> tape \<Rightarrow> bool" 
-where  
-  "inv_end5_loop x (l, r) = 
-     (\<exists> i j. i + j = x \<and> x > 0 \<and> j > 0 \<and> l = Oc\<up>i @ [Bk] \<and> r = Oc\<up>j @ Bk # Oc\<up>x)"
-| "inv_end5_exit x (l, r) = (x > 0 \<and> l = [] \<and> r = Bk # Oc\<up>x @ Bk # Oc\<up>x)"
-
-fun 
-  inv_end0 :: "nat \<Rightarrow> tape \<Rightarrow>  bool" and
-  inv_end1 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_end2 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_end3 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and
-  inv_end4 :: "nat \<Rightarrow> tape \<Rightarrow> bool" and 
-  inv_end5 :: "nat \<Rightarrow> tape \<Rightarrow> bool" 
-where
-  "inv_end0 n (l, r) = (n > 0 \<and> (l, r) = ([Bk], Oc\<up>n @ Bk # Oc\<up>n))"
-| "inv_end1 n (l, r) = (n > 0 \<and> (l, r) = ([Bk], Oc # Bk\<up>n @ Oc\<up>n))"
-| "inv_end2 n (l, r) = (\<exists> i j. i + j = Suc n \<and> n > 0 \<and> l = Oc\<up>i @ [Bk] \<and> r = Bk\<up>j @ Oc\<up>n)"
-| "inv_end3 n (l, r) =
-     (\<exists> i j. n > 0 \<and> i + j = n \<and> l = Oc\<up>i @ [Bk] \<and> r = Oc # Bk\<up>j@ Oc\<up>n)"
-| "inv_end4 n (l, r) = (\<exists> any. n > 0 \<and> l = Oc\<up>n @ [Bk] \<and> r = any#Oc\<up>n)"
-| "inv_end5 n (l, r) = (inv_end5_loop n (l, r) \<or> inv_end5_exit n (l, r))"
-
-fun 
-  inv_end :: "nat \<Rightarrow> config \<Rightarrow> bool"
-where
-  "inv_end n (s, l, r) = (if s = 0 then inv_end0 n (l, r)
-                          else if s = 1 then inv_end1 n (l, r)
-                          else if s = 2 then inv_end2 n (l, r)
-                          else if s = 3 then inv_end3 n (l, r)
-                          else if s = 4 then inv_end4 n (l, r)
-                          else if s = 5 then inv_end5 n (l, r)
-                          else False)"
-
-declare inv_end.simps[simp del] inv_end1.simps[simp del]
-        inv_end0.simps[simp del] inv_end2.simps[simp del]
-        inv_end3.simps[simp del] inv_end4.simps[simp del]
-        inv_end5.simps[simp del]
-
-lemma [simp]:  "inv_end1 x (b, []) = False"
-by (auto simp: inv_end1.simps)
-
-lemma [simp]: "inv_end2 x (b, []) = False"
-by (auto simp: inv_end2.simps)
-
-lemma [simp]: "inv_end3 x (b, []) = False"
-by (auto simp: inv_end3.simps)
-
-lemma [simp]: "inv_end4 x (b, []) = False"
-by (auto simp: inv_end4.simps)
-
-lemma [simp]: "inv_end5 x (b, []) = False"
-by (auto simp: inv_end5.simps)
-
-lemma [simp]: "inv_end1 x ([], list) = False"
-by (auto simp: inv_end1.simps)
-
-lemma [elim]: "\<lbrakk>0 < x; inv_end1 x (b, Bk # list); b \<noteq> []\<rbrakk>
-  \<Longrightarrow> inv_end0 x (tl b, hd b # Bk # list)"
-by (auto simp: inv_end1.simps inv_end0.simps)
-
-lemma [elim]: "\<lbrakk>0 < x; inv_end2 x (b, Bk # list)\<rbrakk> 
-  \<Longrightarrow> inv_end3 x (b, Oc # list)"
-apply(auto simp: inv_end2.simps inv_end3.simps)
-apply(rule_tac x = "j - 1" in exI)
-apply(case_tac j, simp_all)
-apply(case_tac x, simp_all)
-done
-
-lemma [elim]: "\<lbrakk>0 < x; inv_end3 x (b, Bk # list)\<rbrakk> \<Longrightarrow> inv_end2 x (Bk # b, list)"
-by (auto simp: inv_end2.simps inv_end3.simps)
-  
-lemma [elim]: "\<lbrakk>0 < x; inv_end4 x ([], Bk # list)\<rbrakk> \<Longrightarrow> 
-  inv_end5 x ([], Bk # Bk # list)"
-by (auto simp: inv_end4.simps inv_end5.simps)
- 
-lemma [elim]: "\<lbrakk>0 < x; inv_end4 x (b, Bk # list); b \<noteq> []\<rbrakk> \<Longrightarrow> 
-  inv_end5 x (tl b, hd b # Bk # list)"
-apply(auto simp: inv_end4.simps inv_end5.simps)
-apply(rule_tac x = 1 in exI, simp)
-done
-
-lemma [elim]: "\<lbrakk>0 < x; inv_end5 x (b, Bk # list)\<rbrakk> \<Longrightarrow> inv_end0 x (Bk # b, list)"
-apply(auto simp: inv_end5.simps inv_end0.simps)
-apply(case_tac [!] j, simp_all)
-done
-
-lemma [elim]: "\<lbrakk>0 < x; inv_end1 x (b, Oc # list)\<rbrakk> \<Longrightarrow> inv_end2 x (Oc # b, list)"
-by (auto simp: inv_end1.simps inv_end2.simps)
-
-lemma [elim]: "\<lbrakk>0 < x; inv_end2 x ([], Oc # list)\<rbrakk> \<Longrightarrow>
-               inv_end4 x ([], Bk # Oc # list)"
-by (auto simp: inv_end2.simps inv_end4.simps)
-
-lemma [elim]:  "\<lbrakk>0 < x; inv_end2 x (b, Oc # list); b \<noteq> []\<rbrakk> \<Longrightarrow>
-  inv_end4 x (tl b, hd b # Oc # list)"
-apply(auto simp: inv_end2.simps inv_end4.simps)
-apply(case_tac [!] j, simp_all)
-done
-
-lemma [elim]: "\<lbrakk>0 < x; inv_end3 x (b, Oc # list)\<rbrakk> \<Longrightarrow> inv_end2 x (Oc # b, list)"
-by (auto simp: inv_end2.simps inv_end3.simps)
-
-lemma [elim]: "\<lbrakk>0 < x; inv_end4 x (b, Oc # list)\<rbrakk> \<Longrightarrow> inv_end4 x (b, Bk # list)"
-by (auto simp: inv_end2.simps inv_end4.simps)
-
-lemma [elim]: "\<lbrakk>0 < x; inv_end5 x ([], Oc # list)\<rbrakk> \<Longrightarrow> inv_end5 x ([], Bk # Oc # list)"
-by (auto simp: inv_end2.simps inv_end5.simps)
-
-declare inv_end5_loop.simps[simp del]
-        inv_end5_exit.simps[simp del]
-
-lemma [simp]: "inv_end5_exit x (b, Oc # list) = False"
-by (auto simp: inv_end5_exit.simps)
-
-lemma [simp]: "inv_end5_loop x (tl b, Bk # Oc # list) = False"
-apply(auto simp: inv_end5_loop.simps)
-apply(case_tac [!] j, simp_all)
-done
-
-lemma [elim]:
-  "\<lbrakk>0 < x; inv_end5_loop x (b, Oc # list); b \<noteq> []; hd b = Bk\<rbrakk> \<Longrightarrow>
-  inv_end5_exit x (tl b, Bk # Oc # list)"
-apply(auto simp: inv_end5_loop.simps inv_end5_exit.simps)
-apply(case_tac [!] i, simp_all)
-done
-
-lemma [elim]: 
-  "\<lbrakk>0 < x; inv_end5_loop x (b, Oc # list); b \<noteq> []; hd b = Oc\<rbrakk> \<Longrightarrow>
-  inv_end5_loop x (tl b, Oc # Oc # list)"
-apply(simp only: inv_end5_loop.simps inv_end5_exit.simps)
-apply(erule_tac exE)+
-apply(rule_tac x = "i - 1" in exI, 
-      rule_tac x = "Suc j" in exI, auto)
-apply(case_tac [!] i, simp_all)
-done
-
-lemma [elim]: "\<lbrakk>0 < x; inv_end5 x (b, Oc # list); b \<noteq> []\<rbrakk> \<Longrightarrow> 
-  inv_end5 x (tl b, hd b # Oc # list)"
-apply(simp add: inv_end2.simps inv_end5.simps)
-apply(case_tac "hd b", simp_all, auto)
-done
-
-lemma inv_end_step:
-  "\<lbrakk>x > 0; inv_end x cf\<rbrakk> \<Longrightarrow> inv_end x (step cf (tcopy_end, 0))"
-apply(case_tac cf, case_tac c, case_tac [2] aa)
-apply(auto simp: inv_end.simps step.simps tcopy_end_def numeral split: if_splits)
-done
-
-lemma inv_end_steps:
-  "\<lbrakk>x > 0; inv_end x cf\<rbrakk> \<Longrightarrow> inv_end x (steps cf (tcopy_end, 0) stp)"
-apply(induct stp, simp add:steps.simps, simp)
-apply(erule_tac inv_end_step, simp)
-done
-
-fun end_state :: "config \<Rightarrow> nat"
-  where
-  "end_state (s, l, r) = 
-       (if s = 0 then 0
-        else if s = 1 then 5
-        else if s = 2 \<or> s = 3 then 4
-        else if s = 4 then 3 
-        else if s = 5 then 2
-        else 0)"
-
-fun end_stage :: "config \<Rightarrow> nat"
-  where
-  "end_stage (s, l, r) = 
-          (if s = 2 \<or> s = 3 then (length r) else 0)"
-
-fun end_step :: "config \<Rightarrow> nat"
-  where
-  "end_step (s, l, r) = 
-         (if s = 4 then (if hd r = Oc then 1 else 0)
-          else if s = 5 then length l
-          else if s = 2 then 1
-          else if s = 3 then 0
-          else 0)"
-
-definition end_LE :: "(config \<times> config) set"
-  where
-   "end_LE = measures [end_state, end_stage, end_step]"
-
-lemma wf_end_le: "wf end_LE"
-unfolding end_LE_def
-by (auto)
-
-lemma [simp]: "inv_end5 x ([], Oc # list) = False"
-by (auto simp: inv_end5.simps inv_end5_loop.simps)
-
-lemma halt_lemma: 
-  "\<lbrakk>wf LE; \<forall>n. (\<not> P (f n) \<longrightarrow> (f (Suc n), (f n)) \<in> LE)\<rbrakk> \<Longrightarrow> \<exists>n. P (f n)"
-by (metis wf_iff_no_infinite_down_chain)
-
-lemma end_halt: 
-  "\<lbrakk>x > 0; inv_end x (Suc 0, l, r)\<rbrakk> \<Longrightarrow> 
-      \<exists> stp. is_final (steps (Suc 0, l, r) (tcopy_end, 0) stp)"
-proof(rule_tac LE = end_LE in halt_lemma)
-  show "wf end_LE" by(intro wf_end_le)
-next
-  assume great: "0 < x"
-    and inv_start: "inv_end x (Suc 0, l, r)"
-  show "\<forall>n. \<not> is_final (steps (Suc 0, l, r) (tcopy_end, 0) n) \<longrightarrow> 
-    (steps (Suc 0, l, r) (tcopy_end, 0) (Suc n), steps (Suc 0, l, r) (tcopy_end, 0) n) \<in> end_LE"
-  proof(rule_tac allI, rule_tac impI)
-    fix n
-    assume notfinal: "\<not> is_final (steps (Suc 0, l, r) (tcopy_end, 0) n)"
-    obtain s' l' r' where d: "steps (Suc 0, l, r) (tcopy_end, 0) n = (s', l', r')"
-      apply(case_tac "steps (Suc 0, l, r) (tcopy_end, 0) n", auto)
-      done
-    hence "inv_end x (s', l', r') \<and> s' \<noteq> 0"
-      using great inv_start notfinal
-      apply(drule_tac stp = n in inv_end_steps, auto)
-      done
-    hence "(step (s', l', r') (tcopy_end, 0), s', l', r') \<in> end_LE"
-      apply(case_tac r', case_tac [2] a)
-      apply(auto simp: inv_end.simps step.simps tcopy_end_def numeral end_LE_def split: if_splits)
-      done
-    thus "(steps (Suc 0, l, r) (tcopy_end, 0) (Suc n), 
-      steps (Suc 0, l, r) (tcopy_end, 0) n) \<in> end_LE"
-      using d
-      by simp
-  qed
-qed
-
-lemma end_correct:
-  "n > 0 \<Longrightarrow> {inv_end1 n} tcopy_end {inv_end0 n}"
-proof(rule_tac Hoare_haltI)
-  fix l r
-  assume h: "0 < n"
-    "inv_end1 n (l, r)"
-  then have "\<exists> stp. is_final (steps0 (1, l, r) tcopy_end stp)"
-    by (simp add: end_halt inv_end.simps)
-  then obtain stp where "is_final (steps0 (1, l, r) tcopy_end stp)" ..
-  moreover have "inv_end n (steps0 (1, l, r) tcopy_end stp)"
-    apply(rule_tac inv_end_steps)
-    using h by(simp_all add: inv_end.simps)
-  ultimately show
-    "\<exists>stp. is_final (steps (1, l, r) (tcopy_end, 0) stp) \<and> 
-    inv_end0 n holds_for steps (1, l, r) (tcopy_end, 0) stp"        
-    using h
-    apply(rule_tac x = stp in exI)
-    apply(case_tac "(steps0 (1, l, r) tcopy_end stp)") 
-    apply(simp add: inv_end.simps)
-    done
-qed
-
-(* tcopy *)
-
-lemma [intro]: "tm_wf (tcopy_begin, 0)"
-by (auto simp: tm_wf.simps tcopy_begin_def)
-
-lemma [intro]: "tm_wf (tcopy_loop, 0)"
-by (auto simp: tm_wf.simps tcopy_loop_def)
-
-lemma [intro]: "tm_wf (tcopy_end, 0)"
-by (auto simp: tm_wf.simps tcopy_end_def)
-
-lemma tcopy_correct1: 
-  assumes "0 < x"
-  shows "{inv_begin1 x} tcopy {inv_end0 x}"
-proof -
-  have "{inv_begin1 x} tcopy_begin {inv_begin0 x}"
-    by (metis assms begin_correct) 
-  moreover 
-  have "inv_begin0 x \<mapsto> inv_loop1 x"
-    unfolding assert_imp_def
-    unfolding inv_begin0.simps inv_loop1.simps
-    unfolding inv_loop1_loop.simps inv_loop1_exit.simps
-    apply(auto simp add: numeral Cons_eq_append_conv)
-    by (rule_tac x = "Suc 0" in exI, auto)
-  ultimately have "{inv_begin1 x} tcopy_begin {inv_loop1 x}"
-    by (rule_tac Hoare_consequence) (auto)
-  moreover
-  have "{inv_loop1 x} tcopy_loop {inv_loop0 x}"
-    by (metis assms loop_correct) 
-  ultimately 
-  have "{inv_begin1 x} (tcopy_begin |+| tcopy_loop) {inv_loop0 x}"
-    by (rule_tac Hoare_plus_halt) (auto)
-  moreover 
-  have "{inv_end1 x} tcopy_end {inv_end0 x}"
-    by (metis assms end_correct) 
-  moreover 
-  have "inv_loop0 x = inv_end1 x"
-    by(auto simp: inv_end1.simps inv_loop1.simps assert_imp_def)
-  ultimately 
-  show "{inv_begin1 x} tcopy {inv_end0 x}"
-    unfolding tcopy_def
-    by (rule_tac Hoare_plus_halt) (auto)
-qed
-
-abbreviation (input)
-  "pre_tcopy n \<equiv> \<lambda>tp. tp = ([]::cell list, <(n::nat)>)"
-abbreviation (input)
-  "post_tcopy n \<equiv> \<lambda>tp. tp= ([Bk], <(n, n::nat)>)"
-
-lemma tcopy_correct:
-  shows "{pre_tcopy n} tcopy {post_tcopy n}"
-proof -
-  have "{inv_begin1 (Suc n)} tcopy {inv_end0 (Suc n)}"
-    by (rule tcopy_correct1) (simp)
-  moreover
-  have "pre_tcopy n = inv_begin1 (Suc n)" 
-    by (auto simp add: tape_of_nl_abv tape_of_nat_abv)
-  moreover
-  have "inv_end0 (Suc n) = post_tcopy n" 
-    by (auto simp add: inv_end0.simps tape_of_nat_abv tape_of_nat_pair)
-  ultimately
-  show "{pre_tcopy n} tcopy {post_tcopy n}" 
-    by simp
-qed
-
-
-section {* The {\em Dithering} Turing Machine *}
-
-text {*
-  The {\em Dithering} TM, when the input is @{text "1"}, it will loop forever, otherwise, it will
-  terminate.
-*}
-
-definition dither :: "instr list"
-  where
-  "dither \<equiv> [(W0, 1), (R, 2), (L, 1), (L, 0)] "
-
-(* invariants of dither *)
-abbreviation (input)
-  "dither_halt_inv \<equiv> \<lambda>tp. \<exists>k. tp = (Bk \<up> k, <1::nat>)"
-
-abbreviation (input)
-  "dither_unhalt_inv \<equiv> \<lambda>tp. \<exists>k. tp = (Bk \<up> k, <0::nat>)"
-
-lemma dither_loops_aux: 
-  "(steps0 (1, Bk \<up> m, [Oc]) dither stp = (1, Bk \<up> m, [Oc])) \<or> 
-   (steps0 (1, Bk \<up> m, [Oc]) dither stp = (2, Oc # Bk \<up> m, []))"
-  apply(induct stp)
-  apply(auto simp: steps.simps step.simps dither_def numeral tape_of_nat_abv)
-  done
-
-lemma dither_loops:
-  shows "{dither_unhalt_inv} dither \<up>" 
-apply(rule Hoare_unhaltI)
-using dither_loops_aux
-apply(auto simp add: numeral tape_of_nat_abv)
-by (metis Suc_neq_Zero is_final_eq)
-
-lemma dither_halts_aux: 
-  shows "steps0 (1, Bk \<up> m, [Oc, Oc]) dither 2 = (0, Bk \<up> m, [Oc, Oc])"
-unfolding dither_def
-by (simp add: steps.simps step.simps numeral)
-
-lemma dither_halts:
-  shows "{dither_halt_inv} dither {dither_halt_inv}" 
-apply(rule Hoare_haltI)
-using dither_halts_aux
-apply(auto simp add: tape_of_nat_abv)
-by (metis (lifting, mono_tags) holds_for.simps is_final_eq prod.cases)
-
-
-section {* The diagnal argument below shows the undecidability of Halting problem *}
-
-text {*
-  @{text "haltP tp x"} means TM @{text "tp"} terminates on input @{text "x"}
-  and the final configuration is standard.
-*}
-
-definition haltP :: "tprog0 \<Rightarrow> nat list \<Rightarrow> bool"
-  where
-  "haltP p ns \<equiv> {(\<lambda>tp. tp = ([], <ns>))} p {(\<lambda>tp. (\<exists>k n l. tp = (Bk \<up> k,  <n::nat> @ Bk \<up> l)))}"
-
-lemma [intro, simp]: "tm_wf0 tcopy"
-by (auto simp: tcopy_def)
-
-lemma [intro, simp]: "tm_wf0 dither"
-by (auto simp: tm_wf.simps dither_def)
-
-text {*
-  The following locale specifies that TM @{text "H"} can be used to solve 
-  the {\em Halting Problem} and @{text "False"} is going to be derived 
-  under this locale. Therefore, the undecidability of {\em Halting Problem}
-  is established. 
-*}
-
-locale uncomputable = 
-  -- {* The coding function of TM, interestingly, the detailed definition of this 
-  funciton @{text "code"} does not affect the final result. *}
-  fixes code :: "instr list \<Rightarrow> nat" 
-  -- {* 
-  The TM @{text "H"} is the one which is assummed being able to solve the Halting problem.
-  *}
-  and H :: "instr list"
-  assumes h_wf[intro]: "tm_wf0 H"
-  -- {*
-  The following two assumptions specifies that @{text "H"} does solve the Halting problem.
-  *}
-  and h_case: 
-  "\<And> M ns. haltP M ns \<Longrightarrow> {(\<lambda>tp. tp = ([Bk], <(code M, ns)>))} H {(\<lambda>tp. \<exists>k. tp = (Bk \<up> k, <0::nat>))}"
-  and nh_case: 
-  "\<And> M ns. \<not> haltP M ns \<Longrightarrow> {(\<lambda>tp. tp = ([Bk], <(code M, ns)>))} H {(\<lambda>tp. \<exists>k. tp = (Bk \<up> k, <1::nat>))}"
-begin
-
-(* invariants for H *)
-abbreviation (input)
-  "pre_H_inv M ns \<equiv> \<lambda>tp. tp = ([Bk], <(code M, ns::nat list)>)"
-
-abbreviation (input)
-  "post_H_halt_inv \<equiv> \<lambda>tp. \<exists>k. tp = (Bk \<up> k, <1::nat>)"
-
-abbreviation (input)
-  "post_H_unhalt_inv \<equiv> \<lambda>tp. \<exists>k. tp = (Bk \<up> k, <0::nat>)"
-
-
-lemma H_halt_inv:
-  assumes "\<not> haltP M ns" 
-  shows "{pre_H_inv M ns} H {post_H_halt_inv}"
-using assms nh_case by auto
-
-lemma H_unhalt_inv:
-  assumes "haltP M ns" 
-  shows "{pre_H_inv M ns} H {post_H_unhalt_inv}"
-using assms h_case by auto
-   
-(* TM that produces the contradiction and its code *)
-
-definition
-  "tcontra \<equiv> (tcopy |+| H) |+| dither"
-abbreviation
-  "code_tcontra \<equiv> code tcontra"
-
-(* assume tcontra does not halt on its code *)
-lemma tcontra_unhalt: 
-  assumes "\<not> haltP tcontra [code tcontra]"
-  shows "False"
-proof -
-  (* invariants *)
-  def P1 \<equiv> "\<lambda>tp. tp = ([]::cell list, <code_tcontra>)"
-  def P2 \<equiv> "\<lambda>tp. tp = ([Bk], <(code_tcontra, code_tcontra)>)"
-  def P3 \<equiv> "\<lambda>tp. \<exists>k. tp = (Bk \<up> k, <1::nat>)"
-
-  (*
-  {P1} tcopy {P2}  {P2} H {P3} 
-  ----------------------------
-     {P1} (tcopy |+| H) {P3}     {P3} dither {P3}
-  ------------------------------------------------
-                 {P1} tcontra {P3}
-  *)
-
-  have H_wf: "tm_wf0 (tcopy |+| H)" by auto
-
-  (* {P1} (tcopy |+| H) {P3} *)
-  have first: "{P1} (tcopy |+| H) {P3}" 
-  proof (cases rule: Hoare_plus_halt)
-    case A_halt (* of tcopy *)
-    show "{P1} tcopy {P2}" unfolding P1_def P2_def 
-      by (rule tcopy_correct)
-  next
-    case B_halt (* of H *)
-    show "{P2} H {P3}"
-      unfolding P2_def P3_def 
-      using H_halt_inv[OF assms]
-      by (simp add: tape_of_nat_pair tape_of_nl_abv)
-  qed (simp)
-
-  (* {P3} dither {P3} *)
-  have second: "{P3} dither {P3}" unfolding P3_def 
-    by (rule dither_halts)
-  
-  (* {P1} tcontra {P3} *)
-  have "{P1} tcontra {P3}" 
-    unfolding tcontra_def
-    by (rule Hoare_plus_halt[OF first second H_wf])
-
-  with assms show "False"
-    unfolding P1_def P3_def
-    unfolding haltP_def
-    unfolding Hoare_halt_def 
-    apply(auto)    
-    apply(drule_tac x = n in spec)
-    apply(case_tac "steps0 (Suc 0, [], <code tcontra>) tcontra n")
-    apply(auto simp add: tape_of_nl_abv)
-    by (metis append_Nil2 replicate_0)
-qed
-
-(* asumme tcontra halts on its code *)
-lemma tcontra_halt: 
-  assumes "haltP tcontra [code tcontra]"
-  shows "False"
-proof - 
-  (* invariants *)
-  def P1 \<equiv> "\<lambda>tp. tp = ([]::cell list, <code_tcontra>)"
-  def P2 \<equiv> "\<lambda>tp. tp = ([Bk], <(code_tcontra, code_tcontra)>)"
-  def Q3 \<equiv> "\<lambda>tp. \<exists>k. tp = (Bk \<up> k, <0::nat>)"
-
-  (*
-  {P1} tcopy {P2}  {P2} H {Q3} 
-  ----------------------------
-     {P1} (tcopy |+| H) {Q3}     {Q3} dither loops
-  ------------------------------------------------
-               {P1} tcontra loops
-  *)
-
-  have H_wf: "tm_wf0 (tcopy |+| H)" by auto
-
-  (* {P1} (tcopy |+| H) {Q3} *)
-  have first: "{P1} (tcopy |+| H) {Q3}" 
-  proof (cases rule: Hoare_plus_halt)
-    case A_halt (* of tcopy *)
-    show "{P1} tcopy {P2}" unfolding P1_def P2_def
-      by (rule tcopy_correct)
-  next
-    case B_halt (* of H *)
-    then show "{P2} H {Q3}"
-      unfolding P2_def Q3_def using H_unhalt_inv[OF assms]
-      by(simp add: tape_of_nat_pair tape_of_nl_abv)
-  qed (simp)
-
-  (* {P3} dither loops *)
-  have second: "{Q3} dither \<up>" unfolding Q3_def 
-    by (rule dither_loops)
-  
-  (* {P1} tcontra loops *)
-  have "{P1} tcontra \<up>" 
-    unfolding tcontra_def
-    by (rule Hoare_plus_unhalt[OF first second H_wf])
-
-  with assms show "False"
-    unfolding P1_def
-    unfolding haltP_def
-    unfolding Hoare_halt_def Hoare_unhalt_def
-    by (auto simp add: tape_of_nl_abv)
-qed
-
-      
-text {*
-  @{text "False"} can finally derived.
-*}
-
-lemma false: "False"
-using tcontra_halt tcontra_unhalt 
-by auto
-
-end
-
-declare replicate_Suc[simp del]
-
-
-end
-