thys/Recursive.thy
changeset 229 d8e6f0798e23
parent 204 e55c8e5da49f
child 230 49dcc0b9b0b3
equal deleted inserted replaced
228:e9ef4ada308b 229:d8e6f0798e23
     1 (* Title: thys/Recursive.thy
       
     2    Author: Jian Xu, Xingyuan Zhang, and Christian Urban
       
     3 *)
       
     4 
       
     5 header {* Translating Recursive Functions into Abacus Machines *}
       
     6 
       
     7 theory Recursive
     1 theory Recursive
     8 imports Rec_Def Abacus
     2 imports Abacus Rec_Def Abacus_Hoare
     9 begin
     3 begin
    10 
       
    11 
       
    12 text {*
       
    13   Some auxilliary Abacus machines used to construct the result Abacus machines.
       
    14 *}
       
    15 
       
    16 text {*
       
    17   @{text "get_paras_num recf"} returns the arity of recursive function @{text "recf"}.
       
    18 *}
       
    19 fun get_paras_num :: "recf \<Rightarrow> nat"
       
    20   where
       
    21   "get_paras_num z = 1" |
       
    22   "get_paras_num s = 1" |
       
    23   "get_paras_num (id m n) = m" |
       
    24   "get_paras_num (Cn n f gs) = n" |
       
    25   "get_paras_num (Pr n f g) = Suc n"  |
       
    26   "get_paras_num (Mn n f) = n"  
       
    27 
     4 
    28 fun addition :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
     5 fun addition :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
    29   where
     6   where
    30   "addition m n p = [Dec m 4, Inc n, Inc p, Goto 0, Dec p 7, Inc m, Goto 4]"
     7   "addition m n p = [Dec m 4, Inc n, Inc p, Goto 0, Dec p 7, Inc m, Goto 4]"
    31 
     8 
    32 fun mv_box :: "nat \<Rightarrow> nat \<Rightarrow> abc_prog"
     9 fun mv_box :: "nat \<Rightarrow> nat \<Rightarrow> abc_prog"
    33   where
    10   where
    34   "mv_box m n = [Dec m 3, Inc n, Goto 0]"
    11   "mv_box m n = [Dec m 3, Inc n, Goto 0]"
    35 
       
    36 fun abc_inst_shift :: "abc_inst \<Rightarrow> nat \<Rightarrow> abc_inst"
       
    37   where
       
    38   "abc_inst_shift (Inc m) n = Inc m" |
       
    39   "abc_inst_shift (Dec m e) n = Dec m (e + n)" |
       
    40   "abc_inst_shift (Goto m) n = Goto (m + n)"
       
    41 
       
    42 fun abc_shift :: "abc_inst list \<Rightarrow> nat \<Rightarrow> abc_inst list" 
       
    43   where
       
    44   "abc_shift xs n = map (\<lambda> x. abc_inst_shift x n) xs" 
       
    45 
       
    46 fun abc_append :: "abc_inst list \<Rightarrow> abc_inst list \<Rightarrow> 
       
    47                            abc_inst list" (infixl "[+]" 60)
       
    48   where
       
    49   "abc_append al bl = (let al_len = length al in 
       
    50                            al @ abc_shift bl al_len)"
       
    51 
    12 
    52 text {* The compilation of @{text "z"}-operator. *}
    13 text {* The compilation of @{text "z"}-operator. *}
    53 definition rec_ci_z :: "abc_inst list"
    14 definition rec_ci_z :: "abc_inst list"
    54   where
    15   where
    55   "rec_ci_z \<equiv> [Goto 1]"
    16   "rec_ci_z \<equiv> [Goto 1]"
   122             (fprog @ [Dec (Suc n) (len + 5), Dec (Suc n) (len + 3),
    83             (fprog @ [Dec (Suc n) (len + 5), Dec (Suc n) (len + 3),
   123              Goto (len + 1), Inc n, Goto 0], n, max (Suc n) fn))"
    84              Goto (len + 1), Inc n, Goto 0], n, max (Suc n) fn))"
   124 
    85 
   125 declare rec_ci.simps [simp del] rec_ci_s_def[simp del] 
    86 declare rec_ci.simps [simp del] rec_ci_s_def[simp del] 
   126         rec_ci_z_def[simp del] rec_ci_id.simps[simp del]
    87         rec_ci_z_def[simp del] rec_ci_id.simps[simp del]
   127         mv_boxes.simps[simp del] abc_append.simps[simp del]
    88         mv_boxes.simps[simp del] 
   128         mv_box.simps[simp del] addition.simps[simp del]
    89         mv_box.simps[simp del] addition.simps[simp del]
   129   
    90   
   130 declare abc_steps_l.simps[simp del] abc_fetch.simps[simp del] 
    91 declare abc_steps_l.simps[simp del] abc_fetch.simps[simp del] 
   131         abc_step_l.simps[simp del] 
    92         abc_step_l.simps[simp del] 
   132 
    93 
   133 lemma abc_steps_add: 
    94 inductive_cases terminate_pr_reverse: "terminate (Pr n f g) xs"
   134   "abc_steps_l (as, lm) ap (m + n) = 
    95 
   135          abc_steps_l (abc_steps_l (as, lm) ap m) ap n"
    96 inductive_cases terminate_z_reverse[elim!]: "terminate z xs"
   136 apply(induct m arbitrary: n as lm, simp add: abc_steps_l.simps)
    97 
   137 proof -
    98 inductive_cases terminate_s_reverse[elim!]: "terminate s xs"
   138   fix m n as lm
    99 
   139   assume ind: 
   100 inductive_cases terminate_id_reverse[elim!]: "terminate (id m n) xs"
   140     "\<And>n as lm. abc_steps_l (as, lm) ap (m + n) = 
   101 
   141                    abc_steps_l (abc_steps_l (as, lm) ap m) ap n"
   102 inductive_cases terminate_cn_reverse[elim!]: "terminate (Cn n f gs) xs"
   142   show "abc_steps_l (as, lm) ap (Suc m + n) = 
   103 
   143              abc_steps_l (abc_steps_l (as, lm) ap (Suc m)) ap n"
   104 inductive_cases terminate_mn_reverse[elim!]:"terminate (Mn n f) xs"
   144     apply(insert ind[of as lm "Suc n"], simp)
       
   145     apply(insert ind[of as lm "Suc 0"], simp add: abc_steps_l.simps)
       
   146     apply(case_tac "(abc_steps_l (as, lm) ap m)", simp)
       
   147     apply(simp add: abc_steps_l.simps)
       
   148     apply(case_tac "abc_step_l (a, b) (abc_fetch a ap)", 
       
   149           simp add: abc_steps_l.simps)
       
   150     done
       
   151 qed
       
   152 
       
   153 (*lemmas: rec_ci and rec_calc_rel*)
       
   154 
       
   155 lemma rec_calc_inj_case_z: 
       
   156   "\<lbrakk>rec_calc_rel z l x; rec_calc_rel z l y\<rbrakk> \<Longrightarrow> x = y"
       
   157 apply(auto elim: calc_z_reverse)
       
   158 done
       
   159 
       
   160 lemma  rec_calc_inj_case_s: 
       
   161   "\<lbrakk>rec_calc_rel s l x; rec_calc_rel s l y\<rbrakk> \<Longrightarrow> x = y"
       
   162 apply(auto elim: calc_s_reverse)
       
   163 done
       
   164 
       
   165 lemma rec_calc_inj_case_id:
       
   166   "\<lbrakk>rec_calc_rel (recf.id nat1 nat2) l x;
       
   167     rec_calc_rel (recf.id nat1 nat2) l y\<rbrakk> \<Longrightarrow> x = y"
       
   168 apply(auto elim: calc_id_reverse)
       
   169 done
       
   170 
       
   171 lemma rec_calc_inj_case_mn:
       
   172   assumes ind: "\<And> l x y. \<lbrakk>rec_calc_rel f l x; rec_calc_rel f l y\<rbrakk> 
       
   173            \<Longrightarrow> x = y" 
       
   174   and h: "rec_calc_rel (Mn n f) l x" "rec_calc_rel (Mn n f) l y"
       
   175   shows "x = y"
       
   176   apply(insert h)
       
   177   apply(elim  calc_mn_reverse)
       
   178   apply(case_tac "x > y", simp)
       
   179   apply(erule_tac x = "y" in allE, auto)
       
   180 proof -
       
   181   fix v va
       
   182   assume "rec_calc_rel f (l @ [y]) 0" 
       
   183     "rec_calc_rel f (l @ [y]) v"  
       
   184     "0 < v"
       
   185   thus "False"
       
   186     apply(insert ind[of "l @ [y]" 0 v], simp)
       
   187     done
       
   188 next
       
   189   fix v va
       
   190   assume 
       
   191     "rec_calc_rel f (l @ [x]) 0" 
       
   192     "\<forall>x<y. \<exists>v. rec_calc_rel f (l @ [x]) v \<and> 0 < v" "\<not> y < x"
       
   193   thus "x = y"
       
   194     apply(erule_tac x = "x" in allE)
       
   195     apply(case_tac "x = y", auto)
       
   196     apply(drule_tac y = v in ind, simp, simp)
       
   197     done
       
   198 qed 
       
   199 
       
   200 lemma rec_calc_inj_case_pr: 
       
   201   assumes f_ind: 
       
   202   "\<And>l x y. \<lbrakk>rec_calc_rel f l x; rec_calc_rel f l y\<rbrakk> \<Longrightarrow> x = y"
       
   203   and g_ind:
       
   204   "\<And>x xa y xb ya l xc yb. 
       
   205   \<lbrakk>x = rec_ci f; (xa, y) = x; (xb, ya) = y; 
       
   206   rec_calc_rel g l xc; rec_calc_rel g l yb\<rbrakk> \<Longrightarrow> xc = yb"
       
   207   and h: "rec_calc_rel (Pr n f g) l x" "rec_calc_rel (Pr n f g) l y"  
       
   208   shows "x = y"
       
   209   apply(case_tac "rec_ci f")
       
   210 proof -
       
   211   fix a b c
       
   212   assume "rec_ci f = (a, b, c)"
       
   213   hence ng_ind: 
       
   214     "\<And> l xc yb. \<lbrakk>rec_calc_rel g l xc; rec_calc_rel g l yb\<rbrakk>
       
   215     \<Longrightarrow> xc = yb"
       
   216     apply(insert g_ind[of "(a, b, c)" "a" "(b, c)" b c], simp)
       
   217     done
       
   218   from h show "x = y"
       
   219     apply(erule_tac calc_pr_reverse, erule_tac calc_pr_reverse)
       
   220     apply(erule f_ind, simp, simp)
       
   221     apply(erule_tac calc_pr_reverse, simp, simp)
       
   222   proof -
       
   223     fix la ya ry laa yaa rya
       
   224     assume k1:  "rec_calc_rel g (la @ [ya, ry]) x" 
       
   225       "rec_calc_rel g (la @ [ya, rya]) y"
       
   226       and k2: "rec_calc_rel (Pr (length la) f g) (la @ [ya]) ry"
       
   227               "rec_calc_rel (Pr (length la) f g) (la @ [ya]) rya"
       
   228     from k2 have "ry = rya"
       
   229       apply(induct ya arbitrary: ry rya)
       
   230       apply(erule_tac calc_pr_reverse, 
       
   231         erule_tac calc_pr_reverse, simp)
       
   232       apply(erule f_ind, simp, simp, simp)
       
   233       apply(erule_tac calc_pr_reverse, simp)
       
   234       apply(erule_tac rSucy = rya in calc_pr_reverse, simp, simp)
       
   235     proof -
       
   236       fix ya ry rya l y ryb laa yb ryc
       
   237       assume ind:
       
   238         "\<And>ry rya. \<lbrakk>rec_calc_rel (Pr (length l) f g) (l @ [y]) ry; 
       
   239                    rec_calc_rel (Pr (length l) f g) (l @ [y]) rya\<rbrakk> \<Longrightarrow> ry = rya"
       
   240         and j: "rec_calc_rel (Pr (length l) f g) (l @ [y]) ryb"
       
   241         "rec_calc_rel g (l @ [y, ryb]) ry" 
       
   242         "rec_calc_rel (Pr (length l) f g) (l @ [y]) ryc" 
       
   243         "rec_calc_rel g (l @ [y, ryc]) rya"
       
   244       from j show "ry = rya"
       
   245 	apply(insert ind[of ryb ryc], simp)
       
   246 	apply(insert ng_ind[of "l @ [y, ryc]" ry rya], simp)
       
   247 	done
       
   248     qed 
       
   249     from k1 and this show "x = y"
       
   250       apply(simp)
       
   251       apply(insert ng_ind[of "la @ [ya, rya]" x y], simp)
       
   252       done
       
   253   qed  
       
   254 qed
       
   255 
       
   256 lemma Suc_nth_part_eq:
       
   257   "\<forall>k<Suc (length list). (a # xs) ! k = (aa # list) ! k
       
   258        \<Longrightarrow> \<forall>k<(length list). (xs) ! k = (list) ! k"
       
   259 apply(rule allI, rule impI)
       
   260 apply(erule_tac x = "Suc k" in allE, simp)
       
   261 done
       
   262 
       
   263 
       
   264 lemma list_eq_intro:  
       
   265   "\<lbrakk>length xs = length ys; \<forall> k < length xs. xs ! k = ys ! k\<rbrakk> 
       
   266   \<Longrightarrow> xs = ys"
       
   267 apply(induct xs arbitrary: ys, simp)
       
   268 apply(case_tac ys, simp, simp)
       
   269 proof -
       
   270   fix a xs ys aa list
       
   271   assume ind: 
       
   272     "\<And>ys. \<lbrakk>length list = length ys; \<forall>k<length ys. xs ! k = ys ! k\<rbrakk>
       
   273     \<Longrightarrow> xs = ys"
       
   274     and h: "length xs = length list" 
       
   275     "\<forall>k<Suc (length list). (a # xs) ! k = (aa # list) ! k"
       
   276   from h show "a = aa \<and> xs = list"
       
   277     apply(insert ind[of list], simp)
       
   278     apply(frule Suc_nth_part_eq, simp)
       
   279     apply(erule_tac x = "0" in allE, simp)
       
   280     done
       
   281 qed
       
   282 
       
   283 lemma rec_calc_inj_case_cn: 
       
   284   assumes ind: 
       
   285   "\<And>x l xa y.
       
   286   \<lbrakk>x = f \<or> x \<in> set gs; rec_calc_rel x l xa; rec_calc_rel x l y\<rbrakk>
       
   287   \<Longrightarrow> xa = y"
       
   288   and h: "rec_calc_rel (Cn n f gs) l x" 
       
   289          "rec_calc_rel (Cn n f gs) l y"
       
   290   shows "x = y"
       
   291   apply(insert h, elim  calc_cn_reverse)
       
   292   apply(subgoal_tac "rs = rsa")
       
   293   apply(rule_tac x = f and l = rsa and xa = x and y = y in ind, 
       
   294         simp, simp, simp)
       
   295   apply(intro list_eq_intro, simp, rule allI, rule impI)
       
   296   apply(erule_tac x = k in allE, rule_tac x = k in allE, simp, simp)
       
   297   apply(rule_tac x = "gs ! k" in ind, simp, simp, simp)
       
   298   done
       
   299 
       
   300 lemma rec_calc_inj:
       
   301   "\<lbrakk>rec_calc_rel f l x; 
       
   302     rec_calc_rel f l y\<rbrakk> \<Longrightarrow> x = y"
       
   303 apply(induct f arbitrary: l x y rule: rec_ci.induct)
       
   304 apply(simp add: rec_calc_inj_case_z)
       
   305 apply(simp add: rec_calc_inj_case_s)
       
   306 apply(simp add: rec_calc_inj_case_id)
       
   307 apply (metis rec_calc_inj_case_cn)
       
   308 apply(erule rec_calc_inj_case_pr, auto)
       
   309 apply(erule rec_calc_inj_case_mn, auto)
       
   310 done
       
   311 
       
   312 
       
   313 lemma calc_rel_reverse_ind_step_ex: 
       
   314   "\<lbrakk>rec_calc_rel (Pr n f g) (lm @ [Suc x]) rs\<rbrakk> 
       
   315   \<Longrightarrow> \<exists> rs. rec_calc_rel (Pr n f g) (lm @ [x]) rs"
       
   316 apply(erule calc_pr_reverse, simp, simp)
       
   317 apply(rule_tac x = rk in exI, simp)
       
   318 done
       
   319 
       
   320 lemma [simp]: "Suc x \<le> y \<Longrightarrow> Suc (y - Suc x) = y - x"
       
   321 by arith
       
   322 
       
   323 lemma calc_pr_para_not_null: 
       
   324   "rec_calc_rel (Pr n f g) lm rs \<Longrightarrow> lm \<noteq> []"
       
   325 apply(erule calc_pr_reverse, simp, simp)
       
   326 done
       
   327 
       
   328 lemma calc_pr_less_ex: 
       
   329  "\<lbrakk>rec_calc_rel (Pr n f g) lm rs; x \<le> last lm\<rbrakk> \<Longrightarrow> 
       
   330  \<exists>rs. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rs"
       
   331 apply(subgoal_tac "lm \<noteq> []")
       
   332 apply(induct x, rule_tac x = rs in exI, simp, simp, erule exE)
       
   333 apply(rule_tac rs = xa in calc_rel_reverse_ind_step_ex, simp)
       
   334 apply(simp add: calc_pr_para_not_null)
       
   335 done
       
   336 
       
   337 lemma calc_pr_zero_ex:
       
   338   "rec_calc_rel (Pr n f g) lm rs \<Longrightarrow> 
       
   339              \<exists>rs. rec_calc_rel f (butlast lm) rs"
       
   340 apply(drule_tac x = "last lm" in calc_pr_less_ex, simp,
       
   341       erule_tac exE, simp)
       
   342 apply(erule_tac calc_pr_reverse, simp)
       
   343 apply(rule_tac x = rs in exI, simp, simp)
       
   344 done
       
   345 
       
   346 
       
   347 lemma abc_steps_ind: 
       
   348   "abc_steps_l (as, am) ap (Suc stp) =
       
   349           abc_steps_l (abc_steps_l (as, am) ap stp) ap (Suc 0)"
       
   350 apply(insert abc_steps_add[of as am ap stp "Suc 0"], simp)
       
   351 done
       
   352 
       
   353 lemma abc_steps_zero: "abc_steps_l asm ap 0 = asm"
       
   354 apply(case_tac asm, simp add: abc_steps_l.simps)
       
   355 done
       
   356 
       
   357 lemma abc_append_nth: 
       
   358   "n < length ap + length bp \<Longrightarrow> 
       
   359        (ap [+] bp) ! n =
       
   360          (if n < length ap then ap ! n 
       
   361           else abc_inst_shift (bp ! (n - length ap)) (length ap))"
       
   362 apply(simp add: abc_append.simps nth_append map_nth split: if_splits)
       
   363 done
       
   364 
       
   365 lemma abc_state_keep:  
       
   366   "as \<ge> length bp \<Longrightarrow> abc_steps_l (as, lm) bp stp = (as, lm)"
       
   367 apply(induct stp, simp add: abc_steps_zero)
       
   368 apply(simp add: abc_steps_ind)
       
   369 apply(simp add: abc_steps_zero)
       
   370 apply(simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps)
       
   371 done
       
   372 
       
   373 lemma abc_halt_equal: 
       
   374   "\<lbrakk>abc_steps_l (0, lm) bp stpa = (length bp, lm1); 
       
   375     abc_steps_l (0, lm) bp stpb = (length bp, lm2)\<rbrakk> \<Longrightarrow> lm1 = lm2"
       
   376 apply(case_tac "stpa - stpb > 0")
       
   377 apply(insert abc_steps_add[of 0 lm bp stpb "stpa - stpb"], simp)
       
   378 apply(insert abc_state_keep[of bp "length bp" lm2 "stpa - stpb"], 
       
   379       simp, simp add: abc_steps_zero)
       
   380 apply(insert abc_steps_add[of 0 lm bp stpa "stpb - stpa"], simp)
       
   381 apply(insert abc_state_keep[of bp "length bp" lm1 "stpb - stpa"], 
       
   382       simp)
       
   383 done  
       
   384 
       
   385 lemma abc_halt_point_ex: 
       
   386   "\<lbrakk>\<exists>stp. abc_steps_l (0, lm) bp stp = (bs, lm');
       
   387     bs = length bp; bp \<noteq> []\<rbrakk> 
       
   388   \<Longrightarrow> \<exists> stp. (\<lambda> (s, l). s < bs \<and> 
       
   389               (abc_steps_l (s, l) bp (Suc 0)) = (bs, lm')) 
       
   390       (abc_steps_l (0, lm) bp stp) "
       
   391 apply(erule_tac exE)
       
   392 proof -
       
   393   fix stp
       
   394   assume "bs = length bp" 
       
   395          "abc_steps_l (0, lm) bp stp = (bs, lm')" 
       
   396          "bp \<noteq> []"
       
   397   thus 
       
   398     "\<exists>stp. (\<lambda>(s, l). s < bs \<and> 
       
   399       abc_steps_l (s, l) bp (Suc 0) = (bs, lm')) 
       
   400                        (abc_steps_l (0, lm) bp stp)"
       
   401     apply(induct stp, simp add: abc_steps_zero, simp)
       
   402   proof -
       
   403     fix stpa
       
   404     assume ind: 
       
   405      "abc_steps_l (0, lm) bp stpa = (length bp, lm')
       
   406        \<Longrightarrow> \<exists>stp. (\<lambda>(s, l). s < length bp  \<and> abc_steps_l (s, l) bp 
       
   407              (Suc 0) = (length bp, lm')) (abc_steps_l (0, lm) bp stp)"
       
   408     and h: "abc_steps_l (0, lm) bp (Suc stpa) = (length bp, lm')" 
       
   409            "abc_steps_l (0, lm) bp stp = (length bp, lm')" 
       
   410            "bp \<noteq> []"
       
   411     from h show 
       
   412       "\<exists>stp. (\<lambda>(s, l). s < length bp \<and> abc_steps_l (s, l) bp (Suc 0)
       
   413                     = (length bp, lm')) (abc_steps_l (0, lm) bp stp)"
       
   414       apply(case_tac "abc_steps_l (0, lm) bp stpa", 
       
   415             case_tac "a = length bp")
       
   416       apply(insert ind, simp)
       
   417       apply(subgoal_tac "b = lm'", simp)
       
   418       apply(rule_tac abc_halt_equal, simp, simp)
       
   419       apply(rule_tac x = stpa in exI, simp add: abc_steps_ind)
       
   420       apply(simp add: abc_steps_zero)
       
   421       apply(rule classical, simp add: abc_steps_l.simps 
       
   422                              abc_fetch.simps abc_step_l.simps)
       
   423       done
       
   424   qed
       
   425 qed  
       
   426 
       
   427 
       
   428 lemma abc_append_empty_r[simp]: "[] [+] ab = ab"
       
   429 apply(simp add: abc_append.simps abc_inst_shift.simps)
       
   430 apply(induct ab, simp, simp)
       
   431 apply(case_tac a, simp_all add: abc_inst_shift.simps)
       
   432 done
       
   433 
       
   434 lemma abc_append_empty_l[simp]:  "ab [+] [] = ab"
       
   435 apply(simp add: abc_append.simps abc_inst_shift.simps)
       
   436 done
       
   437 
       
   438 
       
   439 lemma abc_append_length[simp]:  
       
   440   "length (ap [+] bp) = length ap + length bp"
       
   441 apply(simp add: abc_append.simps)
       
   442 done
       
   443 
       
   444 declare Let_def[simp]
       
   445 
       
   446 lemma abc_append_commute: "as [+] bs [+] cs = as [+] (bs [+] cs)"
       
   447 apply(simp add: abc_append.simps abc_shift.simps abc_inst_shift.simps)
       
   448 apply(induct cs, simp, simp)
       
   449 apply(case_tac a, auto simp: abc_inst_shift.simps Let_def)
       
   450 done
       
   451 
       
   452 lemma abc_halt_point_step[simp]: 
       
   453   "\<lbrakk>a < length bp; abc_steps_l (a, b) bp (Suc 0) = (length bp, lm')\<rbrakk>
       
   454   \<Longrightarrow> abc_steps_l (length ap + a, b) (ap [+] bp [+] cp) (Suc 0) = 
       
   455                                         (length ap + length bp, lm')"
       
   456 apply(simp add: abc_steps_l.simps abc_fetch.simps abc_append_nth)
       
   457 apply(case_tac "bp ! a", 
       
   458                       auto simp: abc_steps_l.simps abc_step_l.simps)
       
   459 done
       
   460 
       
   461 lemma abc_step_state_in:
       
   462   "\<lbrakk>bs < length bp;  abc_steps_l (a, b) bp (Suc 0) = (bs, l)\<rbrakk>
       
   463   \<Longrightarrow> a < length bp"
       
   464 apply(simp add: abc_steps_l.simps abc_fetch.simps)
       
   465 apply(rule_tac classical, 
       
   466       simp add: abc_step_l.simps abc_steps_l.simps)
       
   467 done
       
   468 
       
   469 
       
   470 lemma abc_append_state_in_exc: 
       
   471   "\<lbrakk>bs < length bp; abc_steps_l (0, lm) bp stpa = (bs, l)\<rbrakk>
       
   472  \<Longrightarrow> abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa = 
       
   473                                              (length ap + bs, l)"
       
   474 apply(induct stpa arbitrary: bs l, simp add: abc_steps_zero)
       
   475 proof -
       
   476   fix stpa bs l
       
   477   assume ind: 
       
   478     "\<And>bs l. \<lbrakk>bs < length bp; abc_steps_l (0, lm) bp stpa = (bs, l)\<rbrakk>
       
   479     \<Longrightarrow> abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa = 
       
   480                                                 (length ap + bs, l)"
       
   481     and h: "bs < length bp" 
       
   482            "abc_steps_l (0, lm) bp (Suc stpa) = (bs, l)"
       
   483   from h show 
       
   484     "abc_steps_l (length ap, lm) (ap [+] bp [+] cp) (Suc stpa) = 
       
   485                                                 (length ap + bs, l)"
       
   486     apply(simp add: abc_steps_ind)
       
   487     apply(case_tac "(abc_steps_l (0, lm) bp stpa)", simp)
       
   488   proof -
       
   489     fix a b
       
   490     assume g: "abc_steps_l (0, lm) bp stpa = (a, b)" 
       
   491               "abc_steps_l (a, b) bp (Suc 0) = (bs, l)"
       
   492     from h and g have k1: "a < length bp"
       
   493       apply(simp add: abc_step_state_in)
       
   494       done
       
   495     from h and g and k1 show 
       
   496    "abc_steps_l (abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa) 
       
   497               (ap [+] bp [+] cp) (Suc 0) = (length ap + bs, l)"
       
   498       apply(insert ind[of a b], simp)
       
   499       apply(simp add: abc_steps_l.simps abc_fetch.simps 
       
   500                       abc_append_nth)
       
   501       apply(case_tac "bp ! a", auto simp: 
       
   502                                  abc_steps_l.simps abc_step_l.simps)
       
   503       done
       
   504   qed
       
   505 qed
       
   506 
       
   507 lemma [simp]: "abc_steps_l (0, am) [] stp = (0, am)"
       
   508 apply(induct stp, simp add: abc_steps_zero)
       
   509 apply(simp add: abc_steps_ind)
       
   510 apply(simp add: abc_steps_zero abc_steps_l.simps 
       
   511                 abc_fetch.simps abc_step_l.simps)
       
   512 done
       
   513 
       
   514 lemma abc_append_exc1:
       
   515   "\<lbrakk>\<exists> stp. abc_steps_l (0, lm) bp stp = (bs, lm');
       
   516     bs = length bp; 
       
   517     as = length ap\<rbrakk>
       
   518     \<Longrightarrow> \<exists> stp. abc_steps_l (as, lm) (ap [+] bp [+] cp) stp 
       
   519                                                  = (as + bs, lm')"
       
   520 apply(case_tac "bp = []", erule_tac exE, simp,
       
   521       rule_tac x = 0 in exI, simp add: abc_steps_zero)
       
   522 apply(frule_tac abc_halt_point_ex, simp, simp,
       
   523       erule_tac exE, erule_tac exE) 
       
   524 apply(rule_tac x = "stpa + Suc 0" in exI)
       
   525 apply(case_tac "(abc_steps_l (0, lm) bp stpa)", 
       
   526       simp add: abc_steps_ind)
       
   527 apply(subgoal_tac 
       
   528   "abc_steps_l (length ap, lm) (ap [+] bp [+] cp) stpa 
       
   529                                    = (length ap + a, b)", simp)
       
   530 apply(simp add: abc_steps_zero)
       
   531 apply(rule_tac abc_append_state_in_exc, simp, simp)
       
   532 done
       
   533 
       
   534 lemma abc_append_exc3: 
       
   535   "\<lbrakk>\<exists> stp. abc_steps_l (0, am) bp stp = (bs, bm); ss = length ap\<rbrakk>
       
   536    \<Longrightarrow>  \<exists> stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
       
   537 apply(erule_tac exE)
       
   538 proof -
       
   539   fix stp
       
   540   assume h: "abc_steps_l (0, am) bp stp = (bs, bm)" "ss = length ap"
       
   541   thus " \<exists>stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
       
   542   proof(induct stp arbitrary: bs bm)
       
   543     fix bs bm
       
   544     assume "abc_steps_l (0, am) bp 0 = (bs, bm)"
       
   545     thus "\<exists>stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
       
   546       apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
       
   547       done
       
   548   next
       
   549     fix stp bs bm
       
   550     assume ind: 
       
   551       "\<And>bs bm. \<lbrakk>abc_steps_l (0, am) bp stp = (bs, bm);
       
   552                  ss = length ap\<rbrakk> \<Longrightarrow> 
       
   553           \<exists>stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
       
   554     and g: "abc_steps_l (0, am) bp (Suc stp) = (bs, bm)"
       
   555     from g show 
       
   556       "\<exists>stp. abc_steps_l (ss, am) (ap [+] bp) stp = (bs + ss, bm)"
       
   557       apply(insert abc_steps_add[of 0 am bp stp "Suc 0"], simp)
       
   558       apply(case_tac "(abc_steps_l (0, am) bp stp)", simp)
       
   559     proof -
       
   560       fix a b
       
   561       assume "(bs, bm) = abc_steps_l (a, b) bp (Suc 0)" 
       
   562              "abc_steps_l (0, am) bp (Suc stp) = 
       
   563                        abc_steps_l (a, b) bp (Suc 0)" 
       
   564               "abc_steps_l (0, am) bp stp = (a, b)"
       
   565       thus "?thesis"
       
   566 	apply(insert ind[of a b], simp add: h, erule_tac exE)
       
   567 	apply(rule_tac x = "Suc stp" in exI)
       
   568 	apply(simp only: abc_steps_ind, simp add: abc_steps_zero)
       
   569       proof -
       
   570 	fix stp
       
   571 	assume "(bs, bm) = abc_steps_l (a, b) bp (Suc 0)"
       
   572 	thus "abc_steps_l (a + length ap, b) (ap [+] bp) (Suc 0)
       
   573                                               = (bs + length ap, bm)"
       
   574 	  apply(simp add: abc_steps_l.simps abc_steps_zero
       
   575                           abc_fetch.simps split: if_splits)
       
   576 	  apply(case_tac "bp ! a", 
       
   577                 simp_all add: abc_inst_shift.simps abc_append_nth
       
   578                    abc_steps_l.simps abc_steps_zero abc_step_l.simps)
       
   579 	  apply(auto)
       
   580 	  done
       
   581       qed
       
   582     qed
       
   583   qed
       
   584 qed
       
   585 
       
   586 lemma abc_add_equal:
       
   587   "\<lbrakk>ap \<noteq> []; 
       
   588     abc_steps_l (0, am) ap astp = (a, b);
       
   589     a < length ap\<rbrakk>
       
   590      \<Longrightarrow> (abc_steps_l (0, am) (ap @ bp) astp) = (a, b)"
       
   591 apply(induct astp arbitrary: a b, simp add: abc_steps_l.simps, simp)
       
   592 apply(simp add: abc_steps_ind)
       
   593 apply(case_tac "(abc_steps_l (0, am) ap astp)")
       
   594 proof -
       
   595   fix astp a b aa ba
       
   596   assume ind: 
       
   597     "\<And>a b. \<lbrakk>abc_steps_l (0, am) ap astp = (a, b); 
       
   598              a < length ap\<rbrakk> \<Longrightarrow> 
       
   599                   abc_steps_l (0, am) (ap @ bp) astp = (a, b)"
       
   600   and h: "abc_steps_l (abc_steps_l (0, am) ap astp) ap (Suc 0)
       
   601                                                             = (a, b)"
       
   602         "a < length ap" 
       
   603         "abc_steps_l (0, am) ap astp = (aa, ba)"
       
   604   from h show "abc_steps_l (abc_steps_l (0, am) (ap @ bp) astp)
       
   605                                         (ap @ bp) (Suc 0) = (a, b)"
       
   606     apply(insert ind[of aa ba], simp)
       
   607     apply(subgoal_tac "aa < length ap", simp)
       
   608     apply(simp add: abc_steps_l.simps abc_fetch.simps
       
   609                      nth_append abc_steps_zero)
       
   610     apply(rule abc_step_state_in, auto)
       
   611     done
       
   612 qed
       
   613 
       
   614 
       
   615 lemma abc_add_exc1: 
       
   616   "\<lbrakk>\<exists> astp. abc_steps_l (0, am) ap astp = (as, bm); as = length ap\<rbrakk>
       
   617   \<Longrightarrow> \<exists> stp. abc_steps_l (0, am) (ap @ bp) stp = (as, bm)"
       
   618 apply(case_tac "ap = []", simp, 
       
   619       rule_tac x = 0 in exI, simp add: abc_steps_zero)
       
   620 apply(drule_tac abc_halt_point_ex, simp, simp)
       
   621 apply(erule_tac exE, case_tac "(abc_steps_l (0, am) ap astp)", simp)
       
   622 apply(rule_tac x = "Suc astp" in exI, simp add: abc_steps_ind, auto)
       
   623 apply(frule_tac bp = bp in abc_add_equal, simp, simp, simp)
       
   624 apply(simp add: abc_steps_l.simps abc_steps_zero 
       
   625                 abc_fetch.simps nth_append)
       
   626 done
       
   627 
       
   628 declare abc_shift.simps[simp del] 
       
   629 
       
   630 lemma abc_append_exc2: 
       
   631   "\<lbrakk>\<exists> astp. abc_steps_l (0, am) ap astp = (as, bm); as = length ap; 
       
   632     \<exists> bstp. abc_steps_l (0, bm) bp bstp = (bs, bm'); bs = length bp;
       
   633     cs = as + bs; bp \<noteq> []\<rbrakk>
       
   634   \<Longrightarrow> \<exists> stp. abc_steps_l (0, am) (ap [+] bp) stp = (cs, bm')"
       
   635 apply(insert abc_append_exc1[of bm bp bs bm' as ap "[]"], simp)
       
   636 apply(drule_tac bp = "abc_shift bp (length ap)" in abc_add_exc1, simp)
       
   637 apply(subgoal_tac "ap @ abc_shift bp (length ap) = ap [+] bp", 
       
   638       simp, auto)
       
   639 apply(rule_tac x = "stpa + stp" in exI, simp add: abc_steps_add)
       
   640 apply(simp add: abc_append.simps)
       
   641 done
       
   642 lemma exponent_add_iff: "a\<up>b @ a\<up>c@ xs = a\<up>(b+c) @ xs"
       
   643 apply(auto simp: replicate_add)
       
   644 done
       
   645 
       
   646 lemma exponent_cons_iff: "a # a\<up>c @ xs = a\<up>(Suc c) @ xs"
       
   647 apply(auto simp: replicate_add)
       
   648 done
       
   649 
       
   650 lemma  [simp]: "length lm = n \<Longrightarrow>  
       
   651   abc_steps_l (Suc 0, lm @ Suc x # 0 # suf_lm) 
       
   652        [Inc n, Dec (Suc n) 3, Goto (Suc 0)] (Suc (Suc 0))
       
   653                                   = (3, lm @ Suc x # 0 # suf_lm)"
       
   654 apply(simp add: abc_steps_l.simps abc_fetch.simps 
       
   655                 abc_step_l.simps abc_lm_v.simps abc_lm_s.simps 
       
   656                 nth_append list_update_append)
       
   657 done
       
   658 
       
   659 lemma [simp]: 
       
   660   "length lm = n \<Longrightarrow> 
       
   661   abc_steps_l (Suc 0, lm @ Suc x # Suc y # suf_lm) 
       
   662      [Inc n, Dec (Suc n) 3, Goto (Suc 0)] (Suc (Suc 0))
       
   663   = (Suc 0, lm @ Suc x # y # suf_lm)"
       
   664 apply(simp add: abc_steps_l.simps abc_fetch.simps 
       
   665                 abc_step_l.simps abc_lm_v.simps abc_lm_s.simps 
       
   666                 nth_append list_update_append)
       
   667 done
       
   668 
       
   669 lemma pr_cycle_part_middle_inv: 
       
   670   "\<lbrakk>length lm = n\<rbrakk> \<Longrightarrow> 
       
   671   \<exists> stp. abc_steps_l (0, lm @ x # y # suf_lm) 
       
   672                          [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp 
       
   673   = (3, lm @ Suc x # 0 # suf_lm)"
       
   674 proof -
       
   675   assume h: "length lm = n"
       
   676   hence k1: "\<exists> stp. abc_steps_l (0, lm @ x # y # suf_lm) 
       
   677                            [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp 
       
   678     = (Suc 0, lm @ Suc x # y # suf_lm)"
       
   679     apply(rule_tac x = "Suc 0" in exI)
       
   680     apply(simp add: abc_steps_l.simps abc_step_l.simps 
       
   681                     abc_lm_v.simps abc_lm_s.simps nth_append 
       
   682                     list_update_append abc_fetch.simps)
       
   683     done
       
   684   from h have k2: 
       
   685     "\<exists> stp. abc_steps_l (Suc 0, lm @ Suc x # y # suf_lm)
       
   686                       [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp 
       
   687     = (3, lm @ Suc x # 0 # suf_lm)"
       
   688     apply(induct y)
       
   689     apply(rule_tac x = "Suc (Suc 0)" in exI, simp, simp, 
       
   690           erule_tac exE)
       
   691     apply(rule_tac x = "Suc (Suc 0) + stp" in exI, 
       
   692           simp only: abc_steps_add, simp)
       
   693     done      
       
   694   from k1 and k2 show 
       
   695     "\<exists> stp. abc_steps_l (0, lm @ x # y # suf_lm) 
       
   696                        [Inc n, Dec (Suc n) 3, Goto (Suc 0)] stp 
       
   697     = (3, lm @ Suc x # 0 # suf_lm)"
       
   698     apply(erule_tac exE, erule_tac exE)
       
   699     apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
       
   700     done
       
   701 qed
       
   702 
       
   703 lemma [simp]: 
       
   704   "length lm = Suc n \<Longrightarrow> 
       
   705   (abc_steps_l (length ap, lm @ x # Suc y # suf_lm) 
       
   706            (ap @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length ap)]) 
       
   707                     (Suc (Suc (Suc 0))))
       
   708   = (length ap, lm @ Suc x # y # suf_lm)"
       
   709 apply(simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps 
       
   710          abc_lm_v.simps list_update_append nth_append abc_lm_s.simps)
       
   711 done
       
   712 
       
   713 lemma switch_para_inv:
       
   714   assumes bp_def:"bp =  ap @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto ss]"
       
   715   and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" 
       
   716          "ss = length ap" 
       
   717          "length lm = Suc n"
       
   718   shows " \<exists>stp. abc_steps_l (ss, lm @ x # y # suf_lm) bp stp =
       
   719                                (0, lm @ (x + y) # 0 # suf_lm)"
       
   720 apply(induct y arbitrary: x)
       
   721 apply(rule_tac x = "Suc 0" in exI,
       
   722   simp add: bp_def mv_box.simps abc_steps_l.simps 
       
   723             abc_fetch.simps h abc_step_l.simps 
       
   724             abc_lm_v.simps list_update_append nth_append
       
   725             abc_lm_s.simps)
       
   726 proof -
       
   727   fix y x
       
   728   assume ind: 
       
   729     "\<And>x. \<exists>stp. abc_steps_l (ss, lm @ x # y # suf_lm) bp stp = 
       
   730                                      (0, lm @ (x + y) # 0 # suf_lm)"
       
   731   show "\<exists>stp. abc_steps_l (ss, lm @ x # Suc y # suf_lm) bp stp = 
       
   732                                   (0, lm @ (x + Suc y) # 0 # suf_lm)"
       
   733     apply(insert ind[of "Suc x"], erule_tac exE)
       
   734     apply(rule_tac x = "Suc (Suc (Suc 0)) + stp" in exI, 
       
   735           simp only: abc_steps_add bp_def h)
       
   736     apply(simp add: h)
       
   737     done
       
   738 qed
       
   739 
       
   740 lemma [simp]:
       
   741   "length lm = rs_pos \<and> Suc (Suc rs_pos) < a_md \<and> 0 < rs_pos \<Longrightarrow> 
       
   742       a_md - Suc 0 < Suc (Suc (Suc (a_md + length suf_lm - 
       
   743                                          Suc (Suc (Suc 0)))))"
       
   744 apply(arith)
       
   745 done
       
   746 
       
   747 lemma [simp]: 
       
   748   "Suc (Suc rs_pos) < a_md \<and> 0 < rs_pos \<Longrightarrow> 
       
   749                            \<not> a_md - Suc 0 < rs_pos - Suc 0"
       
   750 apply(arith)
       
   751 done
       
   752 
       
   753 lemma [simp]: 
       
   754   "Suc (Suc rs_pos) < a_md \<and> 0 < rs_pos \<Longrightarrow> 
       
   755            \<not> a_md - rs_pos < Suc (Suc (a_md - Suc (Suc rs_pos)))"
       
   756 apply(arith)
       
   757 done
       
   758 
       
   759 lemma butlast_append_last: "lm \<noteq> [] \<Longrightarrow> lm = butlast lm @ [last lm]"
       
   760 apply(auto)
       
   761 done
       
   762 
       
   763 lemma [simp]: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)
       
   764            \<Longrightarrow> (Suc (Suc rs_pos)) < a_md"
       
   765 apply(simp add: rec_ci.simps)
       
   766 apply(case_tac "rec_ci f", simp)
       
   767 apply(case_tac "rec_ci g", simp)
       
   768 apply(arith)
       
   769 done
       
   770 
       
   771 lemma ci_pr_para_eq: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)
       
   772                   \<Longrightarrow> rs_pos = Suc n"
       
   773 apply(simp add: rec_ci.simps)
       
   774 apply(case_tac "rec_ci g",  case_tac "rec_ci f", simp)
       
   775 done
       
   776 
       
   777 lemma [intro]:  
       
   778   "\<lbrakk>rec_ci z = (aprog, rs_pos, a_md); rec_calc_rel z lm xs\<rbrakk>
       
   779   \<Longrightarrow> length lm = rs_pos"
       
   780 apply(simp add: rec_ci.simps rec_ci_z_def)
       
   781 apply(erule_tac calc_z_reverse, simp)
       
   782 done
       
   783 
       
   784 lemma [intro]: 
       
   785   "\<lbrakk>rec_ci s = (aprog, rs_pos, a_md); rec_calc_rel s lm xs\<rbrakk>
       
   786   \<Longrightarrow> length lm = rs_pos"
       
   787 apply(simp add: rec_ci.simps rec_ci_s_def)
       
   788 apply(erule_tac calc_s_reverse, simp)
       
   789 done
       
   790 
       
   791 lemma [intro]: 
       
   792   "\<lbrakk>rec_ci (recf.id nat1 nat2) = (aprog, rs_pos, a_md); 
       
   793     rec_calc_rel (recf.id nat1 nat2) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
       
   794 apply(simp add: rec_ci.simps rec_ci_id.simps)
       
   795 apply(erule_tac calc_id_reverse, simp)
       
   796 done
       
   797 
       
   798 lemma [intro]: 
       
   799   "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
       
   800     rec_calc_rel (Cn n f gs) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
       
   801 apply(erule_tac calc_cn_reverse, simp)
       
   802 apply(simp add: rec_ci.simps)
       
   803 apply(case_tac "rec_ci f",  simp)
       
   804 done
       
   805 
       
   806 lemma [intro]:
       
   807   "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
       
   808     rec_calc_rel (Pr n f g) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
       
   809 apply(erule_tac  calc_pr_reverse, simp)
       
   810 apply(drule_tac ci_pr_para_eq, simp, simp)
       
   811 apply(drule_tac ci_pr_para_eq, simp)
       
   812 done
       
   813 
       
   814 lemma [intro]: 
       
   815   "\<lbrakk>rec_ci (Mn n f) = (aprog, rs_pos, a_md);
       
   816     rec_calc_rel (Mn n f) lm xs\<rbrakk> \<Longrightarrow> length lm = rs_pos"
       
   817 apply(erule_tac calc_mn_reverse)
       
   818 apply(simp add: rec_ci.simps)
       
   819 apply(case_tac "rec_ci f",  simp)
       
   820 done
       
   821 
       
   822 lemma para_pattern: 
       
   823   "\<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm xs\<rbrakk>
       
   824   \<Longrightarrow> length lm = rs_pos"
       
   825 apply(case_tac f, auto)
       
   826 done
       
   827 
       
   828 lemma ci_pr_g_paras:
       
   829   "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
       
   830     rec_ci g = (a, aa, ba);
       
   831     rec_calc_rel (Pr n f g) (lm @ [x]) rs; x > 0\<rbrakk> \<Longrightarrow> 
       
   832     aa = Suc rs_pos "
       
   833 apply(erule calc_pr_reverse, simp)
       
   834 apply(subgoal_tac "length (args @ [k, rk]) = aa", simp)
       
   835 apply(subgoal_tac "rs_pos = Suc n", simp)
       
   836 apply(simp add: ci_pr_para_eq)
       
   837 apply(erule para_pattern, simp)
       
   838 done
       
   839 
       
   840 lemma ci_pr_g_md_less: 
       
   841   "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
       
   842     rec_ci g = (a, aa, ba)\<rbrakk> \<Longrightarrow> ba < a_md"
       
   843 apply(simp add: rec_ci.simps)
       
   844 apply(case_tac "rec_ci f",  auto)
       
   845 done
       
   846 
       
   847 lemma [intro]: "rec_ci z = (ap, rp, ad) \<Longrightarrow> rp < ad"
       
   848   by(simp add: rec_ci.simps)
       
   849 
       
   850 lemma [intro]: "rec_ci s = (ap, rp, ad) \<Longrightarrow> rp < ad"
       
   851   by(simp add: rec_ci.simps)
       
   852 
       
   853 lemma [intro]: "rec_ci (recf.id nat1 nat2) = (ap, rp, ad) \<Longrightarrow> rp < ad"
       
   854   by(simp add: rec_ci.simps)
       
   855 
       
   856 lemma [intro]: "rec_ci (Cn n f gs) = (ap, rp, ad) \<Longrightarrow> rp < ad"
       
   857 apply(simp add: rec_ci.simps)
       
   858 apply(case_tac "rec_ci f",  simp)
       
   859 done
       
   860 
       
   861 lemma [intro]: "rec_ci (Pr n f g) = (ap, rp, ad) \<Longrightarrow> rp < ad"
       
   862 apply(simp add: rec_ci.simps)
       
   863 by(case_tac "rec_ci f", case_tac "rec_ci g",  auto)
       
   864 
       
   865 lemma [intro]: "rec_ci (Mn n f) = (ap, rp, ad) \<Longrightarrow> rp < ad"
       
   866 apply(simp add: rec_ci.simps)
       
   867 apply(case_tac "rec_ci f", simp)
       
   868 apply(arith)
       
   869 done
       
   870 
       
   871 lemma ci_ad_ge_paras: "rec_ci f = (ap, rp, ad) \<Longrightarrow> ad > rp"
       
   872 apply(case_tac f, auto)
       
   873 done
       
   874 
       
   875 lemma [elim]: "\<lbrakk>a [+] b = []; a \<noteq> [] \<or> b \<noteq> []\<rbrakk> \<Longrightarrow> RR"
       
   876 apply(auto simp: abc_append.simps abc_shift.simps)
       
   877 done
       
   878 
       
   879 lemma [intro]: "rec_ci z = ([], aa, ba) \<Longrightarrow> False"
       
   880 by(simp add: rec_ci.simps rec_ci_z_def)
       
   881 
       
   882 lemma [intro]: "rec_ci s = ([], aa, ba) \<Longrightarrow> False"
       
   883 by(auto simp: rec_ci.simps rec_ci_s_def addition.simps)
       
   884 
       
   885 lemma [intro]: "rec_ci (id m n) = ([], aa, ba) \<Longrightarrow> False"
       
   886 by(auto simp: rec_ci.simps rec_ci_id.simps addition.simps)
       
   887 
       
   888 lemma [intro]: "rec_ci (Cn n f gs) = ([], aa, ba) \<Longrightarrow> False"
       
   889 apply(case_tac "rec_ci f", auto simp: rec_ci.simps abc_append.simps)
       
   890 apply(simp add: abc_shift.simps mv_box.simps)
       
   891 done
       
   892 
       
   893 lemma [intro]: "rec_ci (Pr n f g) = ([], aa, ba) \<Longrightarrow> False"
       
   894 apply(simp add: rec_ci.simps)
       
   895 apply(case_tac "rec_ci f", case_tac "rec_ci g")
       
   896 by(auto)
       
   897 
       
   898 lemma [intro]: "rec_ci (Mn n f) = ([], aa, ba) \<Longrightarrow> False"
       
   899 apply(case_tac "rec_ci f", auto simp: rec_ci.simps)
       
   900 done
       
   901 
       
   902 lemma rec_ci_not_null:  "rec_ci g = (a, aa, ba) \<Longrightarrow> a \<noteq> []"
       
   903 by(case_tac g, auto)
       
   904 
       
   905 lemma calc_pr_g_def:
       
   906  "\<lbrakk>rec_calc_rel (Pr rs_pos f g) (lm @ [Suc x]) rsa;
       
   907    rec_calc_rel (Pr rs_pos f g) (lm @ [x]) rsxa\<rbrakk>
       
   908  \<Longrightarrow> rec_calc_rel g (lm @ [x, rsxa]) rsa"
       
   909 apply(erule_tac calc_pr_reverse, simp, simp)
       
   910 apply(subgoal_tac "rsxa = rk", simp)
       
   911 apply(erule_tac rec_calc_inj, auto)
       
   912 done
       
   913 
       
   914 lemma ci_pr_md_def: 
       
   915   "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
       
   916     rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\<rbrakk>
       
   917   \<Longrightarrow> a_md = Suc (max (n + 3) (max bc ba))"
       
   918 by(simp add: rec_ci.simps)
       
   919 
       
   920 lemma  ci_pr_f_paras: 
       
   921   "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
       
   922     rec_calc_rel (Pr n f g) lm rs;
       
   923     rec_ci f = (ab, ac, bc)\<rbrakk>  \<Longrightarrow> ac = rs_pos - Suc 0"
       
   924 apply(subgoal_tac "\<exists>rs. rec_calc_rel f (butlast lm) rs", 
       
   925       erule_tac exE)
       
   926 apply(drule_tac f = f and lm = "butlast lm" in para_pattern, 
       
   927       simp, simp)
       
   928 apply(drule_tac para_pattern, simp)
       
   929 apply(subgoal_tac "lm \<noteq> []", simp)
       
   930 apply(erule_tac calc_pr_reverse, simp, simp)
       
   931 apply(erule calc_pr_zero_ex)
       
   932 done
       
   933 
       
   934 lemma ci_pr_md_ge_f:  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
       
   935         rec_ci f = (ab, ac, bc)\<rbrakk> \<Longrightarrow> Suc bc \<le> a_md"
       
   936 apply(case_tac "rec_ci g")
       
   937 apply(simp add: rec_ci.simps, auto)
       
   938 done
       
   939 
       
   940 lemma ci_pr_md_ge_g:  "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
       
   941         rec_ci g = (ab, ac, bc)\<rbrakk> \<Longrightarrow> bc < a_md"
       
   942 apply(case_tac "rec_ci f")
       
   943 apply(simp add: rec_ci.simps, auto)
       
   944 done 
       
   945 
       
   946 lemma rec_calc_rel_def0: 
       
   947   "\<lbrakk>rec_calc_rel (Pr n f g) lm rs; rec_calc_rel f (butlast lm) rsa\<rbrakk>
       
   948   \<Longrightarrow> rec_calc_rel (Pr n f g) (butlast lm @ [0]) rsa"
       
   949   apply(rule_tac calc_pr_zero, simp)
       
   950 apply(erule_tac calc_pr_reverse, simp, simp, simp)
       
   951 done
       
   952 
       
   953 lemma [simp]:  "length (mv_box m n) = 3"
       
   954 by (auto simp: mv_box.simps)
       
   955 
       
   956 
       
   957 lemma [simp]: "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); rec_calc_rel (Pr n f g) lm rs\<rbrakk>
       
   958     \<Longrightarrow> rs_pos = Suc n"
       
   959 apply(simp add: ci_pr_para_eq)
       
   960 done
       
   961 
       
   962 
       
   963 lemma [simp]: "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); rec_calc_rel (Pr n f g) lm rs\<rbrakk>
       
   964     \<Longrightarrow> length lm = Suc n"
       
   965 apply(subgoal_tac "rs_pos = Suc n", rule_tac para_pattern, simp, simp)
       
   966 apply(case_tac "rec_ci f", case_tac "rec_ci g", simp add: rec_ci.simps)
       
   967 done
       
   968 
       
   969 lemma [simp]: "rec_ci (Pr n f g) = (a, rs_pos, a_md) \<Longrightarrow> Suc (Suc n) < a_md"
       
   970 apply(case_tac "rec_ci f", case_tac "rec_ci g", simp add: rec_ci.simps)
       
   971 apply arith
       
   972 done
       
   973 
       
   974 lemma [simp]: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md) \<Longrightarrow> 0 < rs_pos"
       
   975 apply(case_tac "rec_ci f", case_tac "rec_ci g")
       
   976 apply(simp add: rec_ci.simps)
       
   977 done
       
   978 
       
   979 lemma [simp]: "Suc (Suc rs_pos) < a_md \<Longrightarrow> 
       
   980        butlast lm @ (last lm - xa) # (rsa::nat) # 0 # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm =
       
   981        butlast lm @ (last lm - xa) # rsa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm"
       
   982 apply(simp add: replicate_Suc[THEN sym])
       
   983 done
       
   984 
       
   985 lemma pr_cycle_part_ind: 
       
   986   assumes g_ind: 
       
   987   "\<And>lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow> 
       
   988   \<exists>stp. abc_steps_l (0, lm @ 0\<up>(ba - aa) @ suf_lm) a stp = 
       
   989                     (length a, lm @ rs # 0\<up>(ba - Suc aa) @ suf_lm)"
       
   990   and ap_def: 
       
   991   "ap = ([Dec (a_md - Suc 0) (length a + 7)] [+]
       
   992         (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)])) @
       
   993          [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]"
       
   994   and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" 
       
   995          "rec_calc_rel (Pr n f g) 
       
   996                    (butlast lm @ [last lm - Suc xa]) rsxa" 
       
   997          "Suc xa \<le> last lm" 
       
   998          "rec_ci g = (a, aa, ba)"
       
   999          "rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rsa"
       
  1000          "lm \<noteq> []"
       
  1001   shows 
       
  1002   "\<exists>stp. abc_steps_l 
       
  1003      (0, butlast lm @ (last lm - Suc xa) # rsxa # 
       
  1004                0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) ap stp =
       
  1005      (0, butlast lm @ (last lm - xa) # rsa
       
  1006                  # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)"
       
  1007 proof -
       
  1008   have k1: "\<exists>stp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) #
       
  1009     rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) ap stp =
       
  1010          (length a + 4, butlast lm @ (last lm - xa) # 0 # rsa #
       
  1011                            0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)"
       
  1012     apply(simp add: ap_def, rule_tac abc_add_exc1)
       
  1013     apply(rule_tac as = "Suc 0" and 
       
  1014       bm = "butlast lm @ (last lm - Suc xa) # 
       
  1015       rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm" in abc_append_exc2,
       
  1016       auto)
       
  1017   proof -
       
  1018     show 
       
  1019       "\<exists>astp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # rsxa 
       
  1020                    # 0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) 
       
  1021               [Dec (a_md - Suc 0)(length a + 7)] astp =
       
  1022       (Suc 0, butlast lm @ (last lm - Suc xa) # 
       
  1023              rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)"
       
  1024       apply(rule_tac x = "Suc 0" in exI, 
       
  1025           simp add: abc_steps_l.simps abc_step_l.simps
       
  1026                      abc_fetch.simps)
       
  1027       apply(subgoal_tac "length lm = Suc n \<and> rs_pos = Suc n \<and>
       
  1028                               a_md > Suc (Suc rs_pos)")
       
  1029       apply(simp add: abc_lm_v.simps nth_append abc_lm_s.simps)
       
  1030       apply(insert nth_append[of 
       
  1031                  "(last lm - Suc xa) # rsxa # 0\<up>(a_md - Suc (Suc rs_pos))" 
       
  1032                  "Suc xa # suf_lm" "(a_md - rs_pos)"], simp)
       
  1033       apply(simp add: list_update_append del: list_update.simps)
       
  1034       apply(insert list_update_append[of "(last lm - Suc xa) # rsxa # 
       
  1035                                           0\<up>(a_md - Suc (Suc rs_pos))" 
       
  1036                     "Suc xa # suf_lm" "a_md - rs_pos" "xa"], simp)
       
  1037       apply(case_tac a_md, simp, simp)
       
  1038       apply(insert h, simp)
       
  1039       apply(insert para_pattern[of "Pr n f g" aprog rs_pos a_md 
       
  1040                     "(butlast lm @ [last lm - Suc xa])" rsxa], simp)
       
  1041       done
       
  1042   next
       
  1043     show "\<exists>bstp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # 
       
  1044            rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm) (a [+] 
       
  1045             [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)]) bstp =
       
  1046          (3 + length a, butlast lm @ (last lm - xa) # 0 # rsa #
       
  1047                           0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)"
       
  1048       apply(rule_tac as = "length a" and
       
  1049                bm = "butlast lm @ (last lm - Suc xa) # rsxa # rsa #
       
  1050                      0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm" 
       
  1051         in abc_append_exc2, simp_all)
       
  1052     proof -
       
  1053       from h have j1: "aa = Suc rs_pos \<and> a_md > ba \<and> ba > Suc rs_pos"
       
  1054 	apply(insert h)
       
  1055 	apply(insert ci_pr_g_paras[of n f g aprog rs_pos
       
  1056                  a_md a aa ba "butlast lm" "last lm - xa" rsa], simp)
       
  1057 	apply(drule_tac ci_pr_md_ge_g, auto)
       
  1058 	apply(erule_tac ci_ad_ge_paras)
       
  1059 	done
       
  1060       from h have j2: "rec_calc_rel g (butlast lm @ 
       
  1061                                   [last lm - Suc xa, rsxa]) rsa"
       
  1062 	apply(rule_tac  calc_pr_g_def, simp, simp)
       
  1063 	done
       
  1064       from j1 and j2 show 
       
  1065         "\<exists>astp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) #
       
  1066                 rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm) a astp =
       
  1067         (length a, butlast lm @ (last lm - Suc xa) # rsxa # rsa 
       
  1068                          # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)"
       
  1069 	apply(insert g_ind[of
       
  1070           "butlast lm @ (last lm - Suc xa) # [rsxa]" rsa 
       
  1071           "0\<up>(a_md - ba - Suc 0) @ xa # suf_lm"], simp, auto)
       
  1072 	apply(simp add: exponent_add_iff)
       
  1073 	apply(rule_tac x = stp in exI, simp add: numeral_3_eq_3)
       
  1074 	done
       
  1075     next
       
  1076       from h have j3: "length lm = rs_pos \<and> rs_pos > 0"
       
  1077 	apply(rule_tac conjI)
       
  1078 	apply(drule_tac lm = "(butlast lm @ [last lm - Suc xa])"
       
  1079                           and xs = rsxa in para_pattern, simp, simp, simp)
       
  1080         done
       
  1081       from h have j4: "Suc (last lm - Suc xa) = last lm - xa"
       
  1082 	apply(case_tac "last lm", simp, simp)
       
  1083 	done
       
  1084       from j3 and j4 show
       
  1085       "\<exists>bstp. abc_steps_l (0, butlast lm @ (last lm - Suc xa) # rsxa #
       
  1086                      rsa # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)
       
  1087             [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)] bstp =
       
  1088         (3, butlast lm @ (last lm - xa) # 0 # rsa #
       
  1089                        0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm)"
       
  1090 	apply(insert pr_cycle_part_middle_inv[of "butlast lm" 
       
  1091           "rs_pos - Suc 0" "(last lm - Suc xa)" rsxa 
       
  1092           "rsa # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm"], simp)
       
  1093 	done
       
  1094     qed
       
  1095   qed
       
  1096   from h have k2: 
       
  1097     "\<exists>stp. abc_steps_l (length a + 4, butlast lm @ (last lm - xa) # 0 
       
  1098            # rsa # 0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm) ap stp =
       
  1099     (0, butlast lm @ (last lm - xa) # rsa # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)"
       
  1100     apply(insert switch_para_inv[of ap 
       
  1101       "([Dec (a_md - Suc 0) (length a + 7)] [+] 
       
  1102       (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)]))"
       
  1103       n "length a + 4" f g aprog rs_pos a_md 
       
  1104       "butlast lm @ [last lm - xa]" 0 rsa 
       
  1105       "0\<up>(a_md - Suc (Suc (Suc rs_pos))) @ xa # suf_lm"])
       
  1106     apply(simp add: h ap_def)
       
  1107     apply(subgoal_tac "length lm = Suc n \<and> Suc (Suc rs_pos) < a_md", 
       
  1108           simp)
       
  1109     apply(insert h, simp)
       
  1110     apply(frule_tac lm = "(butlast lm @ [last lm - Suc xa])" 
       
  1111       and xs = rsxa in para_pattern, simp, simp)
       
  1112     done   
       
  1113   from k1 and k2 show "?thesis"
       
  1114     apply(auto)
       
  1115     apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
       
  1116     done
       
  1117 qed
       
  1118 
       
  1119 lemma ci_pr_ex1: 
       
  1120   "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
       
  1121     rec_ci g = (a, aa, ba);
       
  1122     rec_ci f = (ab, ac, bc)\<rbrakk>
       
  1123 \<Longrightarrow> \<exists>ap bp. length ap = 6 + length ab \<and>
       
  1124     aprog = ap [+] bp \<and>
       
  1125     bp = ([Dec (a_md - Suc 0) (length a + 7)] [+] (a [+]
       
  1126          [Inc (rs_pos - Suc 0), Dec rs_pos 3, Goto (Suc 0)])) @ 
       
  1127          [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]"
       
  1128 apply(simp add: rec_ci.simps)
       
  1129 apply(rule_tac x = "mv_box n (max (Suc (Suc (Suc n)))
       
  1130     (max bc ba)) [+] ab [+] mv_box n (Suc n)" in exI,
       
  1131      simp)
       
  1132 apply(auto simp add: abc_append_commute numeral_3_eq_3)
       
  1133 done
       
  1134 
       
  1135 lemma pr_cycle_part:
       
  1136   "\<lbrakk>\<And>lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow>
       
  1137      \<exists>stp. abc_steps_l (0, lm @ 0\<up>(ba - aa) @ suf_lm) a stp = 
       
  1138                         (length a, lm @ rs # 0\<up>(ba - Suc aa) @ suf_lm);
       
  1139   rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
       
  1140   rec_calc_rel (Pr n f g) lm rs;
       
  1141   rec_ci g = (a, aa, ba);
       
  1142   rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rsx;
       
  1143   rec_ci f = (ab, ac, bc);
       
  1144   lm \<noteq> [];
       
  1145   x \<le> last lm\<rbrakk> \<Longrightarrow> 
       
  1146   \<exists>stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - x) #
       
  1147               rsx # 0\<up>(a_md - Suc (Suc rs_pos)) @ x # suf_lm) aprog stp =
       
  1148   (6 + length ab, butlast lm @ last lm # rs #
       
  1149                                 0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)"
       
  1150 proof -
       
  1151   assume g_ind:
       
  1152     "\<And>lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow> 
       
  1153     \<exists>stp. abc_steps_l (0, lm @ 0\<up>(ba - aa) @ suf_lm) a stp =
       
  1154                       (length a, lm @ rs # 0\<up>(ba - Suc aa) @ suf_lm)"
       
  1155     and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" 
       
  1156            "rec_calc_rel (Pr n f g) lm rs" 
       
  1157            "rec_ci g = (a, aa, ba)"
       
  1158            "rec_calc_rel (Pr n f g) (butlast lm @ [last lm - x]) rsx" 
       
  1159            "lm \<noteq> []"
       
  1160            "x \<le> last lm" 
       
  1161            "rec_ci f = (ab, ac, bc)" 
       
  1162   from h show 
       
  1163     "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - x) # 
       
  1164             rsx # 0\<up>(a_md - Suc (Suc rs_pos)) @ x # suf_lm) aprog stp =
       
  1165     (6 + length ab, butlast lm @ last lm # rs #
       
  1166                                0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)" 
       
  1167   proof(induct x arbitrary: rsx, simp_all)
       
  1168     fix rsxa
       
  1169     assume "rec_calc_rel (Pr n f g) lm rsxa" 
       
  1170            "rec_calc_rel (Pr n f g) lm rs"
       
  1171     from h and this have "rs = rsxa"
       
  1172       apply(subgoal_tac "lm \<noteq> [] \<and> rs_pos = Suc n", simp)
       
  1173       apply(rule_tac rec_calc_inj, simp, simp)
       
  1174       apply(simp)
       
  1175       done
       
  1176     thus "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @  last lm # 
       
  1177              rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm) aprog stp =
       
  1178       (6 + length ab, butlast lm @ last lm # rs #
       
  1179                                0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)"
       
  1180       by(rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
       
  1181   next
       
  1182     fix xa rsxa
       
  1183     assume ind:
       
  1184    "\<And>rsx. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rsx 
       
  1185   \<Longrightarrow> \<exists>stp. abc_steps_l (6 + length ab, butlast lm @ (last lm - xa) #
       
  1186              rsx # 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm) aprog stp =
       
  1187       (6 + length ab, butlast lm @ last lm # rs # 
       
  1188                                0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)"
       
  1189       and g: "rec_calc_rel (Pr n f g) 
       
  1190                       (butlast lm @ [last lm - Suc xa]) rsxa"
       
  1191       "Suc xa \<le> last lm"
       
  1192       "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)" 
       
  1193       "rec_calc_rel (Pr n f g) lm rs"
       
  1194       "rec_ci g = (a, aa, ba)" 
       
  1195       "rec_ci f = (ab, ac, bc)" "lm \<noteq> []"
       
  1196     from g have k1: 
       
  1197       "\<exists> rs. rec_calc_rel (Pr n f g) (butlast lm @ [last lm - xa]) rs"
       
  1198       apply(rule_tac rs = rs in  calc_pr_less_ex, simp, simp)
       
  1199       done
       
  1200     from g and this show 
       
  1201       "\<exists>stp. abc_steps_l (6 + length ab, 
       
  1202            butlast lm @ (last lm - Suc xa) # rsxa # 
       
  1203               0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) aprog stp =
       
  1204               (6 + length ab, butlast lm @ last lm # rs # 
       
  1205                                 0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)"
       
  1206     proof(erule_tac exE)
       
  1207       fix rsa
       
  1208       assume k2: "rec_calc_rel (Pr n f g) 
       
  1209                            (butlast lm @ [last lm - xa]) rsa"
       
  1210       from g and k2 have
       
  1211       "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @ 
       
  1212        (last lm - Suc xa) # rsxa # 
       
  1213                0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm) aprog stp
       
  1214         = (6 + length ab, butlast lm @ (last lm - xa) # rsa # 
       
  1215                                0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)"
       
  1216 	proof -
       
  1217 	  from g have k2_1: 
       
  1218             "\<exists> ap bp. length ap = 6 + length ab \<and>
       
  1219                    aprog = ap [+] bp \<and> 
       
  1220                    bp = ([Dec (a_md - Suc 0) (length a + 7)] [+]
       
  1221                   (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, 
       
  1222                   Goto (Suc 0)])) @
       
  1223                   [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]"
       
  1224             apply(rule_tac ci_pr_ex1, auto)
       
  1225 	    done
       
  1226 	  from k2_1 and k2 and g show "?thesis"
       
  1227 	    proof(erule_tac exE, erule_tac exE)
       
  1228 	      fix ap bp
       
  1229 	      assume 
       
  1230                 "length ap = 6 + length ab \<and> 
       
  1231                  aprog = ap [+] bp \<and> bp =
       
  1232                 ([Dec (a_md - Suc 0) (length a + 7)] [+] 
       
  1233                 (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3,
       
  1234                 Goto (Suc 0)])) @ 
       
  1235                 [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" 
       
  1236 	      from g and this and k2 and g_ind show "?thesis"
       
  1237 		apply(insert abc_append_exc3[of 
       
  1238                   "butlast lm @ (last lm - Suc xa) # rsxa #
       
  1239                   0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # suf_lm" bp 0
       
  1240                   "butlast lm @ (last lm - xa) # rsa #
       
  1241                 0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm" "length ap" ap],
       
  1242                  simp)
       
  1243 		apply(subgoal_tac 
       
  1244                 "\<exists>stp. abc_steps_l (0, butlast lm @ (last lm - Suc xa)
       
  1245                            # rsxa # 0\<up>(a_md - Suc (Suc rs_pos)) @ Suc xa # 
       
  1246                               suf_lm) bp stp =
       
  1247 	          (0, butlast lm @ (last lm - xa) # rsa #
       
  1248                            0\<up>(a_md - Suc (Suc rs_pos)) @ xa # suf_lm)",
       
  1249                       simp, erule_tac conjE, erule conjE)
       
  1250 		apply(erule pr_cycle_part_ind, auto)
       
  1251 		done
       
  1252 	    qed
       
  1253 	  qed  
       
  1254       from g and k2 and this show "?thesis"
       
  1255 	apply(erule_tac exE)
       
  1256 	apply(insert ind[of rsa], simp)
       
  1257 	apply(erule_tac exE)
       
  1258 	apply(rule_tac x = "stp + stpa" in exI, 
       
  1259               simp add: abc_steps_add)
       
  1260 	done
       
  1261     qed
       
  1262   qed
       
  1263 qed
       
  1264 
       
  1265 lemma ci_pr_length: 
       
  1266   "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
       
  1267     rec_ci g = (a, aa, ba);  
       
  1268     rec_ci f = (ab, ac, bc)\<rbrakk>
       
  1269     \<Longrightarrow>  length aprog = 13 + length ab + length a"
       
  1270 apply(auto simp: rec_ci.simps)
       
  1271 done
       
  1272 
       
  1273 fun mv_box_inv :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> bool"
       
  1274   where
       
  1275   "mv_box_inv (as, lm) m n initlm = 
       
  1276          (let plus = initlm ! m + initlm ! n in
       
  1277            length initlm > max m n \<and> m \<noteq> n \<and> 
       
  1278               (if as = 0 then \<exists> k l. lm = initlm[m := k, n := l] \<and> 
       
  1279                     k + l = plus \<and> k \<le> initlm ! m 
       
  1280               else if as = 1 then \<exists> k l. lm = initlm[m := k, n := l]
       
  1281                              \<and> k + l + 1 = plus \<and> k < initlm ! m 
       
  1282               else if as = 2 then \<exists> k l. lm = initlm[m := k, n := l] 
       
  1283                               \<and> k + l = plus \<and> k \<le> initlm ! m
       
  1284               else if as = 3 then lm = initlm[m := 0, n := plus]
       
  1285               else False))"
       
  1286 
       
  1287 fun mv_box_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
       
  1288   where
       
  1289   "mv_box_stage1 (as, lm) m  = 
       
  1290             (if as = 3 then 0 
       
  1291              else 1)"
       
  1292 
       
  1293 fun mv_box_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
       
  1294   where
       
  1295   "mv_box_stage2 (as, lm) m = (lm ! m)"
       
  1296 
       
  1297 fun mv_box_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
       
  1298   where
       
  1299   "mv_box_stage3 (as, lm) m = (if as = 1 then 3 
       
  1300                                 else if as = 2 then 2
       
  1301                                 else if as = 0 then 1 
       
  1302                                 else 0)"
       
  1303  
       
  1304 fun mv_box_measure :: "((nat \<times> nat list) \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)"
       
  1305   where
       
  1306   "mv_box_measure ((as, lm), m) = 
       
  1307      (mv_box_stage1 (as, lm) m, mv_box_stage2 (as, lm) m,
       
  1308       mv_box_stage3 (as, lm) m)"
       
  1309 
       
  1310 definition lex_pair :: "((nat \<times> nat) \<times> nat \<times> nat) set"
       
  1311   where
       
  1312   "lex_pair = less_than <*lex*> less_than"
       
  1313 
       
  1314 definition lex_triple :: 
       
  1315  "((nat \<times> (nat \<times> nat)) \<times> (nat \<times> (nat \<times> nat))) set"
       
  1316   where
       
  1317   "lex_triple \<equiv> less_than <*lex*> lex_pair"
       
  1318 
       
  1319 definition mv_box_LE :: 
       
  1320  "(((nat \<times> nat list) \<times> nat) \<times> ((nat \<times> nat list) \<times> nat)) set"
       
  1321   where 
       
  1322   "mv_box_LE \<equiv> (inv_image lex_triple mv_box_measure)"
       
  1323 
       
  1324 lemma wf_lex_triple: "wf lex_triple"
       
  1325   by (auto intro:wf_lex_prod simp:lex_triple_def lex_pair_def)
       
  1326 
       
  1327 lemma wf_mv_box_le[intro]: "wf mv_box_LE"
       
  1328 by(auto intro:wf_inv_image wf_lex_triple simp: mv_box_LE_def)
       
  1329 
       
  1330 declare mv_box_inv.simps[simp del]
       
  1331 
       
  1332 lemma mv_box_inv_init:  
       
  1333 "\<lbrakk>m < length initlm; n < length initlm; m \<noteq> n\<rbrakk> \<Longrightarrow> 
       
  1334   mv_box_inv (0, initlm) m n initlm"
       
  1335 apply(simp add: abc_steps_l.simps mv_box_inv.simps)
       
  1336 apply(rule_tac x = "initlm ! m" in exI, 
       
  1337       rule_tac x = "initlm ! n" in exI, simp)
       
  1338 done
       
  1339 
       
  1340 lemma [simp]: "abc_fetch 0 (mv_box m n) = Some (Dec m 3)"
       
  1341 apply(simp add: mv_box.simps abc_fetch.simps)
       
  1342 done
       
  1343 
       
  1344 lemma [simp]: "abc_fetch (Suc 0) (mv_box m n) =
       
  1345                Some (Inc n)"
       
  1346 apply(simp add: mv_box.simps abc_fetch.simps)
       
  1347 done
       
  1348 
       
  1349 lemma [simp]: "abc_fetch 2 (mv_box m n) = Some (Goto 0)"
       
  1350 apply(simp add: mv_box.simps abc_fetch.simps)
       
  1351 done
       
  1352 
       
  1353 lemma [simp]: "abc_fetch 3 (mv_box m n) = None"
       
  1354 apply(simp add: mv_box.simps abc_fetch.simps)
       
  1355 done
       
  1356 
       
  1357 lemma [simp]: 
       
  1358   "\<lbrakk>m \<noteq> n; m < length initlm; n < length initlm;
       
  1359     k + l = initlm ! m + initlm ! n; k \<le> initlm ! m; 0 < k\<rbrakk>
       
  1360  \<Longrightarrow> \<exists>ka la. initlm[m := k, n := l, m := k - Suc 0] = 
       
  1361      initlm[m := ka, n := la] \<and>
       
  1362      Suc (ka + la) = initlm ! m + initlm ! n \<and> 
       
  1363      ka < initlm ! m"
       
  1364 apply(rule_tac x = "k - Suc 0" in exI, rule_tac x = l in exI, 
       
  1365       simp, auto)
       
  1366 apply(subgoal_tac 
       
  1367       "initlm[m := k, n := l, m := k - Suc 0] = 
       
  1368        initlm[n := l, m := k, m := k - Suc 0]")
       
  1369 apply(simp add: list_update_overwrite )
       
  1370 apply(simp add: list_update_swap)
       
  1371 apply(simp add: list_update_swap)
       
  1372 done
       
  1373 
       
  1374 lemma [simp]:
       
  1375   "\<lbrakk>m \<noteq> n; m < length initlm; n < length initlm; 
       
  1376     Suc (k + l) = initlm ! m + initlm ! n;
       
  1377     k < initlm ! m\<rbrakk>
       
  1378     \<Longrightarrow> \<exists>ka la. initlm[m := k, n := l, n := Suc l] = 
       
  1379                 initlm[m := ka, n := la] \<and> 
       
  1380                 ka + la = initlm ! m + initlm ! n \<and> 
       
  1381                 ka \<le> initlm ! m"
       
  1382 apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, auto)
       
  1383 done
       
  1384 
       
  1385 lemma [simp]: 
       
  1386   "\<lbrakk>length initlm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
       
  1387    \<forall>na. \<not> (\<lambda>(as, lm) m. as = 3) 
       
  1388     (abc_steps_l (0, initlm) (mv_box m n) na) m \<and> 
       
  1389   mv_box_inv (abc_steps_l (0, initlm) 
       
  1390            (mv_box m n) na) m n initlm \<longrightarrow>
       
  1391   mv_box_inv (abc_steps_l (0, initlm) 
       
  1392            (mv_box m n) (Suc na)) m n initlm \<and>
       
  1393   ((abc_steps_l (0, initlm) (mv_box m n) (Suc na), m),
       
  1394    abc_steps_l (0, initlm) (mv_box m n) na, m) \<in> mv_box_LE"
       
  1395 apply(rule allI, rule impI, simp add: abc_steps_ind)
       
  1396 apply(case_tac "(abc_steps_l (0, initlm) (mv_box m n) na)",
       
  1397       simp)
       
  1398 apply(auto split:if_splits simp add:abc_steps_l.simps mv_box_inv.simps)
       
  1399 apply(auto simp add: mv_box_LE_def lex_triple_def lex_pair_def 
       
  1400                      abc_step_l.simps abc_steps_l.simps
       
  1401                      mv_box_inv.simps abc_lm_v.simps abc_lm_s.simps
       
  1402                 split: if_splits )
       
  1403 apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, simp)
       
  1404 done
       
  1405 
       
  1406 lemma mv_box_inv_halt: 
       
  1407   "\<lbrakk>length initlm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
       
  1408   \<exists> stp. (\<lambda> (as, lm). as = 3 \<and> 
       
  1409   mv_box_inv (as, lm) m n initlm) 
       
  1410              (abc_steps_l (0::nat, initlm) (mv_box m n) stp)"
       
  1411 apply(insert halt_lemma2[of mv_box_LE
       
  1412     "\<lambda> ((as, lm), m). mv_box_inv (as, lm) m n initlm"
       
  1413     "\<lambda> stp. (abc_steps_l (0, initlm) (mv_box m n) stp, m)"
       
  1414     "\<lambda> ((as, lm), m). as = (3::nat)"
       
  1415     ])
       
  1416 apply(insert wf_mv_box_le)
       
  1417 apply(simp add: mv_box_inv_init abc_steps_zero)
       
  1418 apply(erule_tac exE)
       
  1419 apply(rule_tac x = na in exI)
       
  1420 apply(case_tac "(abc_steps_l (0, initlm) (mv_box m n) na)",
       
  1421       simp, auto)
       
  1422 done
       
  1423 
       
  1424 lemma mv_box_halt_cond:
       
  1425   "\<lbrakk>m \<noteq> n; mv_box_inv (a, b) m n lm; a = 3\<rbrakk> \<Longrightarrow> 
       
  1426   b = lm[n := lm ! m + lm ! n, m := 0]"
       
  1427 apply(simp add: mv_box_inv.simps, auto)
       
  1428 apply(simp add: list_update_swap)
       
  1429 done
       
  1430 
       
  1431 lemma mv_box_ex:
       
  1432   "\<lbrakk>length lm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
       
  1433   \<exists> stp. abc_steps_l (0::nat, lm) (mv_box m n) stp
       
  1434   = (3, (lm[n := (lm ! m + lm ! n)])[m := 0::nat])"
       
  1435 apply(drule mv_box_inv_halt, simp, erule_tac exE)
       
  1436 apply(rule_tac x = stp in exI)
       
  1437 apply(case_tac "abc_steps_l (0, lm) (mv_box m n) stp",
       
  1438       simp)
       
  1439 apply(erule_tac mv_box_halt_cond, auto)
       
  1440 done
       
  1441 
       
  1442 lemma [simp]: 
       
  1443   "\<lbrakk>a_md = Suc (max (Suc (Suc n)) (max bc ba)); 
       
  1444    length lm = rs_pos \<and> rs_pos = n \<and> n > 0\<rbrakk>
       
  1445   \<Longrightarrow> n - Suc 0 < length lm + 
       
  1446   (Suc (max (Suc (Suc n)) (max bc ba)) - rs_pos + length suf_lm) \<and>
       
  1447    Suc (Suc n) < length lm + (Suc (max (Suc (Suc n)) (max bc ba)) -
       
  1448   rs_pos + length suf_lm) \<and> bc < length lm + (Suc (max (Suc (Suc n)) 
       
  1449  (max bc ba)) - rs_pos + length suf_lm) \<and> ba < length lm + 
       
  1450   (Suc (max (Suc (Suc n)) (max bc ba)) - rs_pos + length suf_lm)"
       
  1451 apply(arith)
       
  1452 done
       
  1453 
       
  1454 lemma [simp]:
       
  1455   "\<lbrakk>a_md = Suc (max (Suc (Suc n)) (max bc ba)); 
       
  1456    length lm = rs_pos \<and> rs_pos = n \<and> n > 0\<rbrakk>
       
  1457  \<Longrightarrow> n - Suc 0 < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba)) \<and>
       
  1458      Suc n < length suf_lm + max (Suc (Suc n)) (max bc ba) \<and> 
       
  1459      bc < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba)) \<and> 
       
  1460      ba < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba))"
       
  1461 apply(arith)
       
  1462 done
       
  1463 
       
  1464 lemma [simp]: "n - Suc 0 \<noteq> max (Suc (Suc n)) (max bc ba)"
       
  1465 apply(arith)
       
  1466 done
       
  1467 
       
  1468 lemma [simp]: 
       
  1469   "a_md \<ge> Suc bc \<and> rs_pos > 0 \<and> bc \<ge> rs_pos \<Longrightarrow> 
       
  1470  bc - (rs_pos - Suc 0) + a_md - Suc bc = Suc (a_md - rs_pos - Suc 0)"
       
  1471 apply(arith)
       
  1472 done
       
  1473 
       
  1474 lemma [simp]: "length lm = n \<and> rs_pos = n \<and> 0 < rs_pos \<and> 
       
  1475                                                   Suc rs_pos < a_md 
       
  1476        \<Longrightarrow> n - Suc 0 < Suc (Suc (a_md + length suf_lm - Suc (Suc 0))) 
       
  1477         \<and> n < Suc (Suc (a_md + length suf_lm - Suc (Suc 0)))"
       
  1478 apply(arith)
       
  1479 done
       
  1480      
       
  1481 lemma [simp]: "length lm = n \<and> rs_pos = n \<and> 0 < rs_pos \<and> 
       
  1482                Suc rs_pos < a_md \<Longrightarrow> n - Suc 0 \<noteq> n"
       
  1483 by arith
       
  1484 
       
  1485 lemma ci_pr_ex2: 
       
  1486   "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
       
  1487     rec_calc_rel (Pr n f g) lm rs; 
       
  1488     rec_ci g = (a, aa, ba); 
       
  1489     rec_ci f = (ab, ac, bc)\<rbrakk>
       
  1490   \<Longrightarrow> \<exists>ap bp. aprog = ap [+] bp \<and> 
       
  1491          ap = mv_box n (max (Suc (Suc (Suc n))) (max bc ba))"
       
  1492 apply(simp add: rec_ci.simps)
       
  1493 apply(rule_tac x = "(ab [+] (mv_box n (Suc n) [+]
       
  1494               ([Dec (max (n + 3) (max bc ba)) (length a + 7)] 
       
  1495       [+] (a [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @ 
       
  1496       [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]))" in exI, auto)
       
  1497 apply(simp add: abc_append_commute numeral_3_eq_3)
       
  1498 done
       
  1499 
       
  1500 lemma [simp]: 
       
  1501   "max (Suc (Suc (Suc n))) (max bc ba) - n < 
       
  1502      Suc (max (Suc (Suc (Suc n))) (max bc ba)) - n"
       
  1503 apply(arith)
       
  1504 done
       
  1505 
       
  1506 lemma [simp]: "length lm = n \<and> rs_pos = n \<and> 0 < n \<Longrightarrow> 
       
  1507                       lm[n - Suc 0 := 0::nat] = butlast lm @ [0]"
       
  1508 apply(auto)
       
  1509 apply(insert list_update_append[of "butlast lm" "[last lm]" 
       
  1510                                    "length lm - Suc 0" "0"], simp)
       
  1511 done
       
  1512 
       
  1513 lemma [simp]: "\<lbrakk>length lm = n; 0 < n\<rbrakk>  \<Longrightarrow> lm ! (n - Suc 0) = last lm"
       
  1514 apply(insert nth_append[of "butlast lm" "[last lm]" "n - Suc 0"],
       
  1515       simp)
       
  1516 apply(insert butlast_append_last[of lm], auto)
       
  1517 done
       
  1518 lemma exp_suc_iff: "a\<up>b @ [a] = a\<up>(b + Suc 0)"
       
  1519 apply(simp add: exp_ind del: replicate.simps)
       
  1520 done
       
  1521 
       
  1522 lemma less_not_less[simp]: "n > 0 \<Longrightarrow> \<not> n < n - Suc 0"
       
  1523 by auto
       
  1524 
       
  1525 lemma [simp]:
       
  1526   "Suc n < length suf_lm + max (Suc (Suc n)) (max bc ba) \<and> 
       
  1527   bc < Suc (length suf_lm + max (Suc (Suc n)) 
       
  1528   (max bc ba)) \<and> 
       
  1529   ba < Suc (length suf_lm + max (Suc (Suc n)) (max bc ba))"
       
  1530   by arith
       
  1531 
       
  1532 lemma [simp]: "length lm = n \<and> rs_pos = n \<and> n > 0 \<Longrightarrow> 
       
  1533 (lm @ 0\<up>(Suc (max (Suc (Suc n)) (max bc ba)) - n) @ suf_lm) 
       
  1534   [max (Suc (Suc n)) (max bc ba) :=
       
  1535    (lm @ 0\<up>(Suc (max (Suc (Suc n)) (max bc ba)) - n) @ suf_lm) ! (n - Suc 0) + 
       
  1536        (lm @ 0\<up>(Suc (max (Suc (Suc n)) (max bc ba)) - n) @ suf_lm) ! 
       
  1537                    max (Suc (Suc n)) (max bc ba), n - Suc 0 := 0::nat]
       
  1538  = butlast lm @ 0 # 0\<up>(max (Suc (Suc n)) (max bc ba) - n) @ last lm # suf_lm"
       
  1539 apply(simp add: nth_append nth_replicate list_update_append)
       
  1540 apply(insert list_update_append[of "0\<up>((max (Suc (Suc n)) (max bc ba)) - n)"
       
  1541          "[0]" "max (Suc (Suc n)) (max bc ba) - n" "last lm"], simp)
       
  1542 apply(simp add: exp_suc_iff Suc_diff_le del: list_update.simps)
       
  1543 done
       
  1544 
       
  1545 lemma exp_eq: "(a = b) = (c\<up>a = c\<up>b)"
       
  1546 apply(auto)
       
  1547 done
       
  1548 
       
  1549 lemma [simp]:
       
  1550   "\<lbrakk>length lm = n; 0 < n;  Suc n < a_md\<rbrakk> \<Longrightarrow> 
       
  1551    (butlast lm @ rsa # 0\<up>(a_md - Suc n) @ last lm # suf_lm)
       
  1552     [n := (butlast lm @ rsa # 0\<up>(a_md - Suc n) @ last lm # suf_lm) ! 
       
  1553         (n - Suc 0) + (butlast lm @ rsa # (0::nat)\<up>(a_md - Suc n) @ 
       
  1554                                 last lm # suf_lm) ! n, n - Suc 0 := 0]
       
  1555  = butlast lm @ 0 # rsa # 0\<up>(a_md - Suc (Suc n)) @ last lm # suf_lm"
       
  1556 apply(simp add: nth_append list_update_append)
       
  1557 apply(case_tac "a_md - Suc n", auto)
       
  1558 done
       
  1559 
       
  1560 lemma [simp]: 
       
  1561   "Suc (Suc rs_pos) \<le> a_md \<and> length lm = rs_pos \<and> 0 < rs_pos
       
  1562   \<Longrightarrow> a_md - Suc 0 < 
       
  1563           Suc (Suc (Suc (a_md + length suf_lm - Suc (Suc (Suc 0)))))"
       
  1564 by arith
       
  1565 
       
  1566 lemma [simp]: 
       
  1567   "Suc (Suc rs_pos) \<le> a_md \<and> length lm = rs_pos \<and> 0 < rs_pos \<Longrightarrow> 
       
  1568                                    \<not> a_md - Suc 0 < rs_pos - Suc 0"
       
  1569 by arith
       
  1570 
       
  1571 lemma [simp]: "Suc (Suc rs_pos) \<le> a_md \<Longrightarrow> 
       
  1572                                 \<not> a_md - Suc 0 < rs_pos - Suc 0"
       
  1573 by arith
       
  1574 
       
  1575 lemma [simp]: "\<lbrakk>Suc (Suc rs_pos) \<le> a_md\<rbrakk> \<Longrightarrow> 
       
  1576                \<not> a_md - rs_pos < Suc (Suc (a_md - Suc (Suc rs_pos)))"
       
  1577 by arith 
       
  1578 
       
  1579 lemma [simp]: 
       
  1580   "Suc (Suc rs_pos) \<le> a_md \<and> length lm = rs_pos \<and> 0 < rs_pos
       
  1581  \<Longrightarrow> (abc_lm_v (butlast lm @ last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos)) @
       
  1582         0 # suf_lm) (a_md - Suc 0) = 0 \<longrightarrow>
       
  1583       abc_lm_s (butlast lm @ last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos)) @ 
       
  1584         0 # suf_lm) (a_md - Suc 0) 0 = 
       
  1585          lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) \<and>
       
  1586      abc_lm_v (butlast lm @ last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos)) @ 
       
  1587                0 # suf_lm) (a_md - Suc 0) = 0"
       
  1588 apply(simp add: abc_lm_v.simps nth_append abc_lm_s.simps)
       
  1589 apply(insert nth_append[of "last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos))" 
       
  1590                "0 # suf_lm" "(a_md - rs_pos)"], auto)
       
  1591 apply(simp only: exp_suc_iff)
       
  1592 apply(subgoal_tac "a_md - Suc 0 < a_md + length suf_lm", simp)
       
  1593 apply(case_tac "lm = []", auto)
       
  1594 done
       
  1595 
       
  1596 lemma pr_prog_ex[simp]: "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
       
  1597       rec_ci g = (a, aa, ba); rec_ci f = (ab, ac, bc)\<rbrakk>
       
  1598     \<Longrightarrow> \<exists>cp. aprog = mv_box n (max (n + 3) 
       
  1599                     (max bc ba)) [+] cp"
       
  1600 apply(simp add: rec_ci.simps)
       
  1601 apply(rule_tac x = "(ab [+] (mv_box n (Suc n) [+]
       
  1602               ([Dec (max (n + 3) (max bc ba)) (length a + 7)] 
       
  1603              [+] (a [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]))
       
  1604              @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]))" in exI)
       
  1605 apply(auto simp: abc_append_commute)
       
  1606 done
       
  1607 
       
  1608 lemma [simp]: "mv_box m n \<noteq> []"
       
  1609 by (simp add: mv_box.simps)
       
  1610 (*
       
  1611 lemma [simp]: "\<lbrakk>rs_pos = n; 0 < rs_pos ; Suc rs_pos < a_md\<rbrakk> \<Longrightarrow> 
       
  1612                         n - Suc 0 < a_md + length suf_lm"
       
  1613 by arith
       
  1614 *)
       
  1615 lemma [intro]: 
       
  1616   "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md); 
       
  1617     rec_ci f = (ab, ac, bc)\<rbrakk> \<Longrightarrow> 
       
  1618    \<exists>ap. (\<exists>cp. aprog = ap [+] ab [+] cp) \<and> length ap = 3"
       
  1619 apply(case_tac "rec_ci g", simp add: rec_ci.simps)
       
  1620 apply(rule_tac x = "mv_box n 
       
  1621               (max (n + 3) (max bc c))" in exI, simp)
       
  1622 apply(rule_tac x = "mv_box n (Suc n) [+]
       
  1623                  ([Dec (max (n + 3) (max bc c)) (length a + 7)]
       
  1624                  [+] a [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])
       
  1625                @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" in exI, 
       
  1626       auto)
       
  1627 apply(simp add: abc_append_commute)
       
  1628 done
       
  1629 
       
  1630 lemma [intro]: 
       
  1631   "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
       
  1632     rec_ci g = (a, aa, ba); 
       
  1633     rec_ci f = (ab, ac, bc)\<rbrakk> \<Longrightarrow> 
       
  1634     \<exists>ap. (\<exists>cp. aprog = ap [+] mv_box n (Suc n) [+] cp)
       
  1635       \<and> length ap = 3 + length ab"
       
  1636 apply(simp add: rec_ci.simps)
       
  1637 apply(rule_tac x = "mv_box n (max (n + 3)
       
  1638                                 (max bc ba)) [+] ab" in exI, simp)
       
  1639 apply(rule_tac x = "([Dec (max (n + 3) (max bc ba))
       
  1640   (length a + 7)] [+] a [+] 
       
  1641   [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ 
       
  1642   [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]" in exI)
       
  1643 apply(auto simp: abc_append_commute)
       
  1644 done
       
  1645 
       
  1646 lemma [intro]: 
       
  1647   "\<lbrakk>rec_ci (Pr n f g) = (aprog, rs_pos, a_md);
       
  1648     rec_ci g = (a, aa, ba); 
       
  1649     rec_ci f = (ab, ac, bc)\<rbrakk>
       
  1650     \<Longrightarrow> \<exists>ap. (\<exists>cp. aprog = ap [+] ([Dec (a_md - Suc 0) (length a + 7)]
       
  1651              [+] (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, 
       
  1652              Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n),
       
  1653              Goto (length a + 4)] [+] cp) \<and>
       
  1654              length ap = 6 + length ab"
       
  1655 apply(simp add: rec_ci.simps)
       
  1656 apply(rule_tac x = "mv_box n
       
  1657     (max (n + 3) (max bc ba)) [+] ab [+] 
       
  1658      mv_box n (Suc n)" in exI, simp)
       
  1659 apply(rule_tac x = "[]" in exI, auto)
       
  1660 apply(simp add: abc_append_commute)
       
  1661 done
       
  1662 
       
  1663 lemma [simp]: 
       
  1664   "n < Suc (max (n + 3) (max bc ba) + length suf_lm) \<and> 
       
  1665    Suc (Suc n) < max (n + 3) (max bc ba) + length suf_lm \<and> 
       
  1666    bc < Suc (max (n + 3) (max bc ba) + length suf_lm) \<and> 
       
  1667    ba < Suc (max (n + 3) (max bc ba) + length suf_lm)"
       
  1668 by arith
       
  1669 
       
  1670 lemma [simp]: "n \<noteq> max (n + (3::nat)) (max bc ba)"
       
  1671 by arith
       
  1672 
       
  1673 lemma [simp]:"length lm = Suc n \<Longrightarrow> lm[n := (0::nat)] = butlast lm @ [0]"
       
  1674 apply(subgoal_tac "\<exists> xs x. lm = xs @ [x]", auto simp: list_update_append)
       
  1675 apply(rule_tac x = "butlast lm" in exI, rule_tac x = "last lm" in exI)
       
  1676 apply(case_tac lm, auto)
       
  1677 done
       
  1678 
       
  1679 lemma [simp]:  "length lm = Suc n \<Longrightarrow> lm ! n =last lm"
       
  1680 apply(subgoal_tac "lm \<noteq> []")
       
  1681 apply(simp add: last_conv_nth, case_tac lm, simp_all)
       
  1682 done
       
  1683 
       
  1684 lemma [simp]: "length lm = Suc n \<Longrightarrow> 
       
  1685       (lm @ (0::nat)\<up>(max (n + 3) (max bc ba) - n) @ suf_lm)
       
  1686            [max (n + 3) (max bc ba) := (lm @ 0\<up>(max (n + 3) (max bc ba) - n) @ suf_lm) ! n + 
       
  1687                   (lm @ 0\<up>(max (n + 3) (max bc ba) - n) @ suf_lm) ! max (n + 3) (max bc ba), n := 0]
       
  1688        = butlast lm @ 0 # 0\<up>(max (n + 3) (max bc ba) - Suc n) @ last lm # suf_lm"
       
  1689 apply(auto simp: list_update_append nth_append)
       
  1690 apply(subgoal_tac "(0\<up>(max (n + 3) (max bc ba) - n)) = 0\<up>(max (n + 3) (max bc ba) - Suc n) @ [0::nat]")
       
  1691 apply(simp add: list_update_append)
       
  1692 apply(simp add: exp_suc_iff)
       
  1693 done
       
  1694 
       
  1695 lemma [simp]: "Suc (Suc n) < a_md \<Longrightarrow>  
       
  1696       n < Suc (Suc (a_md + length suf_lm - 2)) \<and>
       
  1697         n < Suc (a_md + length suf_lm - 2)"
       
  1698 by(arith)
       
  1699 
       
  1700 lemma [simp]: "\<lbrakk>length lm = Suc n; Suc (Suc n) < a_md\<rbrakk>
       
  1701         \<Longrightarrow>(butlast lm @ (rsa::nat) # 0\<up>(a_md - Suc (Suc n)) @ last lm # suf_lm)
       
  1702           [Suc n := (butlast lm @ rsa # 0\<up>(a_md - Suc (Suc n)) @ last lm # suf_lm) ! n +
       
  1703                   (butlast lm @ rsa # 0\<up>(a_md - Suc (Suc n)) @ last lm # suf_lm) ! Suc n, n := 0]
       
  1704     = butlast lm @ 0 # rsa # 0\<up>(a_md - Suc (Suc (Suc n))) @ last lm # suf_lm"
       
  1705 apply(auto simp: list_update_append)
       
  1706 apply(subgoal_tac "(0\<up>(a_md - Suc (Suc n))) = (0::nat) # (0\<up>(a_md - Suc (Suc (Suc n))))", simp add: nth_append)
       
  1707 apply(simp add: replicate_Suc[THEN sym])
       
  1708 done
       
  1709 
       
  1710 lemma pr_case:
       
  1711   assumes nf_ind:
       
  1712   "\<And> lm rs suf_lm. rec_calc_rel f lm rs \<Longrightarrow> 
       
  1713   \<exists>stp. abc_steps_l (0, lm @ 0\<up>(bc - ac) @ suf_lm) ab stp = 
       
  1714                 (length ab, lm @ rs # 0\<up>(bc - Suc ac) @ suf_lm)"
       
  1715   and ng_ind: "\<And> lm rs suf_lm. rec_calc_rel g lm rs \<Longrightarrow> 
       
  1716         \<exists>stp. abc_steps_l (0, lm @ 0\<up>(ba - aa) @ suf_lm) a stp = 
       
  1717                        (length a, lm @ rs # 0\<up>(ba - Suc aa) @ suf_lm)"
       
  1718     and h: "rec_ci (Pr n f g) = (aprog, rs_pos, a_md)"  "rec_calc_rel (Pr n f g) lm rs" 
       
  1719            "rec_ci g = (a, aa, ba)" "rec_ci f = (ab, ac, bc)" 
       
  1720   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)"
       
  1721 proof -
       
  1722   from h have k1: "\<exists> stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp
       
  1723     = (3, butlast lm @ 0 # 0\<up>(a_md - rs_pos - 1) @ last lm # suf_lm)"
       
  1724   proof -
       
  1725     have "\<exists>bp cp. aprog = bp [+] cp \<and> bp = mv_box n 
       
  1726                  (max (n + 3) (max bc ba))"
       
  1727       apply(insert h, simp)
       
  1728       apply(erule pr_prog_ex, auto)
       
  1729       done
       
  1730     thus "?thesis"
       
  1731       apply(erule_tac exE, erule_tac exE, simp)
       
  1732       apply(subgoal_tac 
       
  1733            "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm)
       
  1734               ([] [+] mv_box n
       
  1735                   (max (n + 3) (max bc ba)) [+] cp) stp =
       
  1736              (0 + 3, butlast lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ 
       
  1737                                         last lm # suf_lm)", simp)
       
  1738       apply(rule_tac abc_append_exc1, simp_all)
       
  1739       apply(insert mv_box_ex[of "n" "(max (n + 3) 
       
  1740                  (max bc ba))" "lm @ 0\<up>(a_md - rs_pos) @ suf_lm"], simp)
       
  1741       apply(subgoal_tac "a_md = Suc (max (n + 3) (max bc ba))",
       
  1742             simp)
       
  1743       apply(subgoal_tac "length lm = Suc n \<and> rs_pos = Suc n", simp)
       
  1744       apply(insert h)
       
  1745       apply(simp add: para_pattern ci_pr_para_eq)
       
  1746       apply(rule ci_pr_md_def, auto)
       
  1747       done
       
  1748   qed
       
  1749   from h have k2: 
       
  1750   "\<exists> stp. abc_steps_l (3,  butlast lm @ 0 # 0\<up>(a_md - rs_pos - 1) @ 
       
  1751              last lm # suf_lm) aprog stp 
       
  1752     = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  1753   proof -
       
  1754     from h have k2_1: "\<exists> rs. rec_calc_rel f (butlast lm) rs"
       
  1755       apply(erule_tac calc_pr_zero_ex)
       
  1756       done
       
  1757     thus "?thesis"
       
  1758     proof(erule_tac exE)
       
  1759       fix rsa
       
  1760       assume k2_2: "rec_calc_rel f (butlast lm) rsa"
       
  1761       from h and k2_2 have k2_2_1: 
       
  1762        "\<exists> stp. abc_steps_l (3, butlast lm @ 0 # 0\<up>(a_md - rs_pos - 1) 
       
  1763                  @ last lm # suf_lm) aprog stp
       
  1764         = (3 + length ab, butlast lm @ rsa # 0\<up>(a_md - rs_pos - 1) @ 
       
  1765                                              last lm # suf_lm)"
       
  1766       proof -
       
  1767 	from h have j1: "
       
  1768           \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = 3 \<and> 
       
  1769               bp = ab"
       
  1770 	  apply(auto)
       
  1771 	  done
       
  1772 	from h have j2: "ac = rs_pos - 1"
       
  1773 	  apply(drule_tac ci_pr_f_paras, simp, auto)
       
  1774 	  done
       
  1775 	from h and j2 have j3: "a_md \<ge> Suc bc \<and> rs_pos > 0 \<and> bc \<ge> rs_pos"
       
  1776 	  apply(rule_tac conjI)
       
  1777 	  apply(erule_tac ab = ab and ac = ac in ci_pr_md_ge_f, simp)
       
  1778 	  apply(rule_tac context_conjI)
       
  1779           apply(simp_all add: rec_ci.simps)
       
  1780 	  apply(drule_tac ci_ad_ge_paras, drule_tac ci_ad_ge_paras)
       
  1781 	  apply(arith)
       
  1782 	  done	  
       
  1783 	from j1 and j2 show "?thesis"
       
  1784 	  apply(auto simp del: abc_append_commute)
       
  1785 	  apply(rule_tac abc_append_exc1, simp_all)
       
  1786 	  apply(insert nf_ind[of "butlast lm" "rsa" 
       
  1787                 "0\<up>(a_md - bc - Suc 0) @ last lm # suf_lm"], 
       
  1788                simp add: k2_2 j2, erule_tac exE)
       
  1789 	  apply(simp add: exponent_add_iff j3)
       
  1790 	  apply(rule_tac x = "stp" in exI, simp)
       
  1791 	  done
       
  1792       qed
       
  1793       from h have k2_2_2: 
       
  1794       "\<exists> stp. abc_steps_l (3 + length ab, butlast lm @ rsa # 
       
  1795                   0\<up>(a_md - rs_pos - 1) @ last lm # suf_lm) aprog stp
       
  1796         = (6 + length ab, butlast lm @ 0 # rsa # 
       
  1797                        0\<up>(a_md - rs_pos - 2) @ last lm # suf_lm)"
       
  1798       proof -	     
       
  1799 	from h have "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
       
  1800           length ap = 3 + length ab \<and> bp = mv_box n (Suc n)"
       
  1801 	  by auto
       
  1802 	thus "?thesis"
       
  1803 	proof(erule_tac exE, erule_tac exE, erule_tac exE, 
       
  1804               erule_tac exE)
       
  1805 	  fix ap cp bp apa
       
  1806 	  assume "aprog = ap [+] bp [+] cp \<and> length ap = 3 + 
       
  1807                     length ab \<and> bp = mv_box n (Suc n)"
       
  1808 	  thus "?thesis"
       
  1809 	    apply(simp del: abc_append_commute)
       
  1810 	    apply(subgoal_tac 
       
  1811               "\<exists>stp. abc_steps_l (3 + length ab, 
       
  1812                butlast lm @ rsa # 0\<up>(a_md - Suc rs_pos) @
       
  1813                  last lm # suf_lm) (ap [+] 
       
  1814                    mv_box n (Suc n) [+] cp) stp =
       
  1815               ((3 + length ab) + 3, butlast lm @ 0 # rsa # 
       
  1816                   0\<up>(a_md - Suc (Suc rs_pos)) @ last lm # suf_lm)", simp)
       
  1817 	    apply(rule_tac abc_append_exc1, simp_all)
       
  1818 	    apply(insert mv_box_ex[of n "Suc n" 
       
  1819                     "butlast lm @ rsa # 0\<up>(a_md - Suc rs_pos) @ 
       
  1820                           last lm # suf_lm"], simp)
       
  1821 	    apply(subgoal_tac "length lm = Suc n \<and> rs_pos = Suc n \<and> a_md > Suc (Suc n)", simp)
       
  1822 	    apply(insert h, simp)
       
  1823             done
       
  1824 	qed
       
  1825       qed
       
  1826       from h have k2_3: "lm \<noteq> []"
       
  1827 	apply(rule_tac calc_pr_para_not_null, simp)
       
  1828 	done
       
  1829       from h and k2_2 and k2_3 have k2_2_3: 
       
  1830       "\<exists> stp. abc_steps_l (6 + length ab, butlast lm @ 
       
  1831           (last lm - last lm) # rsa # 
       
  1832             0\<up>(a_md - (Suc (Suc rs_pos))) @ last lm # suf_lm) aprog stp
       
  1833         = (6 + length ab, butlast lm @ last lm # rs # 
       
  1834                         0\<up>(a_md - Suc (Suc (rs_pos))) @ 0 # suf_lm)"
       
  1835 	apply(rule_tac x = "last lm" and g = g in pr_cycle_part, auto)
       
  1836 	apply(rule_tac ng_ind, simp)
       
  1837 	apply(rule_tac rec_calc_rel_def0, simp, simp)
       
  1838 	done
       
  1839       from h  have k2_2_4: 
       
  1840        "\<exists> stp. abc_steps_l (6 + length ab,
       
  1841              butlast lm @ last lm # rs # 0\<up>(a_md - rs_pos - 2) @
       
  1842                   0 # suf_lm) aprog stp
       
  1843         = (13 + length ab + length a,
       
  1844                    lm @ rs # 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
       
  1845       proof -
       
  1846 	from h have 
       
  1847         "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
       
  1848                      length ap = 6 + length ab \<and> 
       
  1849                     bp = ([Dec (a_md - Suc 0) (length a + 7)] [+] 
       
  1850                          (a [+] [Inc (rs_pos - Suc 0), 
       
  1851                          Dec rs_pos 3, Goto (Suc 0)])) @ 
       
  1852                         [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length a + 4)]"
       
  1853 	  by auto
       
  1854 	thus "?thesis"
       
  1855 	  apply(auto)
       
  1856 	  apply(subgoal_tac  
       
  1857             "\<exists>stp. abc_steps_l (6 + length ab, butlast lm @ 
       
  1858                 last lm # rs # 0\<up>(a_md - Suc (Suc rs_pos)) @ 0 # suf_lm)
       
  1859                 (ap [+] ([Dec (a_md - Suc 0) (length a + 7)] [+] 
       
  1860                 (a [+] [Inc (rs_pos - Suc 0), Dec rs_pos 3, 
       
  1861                 Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), 
       
  1862                 Goto (length a + 4)] [+] cp) stp =
       
  1863             (6 + length ab + (length a + 7) , 
       
  1864                  lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)", simp)
       
  1865 	  apply(subgoal_tac "13 + (length ab + length a) = 
       
  1866                               13 + length ab + length a", simp)
       
  1867 	  apply(arith)
       
  1868 	  apply(rule abc_append_exc1, simp_all)
       
  1869 	  apply(rule_tac x = "Suc 0" in exI, 
       
  1870                 simp add: abc_steps_l.simps abc_fetch.simps
       
  1871                          nth_append abc_append_nth abc_step_l.simps)
       
  1872 	  apply(subgoal_tac "a_md > Suc (Suc rs_pos) \<and> 
       
  1873                             length lm = rs_pos \<and> rs_pos > 0", simp)
       
  1874 	  apply(insert h, simp)
       
  1875 	  apply(subgoal_tac "rs_pos = Suc n", simp, simp)
       
  1876           done
       
  1877       qed
       
  1878       from h have k2_2_5: "length aprog = 13 + length ab + length a"
       
  1879 	apply(rule_tac ci_pr_length, simp_all)
       
  1880 	done
       
  1881       from k2_2_1 and k2_2_2 and k2_2_3 and k2_2_4 and k2_2_5 
       
  1882       show "?thesis"
       
  1883 	apply(auto)
       
  1884 	apply(rule_tac x = "stp + stpa + stpb + stpc" in exI, 
       
  1885               simp add: abc_steps_add)
       
  1886 	done
       
  1887     qed
       
  1888   qed	
       
  1889   from k1 and k2 show 
       
  1890     "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp 
       
  1891                = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  1892     apply(erule_tac exE)
       
  1893     apply(erule_tac exE)
       
  1894     apply(rule_tac x = "stp + stpa" in exI)
       
  1895     apply(simp add: abc_steps_add)
       
  1896     done
       
  1897 qed
       
  1898 
       
  1899 lemma eq_switch: "x = y \<Longrightarrow> y = x"
       
  1900 by simp
       
  1901 
       
  1902 lemma [simp]: 
       
  1903   "\<lbrakk>rec_ci f = (a, aa, ba); 
       
  1904     rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk> \<Longrightarrow> \<exists>bp. aprog = a @ bp"
       
  1905 apply(simp add: rec_ci.simps)
       
  1906 apply(rule_tac x = "[Dec (Suc n) (length a + 5), 
       
  1907       Dec (Suc n) (length a + 3), Goto (Suc (length a)), 
       
  1908       Inc n, Goto 0]" in exI, auto)
       
  1909 done
       
  1910 
       
  1911 lemma ci_mn_para_eq[simp]: 
       
  1912   "rec_ci (Mn n f) = (aprog, rs_pos, a_md) \<Longrightarrow> rs_pos = n"
       
  1913 apply(case_tac "rec_ci f", simp add: rec_ci.simps)
       
  1914 done
       
  1915 
       
  1916 lemma [simp]: "rec_ci f = (a, aa, ba) \<Longrightarrow> aa < ba"
       
  1917 apply(simp add: ci_ad_ge_paras)
       
  1918 done
       
  1919 
       
  1920 lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); 
       
  1921                 rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
       
  1922     \<Longrightarrow> ba \<le> a_md"
       
  1923 apply(simp add: rec_ci.simps)
       
  1924 by arith
       
  1925 
       
  1926 lemma mn_calc_f: 
       
  1927   assumes ind: 
       
  1928   "\<And>aprog a_md rs_pos rs suf_lm lm.
       
  1929   \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\<rbrakk>  
       
  1930   \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp    
       
  1931            = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
       
  1932   and h: "rec_ci f = (a, aa, ba)" 
       
  1933          "rec_ci (Mn n f) = (aprog, rs_pos, a_md)"  
       
  1934          "rec_calc_rel f (lm @ [x]) rsx" 
       
  1935          "aa = Suc n"
       
  1936   shows "\<exists>stp. abc_steps_l (0, lm @ x # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
       
  1937                   aprog stp = (length a, 
       
  1938                    lm @ x # rsx # 0\<up>(a_md - Suc (Suc rs_pos)) @ suf_lm)"
       
  1939 proof -
       
  1940   from h have k1: "\<exists> ap bp. aprog = ap @ bp \<and> ap = a"
       
  1941     by simp
       
  1942   from h have k2: "rs_pos = n"
       
  1943     apply(erule_tac ci_mn_para_eq)
       
  1944     done
       
  1945   from h and k1 and k2 show "?thesis"
       
  1946   
       
  1947   proof(erule_tac exE, erule_tac exE, simp, 
       
  1948         rule_tac abc_add_exc1, auto)
       
  1949     fix bp
       
  1950     show 
       
  1951       "\<exists>astp. abc_steps_l (0, lm @ x # 0\<up>(a_md - Suc n) @ suf_lm) a astp
       
  1952       = (length a, lm @ x # rsx # 0\<up>(a_md - Suc (Suc n)) @ suf_lm)"
       
  1953       apply(insert ind[of a "Suc n" ba  "lm @ [x]" rsx 
       
  1954              "0\<up>(a_md - ba) @ suf_lm"], simp add: exponent_add_iff h k2)
       
  1955       apply(subgoal_tac "ba > aa \<and> a_md \<ge> ba \<and> aa = Suc n", 
       
  1956             insert h, auto)
       
  1957       done
       
  1958   qed
       
  1959 qed
       
  1960 
       
  1961 fun mn_ind_inv ::
       
  1962   "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> nat list \<Rightarrow> bool"
       
  1963   where
       
  1964   "mn_ind_inv (as, lm') ss x rsx suf_lm lm = 
       
  1965            (if as = ss then lm' = lm @ x # rsx # suf_lm
       
  1966             else if as = ss + 1 then 
       
  1967                  \<exists>y. (lm' = lm @ x # y # suf_lm) \<and> y \<le> rsx
       
  1968             else if as = ss + 2 then 
       
  1969                  \<exists>y. (lm' = lm @ x # y # suf_lm) \<and> y \<le> rsx
       
  1970             else if as = ss + 3 then lm' = lm @ x # 0 # suf_lm
       
  1971             else if as = ss + 4 then lm' = lm @ Suc x # 0 # suf_lm
       
  1972             else if as = 0 then lm' = lm @ Suc x # 0 # suf_lm
       
  1973             else False
       
  1974 )"
       
  1975 
       
  1976 fun mn_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
       
  1977   where
       
  1978   "mn_stage1 (as, lm) ss n = 
       
  1979             (if as = 0 then 0 
       
  1980              else if as = ss + 4 then 1
       
  1981              else if as = ss + 3 then 2
       
  1982              else if as = ss + 2 \<or> as = ss + 1 then 3
       
  1983              else if as = ss then 4
       
  1984              else 0
       
  1985 )"
       
  1986 
       
  1987 fun mn_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
       
  1988   where
       
  1989   "mn_stage2 (as, lm) ss n = 
       
  1990             (if as = ss + 1 \<or> as = ss + 2 then (lm ! (Suc n))
       
  1991              else 0)"
       
  1992 
       
  1993 fun mn_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
       
  1994   where
       
  1995   "mn_stage3 (as, lm) ss n = (if as = ss + 2 then 1 else 0)"
       
  1996 
       
  1997  
       
  1998 fun mn_measure :: "((nat \<times> nat list) \<times> nat \<times> nat) \<Rightarrow>
       
  1999                                                 (nat \<times> nat \<times> nat)"
       
  2000   where
       
  2001   "mn_measure ((as, lm), ss, n) = 
       
  2002      (mn_stage1 (as, lm) ss n, mn_stage2 (as, lm) ss n,
       
  2003                                        mn_stage3 (as, lm) ss n)"
       
  2004 
       
  2005 definition mn_LE :: "(((nat \<times> nat list) \<times> nat \<times> nat) \<times>
       
  2006                      ((nat \<times> nat list) \<times> nat \<times> nat)) set"
       
  2007   where "mn_LE \<equiv> (inv_image lex_triple mn_measure)"
       
  2008 
       
  2009 lemma wf_mn_le[intro]: "wf mn_LE"
       
  2010 by(auto intro:wf_inv_image wf_lex_triple simp: mn_LE_def)
       
  2011 
       
  2012 declare mn_ind_inv.simps[simp del]
       
  2013 
       
  2014 lemma mn_inv_init: 
       
  2015   "mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 0)
       
  2016                                          (length a) x rsx suf_lm lm"
       
  2017 apply(simp add: mn_ind_inv.simps abc_steps_zero)
       
  2018 done
       
  2019 
       
  2020 lemma mn_halt_init: 
       
  2021   "rec_ci f = (a, aa, ba) \<Longrightarrow> 
       
  2022   \<not> (\<lambda>(as, lm') (ss, n). as = 0) 
       
  2023     (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 0) 
       
  2024                                                        (length a, n)"
       
  2025 apply(simp add: abc_steps_zero)
       
  2026 apply(erule_tac rec_ci_not_null)
       
  2027 done
       
  2028 
       
  2029 lemma [simp]: 
       
  2030   "\<lbrakk>rec_ci f = (a, aa, ba); 
       
  2031     rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
       
  2032     \<Longrightarrow> abc_fetch (length a) aprog =
       
  2033                       Some (Dec (Suc n) (length a + 5))"
       
  2034 apply(simp add: rec_ci.simps abc_fetch.simps, 
       
  2035                 erule_tac conjE, erule_tac conjE, simp)
       
  2036 apply(drule_tac eq_switch, drule_tac eq_switch, simp)
       
  2037 done
       
  2038 
       
  2039 lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
       
  2040     \<Longrightarrow> abc_fetch (Suc (length a)) aprog = Some (Dec (Suc n) (length a + 3))"
       
  2041 apply(simp add: rec_ci.simps abc_fetch.simps, erule_tac conjE, erule_tac conjE, simp)
       
  2042 apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append)
       
  2043 done
       
  2044 
       
  2045 lemma [simp]:
       
  2046   "\<lbrakk>rec_ci f = (a, aa, ba);
       
  2047     rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
       
  2048     \<Longrightarrow> abc_fetch (Suc (Suc (length a))) aprog = 
       
  2049                                      Some (Goto (length a + 1))"
       
  2050 apply(simp add: rec_ci.simps abc_fetch.simps,
       
  2051       erule_tac conjE, erule_tac conjE, simp)
       
  2052 apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append)
       
  2053 done
       
  2054 
       
  2055 lemma [simp]: 
       
  2056   "\<lbrakk>rec_ci f = (a, aa, ba);
       
  2057     rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
       
  2058     \<Longrightarrow> abc_fetch (length a + 3) aprog = Some (Inc n)"
       
  2059 apply(simp add: rec_ci.simps abc_fetch.simps, 
       
  2060       erule_tac conjE, erule_tac conjE, simp)
       
  2061 apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append)
       
  2062 done
       
  2063 
       
  2064 lemma [simp]: "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
       
  2065     \<Longrightarrow> abc_fetch (length a + 4) aprog = Some (Goto 0)"
       
  2066 apply(simp add: rec_ci.simps abc_fetch.simps, erule_tac conjE, erule_tac conjE, simp)
       
  2067 apply(drule_tac eq_switch, drule_tac eq_switch, simp add: nth_append)
       
  2068 done
       
  2069 
       
  2070 lemma [simp]: 
       
  2071   "0 < rsx
       
  2072    \<Longrightarrow> \<exists>y. (lm @ x # rsx # suf_lm)[Suc (length lm) := rsx - Suc 0]   
       
  2073     = lm @ x # y # suf_lm \<and> y \<le> rsx"
       
  2074 apply(case_tac rsx, simp, simp)
       
  2075 apply(rule_tac x = nat in exI, simp add: list_update_append)
       
  2076 done
       
  2077 
       
  2078 lemma [simp]: 
       
  2079   "\<lbrakk>y \<le> rsx; 0 < y\<rbrakk>
       
  2080    \<Longrightarrow> \<exists>ya. (lm @ x # y # suf_lm)[Suc (length lm) := y - Suc 0] 
       
  2081           = lm @ x # ya # suf_lm \<and> ya \<le> rsx"
       
  2082 apply(case_tac y, simp, simp)
       
  2083 apply(rule_tac x = nat in exI, simp add: list_update_append)
       
  2084 done
       
  2085 
       
  2086 lemma mn_halt_lemma: 
       
  2087   "\<lbrakk>rec_ci f = (a, aa, ba);
       
  2088     rec_ci (Mn n f) = (aprog, rs_pos, a_md);
       
  2089      0 < rsx; length lm = n\<rbrakk>
       
  2090     \<Longrightarrow>
       
  2091   \<forall>na. \<not> (\<lambda>(as, lm') (ss, n). as = 0)
       
  2092   (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog na) 
       
  2093                                                        (length a, n)
       
  2094  \<and> mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm)
       
  2095                        aprog na) (length a) x rsx suf_lm lm 
       
  2096 \<longrightarrow> mn_ind_inv (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog 
       
  2097                          (Suc na)) (length a) x rsx suf_lm lm
       
  2098  \<and> ((abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog (Suc na), 
       
  2099                                                     length a, n), 
       
  2100     abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog na,
       
  2101                               length a, n) \<in> mn_LE"
       
  2102 apply(rule allI, rule impI, simp add: abc_steps_ind)
       
  2103 apply(case_tac "(abc_steps_l (length a, lm @ x # rsx # suf_lm) 
       
  2104                                                    aprog na)", simp)
       
  2105 apply(auto split:if_splits simp add:abc_steps_l.simps 
       
  2106                            mn_ind_inv.simps abc_steps_zero)
       
  2107 apply(auto simp add: mn_LE_def lex_triple_def lex_pair_def 
       
  2108             abc_step_l.simps abc_steps_l.simps mn_ind_inv.simps
       
  2109             abc_lm_v.simps abc_lm_s.simps nth_append
       
  2110            split: if_splits)
       
  2111 apply(drule_tac  rec_ci_not_null, simp)
       
  2112 done
       
  2113 
       
  2114 lemma mn_halt:
       
  2115   "\<lbrakk>rec_ci f = (a, aa, ba);
       
  2116     rec_ci (Mn n f) = (aprog, rs_pos, a_md);
       
  2117     0 < rsx; length lm = n\<rbrakk>
       
  2118     \<Longrightarrow> \<exists> stp. (\<lambda> (as, lm'). (as = 0 \<and> 
       
  2119            mn_ind_inv (as, lm')  (length a) x rsx suf_lm lm))
       
  2120             (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog stp)"
       
  2121 apply(insert wf_mn_le)	  
       
  2122 apply(insert halt_lemma2[of mn_LE
       
  2123   "\<lambda> ((as, lm'), ss, n). mn_ind_inv (as, lm') ss x rsx suf_lm lm"
       
  2124   "\<lambda> stp. (abc_steps_l (length a, lm @ x # rsx # suf_lm) aprog stp, 
       
  2125   length a, n)"
       
  2126   "\<lambda> ((as, lm'), ss, n). as = 0"], 
       
  2127    simp)
       
  2128 apply(simp add: mn_halt_init mn_inv_init)
       
  2129 apply(drule_tac x = x and suf_lm = suf_lm in mn_halt_lemma, auto)
       
  2130 apply(rule_tac x = n in exI, 
       
  2131       case_tac "(abc_steps_l (length a, lm @ x # rsx # suf_lm)
       
  2132                               aprog n)", simp)
       
  2133 done
       
  2134 
       
  2135 lemma [simp]: "Suc rs_pos < a_md \<Longrightarrow> 
       
  2136                 Suc (a_md - Suc (Suc rs_pos)) = a_md - Suc rs_pos"
       
  2137 by arith
       
  2138 
       
  2139 lemma mn_ind_step: 
       
  2140   assumes ind:  
       
  2141   "\<And>aprog a_md rs_pos rs suf_lm lm.
       
  2142   \<lbrakk>rec_ci f = (aprog, rs_pos, a_md);
       
  2143    rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
       
  2144   \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp
       
  2145             = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
       
  2146   and h: "rec_ci f = (a, aa, ba)" 
       
  2147          "rec_ci (Mn n f) = (aprog, rs_pos, a_md)"  
       
  2148          "rec_calc_rel f (lm @ [x]) rsx" 
       
  2149          "rsx > 0" 
       
  2150          "Suc rs_pos < a_md" 
       
  2151          "aa = Suc rs_pos"
       
  2152   shows "\<exists>stp. abc_steps_l (0, lm @ x # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
       
  2153              aprog stp = (0, lm @ Suc x # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  2154 proof -
       
  2155   have k1: 
       
  2156     "\<exists> stp. abc_steps_l (0, lm @ x #  0\<up>(a_md - Suc (rs_pos)) @ suf_lm)
       
  2157          aprog stp = 
       
  2158        (length a, lm @ x # rsx # 0\<up>(a_md  - Suc (Suc rs_pos)) @ suf_lm)"
       
  2159     apply(insert h)
       
  2160     apply(auto intro: mn_calc_f ind)
       
  2161     done
       
  2162   from h have k2: "length lm = n"
       
  2163     apply(subgoal_tac "rs_pos = n")
       
  2164     apply(drule_tac  para_pattern, simp, simp, simp)
       
  2165     done
       
  2166   from h have k3: "a_md > (Suc rs_pos)"
       
  2167     apply(simp)
       
  2168     done
       
  2169   from k2 and h and k3 have k4: 
       
  2170     "\<exists> stp. abc_steps_l (length a,
       
  2171        lm @ x # rsx # 0\<up>(a_md  - Suc (Suc rs_pos)) @ suf_lm) aprog stp = 
       
  2172         (0, lm @ Suc x # 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
       
  2173     apply(frule_tac x = x and 
       
  2174        suf_lm = "0\<up>(a_md - Suc (Suc rs_pos)) @ suf_lm" in mn_halt, auto)
       
  2175     apply(rule_tac x = "stp" in exI, 
       
  2176           simp add: mn_ind_inv.simps rec_ci_not_null)
       
  2177     apply(simp only: replicate.simps[THEN sym], simp)
       
  2178     done
       
  2179   from k1 and k4 show "?thesis"
       
  2180     apply(auto)
       
  2181     apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
       
  2182     done
       
  2183 qed
       
  2184 
       
  2185 lemma [simp]: 
       
  2186   "\<lbrakk>rec_ci f = (a, aa, ba); rec_ci (Mn n f) = (aprog, rs_pos, a_md);
       
  2187     rec_calc_rel (Mn n f) lm rs\<rbrakk> \<Longrightarrow> aa = Suc rs_pos"
       
  2188 apply(rule_tac calc_mn_reverse, simp)
       
  2189 apply(insert para_pattern [of f a aa ba "lm @ [rs]" 0], simp)
       
  2190 apply(subgoal_tac "rs_pos = length lm", simp)
       
  2191 apply(drule_tac ci_mn_para_eq, simp)
       
  2192 done
       
  2193 
       
  2194 lemma [simp]: "\<lbrakk>rec_ci (Mn n f) = (aprog, rs_pos, a_md);      
       
  2195                 rec_calc_rel (Mn n f) lm rs\<rbrakk>  \<Longrightarrow> Suc rs_pos < a_md"
       
  2196 apply(case_tac "rec_ci f")
       
  2197 apply(subgoal_tac "c > b \<and> b = Suc rs_pos \<and> a_md \<ge> c")
       
  2198 apply(arith, auto)
       
  2199 done
       
  2200 
       
  2201 lemma mn_ind_steps:  
       
  2202   assumes ind:
       
  2203   "\<And>aprog a_md rs_pos rs suf_lm lm. 
       
  2204   \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
       
  2205   \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
       
  2206               (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
       
  2207   and h: "rec_ci f = (a, aa, ba)" 
       
  2208   "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
       
  2209   "rec_calc_rel (Mn n f) lm rs"
       
  2210   "rec_calc_rel f (lm @ [rs]) 0" 
       
  2211   "\<forall>x<rs. (\<exists> v. rec_calc_rel f (lm @ [x]) v \<and> 0 < v)"
       
  2212   "n = length lm" 
       
  2213   "x \<le> rs"
       
  2214   shows "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
       
  2215                  aprog stp = (0, lm @ x # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  2216 apply(insert h, induct x, 
       
  2217       rule_tac x = 0 in exI, simp add: abc_steps_zero, simp)
       
  2218 proof -
       
  2219   fix x
       
  2220   assume k1: 
       
  2221     "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
       
  2222                 aprog stp = (0, lm @ x # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  2223   and k2: "rec_ci (Mn (length lm) f) = (aprog, rs_pos, a_md)" 
       
  2224           "rec_calc_rel (Mn (length lm) f) lm rs" 
       
  2225           "rec_calc_rel f (lm @ [rs]) 0" 
       
  2226           "\<forall>x<rs.(\<exists> v. rec_calc_rel f (lm @ [x]) v \<and> v > 0)" 
       
  2227           "n = length lm" 
       
  2228           "Suc x \<le> rs" 
       
  2229           "rec_ci f = (a, aa, ba)"
       
  2230   hence k2:
       
  2231     "\<exists>stp. abc_steps_l (0, lm @ x # 0\<up>(a_md - rs_pos - 1) @ suf_lm) aprog
       
  2232                stp = (0, lm @ Suc x # 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
       
  2233     apply(erule_tac x = x in allE)
       
  2234     apply(auto)
       
  2235     apply(rule_tac  x = x in mn_ind_step)
       
  2236     apply(rule_tac ind, auto)      
       
  2237     done
       
  2238   from k1 and k2 show 
       
  2239     "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
       
  2240           aprog stp = (0, lm @ Suc x # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  2241     apply(auto)
       
  2242     apply(rule_tac x = "stp + stpa" in exI, simp only: abc_steps_add)
       
  2243     done
       
  2244 qed
       
  2245     
       
  2246 lemma [simp]: 
       
  2247 "\<lbrakk>rec_ci f = (a, aa, ba); 
       
  2248   rec_ci (Mn n f) = (aprog, rs_pos, a_md); 
       
  2249   rec_calc_rel (Mn n f) lm rs;
       
  2250   length lm = n\<rbrakk>
       
  2251  \<Longrightarrow> abc_lm_v (lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) (Suc n) = 0"
       
  2252 apply(auto simp: abc_lm_v.simps nth_append)
       
  2253 done
       
  2254 
       
  2255 lemma [simp]: 
       
  2256   "\<lbrakk>rec_ci f = (a, aa, ba); 
       
  2257     rec_ci (Mn n f) = (aprog, rs_pos, a_md); 
       
  2258     rec_calc_rel (Mn n f) lm rs;
       
  2259      length lm = n\<rbrakk>
       
  2260     \<Longrightarrow> abc_lm_s (lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) (Suc n) 0 =
       
  2261                            lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm"
       
  2262 apply(auto simp: abc_lm_s.simps list_update_append)
       
  2263 done
       
  2264 
       
  2265 lemma mn_length: 
       
  2266   "\<lbrakk>rec_ci f = (a, aa, ba);
       
  2267     rec_ci (Mn n f) = (aprog, rs_pos, a_md)\<rbrakk>
       
  2268   \<Longrightarrow> length aprog = length a + 5"
       
  2269 apply(simp add: rec_ci.simps, erule_tac conjE)
       
  2270 apply(drule_tac eq_switch, drule_tac eq_switch, simp)
       
  2271 done
       
  2272 
       
  2273 lemma mn_final_step:
       
  2274   assumes ind:
       
  2275   "\<And>aprog a_md rs_pos rs suf_lm lm.
       
  2276   \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); 
       
  2277   rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
       
  2278   \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
       
  2279               (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
       
  2280   and h: "rec_ci f = (a, aa, ba)" 
       
  2281          "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
       
  2282          "rec_calc_rel (Mn n f) lm rs" 
       
  2283          "rec_calc_rel f (lm @ [rs]) 0" 
       
  2284   shows "\<exists>stp. abc_steps_l (0, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
       
  2285      aprog stp = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  2286 proof -
       
  2287   from h and ind have k1:
       
  2288     "\<exists>stp.  abc_steps_l (0, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
       
  2289         aprog stp = (length a,  lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  2290     apply(insert mn_calc_f[of f a aa ba n aprog 
       
  2291                                rs_pos a_md lm rs 0 suf_lm], simp)
       
  2292     apply(subgoal_tac "aa = Suc n", simp add: exponent_cons_iff)
       
  2293     apply(subgoal_tac "rs_pos = n", simp, simp)
       
  2294     done
       
  2295   from h have k2: "length lm = n"
       
  2296     apply(subgoal_tac "rs_pos = n")
       
  2297     apply(drule_tac f = "Mn n f" in para_pattern, simp, simp, simp)
       
  2298     done
       
  2299   from h and k2 have k3: 
       
  2300   "\<exists>stp. abc_steps_l (length a, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
       
  2301     aprog stp = (length a + 5, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  2302     apply(rule_tac x = "Suc 0" in exI, 
       
  2303           simp add: abc_step_l.simps abc_steps_l.simps)
       
  2304     done
       
  2305   from h have k4: "length aprog = length a + 5"
       
  2306     apply(simp add: mn_length)
       
  2307     done
       
  2308   from k1 and k3 and k4 show "?thesis"
       
  2309     apply(auto)
       
  2310     apply(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
       
  2311     done
       
  2312 qed
       
  2313 
       
  2314 lemma mn_case: 
       
  2315   assumes ind: 
       
  2316   "\<And>aprog a_md rs_pos rs suf_lm lm.
       
  2317   \<lbrakk>rec_ci f = (aprog, rs_pos, a_md); rec_calc_rel f lm rs\<rbrakk> \<Longrightarrow> 
       
  2318   \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
       
  2319                (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
       
  2320   and h: "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
       
  2321          "rec_calc_rel (Mn n f) lm rs"
       
  2322   shows "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp 
       
  2323   = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
       
  2324 apply(case_tac "rec_ci f", simp)
       
  2325 apply(insert h, rule_tac calc_mn_reverse, simp)
       
  2326 proof -
       
  2327   fix a b c v
       
  2328   assume h: "rec_ci f = (a, b, c)" 
       
  2329             "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" 
       
  2330             "rec_calc_rel (Mn n f) lm rs" 
       
  2331             "rec_calc_rel f (lm @ [rs]) 0" 
       
  2332             "\<forall>x<rs. \<exists>v. rec_calc_rel f (lm @ [x]) v \<and> 0 < v"
       
  2333             "n = length lm"
       
  2334   hence k1:
       
  2335     "\<exists>stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm) aprog
       
  2336                   stp = (0, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  2337     apply(auto intro: mn_ind_steps ind)
       
  2338     done
       
  2339   from h have k2: 
       
  2340     "\<exists>stp. abc_steps_l (0, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm) aprog
       
  2341          stp = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  2342     apply(auto intro: mn_final_step ind)
       
  2343     done
       
  2344   from k1 and k2 show 
       
  2345     "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
       
  2346   (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  2347     apply(auto, insert h)
       
  2348     apply(subgoal_tac "Suc rs_pos < a_md")
       
  2349     apply(rule_tac x = "stp + stpa" in exI, 
       
  2350       simp only: abc_steps_add exponent_cons_iff, simp, simp)
       
  2351     done
       
  2352 qed
       
  2353 
       
  2354 lemma z_rs: "rec_calc_rel z lm rs \<Longrightarrow> rs = 0"
       
  2355 apply(rule_tac calc_z_reverse, auto)
       
  2356 done
       
  2357 
       
  2358 lemma z_case:
       
  2359   "\<lbrakk>rec_ci z = (aprog, rs_pos, a_md); rec_calc_rel z lm rs\<rbrakk>
       
  2360   \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
       
  2361            (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
       
  2362 apply(simp add: rec_ci.simps rec_ci_z_def, auto)
       
  2363 apply(rule_tac x = "Suc 0" in exI, simp add: abc_steps_l.simps 
       
  2364                                abc_fetch.simps abc_step_l.simps z_rs)
       
  2365 done
       
  2366 
   105 
  2367 fun addition_inv :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow>     
   106 fun addition_inv :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow>     
  2368                      nat list \<Rightarrow> bool"
   107                      nat list \<Rightarrow> bool"
  2369   where
   108   where
  2370   "addition_inv (as, lm') m n p lm = 
   109   "addition_inv (as, lm') m n p lm = 
  2425 definition addition_LE :: "(((nat \<times> nat list) \<times> nat \<times> nat) \<times> 
   164 definition addition_LE :: "(((nat \<times> nat list) \<times> nat \<times> nat) \<times> 
  2426                           ((nat \<times> nat list) \<times> nat \<times> nat)) set"
   165                           ((nat \<times> nat list) \<times> nat \<times> nat)) set"
  2427   where "addition_LE \<equiv> (inv_image lex_triple addition_measure)"
   166   where "addition_LE \<equiv> (inv_image lex_triple addition_measure)"
  2428 
   167 
  2429 lemma [simp]: "wf addition_LE"
   168 lemma [simp]: "wf addition_LE"
  2430 by(simp add: wf_inv_image wf_lex_triple addition_LE_def)
   169 by(auto simp: wf_inv_image addition_LE_def lex_triple_def
       
   170              lex_pair_def)
  2431 
   171 
  2432 declare addition_inv.simps[simp del]
   172 declare addition_inv.simps[simp del]
  2433 
   173 
  2434 lemma addition_inv_init: 
   174 lemma addition_inv_init: 
  2435   "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow>
   175   "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow>
  2436                                    addition_inv (0, lm) m n p lm"
   176                                    addition_inv (0, lm) m n p lm"
  2437 apply(simp add: addition_inv.simps)
   177 apply(simp add: addition_inv.simps Let_def)
  2438 apply(rule_tac x = "lm ! m" in exI, simp)
   178 apply(rule_tac x = "lm ! m" in exI, simp)
  2439 done
   179 done
  2440 
   180 
  2441 lemma [simp]: "abc_fetch 0 (addition m n p) = Some (Dec m 4)"
   181 lemma [simp]: "abc_fetch 0 (addition m n p) = Some (Dec m 4)"
  2442 by(simp add: abc_fetch.simps addition.simps)
   182 by(simp add: abc_fetch.simps addition.simps)
  2524                              p := lm ! m - Suc x] 
   264                              p := lm ! m - Suc x] 
  2525                = lm[m := xa, n := lm ! n + lm ! m, p := lm ! m - xa]"
   265                = lm[m := xa, n := lm ! n + lm ! m, p := lm ! m - xa]"
  2526 apply(rule_tac x = "Suc x" in exI, simp)
   266 apply(rule_tac x = "Suc x" in exI, simp)
  2527 done
   267 done
  2528 
   268 
       
   269 lemma abc_steps_zero: "abc_steps_l asm ap 0 = asm"
       
   270 apply(case_tac asm, simp add: abc_steps_l.simps)
       
   271 done
       
   272 
       
   273 declare Let_def[simp]
  2529 lemma addition_halt_lemma: 
   274 lemma addition_halt_lemma: 
  2530   "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow>
   275   "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow>
  2531   \<forall>na. \<not> (\<lambda>(as, lm') (m, p). as = 7) 
   276   \<forall>na. \<not> (\<lambda>(as, lm') (m, p). as = 7) 
  2532         (abc_steps_l (0, lm) (addition m n p) na) (m, p) \<and> 
   277         (abc_steps_l (0, lm) (addition m n p) na) (m, p) \<and> 
  2533   addition_inv (abc_steps_l (0, lm) (addition m n p) na) m n p lm 
   278   addition_inv (abc_steps_l (0, lm) (addition m n p) na) m n p lm 
  2534 \<longrightarrow> addition_inv (abc_steps_l (0, lm) (addition m n p) 
   279 \<longrightarrow> addition_inv (abc_steps_l (0, lm) (addition m n p) 
  2535                                  (Suc na)) m n p lm 
   280                                  (Suc na)) m n p lm 
  2536   \<and> ((abc_steps_l (0, lm) (addition m n p) (Suc na), m, p), 
   281   \<and> ((abc_steps_l (0, lm) (addition m n p) (Suc na), m, p), 
  2537      abc_steps_l (0, lm) (addition m n p) na, m, p) \<in> addition_LE"
   282      abc_steps_l (0, lm) (addition m n p) na, m, p) \<in> addition_LE"
  2538 apply(rule allI, rule impI, simp add: abc_steps_ind)
   283 apply(rule allI, rule impI, simp add: abc_step_red2)
  2539 apply(case_tac "(abc_steps_l (0, lm) (addition m n p) na)", simp)
   284 apply(case_tac "(abc_steps_l (0, lm) (addition m n p) na)", simp)
  2540 apply(auto split:if_splits simp add: addition_inv.simps
   285 apply(auto split:if_splits simp add: addition_inv.simps
  2541                                  abc_steps_zero)
   286                                  abc_steps_zero)
  2542 apply(simp_all add: abc_steps_l.simps abc_steps_zero)
   287 apply(simp_all add: addition.simps abc_steps_l.simps)
  2543 apply(auto simp add: addition_LE_def lex_triple_def lex_pair_def 
   288 apply(auto simp add: addition_LE_def lex_triple_def lex_pair_def 
  2544                      abc_step_l.simps addition_inv.simps 
   289                      abc_step_l.simps addition_inv.simps 
  2545                      abc_lm_v.simps abc_lm_s.simps nth_append
   290                      abc_lm_v.simps abc_lm_s.simps nth_append
  2546                 split: if_splits)
   291                 split: if_splits)
  2547 apply(rule_tac x = x in exI, simp)
   292 apply(rule_tac [!] x = x in exI, simp_all add: list_update_overwrite Suc_diff_Suc)
  2548 done
   293 by (metis list_update_overwrite list_update_swap nat_neq_iff)
  2549 
   294 
  2550 lemma  addition_ex: 
   295 lemma  addition_correct': 
  2551   "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow> 
   296   "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow> 
  2552   \<exists> stp. (\<lambda> (as, lm'). as = 7 \<and> addition_inv (as, lm') m n p lm) 
   297   \<exists> stp. (\<lambda> (as, lm'). as = 7 \<and> addition_inv (as, lm') m n p lm) 
  2553                         (abc_steps_l (0, lm) (addition m n p) stp)"
   298                         (abc_steps_l (0, lm) (addition m n p) stp)"
  2554 apply(insert halt_lemma2[of addition_LE
   299 apply(insert halt_lemma2[of addition_LE
  2555   "\<lambda> ((as, lm'), m, p). addition_inv (as, lm') m n p lm"
   300   "\<lambda> ((as, lm'), m, p). addition_inv (as, lm') m n p lm"
  2560       simp, erule_tac exE)
   305       simp, erule_tac exE)
  2561 apply(rule_tac x = na in exI, 
   306 apply(rule_tac x = na in exI, 
  2562       case_tac "(abc_steps_l (0, lm) (addition m n p) na)", auto)
   307       case_tac "(abc_steps_l (0, lm) (addition m n p) na)", auto)
  2563 done
   308 done
  2564 
   309 
  2565 lemma [simp]: "length (addition m n p) = 7"
   310 lemma length_addition[simp]: "length (addition a b c) = 7"
  2566 by (simp add: addition.simps)
   311 by(auto simp: addition.simps)
  2567 
   312 
  2568 lemma [elim]: "addition 0 (Suc 0) 2 = [] \<Longrightarrow> RR"
   313 lemma addition_correct:
  2569 by(simp add: addition.simps)
   314   "\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk>
  2570 
   315    \<Longrightarrow> {\<lambda> a. a = lm} (addition m n p) {\<lambda> nl. addition_inv (7, nl) m n p lm}"
  2571 lemma [simp]: "(0\<up>2)[0 := n] = [n, 0::nat]"
   316 using assms
  2572 apply(subgoal_tac "2 = Suc 1", 
   317 proof(rule_tac abc_Hoare_haltI, simp)
  2573       simp only: replicate.simps)
   318   fix lma
  2574 apply(auto)
   319   assume "m \<noteq> n" "m < p \<and> n < p" "p < length lm" "lm ! p = 0"
  2575 done
   320   then have "\<exists> stp. (\<lambda> (as, lm'). as = 7 \<and> addition_inv (as, lm') m n p lm) 
  2576 
   321                         (abc_steps_l (0, lm) (addition m n p) stp)"
  2577 lemma [simp]: 
   322     by(rule_tac addition_correct', auto simp: addition_inv.simps)
  2578   "\<exists>stp. abc_steps_l (0, n # 0\<up>2 @ suf_lm) 
   323   thus "\<exists>na. abc_final (abc_steps_l (0, lm) (addition m n p) na) (addition m n p) \<and>
  2579      (addition 0 (Suc 0) 2 [+] [Inc (Suc 0)]) stp = 
   324                   (\<lambda>nl. addition_inv (7, nl) m n p lm) abc_holds_for abc_steps_l (0, lm) (addition m n p) na"
  2580                                       (8, n # Suc n # 0 # suf_lm)"
   325     apply(erule_tac exE)
  2581 apply(rule_tac bm = "n # n # 0 # suf_lm" in abc_append_exc2, auto)
   326     apply(rule_tac x = stp in exI)
  2582 apply(insert addition_ex[of 0 "Suc 0" 2 "n # 0\<up>2 @ suf_lm"], 
   327     apply(auto)
  2583       simp add: nth_append numeral_2_eq_2, erule_tac exE)
   328     done
  2584 apply(rule_tac x = stp in exI,
   329 qed
  2585       case_tac "(abc_steps_l (0, n # 0\<up>2 @ suf_lm)
   330 
  2586                       (addition 0 (Suc 0) 2) stp)", 
   331 lemma compile_s_correct':
  2587       simp add: addition_inv.simps nth_append list_update_append numeral_2_eq_2)
   332   "{\<lambda>nl. nl = n # 0 \<up> 2 @ anything} addition 0 (Suc 0) 2 [+] [Inc (Suc 0)] {\<lambda>nl. nl = n # Suc n # 0 # anything}"
  2588 apply(simp add: nth_append numeral_2_eq_2, erule_tac exE)
   333 proof(rule_tac abc_Hoare_plus_halt)
  2589 apply(rule_tac x = "Suc 0" in exI,
   334   show "{\<lambda>nl. nl = n # 0 \<up> 2 @ anything} addition 0 (Suc 0) 2 {\<lambda> nl. addition_inv (7, nl) 0 (Suc 0) 2 (n # 0 \<up> 2 @ anything)}"
  2590       simp add: abc_steps_l.simps abc_fetch.simps 
   335     by(rule_tac addition_correct, auto simp: numeral_2_eq_2)
  2591       abc_steps_zero abc_step_l.simps abc_lm_s.simps abc_lm_v.simps)
   336 next
  2592 done
   337   show "{\<lambda>nl. addition_inv (7, nl) 0 (Suc 0) 2 (n # 0 \<up> 2 @ anything)} [Inc (Suc 0)] {\<lambda>nl. nl = n # Suc n # 0 # anything}"
  2593 
   338     by(rule_tac abc_Hoare_haltI, rule_tac x = 1 in exI, auto simp: addition_inv.simps 
  2594 lemma s_case:
   339       abc_steps_l.simps abc_step_l.simps abc_fetch.simps numeral_2_eq_2 abc_lm_s.simps abc_lm_v.simps)
  2595   "\<lbrakk>rec_ci s = (aprog, rs_pos, a_md); rec_calc_rel s lm rs\<rbrakk>
   340 qed
  2596   \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
   341   
  2597                (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
   342 declare rec_exec.simps[simp del]
  2598 apply(simp add: rec_ci.simps rec_ci_s_def, auto)
   343 
  2599 apply(rule_tac calc_s_reverse, auto)
   344 lemma abc_comp_commute: "(A [+] B) [+] C = A [+] (B [+] C)"
  2600 done
   345 apply(auto simp: abc_comp.simps abc_shift.simps)
  2601 
   346 apply(case_tac x, auto)
  2602 lemma [simp]: 
   347 done
  2603   "\<lbrakk>n < length lm; lm ! n = rs\<rbrakk>
   348 
  2604     \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0 # 0 #suf_lm)
   349 
  2605                      (addition n (length lm) (Suc (length lm))) stp 
   350 
  2606              = (7, lm @ rs # 0 # suf_lm)"
   351 lemma compile_z_correct: 
  2607 apply(insert addition_ex[of n "length lm"
   352   "\<lbrakk>rec_ci z = (ap, arity, fp); rec_exec z [n] = r\<rbrakk> \<Longrightarrow> 
  2608                            "Suc (length lm)" "lm @ 0 # 0 # suf_lm"])
   353   {\<lambda>nl. nl = n # 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = n # r # 0 \<up> (fp - Suc arity) @ anything}"
  2609 apply(simp add: nth_append, erule_tac exE)
   354 apply(rule_tac abc_Hoare_haltI)
  2610 apply(rule_tac x = stp in exI)
   355 apply(rule_tac x = 1 in exI)
  2611 apply(case_tac "abc_steps_l (0, lm @ 0 # 0 # suf_lm) (addition n (length lm)
   356 apply(auto simp: abc_steps_l.simps rec_ci.simps rec_ci_z_def 
  2612                  (Suc (length lm))) stp", simp)
   357                  numeral_2_eq_2 abc_fetch.simps abc_step_l.simps rec_exec.simps)
  2613 apply(simp add: addition_inv.simps)
   358 done
  2614 apply(insert nth_append[of lm "0 # 0 # suf_lm" "n"], simp)
   359 
  2615 done
   360 lemma compile_s_correct: 
  2616 
   361   "\<lbrakk>rec_ci s = (ap, arity, fp); rec_exec s [n] = r\<rbrakk> \<Longrightarrow> 
  2617 lemma [simp]: "0\<up>2 = [0, 0::nat]"
   362   {\<lambda>nl. nl = n # 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = n # r # 0 \<up> (fp - Suc arity) @ anything}"
  2618 apply(auto simp:numeral_2_eq_2)
   363 apply(auto simp: rec_ci.simps rec_ci_s_def compile_s_correct' rec_exec.simps)
  2619 done
   364 done
  2620 
   365 
  2621 lemma id_case: 
   366 lemma compile_id_correct':
  2622   "\<lbrakk>rec_ci (id m n) = (aprog, rs_pos, a_md); 
   367   assumes "n < length args" 
  2623     rec_calc_rel (id m n) lm rs\<rbrakk>
   368   shows "{\<lambda>nl. nl = args @ 0 \<up> 2 @ anything} addition n (length args) (Suc (length args))
  2624   \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
   369   {\<lambda>nl. nl = args @ rec_exec (recf.id (length args) n) args # 0 # anything}"
  2625                (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
       
  2626 apply(simp add: rec_ci.simps rec_ci_id.simps, auto)
       
  2627 apply(rule_tac calc_id_reverse, simp, simp)
       
  2628 done   
       
  2629 
       
  2630 lemma list_tl_induct:
       
  2631   "\<lbrakk>P []; \<And>a list. P list \<Longrightarrow> P (list @ [a::'a])\<rbrakk> \<Longrightarrow> 
       
  2632                                             P ((list::'a list))"
       
  2633 apply(case_tac "length list", simp)
       
  2634 proof -
   370 proof -
  2635   fix nat
   371   have "{\<lambda>nl. nl = args @ 0 \<up> 2 @ anything} addition n (length args) (Suc (length args))
  2636   assume ind: "\<And>a list. P list \<Longrightarrow> P (list @ [a])"
   372   {\<lambda>nl. addition_inv (7, nl) n (length args) (Suc (length args)) (args @ 0 \<up> 2 @ anything)}"
  2637   and h: "length list = Suc nat" "P []"
   373     using assms
  2638   from h show "P list"
   374     by(rule_tac addition_correct, auto simp: numeral_2_eq_2 nth_append)
  2639   proof(induct nat arbitrary: list, case_tac lista, simp, simp)
   375   thus "?thesis"
  2640     fix lista a listaa
   376     using assms
  2641     from h show "P [a]"
   377     by(simp add: addition_inv.simps rec_exec.simps 
  2642       by(insert ind[of "[]"], simp add: h)
   378       nth_append numeral_2_eq_2 list_update_append)
  2643   next
   379 qed
  2644     fix nat list
   380 
  2645     assume nind: "\<And>list. \<lbrakk>length list = Suc nat; P []\<rbrakk> \<Longrightarrow> P list" 
   381 lemma compile_id_correct: 
  2646     and g: "length (list:: 'a list) = Suc (Suc nat)"
   382   "\<lbrakk>n < m; length xs = m; rec_ci (recf.id m n) = (ap, arity, fp); rec_exec (recf.id m n) xs = r\<rbrakk>
  2647     from g show "P (list::'a list)"
   383        \<Longrightarrow> {\<lambda>nl. nl = xs @ 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ r # 0 \<up> (fp - Suc arity) @ anything}"
  2648       apply(insert nind[of "butlast list"], simp add: h)
   384 apply(auto simp: rec_ci.simps rec_ci_id.simps compile_id_correct')
  2649       apply(insert ind[of "butlast list" "last list"], simp)
       
  2650       apply(subgoal_tac "butlast list @ [last list] = list", simp)
       
  2651       apply(case_tac "list::'a list", simp, simp)
       
  2652       done
       
  2653   qed
       
  2654 qed      
       
  2655   
       
  2656 lemma nth_eq_butlast_nth: "\<lbrakk>length ys > Suc k\<rbrakk> \<Longrightarrow> 
       
  2657                                         ys ! k = butlast ys ! k"
       
  2658 apply(subgoal_tac "\<exists> xs y. ys = xs @ [y]", auto simp: nth_append)
       
  2659 apply(rule_tac x = "butlast ys" in exI, rule_tac x = "last ys" in exI)
       
  2660 apply(case_tac "ys = []", simp, simp)
       
  2661 done
       
  2662 
       
  2663 lemma [simp]: 
       
  2664 "\<lbrakk>\<forall>k<Suc (length list). rec_calc_rel ((list @ [a]) ! k) lm (ys ! k);
       
  2665   length ys = Suc (length list)\<rbrakk>
       
  2666    \<Longrightarrow> \<forall>k<length list. rec_calc_rel (list ! k) lm (butlast ys ! k)"
       
  2667 apply(rule allI, rule impI)
       
  2668 apply(erule_tac  x = k in allE, simp add: nth_append)
       
  2669 apply(subgoal_tac "ys ! k = butlast ys ! k", simp)
       
  2670 apply(rule_tac nth_eq_butlast_nth, arith)
       
  2671 done
   385 done
  2672 
   386 
  2673 lemma cn_merge_gs_tl_app: 
   387 lemma cn_merge_gs_tl_app: 
  2674   "cn_merge_gs (gs @ [g]) pstr = 
   388   "cn_merge_gs (gs @ [g]) pstr = 
  2675         cn_merge_gs gs pstr [+] cn_merge_gs [g] (pstr + length gs)"
   389         cn_merge_gs gs pstr [+] cn_merge_gs [g] (pstr + length gs)"
  2676 apply(induct gs arbitrary: pstr, simp add: cn_merge_gs.simps, simp)
   390 apply(induct gs arbitrary: pstr, simp add: cn_merge_gs.simps, auto)
  2677 apply(case_tac a, simp add: abc_append_commute)
   391 apply(simp add: abc_comp_commute)
  2678 done
   392 done
  2679 
   393 
  2680 lemma cn_merge_gs_length: 
   394 lemma footprint_ge: 
  2681   "length (cn_merge_gs (map rec_ci list) pstr) = 
   395   "rec_ci a = (p, arity, fp) \<Longrightarrow> arity < fp"
  2682       (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci list. length ap) + 3 * length list "
   396 apply(induct a)
  2683 apply(induct list arbitrary: pstr, simp, simp)
   397 apply(auto simp: rec_ci.simps)
  2684 apply(case_tac "rec_ci a", simp)
   398 apply(case_tac "rec_ci a", simp)
  2685 done
   399 apply(case_tac "rec_ci a1", case_tac "rec_ci a2", auto)
  2686 
   400 apply(case_tac "rec_ci a", auto)
  2687 lemma [simp]: "Suc n \<le> pstr \<Longrightarrow> pstr + x - n > 0"
   401 done
  2688 by arith
   402 
       
   403 lemma param_pattern: 
       
   404   "\<lbrakk>terminate f xs; rec_ci f = (p, arity, fp)\<rbrakk> \<Longrightarrow> length xs = arity"
       
   405 apply(induct arbitrary: p arity fp rule: terminate.induct)
       
   406 apply(auto simp: rec_ci.simps)
       
   407 apply(case_tac "rec_ci f", simp)
       
   408 apply(case_tac "rec_ci f", case_tac "rec_ci g", simp)
       
   409 apply(case_tac "rec_ci f", case_tac "rec_ci gs", simp)
       
   410 done
       
   411 
       
   412 lemma replicate_merge_anywhere: 
       
   413   "x\<up>a @ x\<up>b @ ys = x\<up>(a+b) @ ys"
       
   414 by(simp add:replicate_add)
       
   415 
       
   416 fun mv_box_inv :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> bool"
       
   417   where
       
   418   "mv_box_inv (as, lm) m n initlm = 
       
   419          (let plus = initlm ! m + initlm ! n in
       
   420            length initlm > max m n \<and> m \<noteq> n \<and> 
       
   421               (if as = 0 then \<exists> k l. lm = initlm[m := k, n := l] \<and> 
       
   422                     k + l = plus \<and> k \<le> initlm ! m 
       
   423               else if as = 1 then \<exists> k l. lm = initlm[m := k, n := l]
       
   424                              \<and> k + l + 1 = plus \<and> k < initlm ! m 
       
   425               else if as = 2 then \<exists> k l. lm = initlm[m := k, n := l] 
       
   426                               \<and> k + l = plus \<and> k \<le> initlm ! m
       
   427               else if as = 3 then lm = initlm[m := 0, n := plus]
       
   428               else False))"
       
   429 
       
   430 fun mv_box_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
       
   431   where
       
   432   "mv_box_stage1 (as, lm) m  = 
       
   433             (if as = 3 then 0 
       
   434              else 1)"
       
   435 
       
   436 fun mv_box_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
       
   437   where
       
   438   "mv_box_stage2 (as, lm) m = (lm ! m)"
       
   439 
       
   440 fun mv_box_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
       
   441   where
       
   442   "mv_box_stage3 (as, lm) m = (if as = 1 then 3 
       
   443                                 else if as = 2 then 2
       
   444                                 else if as = 0 then 1 
       
   445                                 else 0)"
       
   446  
       
   447 fun mv_box_measure :: "((nat \<times> nat list) \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)"
       
   448   where
       
   449   "mv_box_measure ((as, lm), m) = 
       
   450      (mv_box_stage1 (as, lm) m, mv_box_stage2 (as, lm) m,
       
   451       mv_box_stage3 (as, lm) m)"
       
   452 
       
   453 definition lex_pair :: "((nat \<times> nat) \<times> nat \<times> nat) set"
       
   454   where
       
   455   "lex_pair = less_than <*lex*> less_than"
       
   456 
       
   457 definition lex_triple :: 
       
   458  "((nat \<times> (nat \<times> nat)) \<times> (nat \<times> (nat \<times> nat))) set"
       
   459   where
       
   460   "lex_triple \<equiv> less_than <*lex*> lex_pair"
       
   461 
       
   462 definition mv_box_LE :: 
       
   463  "(((nat \<times> nat list) \<times> nat) \<times> ((nat \<times> nat list) \<times> nat)) set"
       
   464   where 
       
   465   "mv_box_LE \<equiv> (inv_image lex_triple mv_box_measure)"
       
   466 
       
   467 lemma wf_lex_triple: "wf lex_triple"
       
   468   by (auto intro:wf_lex_prod simp:lex_triple_def lex_pair_def)
       
   469 
       
   470 lemma wf_mv_box_le[intro]: "wf mv_box_LE"
       
   471 by(auto intro:wf_inv_image wf_lex_triple simp: mv_box_LE_def)
       
   472 
       
   473 declare mv_box_inv.simps[simp del]
       
   474 
       
   475 lemma mv_box_inv_init:  
       
   476 "\<lbrakk>m < length initlm; n < length initlm; m \<noteq> n\<rbrakk> \<Longrightarrow> 
       
   477   mv_box_inv (0, initlm) m n initlm"
       
   478 apply(simp add: abc_steps_l.simps mv_box_inv.simps)
       
   479 apply(rule_tac x = "initlm ! m" in exI, 
       
   480       rule_tac x = "initlm ! n" in exI, simp)
       
   481 done
       
   482 
       
   483 lemma [simp]: "abc_fetch 0 (mv_box m n) = Some (Dec m 3)"
       
   484 apply(simp add: mv_box.simps abc_fetch.simps)
       
   485 done
       
   486 
       
   487 lemma [simp]: "abc_fetch (Suc 0) (mv_box m n) =
       
   488                Some (Inc n)"
       
   489 apply(simp add: mv_box.simps abc_fetch.simps)
       
   490 done
       
   491 
       
   492 lemma [simp]: "abc_fetch 2 (mv_box m n) = Some (Goto 0)"
       
   493 apply(simp add: mv_box.simps abc_fetch.simps)
       
   494 done
       
   495 
       
   496 lemma [simp]: "abc_fetch 3 (mv_box m n) = None"
       
   497 apply(simp add: mv_box.simps abc_fetch.simps)
       
   498 done
       
   499 
       
   500 lemma replicate_Suc_iff_anywhere: "x # x\<up>b @ ys = x\<up>(Suc b) @ ys"
       
   501 by simp
       
   502 
       
   503 lemma [simp]: 
       
   504   "\<lbrakk>m \<noteq> n; m < length initlm; n < length initlm;
       
   505     k + l = initlm ! m + initlm ! n; k \<le> initlm ! m; 0 < k\<rbrakk>
       
   506  \<Longrightarrow> \<exists>ka la. initlm[m := k, n := l, m := k - Suc 0] = 
       
   507      initlm[m := ka, n := la] \<and>
       
   508      Suc (ka + la) = initlm ! m + initlm ! n \<and> 
       
   509      ka < initlm ! m"
       
   510 apply(rule_tac x = "k - Suc 0" in exI, rule_tac x = l in exI, 
       
   511       simp, auto)
       
   512 apply(subgoal_tac 
       
   513       "initlm[m := k, n := l, m := k - Suc 0] = 
       
   514        initlm[n := l, m := k, m := k - Suc 0]")
       
   515 apply(simp add: list_update_overwrite )
       
   516 apply(simp add: list_update_swap)
       
   517 apply(simp add: list_update_swap)
       
   518 done
  2689 
   519 
  2690 lemma [simp]:
   520 lemma [simp]:
  2691   "\<lbrakk>Suc (pstr + length list) \<le> a_md; 
   521   "\<lbrakk>m \<noteq> n; m < length initlm; n < length initlm; 
  2692     length ys = Suc (length list);
   522     Suc (k + l) = initlm ! m + initlm ! n;
  2693     length lm = n;
   523     k < initlm ! m\<rbrakk>
  2694      Suc n \<le> pstr\<rbrakk>
   524     \<Longrightarrow> \<exists>ka la. initlm[m := k, n := l, n := Suc l] = 
  2695    \<Longrightarrow>  (ys ! length list # 0\<up>(pstr - Suc n) @ butlast ys @
   525                 initlm[m := ka, n := la] \<and> 
  2696              0\<up>(a_md - (pstr + length list)) @ suf_lm) ! 
   526                 ka + la = initlm ! m + initlm ! n \<and> 
  2697                       (pstr + length list - n) = (0 :: nat)"
   527                 ka \<le> initlm ! m"
  2698 apply(insert nth_append[of "ys ! length list # 0\<up>(pstr - Suc n) @
   528 apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, auto)
  2699      butlast ys" "0\<up>(a_md - (pstr + length list)) @ suf_lm"
   529 done
  2700       "(pstr + length list - n)"], simp add: nth_append)
   530 
  2701 done
   531 lemma [simp]: 
       
   532   "\<lbrakk>length initlm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
       
   533    \<forall>na. \<not> (\<lambda>(as, lm) m. as = 3) 
       
   534     (abc_steps_l (0, initlm) (mv_box m n) na) m \<and> 
       
   535   mv_box_inv (abc_steps_l (0, initlm) 
       
   536            (mv_box m n) na) m n initlm \<longrightarrow>
       
   537   mv_box_inv (abc_steps_l (0, initlm) 
       
   538            (mv_box m n) (Suc na)) m n initlm \<and>
       
   539   ((abc_steps_l (0, initlm) (mv_box m n) (Suc na), m),
       
   540    abc_steps_l (0, initlm) (mv_box m n) na, m) \<in> mv_box_LE"
       
   541 apply(rule allI, rule impI, simp add: abc_step_red2)
       
   542 apply(case_tac "(abc_steps_l (0, initlm) (mv_box m n) na)",
       
   543       simp)
       
   544 apply(auto split:if_splits simp add:abc_steps_l.simps mv_box_inv.simps)
       
   545 apply(auto simp add: mv_box_LE_def lex_triple_def lex_pair_def 
       
   546                      abc_step_l.simps abc_steps_l.simps
       
   547                      mv_box_inv.simps abc_lm_v.simps abc_lm_s.simps
       
   548                 split: if_splits )
       
   549 apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, simp)
       
   550 done
       
   551 
       
   552 lemma mv_box_inv_halt: 
       
   553   "\<lbrakk>length initlm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
       
   554   \<exists> stp. (\<lambda> (as, lm). as = 3 \<and> 
       
   555   mv_box_inv (as, lm) m n initlm) 
       
   556              (abc_steps_l (0::nat, initlm) (mv_box m n) stp)"
       
   557 apply(insert halt_lemma2[of mv_box_LE
       
   558     "\<lambda> ((as, lm), m). mv_box_inv (as, lm) m n initlm"
       
   559     "\<lambda> stp. (abc_steps_l (0, initlm) (mv_box m n) stp, m)"
       
   560     "\<lambda> ((as, lm), m). as = (3::nat)"
       
   561     ])
       
   562 apply(insert wf_mv_box_le)
       
   563 apply(simp add: mv_box_inv_init abc_steps_zero)
       
   564 apply(erule_tac exE)
       
   565 apply(rule_tac x = na in exI)
       
   566 apply(case_tac "(abc_steps_l (0, initlm) (mv_box m n) na)",
       
   567       simp, auto)
       
   568 done
       
   569 
       
   570 lemma mv_box_halt_cond:
       
   571   "\<lbrakk>m \<noteq> n; mv_box_inv (a, b) m n lm; a = 3\<rbrakk> \<Longrightarrow> 
       
   572   b = lm[n := lm ! m + lm ! n, m := 0]"
       
   573 apply(simp add: mv_box_inv.simps, auto)
       
   574 apply(simp add: list_update_swap)
       
   575 done
       
   576 
       
   577 lemma mv_box_correct':
       
   578   "\<lbrakk>length lm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow> 
       
   579   \<exists> stp. abc_steps_l (0::nat, lm) (mv_box m n) stp
       
   580   = (3, (lm[n := (lm ! m + lm ! n)])[m := 0::nat])"
       
   581 apply(drule mv_box_inv_halt, simp, erule_tac exE)
       
   582 apply(rule_tac x = stp in exI)
       
   583 apply(case_tac "abc_steps_l (0, lm) (mv_box m n) stp",
       
   584       simp)
       
   585 apply(erule_tac mv_box_halt_cond, auto)
       
   586 done
       
   587 
       
   588 lemma length_mvbox[simp]: "length (mv_box m n) = 3"
       
   589 by(simp add: mv_box.simps)
       
   590 
       
   591 lemma mv_box_correct: 
       
   592   "\<lbrakk>length lm > max m n; m\<noteq>n\<rbrakk> 
       
   593   \<Longrightarrow> {\<lambda> nl. nl = lm} mv_box m n {\<lambda> nl. nl = lm[n := (lm ! m + lm ! n), m:=0]}"
       
   594 apply(drule_tac mv_box_correct', simp)
       
   595 apply(auto simp: abc_Hoare_halt_def)
       
   596 apply(rule_tac x = stp in exI, simp)
       
   597 done
       
   598 
       
   599 declare list_update.simps(2)[simp del]
  2702 
   600 
  2703 lemma [simp]:
   601 lemma [simp]:
  2704   "\<lbrakk>Suc (pstr + length list) \<le> a_md; 
   602   "\<lbrakk>length xs < gf; gf \<le> ft; n < length gs\<rbrakk>
  2705     length ys = Suc (length list);
   603   \<Longrightarrow> (rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything)
  2706     length lm = n;
   604   [ft + n - length xs := rec_exec (gs ! n) xs, 0 := 0] =
  2707      Suc n \<le> pstr\<rbrakk>
   605   0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything"
  2708     \<Longrightarrow> (lm @ last ys # 0\<up>(pstr - Suc n) @ butlast ys @
   606 using list_update_append[of "rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_exec i xs) (take n gs)"
  2709          0\<up>(a_md - (pstr + length list)) @ suf_lm)[pstr + length list := 
   607                              "0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything" "ft + n - length xs" "rec_exec (gs ! n) xs"]
  2710                                         last ys, n := 0] =
   608 apply(auto)
  2711         lm @ (0::nat)\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm"
   609 apply(case_tac "length gs - n", simp, simp add: list_update.simps replicate_Suc_iff_anywhere Suc_diff_Suc del: replicate_Suc)
  2712 apply(insert list_update_length[of 
   610 apply(simp add: list_update.simps)
  2713    "lm @ last ys # 0\<up>(pstr - Suc n) @ butlast ys" 0 
   611 done
  2714    "0\<up>(a_md - Suc (pstr + length list)) @ suf_lm" "last ys"], simp)
   612 
  2715 apply(simp add: exponent_cons_iff)
   613 lemma compile_cn_gs_correct':
  2716 apply(insert list_update_length[of "lm" 
   614   assumes
  2717         "last ys" "0\<up>(pstr - Suc n) @ butlast ys @ 
   615   g_cond: "\<forall>g\<in>set (take n gs). terminate g xs \<and>
  2718       last ys # 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm" 0], simp)
   616   (\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow> (\<forall>xc. {\<lambda>nl. nl = xs @ 0 \<up> (xb - xa) @ xc} x {\<lambda>nl. nl = xs @ rec_exec g xs # 0 \<up> (xb - Suc xa) @ xc}))"
  2719 apply(simp add: exponent_cons_iff)
   617   and ft: "ft = max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
  2720 apply(case_tac "ys = []", simp_all add: append_butlast_last_id)
   618   shows 
  2721 done
   619   "{\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft + length gs) @ anything}
  2722 
   620     cn_merge_gs (map rec_ci (take n gs)) ft
  2723 lemma cn_merge_gs_ex: 
   621   {\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @
  2724   "\<lbrakk>\<And>x aprog a_md rs_pos rs suf_lm lm.
   622                     map (\<lambda>i. rec_exec i xs) (take n gs) @ 0\<up>(length gs - n) @ 0 \<up> Suc (length xs) @ anything}"
  2725     \<lbrakk>x \<in> set gs; rec_ci x = (aprog, rs_pos, a_md);
   623   using g_cond
  2726      rec_calc_rel x lm rs\<rbrakk>
   624 proof(induct n)
  2727      \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp 
   625   case 0
  2728            = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm); 
   626   have "ft > length xs"
  2729    pstr + length gs\<le> a_md;
   627     using ft
  2730    \<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k);
   628     by simp
  2731    length ys = length gs; length lm = n;
   629   thus "?case"
  2732    pstr \<ge> Max (set (Suc n # map (\<lambda>(aprog, p, n). n) (map rec_ci gs)))\<rbrakk>
   630     apply(rule_tac abc_Hoare_haltI)
  2733   \<Longrightarrow> \<exists> stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suf_lm)
   631     apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps replicate_add[THEN sym] 
  2734                    (cn_merge_gs (map rec_ci gs) pstr) stp 
   632       replicate_Suc[THEN sym] del: replicate_Suc)
  2735    = (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) gs) +
       
  2736   3 * length gs, lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - (pstr + length gs)) @ suf_lm)"
       
  2737 apply(induct gs arbitrary: ys rule: list_tl_induct)
       
  2738 apply(simp add: exponent_add_iff, simp)
       
  2739 proof -
       
  2740   fix a list ys
       
  2741   assume ind: "\<And>x aprog a_md rs_pos rs suf_lm lm.
       
  2742     \<lbrakk>x = a \<or> x \<in> set list; rec_ci x = (aprog, rs_pos, a_md); 
       
  2743      rec_calc_rel x lm rs\<rbrakk>
       
  2744     \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
       
  2745                 (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  2746   and ind2: 
       
  2747     "\<And>ys. \<lbrakk>\<And>x aprog a_md rs_pos rs suf_lm lm.
       
  2748     \<lbrakk>x \<in> set list; rec_ci x = (aprog, rs_pos, a_md);
       
  2749      rec_calc_rel x lm rs\<rbrakk>
       
  2750     \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp
       
  2751         = (length aprog, lm @ rs # 0\<up>(a_md - Suc rs_pos) @ suf_lm);
       
  2752     \<forall>k<length list. rec_calc_rel (list ! k) lm (ys ! k); 
       
  2753     length ys = length list\<rbrakk>
       
  2754     \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suf_lm) 
       
  2755                    (cn_merge_gs (map rec_ci list) pstr) stp =
       
  2756     (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) +
       
  2757      3 * length list,
       
  2758                 lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm)"
       
  2759     and h: "Suc (pstr + length list) \<le> a_md" 
       
  2760             "\<forall>k<Suc (length list). 
       
  2761                    rec_calc_rel ((list @ [a]) ! k) lm (ys ! k)" 
       
  2762             "length ys = Suc (length list)" 
       
  2763             "length lm = n"
       
  2764             "Suc n \<le> pstr \<and> (\<lambda>(aprog, p, n). n) (rec_ci a) \<le> pstr \<and> 
       
  2765             (\<forall>a\<in>set list. (\<lambda>(aprog, p, n). n) (rec_ci a) \<le> pstr)"
       
  2766   from h have k1: 
       
  2767     "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suf_lm)
       
  2768                      (cn_merge_gs (map rec_ci list) pstr) stp =
       
  2769     (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) +
       
  2770      3 * length list, lm @ 0\<up>(pstr - n) @ butlast ys @
       
  2771                                0\<up>(a_md - (pstr + length list)) @ suf_lm) "
       
  2772     apply(rule_tac ind2)
       
  2773     apply(rule_tac ind, auto)
       
  2774     done
   633     done
  2775   from k1 and h show 
   634 next
  2776     "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suf_lm) 
   635   case (Suc n)
  2777           (cn_merge_gs (map rec_ci list @ [rec_ci a]) pstr) stp =
   636   have ind': "\<forall>g\<in>set (take n gs).
  2778         (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) + 
   637      terminate g xs \<and> (\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow> 
  2779         (\<lambda>(ap, pos, n). length ap) (rec_ci a) + (3 + 3 * length list),
   638     (\<forall>xc. {\<lambda>nl. nl = xs @ 0 \<up> (xb - xa) @ xc} x {\<lambda>nl. nl = xs @ rec_exec g xs # 0 \<up> (xb - Suc xa) @ xc})) \<Longrightarrow>
  2780              lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm)"
   639     {\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft + length gs) @ anything} cn_merge_gs (map rec_ci (take n gs)) ft 
  2781     apply(simp add: cn_merge_gs_tl_app)
   640     {\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything}"
  2782     apply(rule_tac as = 
   641     by fact
  2783   "(\<Sum>(ap, pos, n)\<leftarrow>map rec_ci list. length ap) + 3 * length list"    
   642   have g_newcond: "\<forall>g\<in>set (take (Suc n) gs).
  2784       and bm = "lm @ 0\<up>(pstr - n) @ butlast ys @ 
   643      terminate g xs \<and> (\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow> (\<forall>xc. {\<lambda>nl. nl = xs @ 0 \<up> (xb - xa) @ xc} x {\<lambda>nl. nl = xs @ rec_exec g xs # 0 \<up> (xb - Suc xa) @ xc}))"
  2785                               0\<up>(a_md - (pstr + length list)) @ suf_lm" 
   644     by fact
  2786       and bs = "(\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3" 
   645   from g_newcond have ind:
  2787       and bm' = "lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ 
   646     "{\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft + length gs) @ anything} cn_merge_gs (map rec_ci (take n gs)) ft 
  2788                                   suf_lm" in abc_append_exc2, simp)
   647     {\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything}"
  2789     apply(simp add: cn_merge_gs_length)
   648     apply(rule_tac ind', rule_tac ballI, erule_tac x = g in ballE, simp_all add: take_Suc)
  2790   proof -
   649     by(case_tac "n < length gs", simp add:take_Suc_conv_app_nth, simp)    
  2791     from h show 
   650   show "?case"
  2792       "\<exists>bstp. abc_steps_l (0, lm @ 0\<up>(pstr - n) @ butlast ys @ 
   651   proof(cases "n < length gs")
  2793                                   0\<up>(a_md - (pstr + length list)) @ suf_lm) 
   652     case True
  2794               ((\<lambda>(gprog, gpara, gn). gprog [+] mv_box gpara 
   653     have h: "n < length gs" by fact
  2795               (pstr + length list)) (rec_ci a)) bstp =
   654     thus "?thesis"
  2796               ((\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3, 
   655     proof(simp add: take_Suc_conv_app_nth cn_merge_gs_tl_app)
  2797              lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm)"
   656       obtain gp ga gf where a: "rec_ci (gs!n) = (gp, ga, gf)"
  2798       apply(case_tac "rec_ci a", simp)
   657         by (metis prod_cases3)
  2799       apply(rule_tac as = "length aa" and 
   658       moreover have "min (length gs) n = n"
  2800                      bm = "lm @ (ys ! (length list)) # 
   659         using h by simp
  2801           0\<up>(pstr - Suc n) @ butlast ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm" 
   660       moreover have 
  2802         and bs = "3" and bm' = "lm @ 0\<up>(pstr - n) @ ys @
   661         "{\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft + length gs) @ anything}
  2803              0\<up>(a_md - Suc (pstr + length list)) @ suf_lm" in abc_append_exc2)
   662         cn_merge_gs (map rec_ci (take n gs)) ft [+] (gp [+] mv_box ga (ft + n))
  2804     proof -
   663         {\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ 
  2805       fix aa b c
   664         rec_exec (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything}"
  2806       assume g: "rec_ci a = (aa, b, c)"
   665       proof(rule_tac abc_Hoare_plus_halt)
  2807       from h and g have k2: "b = n"
   666         show "{\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft + length gs) @ anything} cn_merge_gs (map rec_ci (take n gs)) ft
  2808 	apply(erule_tac x = "length list" in allE, simp)
   667           {\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything}"
  2809 	apply(subgoal_tac "length lm = b", simp)
   668           using ind by simp
  2810 	apply(rule para_pattern, simp, simp)
   669       next
  2811 	done
   670         have x: "gs!n \<in> set (take (Suc n) gs)"
  2812       from h and g and this show 
   671           using h
  2813         "\<exists>astp. abc_steps_l (0, lm @ 0\<up>(pstr - n) @ butlast ys @ 
   672           by(simp add: take_Suc_conv_app_nth)
  2814                          0\<up>(a_md - (pstr + length list)) @ suf_lm) aa astp =
   673         have b: "terminate (gs!n) xs"
  2815         (length aa, lm @ ys ! length list # 0\<up>(pstr - Suc n) @ 
   674           using a g_newcond h x
  2816                        butlast ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm)"
   675           by(erule_tac x = "gs!n" in ballE, simp_all)
  2817 	apply(subgoal_tac "c \<ge> Suc n")
   676         hence c: "length xs = ga"
  2818 	apply(insert ind[of a aa b c lm "ys ! length list" 
   677           using a param_pattern by metis  
  2819           "0\<up>(pstr - c) @ butlast ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm"], simp)
   678         have d: "gf > ga" using footprint_ge a by simp
  2820 	apply(erule_tac x = "length list" in allE, 
   679         have e: "ft \<ge> gf" using ft a h
  2821               simp add: exponent_add_iff)
   680           by(simp,  rule_tac min_max.le_supI2, rule_tac Max_ge, simp, 
  2822 	apply(rule_tac Suc_leI, rule_tac ci_ad_ge_paras, simp)
   681             rule_tac insertI2,  
  2823 	done
   682             rule_tac f = "(\<lambda>(aprog, p, n). n)" and x = "rec_ci (gs!n)" in image_eqI, simp, 
  2824     next
   683             rule_tac x = "gs!n" in image_eqI, simp, simp)
  2825       fix aa b c
   684         show "{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ 
  2826       show "length aa = length aa" by simp 
   685           map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything} gp [+] mv_box ga (ft + n)
  2827     next
   686           {\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) 
  2828       fix aa b c
   687           (take n gs) @ rec_exec (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything}"
  2829       assume "rec_ci a = (aa, b, c)"
   688         proof(rule_tac abc_Hoare_plus_halt)
  2830       from h and this show     
   689           show "{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything} gp 
  2831       "\<exists>bstp. abc_steps_l (0, lm @ ys ! length list #
   690                 {\<lambda>nl. nl = xs @ (rec_exec (gs!n) xs) # 0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_exec i xs) 
  2832           0\<up>(pstr - Suc n) @ butlast ys @ 0\<up>(a_md - (pstr + length list)) @ suf_lm)
   691                               (take n gs) @  0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything}"
  2833                  (mv_box b (pstr + length list)) bstp =
   692           proof -
  2834        (3, lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - Suc (pstr + length list)) @ suf_lm)"
   693             have 
  2835 	apply(insert mv_box_ex [of b "pstr + length list" 
   694               "({\<lambda>nl. nl = xs @ 0 \<up> (gf - ga) @ 0\<up>(ft - gf)@map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything} 
  2836          "lm @ ys ! length list # 0\<up>(pstr - Suc n) @ butlast ys @ 
   695             gp {\<lambda>nl. nl = xs @ (rec_exec (gs!n) xs) # 0 \<up> (gf - Suc ga) @ 
  2837          0\<up>(a_md - (pstr + length list)) @ suf_lm"], simp)
   696               0\<up>(ft - gf)@map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything})"
  2838         apply(subgoal_tac "b = n")
   697               using a g_newcond h x
  2839 	apply(simp add: nth_append split: if_splits)
   698               apply(erule_tac x = "gs!n" in ballE)
  2840 	apply(erule_tac x = "length list" in allE, simp)
   699               apply(simp, simp)
  2841         apply(drule para_pattern, simp, simp)
   700               done            
  2842 	done
   701             thus "?thesis"
  2843     next
   702               using a b c d e
  2844       fix aa b c
   703               by(simp add: replicate_merge_anywhere)
  2845       show "3 = length (mv_box b (pstr + length list))" 
   704           qed
       
   705         next
       
   706           show 
       
   707             "{\<lambda>nl. nl = xs @ rec_exec (gs ! n) xs #
       
   708             0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything}
       
   709             mv_box ga (ft + n)
       
   710             {\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @
       
   711             rec_exec (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything}"
       
   712           proof -
       
   713             have "{\<lambda>nl. nl = xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @
       
   714               map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything}
       
   715               mv_box ga (ft + n) {\<lambda>nl. nl = (xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @
       
   716               map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything)
       
   717               [ft + n := (xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ 
       
   718               0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) ! ga +
       
   719               (xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ 
       
   720               map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) !
       
   721                       (ft + n),  ga := 0]}"
       
   722               using a c d e h
       
   723               apply(rule_tac mv_box_correct)
       
   724               apply(simp, arith, arith)
       
   725               done
       
   726             moreover have "(xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @
       
   727               map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything)
       
   728               [ft + n := (xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ 
       
   729               0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) ! ga +
       
   730               (xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ 
       
   731               map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) !
       
   732                       (ft + n),  ga := 0]= 
       
   733               xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything"
       
   734               using a c d e h
       
   735               by(simp add: list_update_append nth_append length_replicate split: if_splits del: list_update.simps(2), auto)
       
   736             ultimately show "?thesis"
       
   737               by(simp)
       
   738           qed
       
   739         qed  
       
   740       qed
       
   741       ultimately show 
       
   742         "{\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft + length gs) @ anything}
       
   743         cn_merge_gs (map rec_ci (take n gs)) ft [+] (case rec_ci (gs ! n) of (gprog, gpara, gn) \<Rightarrow>
       
   744         gprog [+] mv_box gpara (ft + min (length gs) n))
       
   745         {\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything}"
  2846         by simp
   746         by simp
  2847     next
       
  2848       fix aa b aaa ba
       
  2849       show "length aa + 3 = length aa + 3" by simp
       
  2850     next
       
  2851       fix aa b c
       
  2852       show "mv_box b (pstr + length list) \<noteq> []" 
       
  2853         by(simp add: mv_box.simps)
       
  2854     qed
   747     qed
  2855   next
   748   next
  2856     show "(\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3 = 
   749     case False
  2857         length ((\<lambda>(gprog, gpara, gn). gprog [+]
   750     have h: "\<not> n < length gs" by fact
  2858            mv_box gpara (pstr + length list)) (rec_ci a))"
   751     hence ind': 
  2859       by(case_tac "rec_ci a", simp)
   752       "{\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft + length gs) @ anything} cn_merge_gs (map rec_ci gs) ft
  2860   next
   753         {\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) gs @ 0 \<up> Suc (length xs) @ anything}"
  2861     show "listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) list) +
   754       using ind
  2862       (\<lambda>(ap, pos, n). length ap) (rec_ci a) + (3 + 3 * length list)=
   755       by simp
  2863       (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci list. length ap) + 3 * length list + 
   756     thus "?thesis"
  2864                 ((\<lambda>(ap, pos, n). length ap) (rec_ci a) + 3)" by simp
   757       using h
  2865   next
   758       by(simp)
  2866     show "(\<lambda>(gprog, gpara, gn). gprog [+] 
       
  2867       mv_box gpara (pstr + length list)) (rec_ci a) \<noteq> []"
       
  2868       by(case_tac "rec_ci a", 
       
  2869          simp add: abc_append.simps abc_shift.simps)
       
  2870   qed
   759   qed
  2871 qed
   760 qed
  2872  
   761     
  2873 lemma [simp]: "length (mv_boxes aa ba n) = 3*n"
   762 lemma compile_cn_gs_correct:
       
   763   assumes
       
   764   g_cond: "\<forall>g\<in>set gs. terminate g xs \<and>
       
   765   (\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow> (\<forall>xc. {\<lambda>nl. nl = xs @ 0 \<up> (xb - xa) @ xc} x {\<lambda>nl. nl = xs @ rec_exec g xs # 0 \<up> (xb - Suc xa) @ xc}))"
       
   766   and ft: "ft = max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
       
   767   shows 
       
   768   "{\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft + length gs) @ anything}
       
   769     cn_merge_gs (map rec_ci gs) ft
       
   770   {\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @
       
   771                     map (\<lambda>i. rec_exec i xs) gs @ 0 \<up> Suc (length xs) @ anything}"
       
   772 using assms
       
   773 using compile_cn_gs_correct'[of "length gs" gs xs ft ffp anything ]
       
   774 apply(auto)
       
   775 done
       
   776   
       
   777 lemma length_mvboxes[simp]: "length (mv_boxes aa ba n) = 3*n"
  2874 by(induct n, auto simp: mv_boxes.simps)
   778 by(induct n, auto simp: mv_boxes.simps)
  2875 
   779 
  2876 lemma exp_suc: "a\<up>Suc b = a\<up>b @ [a]"
   780 lemma exp_suc: "a\<up>Suc b = a\<up>b @ [a]"
  2877 by(simp add: exp_ind del: replicate.simps)
   781 by(simp add: exp_ind del: replicate.simps)
  2878 
   782 
  2917   [ba + n := last lm2, aa + n := 0] = 
   821   [ba + n := last lm2, aa + n := 0] = 
  2918   lm1 @ 0 # 0\<up>n @ lm3 @ lm2 @ lm4"
   822   lm1 @ 0 # 0\<up>n @ lm3 @ lm2 @ lm4"
  2919 using list_update_append[of "lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2" "0 # lm4" 
   823 using list_update_append[of "lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2" "0 # lm4" 
  2920                             "ba + n" "last lm2"]
   824                             "ba + n" "last lm2"]
  2921 apply(simp)
   825 apply(simp)
  2922 apply(simp add: list_update_append)
   826 apply(simp add: list_update_append list_update.simps replicate_Suc_iff_anywhere exp_suc
  2923 apply(simp only: exponent_cons_iff exp_suc, simp)
   827   del: replicate_Suc)
  2924 apply(case_tac lm2, simp, simp)
   828 apply(case_tac lm2, simp, simp)
  2925 done
   829 done
  2926 
   830 
  2927 lemma mv_boxes_ex:
   831 lemma [simp]:
  2928   "\<lbrakk>n \<le> ba - aa; ba > aa; length lm1 = aa; 
   832   "\<lbrakk>Suc (length lm1 + n) \<le> ba; length lm2 = Suc n; length lm3 = ba - Suc (length lm1 + n); 
  2929     length (lm2::nat list) = n; length lm3 = ba - aa - n\<rbrakk>
   833   \<not> ba - Suc (length lm1) < ba - Suc (length lm1 + n); \<not> ba + n - length lm1 < n\<rbrakk>
  2930      \<Longrightarrow> \<exists> stp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<up>n @ lm4)
   834     \<Longrightarrow> (0::nat) \<up> n @ (last lm2 # lm3 @ butlast lm2 @ 0 # lm4)[ba - length lm1 := last lm2, 0 := 0] =
  2931        (mv_boxes aa ba n) stp = (3 * n, lm1 @ 0\<up>n @ lm3 @ lm2 @ lm4)"
   835   0 # 0 \<up> n @ lm3 @ lm2 @ lm4"
  2932 apply(induct n arbitrary: lm2 lm3 lm4, simp)
   836 apply(subgoal_tac "ba - length lm1 = Suc n + length lm3", simp add: list_update.simps list_update_append)
  2933 apply(rule_tac x = 0 in exI, simp add: abc_steps_zero, 
   837 apply(simp add: replicate_Suc_iff_anywhere exp_suc del: replicate_Suc)
  2934               simp add: mv_boxes.simps del: exp_suc_iff)
   838 apply(case_tac lm2, simp, simp)
  2935 apply(rule_tac as = "3 *n" and bm = "lm1 @ 0\<up>n @ last lm2 # lm3 @
   839 apply(auto)
  2936                butlast lm2 @ 0 # lm4" in abc_append_exc2, simp_all)
   840 done
  2937 apply(simp only: exponent_cons_iff, simp only: exp_suc, simp)
   841 
  2938 proof -
   842 lemma mv_boxes_correct: 
  2939   fix n lm2 lm3 lm4
   843   "\<lbrakk>aa + n \<le> ba; ba > aa; length lm1 = aa; length lm2 = n; length lm3 = ba - aa - n\<rbrakk>
  2940   assume ind:
   844  \<Longrightarrow> {\<lambda> nl. nl = lm1 @ lm2 @ lm3 @ 0\<up>n @ lm4} (mv_boxes aa ba n) 
  2941     "\<And>lm2 lm3 lm4. \<lbrakk>length lm2 = n; length lm3 = ba - (aa + n)\<rbrakk> \<Longrightarrow>
   845      {\<lambda> nl. nl = lm1 @ 0\<up>n @ lm3 @ lm2 @ lm4}"
  2942     \<exists>stp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<up>n @ lm4) 
   846 proof(induct n arbitrary: lm2 lm3 lm4)
  2943        (mv_boxes aa ba n) stp = (3 * n, lm1 @ 0\<up>n @ lm3 @ lm2 @ lm4)"
   847   case 0
  2944   and h: "Suc n \<le> ba - aa" "aa < ba" "length (lm1::nat list) = aa" 
   848   thus "?case"
  2945          "length (lm2::nat list) = Suc n" 
   849     by(simp add: mv_boxes.simps abc_Hoare_halt_def, rule_tac  x = 0 in exI, simp add: abc_steps_l.simps)
  2946          "length (lm3::nat list) = ba - Suc (aa + n)"
       
  2947   from h show 
       
  2948     "\<exists>astp. abc_steps_l (0, lm1 @ lm2 @ lm3 @ 0\<up>n @ 0 # lm4) 
       
  2949                        (mv_boxes aa ba n) astp = 
       
  2950         (3 * n, lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4)"
       
  2951     apply(insert ind[of "butlast lm2" "last lm2 # lm3" "0 # lm4"], 
       
  2952           simp)
       
  2953     apply(subgoal_tac "lm1 @ butlast lm2 @ last lm2 # lm3 @ 0\<up>n @ 
       
  2954               0 # lm4 = lm1 @ lm2 @ lm3 @ 0\<up>n @ 0 # lm4", simp, simp)
       
  2955     apply(case_tac "lm2 = []", simp, simp)
       
  2956     done
       
  2957 next
   850 next
  2958   fix n lm2 lm3 lm4
   851   case (Suc n)
  2959   assume h: "Suc n \<le> ba - aa"
   852   have ind: 
  2960             "aa < ba" 
   853     "\<And>lm2 lm3 lm4.
  2961             "length (lm1::nat list) = aa" 
   854     \<lbrakk>aa + n \<le> ba; aa < ba; length lm1 = aa; length lm2 = n; length lm3 = ba - aa - n\<rbrakk>
  2962             "length (lm2::nat list) = Suc n" 
   855     \<Longrightarrow> {\<lambda>nl. nl = lm1 @ lm2 @ lm3 @ 0 \<up> n @ lm4} mv_boxes aa ba n {\<lambda>nl. nl = lm1 @ 0 \<up> n @ lm3 @ lm2 @ lm4}"
  2963             "length (lm3::nat list) = ba - Suc (aa + n)"
   856     by fact
  2964   thus " \<exists>bstp. abc_steps_l (0, lm1 @ 0\<up>n @ last lm2 # lm3 @
   857   have h1: "aa + Suc n \<le> ba"  by fact
  2965                        butlast lm2 @ 0 # lm4) 
   858   have h2: "aa < ba" by fact
  2966                          (mv_box (aa + n) (ba + n)) bstp
   859   have h3: "length lm1 = aa" by fact
  2967                = (3, lm1 @ 0 # 0\<up>n @ lm3 @ lm2 @ lm4)"
   860   have h4: "length lm2 = Suc n" by fact 
  2968     apply(insert mv_box_ex[of "aa + n" "ba + n" 
   861   have h5: "length lm3 = ba - aa - Suc n" by fact
  2969        "lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4"], simp)
   862   have "{\<lambda>nl. nl = lm1 @ lm2 @ lm3 @ 0 \<up> Suc n @ lm4} mv_boxes aa ba n [+] mv_box (aa + n) (ba + n)
  2970     done
   863     {\<lambda>nl. nl = lm1 @ 0 \<up> Suc n @ lm3 @ lm2 @ lm4}"
  2971 qed
   864   proof(rule_tac abc_Hoare_plus_halt)
  2972 
   865     have "{\<lambda>nl. nl = lm1 @ butlast lm2 @ (last lm2 # lm3) @ 0 \<up> n @ (0 # lm4)} mv_boxes aa ba n
  2973 lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; ba < aa; length lm1 = ba; 
   866           {\<lambda> nl. nl = lm1 @ 0\<up>n @ (last lm2 # lm3) @ butlast lm2 @ (0 # lm4)}"
  2974         length lm2 = aa - Suc (ba + n); length lm3 = Suc n\<rbrakk>
   867       using h1 h2 h3 h4 h5
  2975    \<Longrightarrow> (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ last lm3 # lm4) ! (aa + n) = last lm3"
   868       by(rule_tac ind, simp_all)
  2976 using nth_append[of "lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n" "last lm3 # lm4" "aa + n"]
   869     moreover have " lm1 @ butlast lm2 @ (last lm2 # lm3) @ 0 \<up> n @ (0 # lm4)
  2977 apply(simp)
   870                   = lm1 @ lm2 @ lm3 @ 0 \<up> Suc n @ lm4"
  2978 done
   871       using h4
  2979 
   872       by(simp add: replicate_Suc[THEN sym] exp_suc del: replicate_Suc, 
  2980 lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; ba < aa; length lm1 = ba; 
   873             case_tac lm2, simp_all)
  2981         length lm2 = aa - Suc (ba + n); length lm3 = Suc n\<rbrakk>
   874     ultimately show "{\<lambda>nl. nl = lm1 @ lm2 @ lm3 @ 0 \<up> Suc n @ lm4} mv_boxes aa ba n
  2982      \<Longrightarrow> (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ last lm3 # lm4) ! (ba + n) = 0"
   875           {\<lambda> nl. nl = lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4}"
  2983 apply(simp add: nth_append)
   876       by (metis append_Cons)
  2984 done
   877   next
  2985 
   878     let ?lm = "lm1 @ 0 \<up> n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4"
  2986 lemma [simp]: "\<lbrakk>Suc n \<le> aa - ba; ba < aa; length lm1 = ba; 
   879     have "{\<lambda>nl. nl = ?lm} mv_box (aa + n) (ba + n)
  2987         length lm2 = aa - Suc (ba + n); length lm3 = Suc n\<rbrakk> 
   880           {\<lambda> nl. nl = ?lm[(ba + n) := ?lm!(aa+n) + ?lm!(ba+n), (aa+n):=0]}"
  2988      \<Longrightarrow> (lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ last lm3 # lm4)[ba + n := last lm3, aa + n := 0]
   881       using h1 h2 h3 h4  h5
  2989       = lm1 @ lm3 @ lm2 @ 0 # 0\<up>n @ lm4"
   882       by(rule_tac mv_box_correct, simp_all)
  2990 using list_update_append[of "lm1 @ butlast lm3" "(0\<Colon>'a) # lm2 @ (0\<Colon>'a)\<up>n @ last lm3 # lm4"]
   883     moreover have "?lm[(ba + n) := ?lm!(aa+n) + ?lm!(ba+n), (aa+n):=0]
  2991 apply(simp)
   884                  =  lm1 @ 0 \<up> Suc n @ lm3 @ lm2 @ lm4"
  2992 using list_update_append[of "lm1 @ butlast lm3 @ last lm3 # lm2 @ (0\<Colon>'a)\<up>n"
   885       using h1 h2 h3 h4 h5
  2993                             "last lm3 # lm4" "aa + n" "0"]
   886       by(auto simp: nth_append list_update_append split: if_splits)
  2994 apply(simp)
   887     ultimately show "{\<lambda>nl. nl = lm1 @ 0 \<up> n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4} mv_box (aa + n) (ba + n)
  2995 apply(simp only: replicate_Suc[THEN sym] exp_suc, simp)
   888           {\<lambda>nl. nl = lm1 @ 0 \<up> Suc n @ lm3 @ lm2 @ lm4}"
  2996 apply(case_tac lm3, simp, simp)
   889      by simp
  2997 done
   890  qed
  2998 
   891  thus "?case"
  2999 lemma mv_boxes_ex2:
   892    by(simp add: mv_boxes.simps)
       
   893 qed
       
   894     
       
   895 lemma [simp]:
       
   896   "\<lbrakk>Suc n \<le> aa - length lm1; length lm1 < aa; 
       
   897   length lm2 = aa - Suc (length lm1 + n); 
       
   898   length lm3 = Suc n; 
       
   899   \<not> aa - Suc (length lm1) < aa - Suc (length lm1 + n);
       
   900   \<not> aa + n - length lm1 < n\<rbrakk>
       
   901   \<Longrightarrow> butlast lm3 @ ((0::nat) # lm2 @ 0 \<up> n @ last lm3 # lm4)[0 := last lm3, aa - length lm1 := 0] = lm3 @ lm2 @ 0 # 0 \<up> n @ lm4"
       
   902   apply(subgoal_tac "aa - length lm1 = length lm2 + Suc n")
       
   903   apply(simp add: list_update.simps list_update_append)
       
   904   apply(simp add: replicate_Suc[THEN sym] exp_suc del: replicate_Suc)
       
   905   apply(case_tac lm3, simp, simp)
       
   906   apply(auto)
       
   907   done
       
   908 
       
   909 lemma mv_boxes_correct2:
  3000   "\<lbrakk>n \<le> aa - ba; 
   910   "\<lbrakk>n \<le> aa - ba; 
  3001     ba < aa; 
   911     ba < aa; 
  3002     length (lm1::nat list) = ba;
   912     length (lm1::nat list) = ba;
  3003     length (lm2::nat list) = aa - ba - n; 
   913     length (lm2::nat list) = aa - ba - n; 
  3004     length (lm3::nat list) = n\<rbrakk>
   914     length (lm3::nat list) = n\<rbrakk>
  3005      \<Longrightarrow> \<exists> stp. abc_steps_l (0, lm1 @ 0\<up>n @ lm2 @ lm3 @ lm4) 
   915   \<Longrightarrow>{\<lambda> nl. nl = lm1 @ 0\<up>n @ lm2 @ lm3 @ lm4}
  3006                 (mv_boxes aa ba n) stp =
   916                 (mv_boxes aa ba n) 
  3007                     (3 * n, lm1 @ lm3 @ lm2 @ 0\<up>n @ lm4)"
   917      {\<lambda> nl. nl = lm1 @ lm3 @ lm2 @ 0\<up>n @ lm4}"
  3008 apply(induct n arbitrary: lm2 lm3 lm4, simp)
   918 proof(induct n arbitrary: lm2 lm3 lm4)
  3009 apply(rule_tac x = 0 in exI, simp add: abc_steps_zero, 
   919   case 0
  3010                    simp add: mv_boxes.simps del: exp_suc_iff)
   920   thus "?case"
  3011 apply(rule_tac as = "3 *n" and bm = "lm1 @ butlast lm3 @ 0 # lm2 @
   921     by(simp add: mv_boxes.simps abc_Hoare_halt_def, rule_tac  x = 0 in exI, simp add: abc_steps_l.simps)
  3012                   0\<up>n @ last lm3 # lm4" in abc_append_exc2, simp_all)
   922 next
  3013 apply(simp only: exponent_cons_iff, simp only: exp_suc, simp)
   923   case (Suc n)
       
   924   have ind:
       
   925     "\<And>lm2 lm3 lm4.
       
   926     \<lbrakk>n \<le> aa - ba; ba < aa; length lm1 = ba; length lm2 = aa - ba - n; length lm3 = n\<rbrakk>
       
   927     \<Longrightarrow> {\<lambda>nl. nl = lm1 @ 0 \<up> n @ lm2 @ lm3 @ lm4} mv_boxes aa ba n {\<lambda>nl. nl = lm1 @ lm3 @ lm2 @ 0 \<up> n @ lm4}"
       
   928     by fact
       
   929   have h1: "Suc n \<le> aa - ba" by fact
       
   930   have h2: "ba < aa" by fact
       
   931   have h3: "length lm1 = ba" by fact 
       
   932   have h4: "length lm2 = aa - ba - Suc n" by fact
       
   933   have h5: "length lm3 = Suc n" by fact
       
   934   have "{\<lambda>nl. nl = lm1 @ 0 \<up> Suc n @ lm2 @ lm3 @ lm4}  mv_boxes aa ba n [+] mv_box (aa + n) (ba + n) 
       
   935     {\<lambda>nl. nl = lm1 @ lm3 @ lm2 @ 0 \<up> Suc n @ lm4}"
       
   936   proof(rule_tac abc_Hoare_plus_halt)
       
   937     have "{\<lambda> nl. nl = lm1 @ 0 \<up> n @ (0 # lm2) @ (butlast lm3) @ (last lm3 # lm4)} mv_boxes aa ba n
       
   938            {\<lambda> nl. nl = lm1 @ butlast lm3 @ (0 # lm2) @ 0\<up>n @ (last lm3 # lm4)}"
       
   939       using h1 h2 h3 h4 h5
       
   940       by(rule_tac ind, simp_all)
       
   941     moreover have "lm1 @ 0 \<up> n @ (0 # lm2) @ (butlast lm3) @ (last lm3 # lm4) 
       
   942                    = lm1 @ 0 \<up> Suc n @ lm2 @ lm3 @ lm4"
       
   943       using h5
       
   944      by(simp add: replicate_Suc_iff_anywhere exp_suc 
       
   945         del: replicate_Suc, case_tac lm3, simp_all)
       
   946    ultimately show "{\<lambda>nl. nl = lm1 @ 0 \<up> Suc n @ lm2 @ lm3 @ lm4} mv_boxes aa ba n
       
   947      {\<lambda> nl. nl = lm1 @ butlast lm3 @ (0 # lm2) @ 0\<up>n @ (last lm3 # lm4)}"
       
   948      by metis
       
   949  next
       
   950    thm mv_box_correct
       
   951    let ?lm = "lm1 @ butlast lm3 @ (0 # lm2) @ 0 \<up> n @ last lm3 # lm4"
       
   952    have "{\<lambda>nl. nl = ?lm} mv_box (aa + n) (ba + n)
       
   953          {\<lambda>nl. nl = ?lm[ba+n := ?lm!(aa+n)+?lm!(ba+n), (aa+n):=0]}"
       
   954      using h1 h2 h3 h4 h5
       
   955      by(rule_tac mv_box_correct, simp_all)
       
   956    moreover have "?lm[ba+n := ?lm!(aa+n)+?lm!(ba+n), (aa+n):=0]
       
   957                = lm1 @ lm3 @ lm2 @ 0 \<up> Suc n @ lm4"
       
   958      using h1 h2 h3 h4 h5
       
   959      by(auto simp: nth_append list_update_append split: if_splits)
       
   960    ultimately show "{\<lambda>nl. nl = lm1 @ butlast lm3 @ (0 # lm2) @ 0 \<up> n @ last lm3 # lm4} mv_box (aa + n) (ba + n)
       
   961      {\<lambda>nl. nl = lm1 @ lm3 @ lm2 @ 0 \<up> Suc n @ lm4}"
       
   962      by simp
       
   963  qed
       
   964  thus "?case"
       
   965    by(simp add: mv_boxes.simps)
       
   966 qed    
       
   967      
       
   968 lemma save_paras: 
       
   969   "{\<lambda>nl. nl = xs @ 0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - length xs) @
       
   970   map (\<lambda>i. rec_exec i xs) gs @ 0 \<up> Suc (length xs) @ anything}
       
   971   mv_boxes 0 (Suc (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) (length xs)
       
   972   {\<lambda>nl. nl = 0 \<up> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @ map (\<lambda>i. rec_exec i xs) gs @ 0 # xs @ anything}"
  3014 proof -
   973 proof -
  3015   fix n lm2 lm3 lm4
   974   let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
  3016   assume ind: 
   975   have "{\<lambda>nl. nl = [] @ xs @ (0\<up>(?ft - length xs) @  map (\<lambda>i. rec_exec i xs) gs @ [0]) @ 
  3017 "\<And>lm2 lm3 lm4. \<lbrakk>length lm2 = aa - (ba + n); length lm3 = n\<rbrakk> \<Longrightarrow> 
   976           0 \<up> (length xs) @ anything} mv_boxes 0 (Suc ?ft + length gs) (length xs) 
  3018   \<exists>stp. abc_steps_l (0, lm1 @ 0\<up>n @ lm2 @ lm3 @ lm4) 
   977         {\<lambda>nl. nl = [] @ 0 \<up> (length xs) @ (0\<up>(?ft - length xs) @  map (\<lambda>i. rec_exec i xs) gs @ [0]) @ xs @ anything}"
  3019                  (mv_boxes aa ba n) stp = 
   978     by(rule_tac mv_boxes_correct, auto)
  3020                             (3 * n, lm1 @ lm3 @ lm2 @ 0\<up>n @ lm4)"
   979   thus "?thesis"
  3021   and h: "Suc n \<le> aa - ba" 
   980     by(simp add: replicate_merge_anywhere)
  3022          "ba < aa"  
   981 qed
  3023          "length (lm1::nat list) = ba" 
   982 
  3024          "length (lm2::nat list) = aa - Suc (ba + n)" 
   983 lemma [intro]: 
  3025          "length (lm3::nat list) = Suc n"
   984   "length gs \<le> ffp \<Longrightarrow> length gs \<le> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
  3026   from h show
   985  apply(rule_tac min_max.le_supI2)
  3027     "\<exists>astp. abc_steps_l (0, lm1 @ 0\<up>n @ 0 # lm2 @ lm3 @ lm4)
   986  apply(simp add: Max_ge_iff)
  3028         (mv_boxes aa ba n) astp = 
   987 done
  3029           (3 * n, lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ last lm3 # lm4)"
   988 
  3030     apply(insert ind[of "0 # lm2" "butlast lm3" "last lm3 # lm4"],
   989 lemma restore_new_paras:
  3031           simp)
   990   "ffp \<ge> length gs 
  3032     apply(subgoal_tac
   991  \<Longrightarrow> {\<lambda>nl. nl = 0 \<up> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @ map (\<lambda>i. rec_exec i xs) gs @ 0 # xs @ anything}
  3033       "lm1 @ 0\<up>n @ 0 # lm2 @ butlast lm3 @ last lm3 # lm4 =
   992     mv_boxes (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) 0 (length gs)
  3034            lm1 @ 0\<up>n @ 0 # lm2 @ lm3 @ lm4", simp, simp)
   993   {\<lambda>nl. nl = map (\<lambda>i. rec_exec i xs) gs @ 0 \<up> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @ 0 # xs @ anything}"
  3035     apply(case_tac "lm3 = []", simp, simp)
   994 proof -
       
   995   let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
       
   996   assume j: "ffp \<ge> length gs"
       
   997   hence "{\<lambda> nl. nl = [] @ 0\<up>length gs @ 0\<up>(?ft - length gs) @  map (\<lambda>i. rec_exec i xs) gs @ ((0 # xs) @ anything)}
       
   998        mv_boxes ?ft 0 (length gs)
       
   999         {\<lambda> nl. nl = [] @ map (\<lambda>i. rec_exec i xs) gs @ 0\<up>(?ft - length gs) @ 0\<up>length gs @ ((0 # xs) @ anything)}"
       
  1000     by(rule_tac mv_boxes_correct2, auto)
       
  1001   moreover have "?ft \<ge> length gs"
       
  1002     using j
       
  1003     by(auto)
       
  1004   ultimately show "?thesis"
       
  1005     using j
       
  1006     by(simp add: replicate_merge_anywhere le_add_diff_inverse)
       
  1007 qed
       
  1008    
       
  1009 lemma [intro]: "ffp \<le> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
       
  1010 apply(rule_tac min_max.le_supI2)
       
  1011 apply(rule_tac Max_ge, auto)
       
  1012 done
       
  1013 
       
  1014 declare max_less_iff_conj[simp del]
       
  1015 
       
  1016 lemma save_rs:
       
  1017   "\<lbrakk>far = length gs;
       
  1018   ffp \<le> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)));
       
  1019   far < ffp\<rbrakk>
       
  1020 \<Longrightarrow>  {\<lambda>nl. nl = map (\<lambda>i. rec_exec i xs) gs @
       
  1021   rec_exec (Cn (length xs) f gs) xs # 0 \<up> max (Suc (length xs))
       
  1022   (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @ xs @ anything}
       
  1023     mv_box far (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))))
       
  1024     {\<lambda>nl. nl = map (\<lambda>i. rec_exec i xs) gs @
       
  1025                0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - length gs) @
       
  1026                rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything}"
       
  1027 proof -
       
  1028   let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
       
  1029   thm mv_box_correct
       
  1030   let ?lm= " map (\<lambda>i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> ?ft @ xs @ anything"
       
  1031   assume h: "far = length gs" "ffp \<le> ?ft" "far < ffp"
       
  1032   hence "{\<lambda> nl. nl = ?lm} mv_box far ?ft {\<lambda> nl. nl = ?lm[?ft := ?lm!far + ?lm!?ft, far := 0]}"
       
  1033     apply(rule_tac mv_box_correct)
       
  1034     by(case_tac "rec_ci a", auto)  
       
  1035   moreover have "?lm[?ft := ?lm!far + ?lm!?ft, far := 0]
       
  1036     = map (\<lambda>i. rec_exec i xs) gs @
       
  1037     0 \<up> (?ft - length gs) @
       
  1038     rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything"
       
  1039     using h
       
  1040     apply(simp add: nth_append)
       
  1041     using list_update_length[of "map (\<lambda>i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs #
       
  1042        0 \<up> (?ft - Suc (length gs))" 0 "0 \<up> length gs @ xs @ anything" "rec_exec (Cn (length xs) f gs) xs"]
       
  1043     apply(simp add: replicate_merge_anywhere replicate_Suc_iff_anywhere del: replicate_Suc)
       
  1044     by(simp add: list_update_append list_update.simps replicate_Suc_iff_anywhere del: replicate_Suc)
       
  1045   ultimately show "?thesis"
       
  1046     by(simp)
       
  1047 qed
       
  1048 
       
  1049 lemma [simp]: "length (empty_boxes n) = 2*n"
       
  1050 apply(induct n, simp, simp)
       
  1051 done
       
  1052 
       
  1053 lemma empty_one_box_correct:
       
  1054   "{\<lambda>nl. nl = 0 \<up> n @ x # lm} [Dec n 2, Goto 0] {\<lambda>nl. nl = 0 # 0 \<up> n @ lm}"
       
  1055 proof(induct x)
       
  1056   case 0
       
  1057   thus "?case"
       
  1058     by(simp add: abc_Hoare_halt_def, 
       
  1059           rule_tac x = 1 in exI, simp add: abc_steps_l.simps 
       
  1060           abc_step_l.simps abc_fetch.simps abc_lm_v.simps nth_append abc_lm_s.simps
       
  1061           replicate_Suc[THEN sym] exp_suc del: replicate_Suc)
       
  1062 next
       
  1063   case (Suc x)
       
  1064   have "{\<lambda>nl. nl = 0 \<up> n @ x # lm} [Dec n 2, Goto 0] {\<lambda>nl. nl = 0 # 0 \<up> n @ lm}"
       
  1065     by fact
       
  1066   then obtain stp where "abc_steps_l (0, 0 \<up> n @ x # lm) [Dec n 2, Goto 0] stp
       
  1067                       = (Suc (Suc 0), 0 # 0 \<up> n @ lm)"
       
  1068     apply(auto simp: abc_Hoare_halt_def)
       
  1069     by(case_tac "abc_steps_l (0, 0 \<up> n @ x # lm) [Dec n 2, Goto 0] na", simp)
       
  1070   moreover have "abc_steps_l (0, 0\<up>n @ Suc x # lm) [Dec n 2, Goto 0] (Suc (Suc 0)) 
       
  1071         = (0,  0 \<up> n @ x # lm)"
       
  1072     by(auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps abc_lm_v.simps
       
  1073       nth_append abc_lm_s.simps list_update.simps list_update_append)
       
  1074   ultimately have "abc_steps_l (0, 0\<up>n @ Suc x # lm) [Dec n 2, Goto 0] (Suc (Suc 0) + stp) 
       
  1075                 = (Suc (Suc 0), 0 # 0\<up>n @ lm)"
       
  1076     by(simp only: abc_steps_add)
       
  1077   thus "?case"
       
  1078     apply(simp add: abc_Hoare_halt_def)
       
  1079     apply(rule_tac x = "Suc (Suc stp)" in exI, simp)
       
  1080     done
       
  1081 qed
       
  1082 
       
  1083 lemma empty_boxes_correct: 
       
  1084   "length lm \<ge> n \<Longrightarrow>
       
  1085   {\<lambda> nl. nl = lm} empty_boxes n {\<lambda> nl. nl = 0\<up>n @ drop n lm}"
       
  1086 proof(induct n)
       
  1087   case 0
       
  1088   thus "?case"
       
  1089     by(simp add: empty_boxes.simps abc_Hoare_halt_def, 
       
  1090           rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
       
  1091 next
       
  1092   case (Suc n)
       
  1093   have ind: "n \<le> length lm \<Longrightarrow> {\<lambda>nl. nl = lm} empty_boxes n {\<lambda>nl. nl = 0 \<up> n @ drop n lm}" by fact
       
  1094   have h: "Suc n \<le> length lm" by fact
       
  1095   have "{\<lambda>nl. nl = lm} empty_boxes n [+] [Dec n 2, Goto 0] {\<lambda>nl. nl = 0 # 0 \<up> n @ drop (Suc n) lm}"
       
  1096   proof(rule_tac abc_Hoare_plus_halt)
       
  1097     show "{\<lambda>nl. nl = lm} empty_boxes n {\<lambda>nl. nl = 0 \<up> n @ drop n lm}"
       
  1098       using h
       
  1099       by(rule_tac ind, simp)
       
  1100   next
       
  1101     show "{\<lambda>nl. nl = 0 \<up> n @ drop n lm} [Dec n 2, Goto 0] {\<lambda>nl. nl = 0 # 0 \<up> n @ drop (Suc n) lm}"
       
  1102       using empty_one_box_correct[of n "lm ! n" "drop (Suc n) lm"]
       
  1103       using h
       
  1104       by(simp add: nth_drop')
       
  1105   qed
       
  1106   thus "?case"
       
  1107     by(simp add: empty_boxes.simps)
       
  1108 qed
       
  1109 
       
  1110 lemma [simp]: "length gs \<le> ffp \<Longrightarrow>
       
  1111     length gs + (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - length gs) =
       
  1112     max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
       
  1113 apply(rule_tac le_add_diff_inverse)
       
  1114 apply(rule_tac min_max.le_supI2)
       
  1115 apply(simp add: Max_ge_iff)
       
  1116 done
       
  1117 
       
  1118 
       
  1119 lemma clean_paras: 
       
  1120   "ffp \<ge> length gs \<Longrightarrow>
       
  1121   {\<lambda>nl. nl = map (\<lambda>i. rec_exec i xs) gs @
       
  1122   0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - length gs) @
       
  1123   rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything}
       
  1124   empty_boxes (length gs)
       
  1125   {\<lambda>nl. nl = 0 \<up> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @ 
       
  1126   rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything}"
       
  1127 proof-
       
  1128   let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
       
  1129   assume h: "length gs \<le> ffp"
       
  1130   let ?lm = "map (\<lambda>i. rec_exec i xs) gs @ 0 \<up> (?ft - length gs) @
       
  1131     rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything"
       
  1132   have "{\<lambda> nl. nl = ?lm} empty_boxes (length gs) {\<lambda> nl. nl = 0\<up>length gs @ drop (length gs) ?lm}"
       
  1133     by(rule_tac empty_boxes_correct, simp)
       
  1134   moreover have "0\<up>length gs @ drop (length gs) ?lm 
       
  1135            =  0 \<up> ?ft @  rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything"
       
  1136     using h
       
  1137     by(simp add: replicate_merge_anywhere)
       
  1138   ultimately show "?thesis"
       
  1139     by metis
       
  1140 qed
       
  1141  
       
  1142 
       
  1143 lemma restore_rs:
       
  1144   "{\<lambda>nl. nl = 0 \<up> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @ 
       
  1145   rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything}
       
  1146   mv_box (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) (length xs)
       
  1147   {\<lambda>nl. nl = 0 \<up> length xs @
       
  1148   rec_exec (Cn (length xs) f gs) xs #
       
  1149   0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - (length xs)) @
       
  1150   0 \<up> length gs @ xs @ anything}"
       
  1151 proof -
       
  1152   let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
       
  1153   let ?lm = "0\<up>(length xs) @  0\<up>(?ft - (length xs)) @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything"
       
  1154   thm mv_box_correct
       
  1155   have "{\<lambda> nl. nl = ?lm} mv_box ?ft (length xs) {\<lambda> nl. nl = ?lm[length xs := ?lm!?ft + ?lm!(length xs), ?ft := 0]}"
       
  1156     by(rule_tac mv_box_correct, simp, simp)
       
  1157   moreover have "?lm[length xs := ?lm!?ft + ?lm!(length xs), ?ft := 0]
       
  1158                =  0 \<up> length xs @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> (?ft - (length xs)) @ 0 \<up> length gs @ xs @ anything"
       
  1159     apply(auto simp: list_update_append nth_append)
       
  1160     apply(case_tac ?ft, simp_all add: Suc_diff_le list_update.simps)
       
  1161     apply(simp add: exp_suc replicate_Suc[THEN sym] del: replicate_Suc)
       
  1162     done
       
  1163   ultimately show "?thesis"
       
  1164     by(simp add: replicate_merge_anywhere)
       
  1165 qed
       
  1166 
       
  1167 lemma restore_orgin_paras:
       
  1168   "{\<lambda>nl. nl = 0 \<up> length xs @
       
  1169   rec_exec (Cn (length xs) f gs) xs #
       
  1170   0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - length xs) @ 0 \<up> length gs @ xs @ anything}
       
  1171   mv_boxes (Suc (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) 0 (length xs)
       
  1172   {\<lambda>nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> 
       
  1173   (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything}"
       
  1174 proof -
       
  1175   let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
       
  1176   thm mv_boxes_correct2
       
  1177   have "{\<lambda> nl. nl = [] @ 0\<up>(length xs) @ (rec_exec (Cn (length xs) f gs) xs # 0 \<up> (?ft - length xs) @ 0 \<up> length gs) @ xs @ anything}
       
  1178         mv_boxes (Suc ?ft + length gs) 0 (length xs)
       
  1179         {\<lambda> nl. nl = [] @ xs @ (rec_exec (Cn (length xs) f gs) xs # 0 \<up> (?ft - length xs) @ 0 \<up> length gs) @ 0\<up>length xs @ anything}"
       
  1180     by(rule_tac mv_boxes_correct2, auto)
       
  1181   thus "?thesis"
       
  1182     by(simp add: replicate_merge_anywhere)
       
  1183 qed
       
  1184 
       
  1185 lemma compile_cn_correct':
       
  1186   assumes f_ind: 
       
  1187   "\<And> anything r. rec_exec f (map (\<lambda>g. rec_exec g xs) gs) = rec_exec (Cn (length xs) f gs) xs \<Longrightarrow>
       
  1188   {\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ 0 \<up> (ffp - far) @ anything} fap
       
  1189                 {\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> (ffp - Suc far) @ anything}"
       
  1190   and compile: "rec_ci f = (fap, far, ffp)"
       
  1191   and term_f: "terminate f (map (\<lambda>g. rec_exec g xs) gs)"
       
  1192   and g_cond: "\<forall>g\<in>set gs. terminate g xs \<and>
       
  1193   (\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow> 
       
  1194   (\<forall>xc. {\<lambda>nl. nl = xs @ 0 \<up> (xb - xa) @ xc} x {\<lambda>nl. nl = xs @ rec_exec g xs # 0 \<up> (xb - Suc xa) @ xc}))"
       
  1195   shows 
       
  1196   "{\<lambda>nl. nl = xs @ 0 # 0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything}
       
  1197   cn_merge_gs (map rec_ci gs) (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) [+]
       
  1198   (mv_boxes 0 (Suc (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) (length xs) [+]
       
  1199   (mv_boxes (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) 0 (length gs) [+]
       
  1200   (fap [+] (mv_box far (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) [+]
       
  1201   (empty_boxes (length gs) [+]
       
  1202   (mv_box (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) (length xs) [+]
       
  1203   mv_boxes (Suc (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) 0 (length xs)))))))
       
  1204   {\<lambda>nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # 
       
  1205 0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything}"
       
  1206 proof -
       
  1207   let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
       
  1208   let ?A = "cn_merge_gs (map rec_ci gs) ?ft"
       
  1209   let ?B = "mv_boxes 0 (Suc (?ft+length gs)) (length xs)"
       
  1210   let ?C = "mv_boxes ?ft 0 (length gs)"
       
  1211   let ?D = fap
       
  1212   let ?E = "mv_box far ?ft"
       
  1213   let ?F = "empty_boxes (length gs)"
       
  1214   let ?G = "mv_box ?ft (length xs)"
       
  1215   let ?H = "mv_boxes (Suc (?ft + length gs)) 0 (length xs)"
       
  1216   let ?P1 = "\<lambda>nl. nl = xs @ 0 # 0 \<up> (?ft + length gs) @ anything"
       
  1217   let ?S = "\<lambda>nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> (?ft + length gs) @ anything"
       
  1218   let ?Q1 = "\<lambda> nl. nl = xs @ 0\<up>(?ft - length xs) @ map (\<lambda> i. rec_exec i xs) gs @ 0\<up>(Suc (length xs)) @ anything"
       
  1219   show "{?P1} (?A [+] (?B [+] (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H))))))) {?S}"
       
  1220   proof(rule_tac abc_Hoare_plus_halt)
       
  1221     show "{?P1} ?A {?Q1}"
       
  1222       using g_cond
       
  1223       by(rule_tac compile_cn_gs_correct, auto)
       
  1224   next
       
  1225     let ?Q2 = "\<lambda>nl. nl = 0 \<up> ?ft @
       
  1226                     map (\<lambda>i. rec_exec i xs) gs @ 0 # xs @ anything"
       
  1227     show "{?Q1} (?B [+] (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H)))))) {?S}"
       
  1228     proof(rule_tac abc_Hoare_plus_halt)
       
  1229       show "{?Q1} ?B {?Q2}"
       
  1230         by(rule_tac save_paras)
       
  1231     next
       
  1232       let ?Q3 = "\<lambda> nl. nl = map (\<lambda>i. rec_exec i xs) gs @ 0\<up>?ft @ 0 # xs @ anything" 
       
  1233       show "{?Q2} (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H))))) {?S}"
       
  1234       proof(rule_tac abc_Hoare_plus_halt)
       
  1235         have "ffp \<ge> length gs"
       
  1236           using compile term_f
       
  1237           apply(subgoal_tac "length gs = far")
       
  1238           apply(drule_tac footprint_ge, simp)
       
  1239           by(drule_tac param_pattern, auto)          
       
  1240         thus "{?Q2} ?C {?Q3}"
       
  1241           by(erule_tac restore_new_paras)
       
  1242       next
       
  1243         let ?Q4 = "\<lambda> nl. nl = map (\<lambda>i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0\<up>?ft @ xs @ anything"
       
  1244         have a: "far = length gs"
       
  1245           using compile term_f
       
  1246           by(drule_tac param_pattern, auto)
       
  1247         have b:"?ft \<ge> ffp"
       
  1248           by auto
       
  1249         have c: "ffp > far"
       
  1250           using compile
       
  1251           by(erule_tac footprint_ge)
       
  1252         show "{?Q3} (?D [+] (?E [+] (?F [+] (?G [+] ?H)))) {?S}"
       
  1253         proof(rule_tac abc_Hoare_plus_halt)
       
  1254           have "{\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ 0 \<up> (ffp - far) @ 0\<up>(?ft - ffp + far) @ 0 # xs @ anything} fap
       
  1255             {\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ rec_exec (Cn (length xs) f gs) xs # 
       
  1256             0 \<up> (ffp - Suc far) @ 0\<up>(?ft - ffp + far) @ 0 # xs @ anything}"
       
  1257             by(rule_tac f_ind, simp add: rec_exec.simps)
       
  1258           thus "{?Q3} fap {?Q4}"
       
  1259             using a b c
       
  1260             by(simp add: replicate_merge_anywhere,
       
  1261                case_tac "?ft", simp_all add: exp_suc del: replicate_Suc)
       
  1262         next
       
  1263           let ?Q5 = "\<lambda>nl. nl = map (\<lambda>i. rec_exec i xs) gs @
       
  1264                0\<up>(?ft - length gs) @ rec_exec (Cn (length xs) f gs) xs # 0\<up>(length gs)@ xs @ anything"
       
  1265           show "{?Q4} (?E [+] (?F [+] (?G [+] ?H))) {?S}"
       
  1266           proof(rule_tac abc_Hoare_plus_halt)
       
  1267             from a b c show "{?Q4} ?E {?Q5}"
       
  1268               by(erule_tac save_rs, simp_all)
       
  1269           next
       
  1270             let ?Q6 = "\<lambda>nl. nl = 0\<up>?ft @ rec_exec (Cn (length xs) f gs) xs # 0\<up>(length gs)@ xs @ anything"
       
  1271             show "{?Q5} (?F [+] (?G [+] ?H)) {?S}"
       
  1272             proof(rule_tac abc_Hoare_plus_halt)
       
  1273               have "length gs \<le> ffp" using a b c
       
  1274                 by simp
       
  1275               thus "{?Q5} ?F {?Q6}"
       
  1276                 by(erule_tac clean_paras)
       
  1277             next
       
  1278               let ?Q7 = "\<lambda>nl. nl = 0\<up>length xs @ rec_exec (Cn (length xs) f gs) xs # 0\<up>(?ft - (length xs)) @ 0\<up>(length gs)@ xs @ anything"
       
  1279               show "{?Q6} (?G [+] ?H) {?S}"
       
  1280               proof(rule_tac abc_Hoare_plus_halt)
       
  1281                 show "{?Q6} ?G {?Q7}"
       
  1282                   by(rule_tac restore_rs)
       
  1283               next
       
  1284                 show "{?Q7} ?H {?S}"
       
  1285                   by(rule_tac restore_orgin_paras)
       
  1286               qed
       
  1287             qed
       
  1288           qed
       
  1289         qed        
       
  1290       qed
       
  1291     qed
       
  1292   qed
       
  1293 qed
       
  1294 
       
  1295 lemma compile_cn_correct:
       
  1296   assumes termi_f: "terminate f (map (\<lambda>g. rec_exec g xs) gs)"
       
  1297   and f_ind: "\<And>ap arity fp anything.
       
  1298   rec_ci f = (ap, arity, fp)
       
  1299   \<Longrightarrow> {\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ 0 \<up> (fp - arity) @ anything} ap
       
  1300   {\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ rec_exec f (map (\<lambda>g. rec_exec g xs) gs) # 0 \<up> (fp - Suc arity) @ anything}"
       
  1301   and g_cond: 
       
  1302   "\<forall>g\<in>set gs. terminate g xs \<and>
       
  1303   (\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow>   (\<forall>xc. {\<lambda>nl. nl = xs @ 0 \<up> (xb - xa) @ xc} x {\<lambda>nl. nl = xs @ rec_exec g xs # 0 \<up> (xb - Suc xa) @ xc}))"
       
  1304   and compile: "rec_ci (Cn n f gs) = (ap, arity, fp)"
       
  1305   and len: "length xs = n"
       
  1306   shows "{\<lambda>nl. nl = xs @ 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ rec_exec (Cn n f gs) xs # 0 \<up> (fp - Suc arity) @ anything}"
       
  1307 proof(case_tac "rec_ci f")
       
  1308   fix fap far ffp
       
  1309   assume h: "rec_ci f = (fap, far, ffp)"
       
  1310   then have f_newind: "\<And> anything .{\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ 0 \<up> (ffp - far) @ anything} fap
       
  1311     {\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ rec_exec f (map (\<lambda>g. rec_exec g xs) gs) # 0 \<up> (ffp - Suc far) @ anything}"
       
  1312     by(rule_tac f_ind, simp_all)
       
  1313   thus "{\<lambda>nl. nl = xs @ 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ rec_exec (Cn n f gs) xs # 0 \<up> (fp - Suc arity) @ anything}"
       
  1314     using compile len h termi_f g_cond
       
  1315     apply(auto simp: rec_ci.simps abc_comp_commute)
       
  1316     apply(rule_tac compile_cn_correct', simp_all)
       
  1317     done
       
  1318 qed
       
  1319 
       
  1320 lemma [simp]: 
       
  1321   "\<lbrakk>length xs = n; ft = max (n+3) (max fft gft)\<rbrakk> 
       
  1322  \<Longrightarrow> {\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft - n) @ anything} mv_box n ft 
       
  1323        {\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft - n) @ anything}"
       
  1324 using mv_box_correct[of n ft "xs @ 0 # 0 \<up> (ft - n) @ anything"]
       
  1325 by(auto)
       
  1326 
       
  1327 lemma [simp]: "length xs < max (length xs + 3) (max fft gft)"
       
  1328 by auto
       
  1329 
       
  1330 lemma save_init_rs: 
       
  1331   "\<lbrakk>length xs = n; ft = max (n+3) (max fft gft)\<rbrakk> 
       
  1332      \<Longrightarrow>  {\<lambda>nl. nl = xs @ rec_exec f xs # 0 \<up> (ft - n) @ anything} mv_box n (Suc n) 
       
  1333        {\<lambda>nl. nl = xs @ 0 # rec_exec f xs # 0 \<up> (ft - Suc n) @ anything}"
       
  1334 using mv_box_correct[of n "Suc n" "xs @ rec_exec f xs # 0 \<up> (ft - n) @ anything"]
       
  1335 apply(auto simp: list_update_append list_update.simps nth_append split: if_splits)
       
  1336 apply(case_tac "(max (length xs + 3) (max fft gft))", simp_all add: list_update.simps Suc_diff_le)
       
  1337 done
       
  1338 
       
  1339 lemma [simp]: "n + (2::nat) < max (n + 3) (max fft gft)"
       
  1340 by auto
       
  1341 
       
  1342 lemma [simp]: "n < max (n + (3::nat)) (max fft gft)"
       
  1343 by auto
       
  1344 
       
  1345 lemma [simp]:
       
  1346   "length xs = n \<Longrightarrow> 
       
  1347   {\<lambda>nl. nl = xs @ x # 0 \<up> (max (n + (3::nat)) (max fft gft) - n) @ anything} mv_box n (max (n + 3) (max fft gft))
       
  1348   {\<lambda>nl. nl = xs @ 0 \<up> (max (n + 3) (max fft gft) - n) @ x # anything}"
       
  1349 proof -
       
  1350   assume h: "length xs = n"
       
  1351   let ?ft = "max (n+3) (max fft gft)"
       
  1352   let ?lm = "xs @ x # 0\<up>(?ft - Suc n) @ 0 # anything"
       
  1353   have g: "?ft > n + 2"
       
  1354     by simp
       
  1355   thm mv_box_correct
       
  1356   have a: "{\<lambda> nl. nl = ?lm} mv_box n ?ft {\<lambda> nl. nl = ?lm[?ft := ?lm!n + ?lm!?ft, n := 0]}"
       
  1357     using h
       
  1358     by(rule_tac mv_box_correct, auto)
       
  1359   have b:"?lm = xs @ x # 0 \<up> (max (n + 3) (max fft gft) - n) @ anything"
       
  1360     by(case_tac ?ft, simp_all add: Suc_diff_le exp_suc del: replicate_Suc)
       
  1361   have c: "?lm[?ft := ?lm!n + ?lm!?ft, n := 0] = xs @ 0 \<up> (max (n + 3) (max fft gft) - n) @ x # anything"
       
  1362     using h g
       
  1363     apply(auto simp: nth_append list_update_append split: if_splits)
       
  1364     using list_update_append[of "x # 0 \<up> (max (length xs + 3) (max fft gft) - Suc (length xs))" "0 # anything"
       
  1365                                  "max (length xs + 3) (max fft gft) - length xs" "x"]
       
  1366     apply(auto simp: if_splits)
       
  1367     apply(simp add: list_update.simps replicate_Suc[THEN sym] del: replicate_Suc)
       
  1368     done
       
  1369   from a c show "?thesis"
       
  1370     using h
       
  1371     apply(simp)
       
  1372     using b
       
  1373     by simp
       
  1374 qed
       
  1375 
       
  1376 lemma [simp]: "max n (Suc n) < Suc (Suc (max (n + 3) (max fft gft) + length anything - Suc 0))"
       
  1377 by arith    
       
  1378 
       
  1379 lemma [simp]: "Suc n < max (n + 3) (max fft gft)"
       
  1380 by arith
       
  1381 
       
  1382 lemma [simp]:
       
  1383   "length xs = n
       
  1384  \<Longrightarrow> {\<lambda>nl. nl = xs @ rec_exec f xs # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ x # anything} mv_box n (Suc n)
       
  1385     {\<lambda>nl. nl = xs @ 0 # rec_exec f xs # 0 \<up> (max (n + 3) (max fft gft) - Suc (Suc n)) @ x # anything}"
       
  1386 using mv_box_correct[of n "Suc n" "xs @ rec_exec f xs # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ x # anything"]
       
  1387 apply(simp add: nth_append list_update_append list_update.simps)
       
  1388 apply(case_tac "max (n + 3) (max fft gft)", simp_all)
       
  1389 apply(case_tac nat, simp_all add: Suc_diff_le list_update.simps)
       
  1390 done
       
  1391 
       
  1392 lemma abc_append_frist_steps_eq_pre: 
       
  1393   assumes notfinal: "abc_notfinal (abc_steps_l (0, lm)  A n) A"
       
  1394   and notnull: "A \<noteq> []"
       
  1395   shows "abc_steps_l (0, lm) (A @ B) n = abc_steps_l (0, lm) A n"
       
  1396 using notfinal
       
  1397 proof(induct n)
       
  1398   case 0
       
  1399   thus "?case"
       
  1400     by(simp add: abc_steps_l.simps)
       
  1401 next
       
  1402   case (Suc n)
       
  1403   have ind: "abc_notfinal (abc_steps_l (0, lm) A n) A \<Longrightarrow> abc_steps_l (0, lm) (A @ B) n = abc_steps_l (0, lm) A n"
       
  1404     by fact
       
  1405   have h: "abc_notfinal (abc_steps_l (0, lm) A (Suc n)) A" by fact
       
  1406   then have a: "abc_notfinal (abc_steps_l (0, lm) A n) A"
       
  1407     by(simp add: notfinal_Suc)
       
  1408   then have b: "abc_steps_l (0, lm) (A @ B) n = abc_steps_l (0, lm) A n"
       
  1409     using ind by simp
       
  1410   obtain s lm' where c: "abc_steps_l (0, lm) A n = (s, lm')"
       
  1411     by (metis prod.exhaust)
       
  1412   then have d: "s < length A \<and> abc_steps_l (0, lm) (A @ B) n = (s, lm')" 
       
  1413     using a b by simp
       
  1414   thus "?case"
       
  1415     using c
       
  1416     by(simp add: abc_step_red2 abc_fetch.simps abc_step_l.simps nth_append)
       
  1417 qed
       
  1418 
       
  1419 lemma abc_append_first_step_eq_pre: 
       
  1420   "st < length A
       
  1421  \<Longrightarrow> abc_step_l (st, lm) (abc_fetch st (A @ B)) = 
       
  1422     abc_step_l (st, lm) (abc_fetch st A)"
       
  1423 by(simp add: abc_step_l.simps abc_fetch.simps nth_append)
       
  1424 
       
  1425 lemma abc_append_frist_steps_halt_eq': 
       
  1426   assumes final: "abc_steps_l (0, lm) A n = (length A, lm')"
       
  1427     and notnull: "A \<noteq> []"
       
  1428   shows "\<exists> n'. abc_steps_l (0, lm) (A @ B) n' = (length A, lm')"
       
  1429 proof -
       
  1430   have "\<exists> n'. abc_notfinal (abc_steps_l (0, lm) A n') A \<and> 
       
  1431     abc_final (abc_steps_l (0, lm) A (Suc n')) A"
       
  1432     using assms
       
  1433     by(rule_tac n = n in abc_before_final, simp_all)
       
  1434   then obtain na where a:
       
  1435     "abc_notfinal (abc_steps_l (0, lm) A na) A \<and> 
       
  1436             abc_final (abc_steps_l (0, lm) A (Suc na)) A" ..
       
  1437   obtain sa lma where b: "abc_steps_l (0, lm) A na = (sa, lma)"
       
  1438     by (metis prod.exhaust)
       
  1439   then have c: "abc_steps_l (0, lm) (A @ B) na = (sa, lma)"
       
  1440     using a abc_append_frist_steps_eq_pre[of lm A na B] assms 
       
  1441     by simp
       
  1442   have d: "sa < length A" using b a by simp
       
  1443   then have e: "abc_step_l (sa, lma) (abc_fetch sa (A @ B)) = 
       
  1444     abc_step_l (sa, lma) (abc_fetch sa A)"
       
  1445     by(rule_tac abc_append_first_step_eq_pre)
       
  1446   from a have "abc_steps_l (0, lm) A (Suc na) = (length A, lm')"
       
  1447     using final equal_when_halt
       
  1448     by(case_tac "abc_steps_l (0, lm) A (Suc na)" , simp)
       
  1449   then have "abc_steps_l (0, lm) (A @ B) (Suc na) = (length A, lm')"
       
  1450     using a b c e
       
  1451     by(simp add: abc_step_red2)
       
  1452   thus "?thesis"
       
  1453     by blast
       
  1454 qed
       
  1455 
       
  1456 lemma abc_append_frist_steps_halt_eq: 
       
  1457   assumes final: "abc_steps_l (0, lm) A n = (length A, lm')"
       
  1458   shows "\<exists> n'. abc_steps_l (0, lm) (A @ B) n' = (length A, lm')"
       
  1459 using final
       
  1460 apply(case_tac "A = []")
       
  1461 apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps abc_exec_null)
       
  1462 apply(rule_tac abc_append_frist_steps_halt_eq', simp_all)
       
  1463 done
       
  1464 
       
  1465 lemma [simp]: "Suc (Suc (max (length xs + 3) (max fft gft) - Suc (Suc (length xs))))
       
  1466            = max (length xs + 3) (max fft gft) - (length xs)"
       
  1467 by arith
       
  1468 
       
  1469 lemma [simp]: "\<lbrakk>ft = max (n + 3) (max fft gft); length xs = n\<rbrakk> \<Longrightarrow>
       
  1470      {\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything}
       
  1471      [Dec ft (length gap + 7)] 
       
  1472      {\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ y # anything}"
       
  1473 apply(simp add: abc_Hoare_halt_def)
       
  1474 apply(rule_tac x = 1 in exI)
       
  1475 apply(auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_lm_v.simps abc_lm_s.simps list_update_append)
       
  1476 using list_update_length
       
  1477 [of "(x - Suc y) # rec_exec (Pr (length xs) f g) (xs @ [x - Suc y]) #
       
  1478           0 \<up> (max (length xs + 3) (max fft gft) - Suc (Suc (length xs)))" "Suc y" anything y]
       
  1479 apply(simp)
       
  1480 done
       
  1481 
       
  1482 lemma adjust_paras': 
       
  1483   "length xs = n \<Longrightarrow> {\<lambda>nl. nl = xs @ x # y # anything}  [Inc n] [+] [Dec (Suc n) 2, Goto 0]
       
  1484        {\<lambda>nl. nl = xs @ Suc x # 0 # anything}"
       
  1485 proof(rule_tac abc_Hoare_plus_halt)
       
  1486   assume "length xs = n"
       
  1487   thus "{\<lambda>nl. nl = xs @ x # y # anything} [Inc n] {\<lambda> nl. nl = xs @ Suc x # y # anything}"
       
  1488     apply(simp add: abc_Hoare_halt_def)
       
  1489     apply(rule_tac x = 1 in exI, simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps abc_comp.simps
       
  1490                                            abc_lm_v.simps abc_lm_s.simps nth_append list_update_append list_update.simps)
  3036     done
  1491     done
  3037 next
  1492 next
  3038   fix n lm2 lm3 lm4
  1493   assume h: "length xs = n"
  3039   assume h:
  1494   thus "{\<lambda>nl. nl = xs @ Suc x # y # anything} [Dec (Suc n) 2, Goto 0] {\<lambda>nl. nl = xs @ Suc x # 0 # anything}"
  3040     "Suc n \<le> aa - ba" 
  1495   proof(induct y)
  3041     "ba < aa"
  1496     case 0
  3042     "length lm1 = ba"
  1497     thus "?case"
  3043     "length (lm2::nat list) = aa - Suc (ba + n)" 
  1498       apply(simp add: abc_Hoare_halt_def)
  3044     "length (lm3::nat list) = Suc n"
  1499       apply(rule_tac x = 1 in exI, simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps abc_comp.simps
  3045   thus
  1500                                            abc_lm_v.simps abc_lm_s.simps nth_append list_update_append list_update.simps)
  3046     "\<exists>bstp. abc_steps_l (0, lm1 @ butlast lm3 @ 0 # lm2 @ 0\<up>n @ 
       
  3047                                last lm3 # lm4) 
       
  3048            (mv_box (aa + n) (ba + n)) bstp =
       
  3049                  (3, lm1 @ lm3 @ lm2 @ 0 # 0\<up>n @ lm4)"
       
  3050     apply(insert mv_box_ex[of "aa + n" "ba + n" "lm1 @ butlast lm3 @ 
       
  3051                           0 # lm2 @ 0\<up>n @ last lm3 # lm4"], simp)
       
  3052     done
       
  3053 qed
       
  3054 
       
  3055 lemma cn_merge_gs_len: 
       
  3056   "length (cn_merge_gs (map rec_ci gs) pstr) = 
       
  3057       (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs"
       
  3058 apply(induct gs arbitrary: pstr, simp, simp)
       
  3059 apply(case_tac "rec_ci a", simp)
       
  3060 done
       
  3061 
       
  3062 lemma [simp]: "n < pstr \<Longrightarrow>
       
  3063      Suc (pstr + length ys - n) = Suc (pstr + length ys) - n"
       
  3064 by arith
       
  3065 
       
  3066 lemma save_paras':  
       
  3067   "\<lbrakk>length lm = n; pstr > n; a_md > pstr + length ys + n\<rbrakk>
       
  3068   \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(pstr - n) @ ys @
       
  3069                0\<up>(a_md - pstr - length ys) @ suf_lm) 
       
  3070                  (mv_boxes 0 (pstr + Suc (length ys)) n) stp
       
  3071         = (3 * n, 0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
       
  3072 apply(insert mv_boxes_ex[of n "pstr + Suc (length ys)" 0 "[]" "lm" 
       
  3073          "0\<up>(pstr - n) @ ys @ [0]" "0\<up>(a_md - pstr - length ys - n - Suc 0) @ suf_lm"], simp)
       
  3074 apply(erule_tac exE, rule_tac x = stp in exI,
       
  3075                             simp add: exponent_add_iff)
       
  3076 apply(simp only: exponent_cons_iff, simp)
       
  3077 done
       
  3078 
       
  3079 lemma [simp]:
       
  3080  "(max ba (Max (insert ba (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs))))
       
  3081  = (Max (insert ba (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs)))"
       
  3082 apply(rule min_max.sup_absorb2, auto)
       
  3083 done
       
  3084 
       
  3085 lemma [simp]:
       
  3086   "((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs) = 
       
  3087                   (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs)"
       
  3088 apply(induct gs)
       
  3089 apply(simp, simp)
       
  3090 done
       
  3091 
       
  3092 lemma ci_cn_md_def:  
       
  3093   "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
       
  3094   rec_ci f = (a, aa, ba)\<rbrakk>
       
  3095     \<Longrightarrow> a_md = max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) o 
       
  3096   rec_ci) ` set gs))) + Suc (length gs) + n"
       
  3097 apply(simp add: rec_ci.simps, auto)
       
  3098 done
       
  3099 
       
  3100 lemma save_paras_prog_ex:
       
  3101   "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
       
  3102     rec_ci f = (a, aa, ba); 
       
  3103     pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
       
  3104                                     (map rec_ci (gs))))\<rbrakk>
       
  3105     \<Longrightarrow> \<exists>ap bp cp. 
       
  3106       aprog = ap [+] bp [+] cp \<and>
       
  3107       length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
       
  3108               3 * length gs \<and> bp = mv_boxes 0 (pstr + Suc (length gs)) n"
       
  3109 apply(simp add: rec_ci.simps)
       
  3110 apply(rule_tac x = 
       
  3111   "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba 
       
  3112       (((\<lambda>(aprog, p, n). n) o rec_ci) ` set gs))))" in exI,
       
  3113       simp add: cn_merge_gs_len)
       
  3114 apply(rule_tac x = 
       
  3115   "mv_boxes (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
       
  3116    0 (length gs) [+] a [+]mv_box aa (max (Suc n) 
       
  3117    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
       
  3118    empty_boxes (length gs) [+] mv_box (max (Suc n) 
       
  3119   (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
       
  3120    mv_boxes (Suc (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) 
       
  3121    ` set gs))) + length gs)) 0 n" in exI, auto)
       
  3122 apply(simp add: abc_append_commute)
       
  3123 done
       
  3124 
       
  3125 lemma save_paras: 
       
  3126   "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
       
  3127     rs_pos = n;
       
  3128     \<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k);
       
  3129     length ys = length gs;
       
  3130     length lm = n;
       
  3131     rec_ci f = (a, aa, ba);
       
  3132     pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
       
  3133                                           (map rec_ci (gs))))\<rbrakk>
       
  3134   \<Longrightarrow> \<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
       
  3135           3 * length gs, lm @ 0\<up>(pstr - n) @ ys @
       
  3136                  0\<up>(a_md - pstr - length ys) @ suf_lm) aprog stp = 
       
  3137            ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
       
  3138                       3 * length gs + 3 * n, 
       
  3139              0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
       
  3140 proof -
       
  3141   assume h:
       
  3142     "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
       
  3143     "rs_pos = n" 
       
  3144     "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)" 
       
  3145     "length ys = length gs"  
       
  3146     "length lm = n"    
       
  3147     "rec_ci f = (a, aa, ba)"
       
  3148     and g: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
       
  3149                                         (map rec_ci (gs))))"
       
  3150   from h and g have k1: 
       
  3151     "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
       
  3152     length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
       
  3153                 3 *length gs \<and> bp = mv_boxes 0 (pstr + Suc (length ys)) n"
       
  3154     apply(drule_tac save_paras_prog_ex, auto)
       
  3155     done
       
  3156   from h have k2: 
       
  3157     "\<exists> stp. abc_steps_l (0, lm @ 0\<up>(pstr - n) @ ys @ 
       
  3158                          0\<up>(a_md - pstr - length ys) @ suf_lm)
       
  3159          (mv_boxes 0 (pstr + Suc (length ys)) n) stp = 
       
  3160         (3 * n, 0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
       
  3161     apply(rule_tac save_paras', simp, simp_all add: g)
       
  3162     apply(drule_tac a = a and aa = aa and ba = ba in 
       
  3163                                         ci_cn_md_def, simp, simp)
       
  3164     done
       
  3165   from k1 show 
       
  3166     "\<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
       
  3167          3 * length gs, lm @ 0\<up>(pstr - n) @ ys @ 
       
  3168                  0\<up>(a_md - pstr - length ys) @ suf_lm) aprog stp =
       
  3169              ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
       
  3170                3 * length gs + 3 * n, 
       
  3171                 0\<up> pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
       
  3172   proof(erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
       
  3173     fix ap bp apa cp
       
  3174     assume "aprog = ap [+] bp [+] cp \<and> length ap = 
       
  3175             (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs
       
  3176             \<and> bp = mv_boxes 0 (pstr + Suc (length ys)) n"
       
  3177     from this and k2 show "?thesis"
       
  3178       apply(simp)
       
  3179       apply(rule_tac abc_append_exc1, simp, simp, simp)
       
  3180       done
       
  3181   qed
       
  3182 qed
       
  3183  
       
  3184 lemma ci_cn_para_eq:
       
  3185   "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md) \<Longrightarrow> rs_pos = n"
       
  3186 apply(simp add: rec_ci.simps, case_tac "rec_ci f", simp)
       
  3187 done
       
  3188 
       
  3189 lemma calc_gs_prog_ex: 
       
  3190   "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
       
  3191     rec_ci f = (a, aa, ba);
       
  3192     Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
       
  3193                          (map rec_ci (gs)))) = pstr\<rbrakk>
       
  3194    \<Longrightarrow> \<exists>ap bp. aprog = ap [+] bp \<and> 
       
  3195                  ap = cn_merge_gs (map rec_ci gs) pstr"
       
  3196 apply(simp add: rec_ci.simps)
       
  3197 apply(rule_tac x = "mv_boxes 0 (Suc (max (Suc n)  
       
  3198    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n [+]
       
  3199    mv_boxes (max (Suc n) (Max (insert ba 
       
  3200   (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
       
  3201    a [+] mv_box aa (max (Suc n)
       
  3202     (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
       
  3203    empty_boxes (length gs) [+] mv_box (max (Suc n)
       
  3204     (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
       
  3205     mv_boxes (Suc (max (Suc n) (Max 
       
  3206     (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n"
       
  3207    in exI)
       
  3208 apply(auto simp: abc_append_commute)
       
  3209 done
       
  3210 
       
  3211 lemma cn_calc_gs: 
       
  3212   assumes ind: 
       
  3213   "\<And>x aprog a_md rs_pos rs suf_lm lm.
       
  3214   \<lbrakk>x \<in> set gs; 
       
  3215    rec_ci x = (aprog, rs_pos, a_md); 
       
  3216    rec_calc_rel x lm rs\<rbrakk>
       
  3217   \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
       
  3218      (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
       
  3219   and h:  "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"  
       
  3220           "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
       
  3221           "length ys = length gs" 
       
  3222           "length lm = n" 
       
  3223           "rec_ci f = (a, aa, ba)" 
       
  3224           "Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
       
  3225                                (map rec_ci (gs)))) = pstr"
       
  3226   shows  
       
  3227   "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp =
       
  3228   ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs, 
       
  3229    lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md -pstr - length ys) @ suf_lm) "
       
  3230 proof -
       
  3231   from h have k1:
       
  3232     "\<exists> ap bp. aprog = ap [+] bp \<and> ap = 
       
  3233                         cn_merge_gs (map rec_ci gs) pstr"
       
  3234     by(erule_tac calc_gs_prog_ex, auto)
       
  3235   from h have j1: "rs_pos = n"
       
  3236     by(simp add: ci_cn_para_eq)
       
  3237   from h have j2: "a_md \<ge> pstr"
       
  3238     by(drule_tac a = a and aa = aa and ba = ba in 
       
  3239                                 ci_cn_md_def, simp, simp)
       
  3240   from h have j3: "pstr > n"
       
  3241     by(auto)    
       
  3242   from j1 and j2 and j3 and h have k2:
       
  3243     "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) 
       
  3244                          (cn_merge_gs (map rec_ci gs) pstr) stp 
       
  3245     = ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs, 
       
  3246                   lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - pstr - length ys) @ suf_lm)"
       
  3247     apply(simp)
       
  3248     apply(rule_tac cn_merge_gs_ex, rule_tac ind, simp, simp, auto)
       
  3249     apply(drule_tac a = a and aa = aa and ba = ba in 
       
  3250                                  ci_cn_md_def, simp, simp)
       
  3251     apply(rule min_max.le_supI2, auto)
       
  3252     done
       
  3253   from k1 show "?thesis"
       
  3254   proof(erule_tac exE, erule_tac exE, simp)
       
  3255     fix ap bp
       
  3256     from k2 show 
       
  3257       "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm)
       
  3258            (cn_merge_gs (map rec_ci gs) pstr [+] bp) stp =
       
  3259       (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) gs) +
       
  3260          3 * length gs, 
       
  3261          lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - (pstr + length ys)) @ suf_lm)"
       
  3262       apply(insert abc_append_exc1[of 
       
  3263         "lm @ 0\<up>(a_md - rs_pos) @ suf_lm" 
       
  3264         "(cn_merge_gs (map rec_ci gs) pstr)" 
       
  3265         "length (cn_merge_gs (map rec_ci gs) pstr)" 
       
  3266         "lm @ 0\<up>(pstr - n) @ ys @ 0\<up>(a_md - pstr - length ys) @ suf_lm" 0 
       
  3267         "[]" bp], simp add: cn_merge_gs_len)
       
  3268       done      
       
  3269   qed
       
  3270 qed
       
  3271 
       
  3272 lemma reset_new_paras': 
       
  3273   "\<lbrakk>length lm = n; 
       
  3274     pstr > 0; 
       
  3275     a_md \<ge> pstr + length ys + n;
       
  3276      pstr > length ys\<rbrakk> \<Longrightarrow>
       
  3277    \<exists>stp. abc_steps_l (0, 0\<up>pstr @ ys @ 0 # lm @  0\<up>(a_md - Suc (pstr + length ys + n)) @
       
  3278           suf_lm) (mv_boxes pstr 0 (length ys)) stp =
       
  3279   (3 * length ys, ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
       
  3280 apply(insert mv_boxes_ex2[of "length ys" "pstr" 0 "[]"
       
  3281      "0\<up>(pstr - length ys)" "ys" 
       
  3282      "0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm"], 
       
  3283      simp add: exponent_add_iff)
       
  3284 done
       
  3285 
       
  3286 lemma [simp]:  
       
  3287   "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
       
  3288   rec_calc_rel f ys rs; rec_ci f = (a, aa, ba);
       
  3289   pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
       
  3290                (map rec_ci (gs))))\<rbrakk>
       
  3291   \<Longrightarrow> length ys < pstr"
       
  3292 apply(subgoal_tac "length ys = aa", simp)
       
  3293 apply(subgoal_tac "aa < ba \<and> ba \<le> pstr", 
       
  3294       rule basic_trans_rules(22), auto)
       
  3295 apply(rule min_max.le_supI2)
       
  3296 apply(auto)
       
  3297 apply(erule_tac para_pattern, simp)
       
  3298 done
       
  3299 
       
  3300 lemma reset_new_paras_prog_ex: 
       
  3301   "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
       
  3302    rec_ci f = (a, aa, ba);
       
  3303    Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
       
  3304   (map rec_ci (gs)))) = pstr\<rbrakk>
       
  3305   \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
       
  3306   length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
       
  3307            3 *length gs + 3 * n \<and> bp = mv_boxes pstr 0 (length gs)"
       
  3308 apply(simp add: rec_ci.simps)
       
  3309 apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) 
       
  3310           (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
       
  3311           mv_boxes 0 (Suc (max (Suc n) (Max (insert ba 
       
  3312            (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n" in exI, 
       
  3313        simp add: cn_merge_gs_len)
       
  3314 apply(rule_tac x = "a [+]
       
  3315      mv_box aa (max (Suc n) (Max (insert ba 
       
  3316      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
       
  3317      empty_boxes (length gs) [+] mv_box 
       
  3318      (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n
       
  3319       [+] mv_boxes (Suc (max (Suc n) (Max (insert ba 
       
  3320      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI,
       
  3321        auto simp: abc_append_commute)
       
  3322 done
       
  3323 
       
  3324 lemma reset_new_paras:
       
  3325        "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
       
  3326         rs_pos = n;
       
  3327         \<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k);
       
  3328         length ys = length gs;
       
  3329         length lm = n;
       
  3330         length ys = aa;
       
  3331         rec_ci f = (a, aa, ba);
       
  3332         pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
       
  3333                                     (map rec_ci (gs))))\<rbrakk>
       
  3334 \<Longrightarrow> \<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
       
  3335                                                3 * length gs + 3 * n,
       
  3336         0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
       
  3337   ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n,
       
  3338            ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
       
  3339 proof -
       
  3340   assume h:
       
  3341     "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
       
  3342     "rs_pos = n" 
       
  3343     "length ys = aa"
       
  3344     "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
       
  3345     "length ys = length gs"  "length lm = n"    
       
  3346     "rec_ci f = (a, aa, ba)"
       
  3347     and g: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
       
  3348                                          (map rec_ci (gs))))"
       
  3349   from h and g have k1:
       
  3350     "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = 
       
  3351     (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
       
  3352           3 *length gs + 3 * n \<and> bp = mv_boxes pstr 0 (length ys)"
       
  3353     by(drule_tac reset_new_paras_prog_ex, auto)
       
  3354   from h have k2:
       
  3355     "\<exists> stp. abc_steps_l (0, 0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @
       
  3356               suf_lm) (mv_boxes pstr 0 (length ys)) stp = 
       
  3357     (3 * (length ys), 
       
  3358      ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
       
  3359     apply(rule_tac reset_new_paras', simp)
       
  3360     apply(simp add: g)
       
  3361     apply(drule_tac a = a and aa = aa and ba = ba in ci_cn_md_def,
       
  3362       simp, simp add: g, simp)
       
  3363     apply(subgoal_tac "length gs = aa \<and> aa < ba \<and> ba \<le> pstr", arith,
       
  3364           simp add: para_pattern)
       
  3365     apply(insert g, auto intro: min_max.le_supI2)
       
  3366     done
       
  3367   from k1 show 
       
  3368     "\<exists>stp. abc_steps_l ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3
       
  3369     * length gs + 3 * n, 0\<up>pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ 
       
  3370      suf_lm) aprog stp =
       
  3371     ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs +
       
  3372       3 * n, ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
       
  3373   proof(erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
       
  3374     fix ap bp apa cp
       
  3375     assume "aprog = ap [+] bp [+] cp \<and> length ap = 
       
  3376       (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs +
       
  3377                   3 * n \<and> bp = mv_boxes pstr 0 (length ys)"
       
  3378     from this and k2 show "?thesis"
       
  3379       apply(simp)
       
  3380       apply(drule_tac as = 
       
  3381         "(\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 3 * length gs +
       
  3382         3 * n" and ap = ap and cp = cp in abc_append_exc1, auto)
       
  3383       apply(rule_tac x = stp in exI, simp add: h)
       
  3384       using h
       
  3385       apply(simp)
       
  3386       done
       
  3387   qed
       
  3388 qed
       
  3389 
       
  3390 lemma calc_f_prog_ex: 
       
  3391   "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
       
  3392     rec_ci f = (a, aa, ba);
       
  3393     Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
       
  3394                    (map rec_ci (gs)))) = pstr\<rbrakk>
       
  3395    \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
       
  3396   length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
       
  3397                                 6 *length gs + 3 * n \<and> bp = a"
       
  3398 apply(simp add: rec_ci.simps)
       
  3399 apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba
       
  3400      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
       
  3401      mv_boxes 0 (Suc (max (Suc n) (Max (insert ba 
       
  3402             (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n [+]
       
  3403      mv_boxes (max (Suc n) (Max (insert ba 
       
  3404       (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs)" in exI,
       
  3405      simp add: cn_merge_gs_len)
       
  3406 apply(rule_tac x = "mv_box aa (max (Suc n) (Max (insert ba 
       
  3407      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
       
  3408      empty_boxes (length gs) [+] mv_box (max (Suc n) (
       
  3409      Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
       
  3410      mv_boxes (Suc (max (Suc n) (Max (insert ba 
       
  3411      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI,
       
  3412   auto simp: abc_append_commute)
       
  3413 done
       
  3414 
       
  3415 lemma calc_cn_f:
       
  3416   assumes ind: 
       
  3417   "\<And>x ap fp arity r anything args.
       
  3418     \<lbrakk>x = map rec_ci gs; rec_ci f = (ap, arity, fp); rec_calc_rel f args r\<rbrakk>
       
  3419     \<Longrightarrow> \<exists>stp. abc_steps_l (0, args @ 0 \<up> (fp - arity) @ anything) ap stp =
       
  3420     (length ap, args @ [r] @ 0 \<up> (fp - arity - 1) @ anything)"  
       
  3421   and h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
       
  3422   "rec_calc_rel (Cn n f gs) lm rs"
       
  3423   "length ys = length gs"
       
  3424   "rec_calc_rel f ys rs"
       
  3425   "length lm = n"
       
  3426   "rec_ci f = (a, aa, ba)" 
       
  3427   and p: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
       
  3428                                 (map rec_ci (gs))))"
       
  3429   shows "\<exists>stp. abc_steps_l   
       
  3430   ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n,
       
  3431   ys @ 0\<up>pstr @ 0 # lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
       
  3432   ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 
       
  3433                 3 * n + length a,
       
  3434   ys @ rs # 0\<up>pstr @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
       
  3435 proof -
       
  3436   from h have k1: 
       
  3437     "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
       
  3438     length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
       
  3439     6 *length gs + 3 * n \<and> bp = a"
       
  3440     by(drule_tac calc_f_prog_ex, auto)
       
  3441   from h and k1 show "?thesis"
       
  3442   proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
       
  3443     fix ap bp apa cp
       
  3444     assume
       
  3445       "aprog = ap [+] bp [+] cp \<and> 
       
  3446       length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
       
  3447       6 * length gs + 3 * n \<and> bp = a"
       
  3448     from h and this show "?thesis"
       
  3449       apply(simp, rule_tac abc_append_exc1, simp_all)
       
  3450       thm ind
       
  3451       apply(insert ind[of "map rec_ci gs" a aa ba ys rs 
       
  3452         "0\<up>(pstr - ba + length gs) @ 0 # lm @ 
       
  3453         0\<up>(a_md - Suc (pstr + length gs + n)) @ suf_lm"], simp)
       
  3454       apply(subgoal_tac "ba > aa \<and> aa = length gs\<and> pstr \<ge> ba", simp)
       
  3455       apply(simp add: exponent_add_iff)
       
  3456       apply(case_tac pstr, simp add: p)
       
  3457       apply(simp only: exp_suc, simp)
       
  3458       apply(rule conjI, rule ci_ad_ge_paras, simp, rule conjI)
       
  3459       apply(subgoal_tac "length ys = aa", simp,
       
  3460         rule para_pattern, simp, simp)
       
  3461       apply(insert p, simp)
       
  3462       apply(auto intro: min_max.le_supI2)
       
  3463       done
       
  3464   qed
       
  3465 qed
       
  3466 
       
  3467 lemma [simp]: 
       
  3468   "\<lbrakk>pstr + length ys + n \<le> a_md; ys \<noteq> []\<rbrakk> \<Longrightarrow> 
       
  3469                           pstr < a_md + length suf_lm"
       
  3470 apply(case_tac "length ys", simp)
       
  3471 apply(arith)
       
  3472 done
       
  3473 
       
  3474 lemma [simp]: 
       
  3475   "pstr > length ys 
       
  3476   \<Longrightarrow> (ys @ rs # 0\<up>pstr @ lm @
       
  3477   0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) ! pstr = (0::nat)"
       
  3478 apply(simp add: nth_append)
       
  3479 done
       
  3480 
       
  3481 (*
       
  3482 lemma [simp]: "\<lbrakk>length ys < pstr; pstr - length ys = Suc x\<rbrakk>
       
  3483   \<Longrightarrow> pstr - Suc (length ys) = x"
       
  3484 by arith
       
  3485 *)
       
  3486 
       
  3487 lemma [simp]: "pstr > length ys \<Longrightarrow> 
       
  3488       (ys @ rs # 0\<up>pstr @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)
       
  3489                                          [pstr := rs, length ys := 0] =
       
  3490        ys @ 0\<up>(pstr - length ys) @ (rs::nat) # 0\<up>length ys @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm"
       
  3491 apply(auto simp: list_update_append)
       
  3492 apply(case_tac "pstr - length ys",simp_all)
       
  3493 using list_update_length[of 
       
  3494   "0\<up>(pstr - Suc (length ys))" "0" "0\<up>length ys @ lm @ 
       
  3495   0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm" rs]
       
  3496 apply(simp only: exponent_cons_iff exponent_add_iff, simp)
       
  3497 apply(subgoal_tac "pstr - Suc (length ys) = nat", simp, simp)
       
  3498 done
       
  3499 
       
  3500 lemma save_rs': 
       
  3501   "\<lbrakk>pstr > length ys\<rbrakk>
       
  3502   \<Longrightarrow> \<exists>stp. abc_steps_l (0, ys @ rs # 0\<up>pstr @ lm @ 
       
  3503   0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) 
       
  3504   (mv_box (length ys) pstr) stp =
       
  3505   (3, ys @ 0\<up>(pstr - (length ys)) @ rs # 
       
  3506   0\<up>length ys  @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
       
  3507 apply(insert mv_box_ex[of "length ys" pstr 
       
  3508   "ys @ rs # 0\<up>pstr @ lm @ 0\<up>(a_md - Suc(pstr + length ys + n)) @ suf_lm"], 
       
  3509   simp)
       
  3510 done
       
  3511 
       
  3512 
       
  3513 lemma save_rs_prog_ex:
       
  3514   "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
       
  3515   rec_ci f = (a, aa, ba);
       
  3516   Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
       
  3517                         (map rec_ci (gs)))) = pstr\<rbrakk>
       
  3518   \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
       
  3519   length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
       
  3520               6 *length gs + 3 * n + length a
       
  3521   \<and> bp = mv_box aa pstr"
       
  3522 apply(simp add: rec_ci.simps)
       
  3523 apply(rule_tac x =
       
  3524   "cn_merge_gs (map rec_ci gs) (max (Suc n) (Max (insert ba 
       
  3525    (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
       
  3526    [+] mv_boxes 0 (Suc (max (Suc n) (Max (insert ba 
       
  3527    (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n [+]
       
  3528    mv_boxes (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
       
  3529     0 (length gs) [+] a" 
       
  3530   in exI, simp add: cn_merge_gs_len)
       
  3531 apply(rule_tac x = 
       
  3532   "empty_boxes (length gs) [+]
       
  3533    mv_box (max (Suc n) (Max (insert ba 
       
  3534     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
       
  3535    mv_boxes (Suc (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))
       
  3536     + length gs)) 0 n" in exI, 
       
  3537   auto simp: abc_append_commute)
       
  3538 done
       
  3539 
       
  3540 lemma save_rs:  
       
  3541   assumes h: 
       
  3542   "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
       
  3543   "rec_calc_rel (Cn n f gs) lm rs"
       
  3544   "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
       
  3545   "length ys = length gs" 
       
  3546   "rec_calc_rel f ys rs" 
       
  3547   "rec_ci f = (a, aa, ba)"  
       
  3548   "length lm = n"
       
  3549   and pdef: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
       
  3550                                             (map rec_ci (gs))))"
       
  3551   shows "\<exists>stp. abc_steps_l
       
  3552            ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs
       
  3553           + 3 * n + length a, ys @ rs # 0\<up>pstr @ lm @
       
  3554              0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
       
  3555   ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs 
       
  3556   + 3 * n + length a + 3,
       
  3557   ys @ 0\<up>(pstr - length ys) @ rs # 0\<up>length ys @ lm @ 
       
  3558                                0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
       
  3559 proof -
       
  3560   from h and pdef have k1: 
       
  3561     "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
       
  3562     length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
       
  3563     6 *length gs + 3 * n + length a \<and> bp = mv_box (length ys) pstr "
       
  3564     apply(subgoal_tac "length ys = aa")
       
  3565     apply(drule_tac a = a and aa = aa and ba = ba in save_rs_prog_ex, 
       
  3566       simp, simp, simp)
       
  3567     by(rule_tac para_pattern, simp, simp)
       
  3568   from k1 show 
       
  3569     "\<exists>stp. abc_steps_l
       
  3570     ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n
       
  3571     + length a, ys @ rs # 0\<up>pstr @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) 
       
  3572     @ suf_lm) aprog stp =
       
  3573     ((\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 3 * n
       
  3574     + length a + 3, ys @ 0\<up>(pstr - length ys) @ rs # 
       
  3575     0\<up>length ys @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
       
  3576   proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
       
  3577     fix ap bp apa cp
       
  3578     assume "aprog = ap [+] bp [+] cp \<and> length ap = 
       
  3579       (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 6 * length gs + 
       
  3580       3 * n + length a \<and> bp = mv_box (length ys) pstr"
       
  3581     thus"?thesis"
       
  3582       apply(simp, rule_tac abc_append_exc1, simp_all)
       
  3583       apply(rule_tac save_rs', insert h)
       
  3584       apply(subgoal_tac "length gs = aa \<and> pstr \<ge> ba \<and> ba > aa",
       
  3585             arith)
       
  3586       apply(simp add: para_pattern, insert pdef, auto)
       
  3587       apply(rule_tac min_max.le_supI2, simp)
       
  3588       done
       
  3589   qed
       
  3590 qed
       
  3591 
       
  3592 lemma [simp]: "length (empty_boxes n) = 2*n"
       
  3593 apply(induct n, simp, simp)
       
  3594 done
       
  3595 
       
  3596 lemma mv_box_step_ex: "length lm = n \<Longrightarrow> 
       
  3597       \<exists>stp. abc_steps_l (0, lm @ Suc x # suf_lm) [Dec n 2, Goto 0] stp
       
  3598   = (0, lm @ x # suf_lm)"
       
  3599 apply(rule_tac x = "Suc (Suc 0)" in exI, 
       
  3600   simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps 
       
  3601          abc_lm_v.simps abc_lm_s.simps nth_append list_update_append)
       
  3602 done
       
  3603 
       
  3604 lemma mv_box_ex': 
       
  3605   "\<lbrakk>length lm = n\<rbrakk> \<Longrightarrow> 
       
  3606   \<exists> stp. abc_steps_l (0, lm @ x # suf_lm) [Dec n 2, Goto 0] stp =
       
  3607   (Suc (Suc 0), lm @ 0 # suf_lm)"
       
  3608 apply(induct x)
       
  3609 apply(rule_tac x = "Suc 0" in exI, 
       
  3610   simp add: abc_steps_l.simps abc_fetch.simps abc_step_l.simps
       
  3611             abc_lm_v.simps nth_append abc_lm_s.simps, simp)
       
  3612 apply(drule_tac x = x and suf_lm = suf_lm in mv_box_step_ex, 
       
  3613       erule_tac exE, erule_tac exE)
       
  3614 apply(rule_tac x = "stpa + stp" in exI, simp add: abc_steps_add)
       
  3615 done
       
  3616 
       
  3617 lemma [simp]: "drop n lm = a # list \<Longrightarrow> list = drop (Suc n) lm"
       
  3618 apply(induct n arbitrary: lm a list, simp)
       
  3619 apply(case_tac "lm", simp, simp)
       
  3620 done
       
  3621 
       
  3622 lemma empty_boxes_ex: "\<lbrakk>length lm \<ge> n\<rbrakk>
       
  3623      \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm) (empty_boxes n) stp = 
       
  3624                                           (2*n, 0\<up>n @ drop n lm)"
       
  3625 apply(induct n, simp, simp)
       
  3626 apply(rule_tac abc_append_exc2, auto)
       
  3627 apply(case_tac "drop n lm", simp, simp)
       
  3628 proof -
       
  3629   fix n stp a list
       
  3630   assume h: "Suc n \<le> length lm"  "drop n lm = a # list"
       
  3631   thus "\<exists>bstp. abc_steps_l (0, 0\<up>n @ a # list) [Dec n 2, Goto 0] bstp =
       
  3632                        (Suc (Suc 0), 0 # 0\<up>n @ drop (Suc n) lm)"
       
  3633     apply(insert mv_box_ex'[of "0\<up>n" n a list], simp, erule_tac exE)
       
  3634     apply(rule_tac x = stp in exI, simp, simp only: exponent_cons_iff)
       
  3635     apply(simp add:exp_ind del: replicate.simps)
       
  3636     done
       
  3637 qed
       
  3638 
       
  3639 
       
  3640 lemma mv_box_paras_prog_ex:
       
  3641   "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
       
  3642   rec_ci f = (a, aa, ba); 
       
  3643   Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
       
  3644                     (map rec_ci (gs)))) = pstr\<rbrakk>
       
  3645   \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> 
       
  3646   length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
       
  3647   6 *length gs + 3 * n + length a + 3 \<and> bp = empty_boxes (length gs)"
       
  3648 apply(simp add: rec_ci.simps)
       
  3649 apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) 
       
  3650     (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
       
  3651     mv_boxes 0 (Suc (max (Suc n) (Max 
       
  3652      (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) n
       
  3653     [+] mv_boxes (max (Suc n) (Max (insert ba 
       
  3654     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
       
  3655      a [+] mv_box aa (max (Suc n) 
       
  3656    (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))" 
       
  3657     in exI, simp add: cn_merge_gs_len)
       
  3658 apply(rule_tac x = " mv_box (max (Suc n) (Max (insert ba
       
  3659      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
       
  3660      mv_boxes (Suc (max (Suc n) (Max (insert ba 
       
  3661      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI, 
       
  3662   auto simp: abc_append_commute)
       
  3663 done
       
  3664 
       
  3665 lemma mv_box_paras: 
       
  3666  assumes h: 
       
  3667   "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
       
  3668   "rec_calc_rel (Cn n f gs) lm rs" 
       
  3669   "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
       
  3670   "length ys = length gs" 
       
  3671   "rec_calc_rel f ys rs" 
       
  3672   "rec_ci f = (a, aa, ba)" 
       
  3673   and pdef: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
       
  3674                                              (map rec_ci (gs))))"
       
  3675   and starts: "ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
       
  3676                               6 * length gs + 3 * n + length a + 3"
       
  3677   shows "\<exists>stp. abc_steps_l
       
  3678            (ss, ys @ 0\<up>(pstr - length ys) @ rs # 0\<up>length ys 
       
  3679                @ lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
       
  3680    (ss + 2 * length gs, 0\<up>pstr @ rs # 0\<up>length ys  @ lm @ 
       
  3681                                 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
       
  3682 proof -
       
  3683   from h and pdef and starts have k1: 
       
  3684     "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and>
       
  3685     length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
       
  3686                                6 *length gs + 3 * n + length a + 3
       
  3687     \<and> bp = empty_boxes (length ys)"
       
  3688     by(drule_tac mv_box_paras_prog_ex, auto)
       
  3689   from h have j1: "aa < ba"
       
  3690     by(simp add: ci_ad_ge_paras)
       
  3691   from h have j2: "length gs = aa"
       
  3692     by(drule_tac f = f in para_pattern, simp, simp)
       
  3693   from h and pdef have j3: "ba \<le> pstr"
       
  3694     apply simp 
       
  3695     apply(rule_tac min_max.le_supI2, simp)
       
  3696     done
       
  3697   from k1 show "?thesis"
       
  3698   proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
       
  3699     fix ap bp apa cp
       
  3700     assume "aprog = ap [+] bp [+] cp \<and> 
       
  3701       length ap = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
       
  3702       6 * length gs + 3 * n + length a + 3 \<and> 
       
  3703       bp = empty_boxes (length ys)"
       
  3704     thus"?thesis"
       
  3705       apply(simp, rule_tac abc_append_exc1, simp_all add: starts h)
       
  3706       apply(insert empty_boxes_ex[of 
       
  3707         "length gs" "ys @ 0\<up>(pstr - (length gs)) @ rs #
       
  3708         0\<up>length gs @ lm @ 0\<up>(a_md - Suc (pstr + length gs + n)) @ suf_lm"], 
       
  3709         simp add: h)
       
  3710       apply(erule_tac exE, rule_tac x = stp in exI, 
       
  3711         simp add: replicate.simps[THEN sym]
       
  3712         replicate_add[THEN sym] del: replicate.simps)
       
  3713       apply(subgoal_tac "pstr >(length gs)", simp)
       
  3714       apply(subgoal_tac "ba > aa \<and> length gs = aa \<and> pstr \<ge> ba", simp)
       
  3715       apply(simp add: j1 j2 j3)
       
  3716       done     
       
  3717   qed
       
  3718 qed
       
  3719 
       
  3720 lemma restore_rs_prog_ex:
       
  3721   "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md);
       
  3722   rec_ci f = (a, aa, ba);
       
  3723   Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
       
  3724   (map rec_ci (gs)))) = pstr;
       
  3725   ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
       
  3726   8 * length gs + 3 * n + length a + 3\<rbrakk>
       
  3727   \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
       
  3728                                            bp = mv_box pstr n"
       
  3729 apply(simp add: rec_ci.simps)
       
  3730 apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) 
       
  3731       (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+] 
       
  3732       mv_boxes 0 (Suc (max (Suc n) (Max (insert ba (((\<lambda>(aprog, p, n). n)
       
  3733         \<circ> rec_ci) ` set gs))) + length gs)) n [+]
       
  3734      mv_boxes (max (Suc n) (Max (insert ba 
       
  3735       (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
       
  3736      a [+] mv_box aa (max (Suc n)
       
  3737        (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
       
  3738      empty_boxes (length gs)" in exI, simp add: cn_merge_gs_len)
       
  3739 apply(rule_tac x = "mv_boxes (Suc (max (Suc n) 
       
  3740        (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) 
       
  3741         + length gs)) 0 n" 
       
  3742   in exI, auto simp: abc_append_commute)
       
  3743 done
       
  3744 
       
  3745 lemma exp_add: "a\<up>(b+c) = a\<up>b @ a\<up>c"
       
  3746 apply(simp add:replicate_add)
       
  3747 done
       
  3748 
       
  3749 lemma [simp]: "n < pstr \<Longrightarrow> (0\<up>pstr)[n := rs] @ [0::nat] = 0\<up>n @ rs # 0\<up>(pstr - n)"
       
  3750 using list_update_length[of "0\<up>n" "0::nat" "0\<up>(pstr - Suc n)" rs]
       
  3751 apply(simp add: replicate_Suc[THEN sym] exp_add[THEN sym] exp_suc[THEN sym])
       
  3752 done
       
  3753 
       
  3754 lemma restore_rs:
       
  3755   assumes h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
       
  3756   "rec_calc_rel (Cn n f gs) lm rs" 
       
  3757   "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
       
  3758   "length ys = length gs"
       
  3759   "rec_calc_rel f ys rs" 
       
  3760   "rec_ci f = (a, aa, ba)" 
       
  3761   and pdef: "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
       
  3762                                         (map rec_ci (gs))))"
       
  3763   and starts: "ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
       
  3764                               8 * length gs + 3 * n + length a + 3" 
       
  3765   shows "\<exists>stp. abc_steps_l
       
  3766            (ss, 0\<up>pstr @ rs # 0\<up>length ys  @ lm @
       
  3767                     0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm) aprog stp =
       
  3768   (ss + 3, 0\<up>n @ rs # 0\<up>(pstr - n) @ 0\<up>length ys  @ lm @ 
       
  3769                                    0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)"
       
  3770 proof -
       
  3771  from h and pdef and starts have k1:
       
  3772    "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
       
  3773                                             bp = mv_box pstr n"
       
  3774    by(drule_tac restore_rs_prog_ex, auto)
       
  3775  from k1 show "?thesis"
       
  3776  proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
       
  3777    fix ap bp apa cp
       
  3778    assume "aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
       
  3779                                  bp = mv_box pstr n"
       
  3780    thus"?thesis"
       
  3781      apply(simp, rule_tac abc_append_exc1, simp_all add: starts h)
       
  3782      apply(insert mv_box_ex[of pstr n "0\<up>pstr @ rs # 0\<up>length gs @
       
  3783                      lm @ 0\<up>(a_md - Suc (pstr + length gs + n)) @ suf_lm"], simp)
       
  3784      apply(subgoal_tac "pstr > n", simp)
       
  3785      apply(erule_tac exE, rule_tac x = stp in exI, 
       
  3786                          simp add: nth_append list_update_append)
       
  3787      apply(simp add: pdef)
       
  3788      done
       
  3789   qed
       
  3790 qed
       
  3791 
       
  3792 lemma [simp]:"xs \<noteq> [] \<Longrightarrow> length xs \<ge> Suc 0"
       
  3793 by(case_tac xs, auto)
       
  3794 
       
  3795 lemma  [simp]: "n < max (Suc n) (max ba (Max (((\<lambda>(aprog, p, n). n) o 
       
  3796                                                   rec_ci) ` set gs)))"
       
  3797 by(simp)
       
  3798 
       
  3799 lemma restore_paras_prog_ex: 
       
  3800   "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
       
  3801   rec_ci f = (a, aa, ba);
       
  3802   Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n)
       
  3803                           (map rec_ci (gs)))) = pstr;
       
  3804   ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
       
  3805                          8 * length gs + 3 * n + length a + 6\<rbrakk>
       
  3806   \<Longrightarrow> \<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
       
  3807                       bp = mv_boxes (pstr + Suc (length gs)) (0::nat) n"
       
  3808 apply(simp add: rec_ci.simps)
       
  3809 apply(rule_tac x = "cn_merge_gs (map rec_ci gs) (max (Suc n) 
       
  3810       (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))
       
  3811       [+] mv_boxes 0 (Suc (max (Suc n) 
       
  3812        (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) 
       
  3813      + length gs)) n [+] mv_boxes (max (Suc n) 
       
  3814     (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
       
  3815      a [+] mv_box aa (max (Suc n) 
       
  3816       (Max (insert ba (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
       
  3817      empty_boxes (length gs) [+]
       
  3818      mv_box (max (Suc n) (Max (insert ba 
       
  3819      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n" in exI, simp add: cn_merge_gs_len)
       
  3820 apply(rule_tac x = "[]" in exI, auto simp: abc_append_commute)
       
  3821 done
       
  3822 
       
  3823 lemma restore_paras: 
       
  3824   assumes h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
       
  3825   "rec_calc_rel (Cn n f gs) lm rs" 
       
  3826   "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)"
       
  3827   "length ys = length gs"
       
  3828   "rec_calc_rel f ys rs" 
       
  3829   "rec_ci f = (a, aa, ba)"
       
  3830   and pdef: 
       
  3831   "pstr = Max (set (Suc n # ba # map (\<lambda>(aprog, p, n). n) 
       
  3832                          (map rec_ci (gs))))"
       
  3833   and starts: "ss = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) + 
       
  3834                               8 * length gs + 3 * n + length a + 6" 
       
  3835   shows "\<exists>stp. abc_steps_l (ss, 0\<up>n @ rs # 0\<up>(pstr - n+ length ys) @
       
  3836                          lm @ 0\<up>(a_md - Suc (pstr + length ys + n)) @ suf_lm)
       
  3837   aprog stp = (ss + 3 * n, lm @ rs # 0\<up>(a_md - Suc n) @ suf_lm)"
       
  3838 proof -
       
  3839   from h and pdef and starts have k1:
       
  3840     "\<exists> ap bp cp. aprog = ap [+] bp [+] cp \<and> length ap = ss \<and>
       
  3841                      bp = mv_boxes (pstr + Suc (length gs)) (0::nat) n"
       
  3842     by(drule_tac restore_paras_prog_ex, auto)
       
  3843   from k1 show "?thesis"
       
  3844   proof (erule_tac exE, erule_tac exE, erule_tac exE, erule_tac exE)
       
  3845     fix ap bp apa cp
       
  3846     assume "aprog = ap [+] bp [+] cp \<and> length ap = ss \<and> 
       
  3847                               bp = mv_boxes (pstr + Suc (length gs)) 0 n"
       
  3848     thus"?thesis"
       
  3849       apply(simp, rule_tac abc_append_exc1, simp_all add: starts h)
       
  3850       apply(insert mv_boxes_ex2[of n "pstr + Suc (length gs)" 0 "[]" 
       
  3851         "rs # 0\<up>(pstr - n + length gs)" "lm" 
       
  3852         "0\<up>(a_md - Suc (pstr + length gs + n)) @ suf_lm"], simp)
       
  3853       apply(subgoal_tac "pstr > n \<and> 
       
  3854         a_md > pstr + length gs + n \<and> length lm = n" , simp add: exponent_add_iff h)
       
  3855       using h pdef
       
  3856       apply(simp)     
       
  3857       apply(frule_tac a = a and 
       
  3858         aa = aa and ba = ba in ci_cn_md_def, simp, simp)
       
  3859       apply(subgoal_tac "length lm = rs_pos",
       
  3860         simp add: ci_cn_para_eq, erule_tac para_pattern, simp)
       
  3861       done
       
  3862   qed
       
  3863 qed
       
  3864 
       
  3865 lemma ci_cn_length:
       
  3866   "\<lbrakk>rec_ci (Cn n f gs) = (aprog, rs_pos, a_md); 
       
  3867   rec_calc_rel (Cn n f gs) lm rs;
       
  3868   rec_ci f = (a, aa, ba)\<rbrakk>
       
  3869   \<Longrightarrow> length aprog = (\<Sum>(ap, pos, n)\<leftarrow>map rec_ci gs. length ap) +
       
  3870                              8 * length gs + 6 * n + length a + 6"
       
  3871 apply(simp add: rec_ci.simps, auto simp: cn_merge_gs_len)
       
  3872 done
       
  3873 
       
  3874 lemma  cn_case: 
       
  3875   assumes ind1:
       
  3876   "\<And>x aprog a_md rs_pos rs suf_lm lm.
       
  3877   \<lbrakk>x \<in> set gs;
       
  3878   rec_ci x = (aprog, rs_pos, a_md);
       
  3879   rec_calc_rel x lm rs\<rbrakk>
       
  3880   \<Longrightarrow> \<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
       
  3881                (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
       
  3882   and ind2: 
       
  3883   "\<And>x ap fp arity r anything args.
       
  3884     \<lbrakk>x = map rec_ci gs; rec_ci f = (ap, arity, fp); rec_calc_rel f args r\<rbrakk>
       
  3885     \<Longrightarrow> \<exists>stp. abc_steps_l (0, args @ 0 \<up> (fp - arity) @ anything) ap stp =
       
  3886     (length ap, args @ [r] @ 0 \<up> (fp - arity - 1) @ anything)"
       
  3887   and h: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
       
  3888          "rec_calc_rel (Cn n f gs) lm rs"
       
  3889   shows "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp 
       
  3890   = (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
       
  3891 apply(insert h, case_tac "rec_ci f",  rule_tac calc_cn_reverse, simp)
       
  3892 proof -
       
  3893   fix a b c ys
       
  3894   let ?pstr = "Max (set (Suc n # c # (map (\<lambda>(aprog, p, n). n) 
       
  3895                                          (map rec_ci (gs)))))"  
       
  3896   let ?gs_len = "listsum (map (\<lambda> (ap, pos, n). length ap) 
       
  3897                                                 (map rec_ci (gs)))"
       
  3898   assume g: "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)"
       
  3899     "rec_calc_rel (Cn n f gs) lm rs"
       
  3900     "\<forall>k<length gs. rec_calc_rel (gs ! k) lm (ys ! k)" 
       
  3901     "length ys = length gs" 
       
  3902     "rec_calc_rel f ys rs"
       
  3903     "n = length lm"
       
  3904     "rec_ci f = (a, b, c)"  
       
  3905   hence k1:
       
  3906     "\<exists> stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
       
  3907     (?gs_len + 3 * length gs, lm @ 0\<up>(?pstr - n) @ ys @
       
  3908                                0\<up>(a_md - ?pstr - length ys) @ suf_lm)"	
       
  3909     apply(rule_tac a = a and aa = b and ba = c in cn_calc_gs)
       
  3910     apply(rule_tac ind1, auto)
       
  3911     done  
       
  3912   from g have k2: 
       
  3913     "\<exists> stp. abc_steps_l (?gs_len + 3 * length gs,  lm @ 
       
  3914         0\<up>(?pstr - n) @ ys @ 0\<up>(a_md - ?pstr - length ys) @ suf_lm) aprog stp = 
       
  3915     (?gs_len + 3 * length gs + 3 * n, 0\<up>?pstr @ ys @ 0 # lm @ 
       
  3916                               0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm)"
       
  3917     apply(erule_tac ba = c in save_paras, auto intro: ci_cn_para_eq)
       
  3918     done
       
  3919   from g have k3: 
       
  3920     "\<exists> stp. abc_steps_l (?gs_len + 3 * length gs + 3 * n,
       
  3921     0\<up>?pstr @ ys @ 0 # lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm) aprog stp =
       
  3922     (?gs_len + 6 * length gs + 3 * n,  
       
  3923            ys @ 0\<up>?pstr @ 0 # lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm)"
       
  3924     apply(erule_tac ba = c in reset_new_paras, 
       
  3925           auto intro: ci_cn_para_eq)
       
  3926     using para_pattern[of f a b c ys rs]
       
  3927     apply(simp)
       
  3928     done
       
  3929   from g have k4: 
       
  3930     "\<exists>stp. abc_steps_l  (?gs_len + 6 * length gs + 3 * n,  
       
  3931     ys @ 0\<up>?pstr @ 0 # lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm) aprog stp =
       
  3932     (?gs_len + 6 * length gs + 3 * n + length a, 
       
  3933    ys @ rs # 0\<up>?pstr  @ lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm)"
       
  3934     apply(rule_tac ba = c in calc_cn_f, rule_tac ind2, auto)
       
  3935     done
       
  3936   from g h have k5:
       
  3937     "\<exists> stp. abc_steps_l (?gs_len + 6 * length gs + 3 * n + length a,
       
  3938     ys @ rs # 0\<up>?pstr @ lm @ 0\<up>(a_md - Suc (?pstr + length ys + n)) @ suf_lm)
       
  3939     aprog stp =
       
  3940     (?gs_len + 6 * length gs + 3 * n + length a + 3,
       
  3941     ys @ 0\<up>(?pstr - length ys) @ rs # 0\<up>length ys @ lm @ 
       
  3942     0\<up>(a_md  - Suc (?pstr + length ys + n)) @ suf_lm)"
       
  3943     apply(rule_tac save_rs, auto simp: h)
       
  3944     done
       
  3945   from g have k6: 
       
  3946     "\<exists> stp. abc_steps_l (?gs_len + 6 * length gs + 3 * n + 
       
  3947     length a + 3, ys @ 0\<up>(?pstr - length ys) @ rs # 0\<up>length ys @ lm @ 
       
  3948     0\<up>(a_md  - Suc (?pstr + length ys + n)) @ suf_lm) 
       
  3949     aprog stp =
       
  3950     (?gs_len + 8 * length gs + 3 *n + length a + 3,
       
  3951     0\<up>?pstr @ rs # 0\<up>length ys @ lm @ 
       
  3952                         0\<up>(a_md -Suc (?pstr + length ys + n)) @ suf_lm)"
       
  3953     apply(drule_tac suf_lm = suf_lm in mv_box_paras, auto)
       
  3954     apply(rule_tac x = stp in exI, simp)
       
  3955     done
       
  3956   from g have k7: 
       
  3957     "\<exists> stp. abc_steps_l (?gs_len + 8 * length gs + 3 *n + 
       
  3958     length a + 3, 0\<up>?pstr  @ rs # 0\<up>length ys @ lm @ 
       
  3959     0\<up>(a_md -Suc (?pstr + length ys + n)) @ suf_lm) aprog stp =
       
  3960     (?gs_len + 8 * length gs + 3 * n + length a + 6, 
       
  3961     0\<up>n @ rs # 0\<up>(?pstr  - n) @ 0\<up>length ys @ lm @
       
  3962                         0\<up>(a_md -Suc (?pstr + length ys + n)) @ suf_lm)"
       
  3963     apply(drule_tac suf_lm = suf_lm in restore_rs, auto)
       
  3964     apply(rule_tac x = stp in exI, simp)
       
  3965     done
       
  3966   from g have k8: "\<exists> stp. abc_steps_l (?gs_len + 8 * length gs + 
       
  3967     3 * n + length a + 6,
       
  3968     0\<up>n @ rs # 0\<up>(?pstr  - n) @ 0\<up>length ys @ lm @
       
  3969                       0\<up>(a_md -Suc (?pstr + length ys + n)) @ suf_lm) aprog stp =
       
  3970     (?gs_len + 8 * length gs + 6 * n + length a + 6,
       
  3971                            lm @ rs # 0\<up>(a_md - Suc n) @ suf_lm)"
       
  3972     apply(drule_tac suf_lm = suf_lm in restore_paras, auto)
       
  3973     apply(simp add: exponent_add_iff)
       
  3974     apply(rule_tac x = stp in exI, simp)
       
  3975     done
       
  3976   from g have j1: 
       
  3977     "length aprog = ?gs_len + 8 * length gs + 6 * n + length a + 6"
       
  3978     by(drule_tac a = a and aa = b and ba = c in ci_cn_length,
       
  3979       simp, simp, simp)
       
  3980   from g have j2: "rs_pos = n"
       
  3981     by(simp add: ci_cn_para_eq)
       
  3982   from k1 and k2 and k3 and k4 and k5 and k6 and k7 and k8
       
  3983     and j1 and j2 show 
       
  3984     "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp = 
       
  3985     (length aprog, lm @ [rs] @ 0\<up>(a_md - rs_pos - 1) @ suf_lm)"
       
  3986     apply(auto)
       
  3987     apply(rule_tac x = "stp + stpa + stpb + stpc +
       
  3988       stpd + stpe + stpf + stpg" in exI, simp add: abc_steps_add)
       
  3989     done
       
  3990 qed
       
  3991 
       
  3992 text {*
       
  3993   Correctness of the complier (terminate case), which says if the execution of 
       
  3994   a recursive function @{text "recf"} terminates and gives result, then 
       
  3995   the Abacus program compiled from @{text "recf"} termintes and gives the same result.
       
  3996   Additionally, to facilitate induction proof, we append @{text "anything"} to the
       
  3997   end of Abacus memory.
       
  3998 *}
       
  3999 
       
  4000 lemma recursive_compile_correct:
       
  4001   "\<lbrakk>rec_ci recf = (ap, arity, fp);
       
  4002     rec_calc_rel recf args r\<rbrakk>
       
  4003   \<Longrightarrow> (\<exists> stp. (abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp) = 
       
  4004               (length ap, args@[r]@0\<up>(fp - arity - 1) @ anything))"
       
  4005 apply(induct arbitrary: ap fp arity r anything args
       
  4006   rule: rec_ci.induct)
       
  4007 prefer 5
       
  4008 proof(case_tac "rec_ci g", case_tac "rec_ci f", simp)
       
  4009   fix n f g ap fp arity r anything args  a b c aa ba ca
       
  4010   assume f_ind:
       
  4011     "\<And>ap fp arity r anything args.
       
  4012     \<lbrakk>aa = ap \<and> ba = arity \<and> ca = fp; rec_calc_rel f args r\<rbrakk> \<Longrightarrow> 
       
  4013     \<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
       
  4014     (length ap, args @ r # 0\<up>(fp - Suc arity) @ anything)"
       
  4015     and g_ind:
       
  4016     "\<And>x xa y xb ya ap fp arity r anything args.
       
  4017     \<lbrakk>x = (aa, ba, ca); xa = aa \<and> y = (ba, ca); xb = ba \<and> ya = ca; 
       
  4018     a = ap \<and> b = arity \<and> c = fp; rec_calc_rel g args r\<rbrakk>
       
  4019     \<Longrightarrow> \<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
       
  4020     (length ap, args @ r # 0\<up>(fp - Suc arity) @ anything)"
       
  4021     and h: "rec_ci (Pr n f g) = (ap, arity, fp)" 
       
  4022     "rec_calc_rel (Pr n f g) args r" 
       
  4023     "rec_ci g = (a, b, c)" 
       
  4024     "rec_ci f = (aa, ba, ca)"
       
  4025   from h have nf_ind: 
       
  4026     "\<And> args r anything. rec_calc_rel f args r \<Longrightarrow> 
       
  4027     \<exists>stp. abc_steps_l (0, args @ 0\<up>(ca - ba) @ anything) aa stp = 
       
  4028     (length aa, args @ r # 0\<up>(ca - Suc ba) @ anything)"
       
  4029     and ng_ind: 
       
  4030     "\<And> args r anything. rec_calc_rel g args r \<Longrightarrow> 
       
  4031     \<exists>stp. abc_steps_l (0, args @ 0\<up>(c - b) @ anything) a stp = 
       
  4032          (length a, args @ r # 0\<up>(c - Suc b)  @ anything)"
       
  4033     apply(insert f_ind[of aa ba ca], simp)
       
  4034     apply(insert g_ind[of "(aa, ba, ca)" aa "(ba, ca)" ba ca a b c],
       
  4035       simp)
       
  4036     done
       
  4037   from nf_ind and ng_ind and h show 
       
  4038     "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp = 
       
  4039     (length ap, args @ r # 0\<up>(fp - Suc arity) @ anything)"
       
  4040     apply(auto intro: nf_ind ng_ind pr_case)
       
  4041     done
       
  4042 next
       
  4043   fix ap fp arity r anything args
       
  4044   assume h:
       
  4045     "rec_ci z = (ap, arity, fp)" "rec_calc_rel z args r"
       
  4046   thus "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
       
  4047     (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
       
  4048     by (rule_tac z_case)    
       
  4049 next
       
  4050   fix ap fp arity r anything args
       
  4051   assume h: 
       
  4052     "rec_ci s = (ap, arity, fp)" 
       
  4053     "rec_calc_rel s args r"
       
  4054   thus 
       
  4055     "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
       
  4056     (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
       
  4057     by(erule_tac s_case, simp)
       
  4058 next
       
  4059   fix m n ap fp arity r anything args
       
  4060   assume h: "rec_ci (id m n) = (ap, arity, fp)" 
       
  4061     "rec_calc_rel (id m n) args r"
       
  4062   thus "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp 
       
  4063     = (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
       
  4064     by(erule_tac id_case)
       
  4065 next
       
  4066   fix n f gs ap fp arity r anything args
       
  4067   assume ind1: "\<And>x ap fp arity r anything args.
       
  4068     \<lbrakk>x \<in> set gs; 
       
  4069     rec_ci x = (ap, arity, fp); 
       
  4070     rec_calc_rel x args r\<rbrakk>
       
  4071     \<Longrightarrow> \<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
       
  4072     (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
       
  4073   and ind2: 
       
  4074     "\<And>x ap fp arity r anything args.
       
  4075     \<lbrakk>x = map rec_ci gs; rec_ci f = (ap, arity, fp); rec_calc_rel f args r\<rbrakk>
       
  4076     \<Longrightarrow> \<exists>stp. abc_steps_l (0, args @ 0 \<up> (fp - arity) @ anything) ap stp =
       
  4077     (length ap, args @ [r] @ 0 \<up> (fp - arity - 1) @ anything)"
       
  4078   and h: "rec_ci (Cn n f gs) = (ap, arity, fp)" 
       
  4079     "rec_calc_rel (Cn n f gs) args r"
       
  4080   from h show
       
  4081     "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) 
       
  4082        ap stp = (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
       
  4083     apply(rule_tac cn_case)
       
  4084     apply(rule_tac ind1, simp, simp, simp)
       
  4085     apply(rule_tac ind2, auto)    
       
  4086     done
       
  4087 next
       
  4088   fix n f ap fp arity r anything args
       
  4089   assume ind:
       
  4090     "\<And>ap fp arity r anything args.
       
  4091     \<lbrakk>rec_ci f = (ap, arity, fp); rec_calc_rel f args r\<rbrakk> \<Longrightarrow> 
       
  4092     \<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp =
       
  4093     (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
       
  4094   and h: "rec_ci (Mn n f) = (ap, arity, fp)" 
       
  4095     "rec_calc_rel (Mn n f) args r"
       
  4096   from h show 
       
  4097     "\<exists>stp. abc_steps_l (0, args @ 0\<up>(fp - arity) @ anything) ap stp = 
       
  4098               (length ap, args @ [r] @ 0\<up>(fp - arity - 1) @ anything)"
       
  4099     apply(rule_tac mn_case, rule_tac ind, auto)
       
  4100     done    
       
  4101 qed
       
  4102 
       
  4103 lemma abc_append_uhalt1:
       
  4104   "\<lbrakk>\<forall> stp. (\<lambda> (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp);
       
  4105     p = ap [+] bp [+] cp\<rbrakk>
       
  4106   \<Longrightarrow> \<forall> stp. (\<lambda> (ss, e). ss < length p) 
       
  4107                      (abc_steps_l (length ap, lm) p stp)"
       
  4108 apply(auto)
       
  4109 apply(erule_tac x = stp in allE, auto)
       
  4110 apply(frule_tac ap = ap and cp = cp in abc_append_state_in_exc, auto)
       
  4111 done
       
  4112 
       
  4113 
       
  4114 lemma abc_append_unhalt2:
       
  4115   "\<lbrakk>abc_steps_l (0, am) ap stp = (length ap, lm); bp \<noteq> [];
       
  4116   \<forall> stp. (\<lambda> (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp);
       
  4117   p = ap [+] bp [+] cp\<rbrakk>
       
  4118   \<Longrightarrow> \<forall> stp. (\<lambda> (ss, e). ss < length p) (abc_steps_l (0, am) p stp)"
       
  4119 proof -
       
  4120   assume h: 
       
  4121     "abc_steps_l (0, am) ap stp = (length ap, lm)" 
       
  4122     "bp \<noteq> []"
       
  4123     "\<forall> stp. (\<lambda> (ss, e). ss < length bp) (abc_steps_l (0, lm) bp stp)"
       
  4124     "p = ap [+] bp [+] cp"
       
  4125   have "\<exists> stp. (abc_steps_l (0, am) p stp) = (length ap, lm)"
       
  4126     using h
       
  4127     apply(simp add: abc_append.simps)
       
  4128     apply(rule_tac abc_add_exc1, auto)
       
  4129     done
       
  4130   from this obtain stpa where g1: 
       
  4131     "(abc_steps_l (0, am) p stpa) = (length ap, lm)" ..
       
  4132   moreover have g2: "\<forall> stp. (\<lambda> (ss, e). ss < length p) 
       
  4133                           (abc_steps_l (length ap, lm) p stp)"
       
  4134     using h
       
  4135     apply(erule_tac abc_append_uhalt1, simp)
       
  4136     done
       
  4137   moreover from g1 and g2 have
       
  4138     "\<forall> stp. (\<lambda> (ss, e). ss < length p) 
       
  4139                     (abc_steps_l (0, am) p (stpa + stp))"
       
  4140     apply(simp add: abc_steps_add)
       
  4141     done
       
  4142   thus "\<forall> stp. (\<lambda> (ss, e). ss < length p) 
       
  4143                            (abc_steps_l (0, am) p stp)"
       
  4144     apply(rule_tac allI, auto)
       
  4145     apply(case_tac "stp \<ge>  stpa")
       
  4146     apply(erule_tac x = "stp - stpa" in allE, simp)
       
  4147   proof - 	
       
  4148     fix stp a b
       
  4149     assume g3:  "abc_steps_l (0, am) p stp = (a, b)" 
       
  4150                 "\<not> stpa \<le> stp"
       
  4151     thus "a < length p"
       
  4152       using g1 h
       
  4153       apply(case_tac "a < length p", simp, simp)
       
  4154       apply(subgoal_tac "\<exists> d. stpa = stp + d")
       
  4155       using  abc_state_keep[of p a b "stpa - stp"]
       
  4156       apply(erule_tac exE, simp add: abc_steps_add)
       
  4157       apply(rule_tac x = "stpa - stp" in exI, simp)
       
  4158       done
       
  4159   qed
       
  4160 qed
       
  4161 
       
  4162 text {*
       
  4163   Correctness of the complier (non-terminating case for Mn). There are many cases when a 
       
  4164   recursive function does not terminate. For the purpose of Uiversal Turing Machine, we only 
       
  4165   need to prove the case for @{text "Mn"} and @{text "Cn"}.
       
  4166   This lemma is for @{text "Mn"}. For @{text "Mn n f"}, this lemma describes what 
       
  4167   happens when @{text "f"} always terminates but always does not return zero, so that
       
  4168   @{text "Mn"} has to loop forever.
       
  4169   *}
       
  4170 
       
  4171 lemma Mn_unhalt:
       
  4172   assumes mn_rf: "rf = Mn n f"
       
  4173   and compiled_mnrf: "rec_ci rf = (aprog, rs_pos, a_md)"
       
  4174   and compiled_f: "rec_ci f = (aprog', rs_pos', a_md')"
       
  4175   and args: "length lm = n"
       
  4176   and unhalt_condition: "\<forall> y. (\<exists> rs. rec_calc_rel f (lm @ [y]) rs \<and> rs \<noteq> 0)"
       
  4177   shows "\<forall> stp. case abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm)
       
  4178                aprog stp of (ss, e) \<Rightarrow> ss < length aprog"
       
  4179   using mn_rf compiled_mnrf compiled_f args unhalt_condition
       
  4180 proof(rule_tac allI)
       
  4181   fix stp
       
  4182   assume h: "rf = Mn n f" 
       
  4183             "rec_ci rf = (aprog, rs_pos, a_md)"
       
  4184             "rec_ci f = (aprog', rs_pos', a_md')" 
       
  4185             "\<forall>y. \<exists>rs. rec_calc_rel f (lm @ [y]) rs \<and> rs \<noteq> 0" "length lm = n"
       
  4186   have "\<exists>stpa \<ge> stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm) aprog stpa 
       
  4187          = (0, lm @ stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  4188   proof(induct stp, auto)
       
  4189     show "\<exists>stpa. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm) 
       
  4190           aprog stpa = (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  4191       apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
       
  4192       done
  1501       done
  4193   next
  1502   next
  4194     fix stp stpa
  1503     case (Suc y)
  4195     assume g1: "stp \<le> stpa"
  1504     have "length xs = n \<Longrightarrow> 
  4196       and g2: "abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
  1505       {\<lambda>nl. nl = xs @ Suc x # y # anything} [Dec (Suc n) 2, Goto 0] {\<lambda>nl. nl = xs @ Suc x # 0 # anything}" by fact
  4197                             aprog stpa
  1506     then obtain stp where 
  4198                = (0, lm @ stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
  1507       "abc_steps_l (0, xs @ Suc x # y # anything) [Dec (Suc n) 2, Goto 0] stp = (2, xs @ Suc x # 0 # anything)"
  4199     have "\<exists>rs. rec_calc_rel f (lm @ [stp]) rs \<and> rs \<noteq> 0"
       
  4200       using h
  1508       using h
  4201       apply(erule_tac x = stp in allE, simp)
  1509       apply(auto simp: abc_Hoare_halt_def)
  4202       done
  1510       by(case_tac "abc_steps_l (0, xs @ Suc x # y # anything) [Dec (Suc (length xs)) 2, Goto 0] n",   
  4203     from this obtain rs where g3:
  1511             simp_all add: numeral_2_eq_2)   
  4204       "rec_calc_rel f (lm @ [stp]) rs \<and> rs \<noteq> 0" ..
  1512     moreover have "abc_steps_l (0, xs @ Suc x # Suc y # anything) [Dec (Suc n) 2, Goto 0] 2 = 
  4205     hence "\<exists> stpb. abc_steps_l (0, lm @ stp # 0\<up>(a_md - Suc rs_pos) @
  1513                  (0, xs @ Suc x # y # anything)"
  4206                      suf_lm) aprog stpb 
       
  4207       = (0, lm @ Suc stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  4208       using h
  1514       using h
  4209       apply(rule_tac mn_ind_step)
  1515       by(simp add: abc_steps_l.simps numeral_2_eq_2 abc_step_l.simps abc_fetch.simps
  4210       apply(rule_tac recursive_compile_correct, simp, simp)
  1516         abc_lm_v.simps abc_lm_s.simps nth_append list_update_append list_update.simps)
       
  1517     ultimately show "?case"
       
  1518       apply(simp add: abc_Hoare_halt_def)
       
  1519       by(rule_tac x = "2 + stp" in exI, simp only: abc_steps_add, simp)
       
  1520   qed
       
  1521 qed
       
  1522 
       
  1523 lemma adjust_paras: 
       
  1524   "length xs = n \<Longrightarrow> {\<lambda>nl. nl = xs @ x # y # anything}  [Inc n, Dec (Suc n) 3, Goto (Suc 0)]
       
  1525        {\<lambda>nl. nl = xs @ Suc x # 0 # anything}"
       
  1526 using adjust_paras'[of xs n x y anything]
       
  1527 by(simp add: abc_comp.simps abc_shift.simps numeral_2_eq_2 numeral_3_eq_3)
       
  1528 
       
  1529 lemma [simp]: "\<lbrakk>rec_ci g = (gap, gar, gft); \<forall>y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]);
       
  1530         length xs = n; Suc y\<le>x\<rbrakk> \<Longrightarrow> gar = Suc (Suc n)"
       
  1531   apply(erule_tac x = y in allE, simp)
       
  1532   apply(drule_tac param_pattern, auto)
       
  1533   done
       
  1534 
       
  1535 lemma loop_back':  
       
  1536   assumes h: "length A = length gap + 4" "length xs = n"
       
  1537   and le: "y \<ge> x"
       
  1538   shows "\<exists> stp. abc_steps_l (length A, xs @ m # (y - x) # x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp
       
  1539      = (length A, xs @ m # y # 0 # anything)"
       
  1540   using le
       
  1541 proof(induct x)
       
  1542   case 0
       
  1543   thus "?case"
       
  1544     using h
       
  1545     by(rule_tac x = 0 in exI,
       
  1546     auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_lm_s.simps abc_lm_v.simps)
       
  1547 next
       
  1548   case (Suc x)
       
  1549   have "x \<le> y \<Longrightarrow> \<exists>stp. abc_steps_l (length A, xs @ m # (y - x) # x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp =
       
  1550               (length A, xs @ m # y # 0 # anything)" by fact
       
  1551   moreover have "Suc x \<le> y" by fact
       
  1552   moreover then have "\<exists> stp. abc_steps_l (length A, xs @ m # (y - Suc x) # Suc x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp
       
  1553                 = (length A, xs @ m # (y - x) # x # anything)"
       
  1554     using h
       
  1555     apply(rule_tac x = 3 in exI)
       
  1556     by(simp add: abc_steps_l.simps numeral_3_eq_3 abc_step_l.simps abc_fetch.simps nth_append abc_lm_v.simps abc_lm_s.simps
       
  1557                     list_update_append list_update.simps)
       
  1558   ultimately show "?case"
       
  1559     apply(auto)
       
  1560     apply(rule_tac x = "stpa + stp" in exI)
       
  1561     by(simp add: abc_steps_add)
       
  1562 qed
       
  1563 
       
  1564 
       
  1565 lemma loop_back:  
       
  1566   assumes h: "length A = length gap + 4" "length xs = n"
       
  1567   shows "\<exists> stp. abc_steps_l (length A, xs @ m # 0 # x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp
       
  1568      = (0, xs @ m # x # 0 # anything)"
       
  1569 using loop_back'[of A gap xs n x x m anything] assms
       
  1570 apply(auto)
       
  1571 apply(rule_tac x = "stp + 1" in exI)
       
  1572 apply(simp only: abc_steps_add, simp)
       
  1573 apply(simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_lm_v.simps abc_lm_s.simps)
       
  1574 done
       
  1575 
       
  1576 lemma rec_exec_pr_0_simps: "rec_exec (Pr n f g) (xs @ [0]) = rec_exec f xs"
       
  1577  by(simp add: rec_exec.simps)
       
  1578 
       
  1579 lemma rec_exec_pr_Suc_simps: "rec_exec (Pr n f g) (xs @ [Suc y])
       
  1580           = rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])"
       
  1581 apply(induct y)
       
  1582 apply(simp add: rec_exec.simps)
       
  1583 apply(simp add: rec_exec.simps)
       
  1584 done
       
  1585 
       
  1586 lemma [simp]: "Suc (max (n + 3) (max fft gft) - Suc (Suc (Suc n))) = max (n + 3) (max fft gft) - Suc (Suc n)"
       
  1587 by arith
       
  1588 
       
  1589 lemma pr_loop:
       
  1590   assumes code: "code = ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @
       
  1591     [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]"
       
  1592   and len: "length xs = n"
       
  1593   and g_ind: "\<forall> y<x. (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap
       
  1594   {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \<up> (gft - Suc gar) @ anything})"
       
  1595   and compile_g: "rec_ci g = (gap, gar, gft)"
       
  1596   and termi_g: "\<forall> y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])"
       
  1597   and ft: "ft = max (n + 3) (max fft gft)"
       
  1598   and less: "Suc y \<le> x"
       
  1599   shows 
       
  1600   "\<exists>stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything)
       
  1601   code stp = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (ft - Suc (Suc n)) @ y # anything)"
       
  1602 proof -
       
  1603   let ?A = "[Dec  ft (length gap + 7)]"
       
  1604   let ?B = "gap"
       
  1605   let ?C = "[Inc n, Dec (Suc n) 3, Goto (Suc 0)]"
       
  1606   let ?D = "[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]"
       
  1607   have "\<exists> stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything)
       
  1608             ((?A [+] (?B [+] ?C)) @ ?D) stp = (length (?A [+] (?B [+] ?C)), 
       
  1609           xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])])
       
  1610                   # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)"
       
  1611   proof -
       
  1612     have "\<exists> stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything)
       
  1613       ((?A [+] (?B [+] ?C))) stp = (length (?A [+] (?B [+] ?C)),  xs @ (x - y) # 0 # 
       
  1614       rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)"
  4211     proof -
  1615     proof -
  4212       show "rec_ci f = ((aprog', rs_pos', a_md'))" using h by simp
  1616       have "{\<lambda> nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything}
       
  1617         (?A [+] (?B [+] ?C)) 
       
  1618         {\<lambda> nl. nl = xs @ (x - y) # 0 # 
       
  1619         rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything}"
       
  1620       proof(rule_tac abc_Hoare_plus_halt)
       
  1621         show "{\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything}
       
  1622           [Dec ft (length gap + 7)] 
       
  1623           {\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ y # anything}"
       
  1624           using ft len
       
  1625           by(simp)
       
  1626       next
       
  1627         show 
       
  1628           "{\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ y # anything} 
       
  1629           ?B [+] ?C
       
  1630           {\<lambda>nl. nl = xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything}"
       
  1631         proof(rule_tac abc_Hoare_plus_halt)
       
  1632           have a: "gar = Suc (Suc n)" 
       
  1633             using compile_g termi_g len less
       
  1634             by simp
       
  1635           have b: "gft > gar"
       
  1636             using compile_g
       
  1637             by(erule_tac footprint_ge)
       
  1638           show "{\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ y # anything} gap 
       
  1639                 {\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 
       
  1640                       rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything}"
       
  1641           proof -
       
  1642             have 
       
  1643               "{\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (gft - gar) @ 0\<up>(ft - gft) @ y # anything} gap
       
  1644               {\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 
       
  1645               rec_exec g (xs @ [(x - Suc y), rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (gft - Suc gar) @ 0\<up>(ft - gft) @ y # anything}"
       
  1646               using g_ind less by simp
       
  1647             thus "?thesis"
       
  1648               using a b ft
       
  1649               by(simp add: replicate_merge_anywhere numeral_3_eq_3)
       
  1650           qed
       
  1651         next
       
  1652           show "{\<lambda>nl. nl = xs @ (x - Suc y) #
       
  1653                     rec_exec (Pr n f g) (xs @ [x - Suc y]) #
       
  1654             rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything}
       
  1655             [Inc n, Dec (Suc n) 3, Goto (Suc 0)]
       
  1656             {\<lambda>nl. nl = xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) 
       
  1657                     (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything}"
       
  1658             using len less
       
  1659             using adjust_paras[of xs n "x - Suc y" " rec_exec (Pr n f g) (xs @ [x - Suc y])"
       
  1660               " rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 
       
  1661               0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything"]
       
  1662             by(simp add: Suc_diff_Suc)
       
  1663         qed
       
  1664       qed
       
  1665       thus "?thesis"
       
  1666         by(simp add: abc_Hoare_halt_def, auto, rule_tac x = na in exI, case_tac "abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 
       
  1667           0 \<up> (ft - Suc (Suc n)) @ Suc y # anything)
       
  1668              ([Dec ft (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) na", simp)
       
  1669     qed
       
  1670     then obtain stpa where "abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything)
       
  1671             ((?A [+] (?B [+] ?C))) stpa = (length (?A [+] (?B [+] ?C)), 
       
  1672           xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])])
       
  1673                   # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)" ..
       
  1674     thus "?thesis"
       
  1675       by(erule_tac abc_append_frist_steps_halt_eq)
       
  1676   qed
       
  1677   moreover have 
       
  1678     "\<exists> stp. abc_steps_l (length (?A [+] (?B [+] ?C)),
       
  1679     xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)
       
  1680     ((?A [+] (?B [+] ?C)) @ ?D) stp  = (0, xs @ (x - y) # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 
       
  1681     0 # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)"
       
  1682     using len
       
  1683     by(rule_tac loop_back, simp_all)
       
  1684   moreover have "rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) = rec_exec (Pr n f g) (xs @ [x - y])"
       
  1685     using less
       
  1686     thm rec_exec.simps
       
  1687     apply(case_tac "x - y", simp_all add: rec_exec_pr_Suc_simps)
       
  1688     by(subgoal_tac "nat = x - Suc y", simp, arith)    
       
  1689   ultimately show "?thesis"
       
  1690     using code ft
       
  1691     by(auto, rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add replicate_Suc_iff_anywhere del: replicate_Suc)
       
  1692 qed
       
  1693 
       
  1694 lemma [simp]: 
       
  1695   "length xs = n \<Longrightarrow> abc_lm_s (xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) 
       
  1696   (max fft gft) - Suc (Suc n)) @ 0 # anything) (max (n + 3) (max fft gft)) 0 =
       
  1697     xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ anything"
       
  1698 apply(simp add: abc_lm_s.simps)
       
  1699 using list_update_length[of "xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) (max fft gft) - Suc (Suc n))"
       
  1700                         0 anything 0]
       
  1701 apply(auto simp: Suc_diff_Suc)
       
  1702 apply(simp add: exp_suc[THEN sym] Suc_diff_Suc  del: replicate_Suc)
       
  1703 done
       
  1704 
       
  1705 lemma [simp]:
       
  1706   "(xs @ x # rec_exec (Pr (length xs) f g) (xs @ [x]) # 0 \<up> (max (length xs + 3)
       
  1707   (max fft gft) - Suc (Suc (length xs))) @ 0 # anything) !
       
  1708     max (length xs + 3) (max fft gft) = 0"
       
  1709 using nth_append_length[of "xs @ x # rec_exec (Pr (length xs) f g) (xs @ [x]) #
       
  1710   0 \<up> (max (length xs + 3) (max fft gft) - Suc (Suc (length xs)))" 0  anything]
       
  1711 by(simp)
       
  1712 
       
  1713 lemma pr_loop_correct:
       
  1714   assumes less: "y \<le> x" 
       
  1715   and len: "length xs = n"
       
  1716   and compile_g: "rec_ci g = (gap, gar, gft)"
       
  1717   and termi_g: "\<forall> y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])"
       
  1718   and g_ind: "\<forall> y<x. (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap
       
  1719   {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \<up> (gft - Suc gar) @ anything})"
       
  1720   shows "{\<lambda>nl. nl = xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (max (n + 3) (max fft gft) - Suc (Suc n)) @ y # anything}
       
  1721    ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]
       
  1722    {\<lambda>nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ anything}" 
       
  1723   using less
       
  1724 proof(induct y)
       
  1725   case 0
       
  1726   thus "?case"
       
  1727     using len
       
  1728     apply(simp add: abc_Hoare_halt_def)
       
  1729     apply(rule_tac x = 1 in exI)
       
  1730     by(auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps 
       
  1731       nth_append abc_comp.simps abc_shift.simps, simp add: abc_lm_v.simps)
       
  1732 next
       
  1733   case (Suc y)
       
  1734   let ?ft = "max (n + 3) (max fft gft)"
       
  1735   let ?C = "[Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] 
       
  1736     [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]"
       
  1737   have ind: "y \<le> x \<Longrightarrow>
       
  1738          {\<lambda>nl. nl = xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything}
       
  1739          ?C {\<lambda>nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (?ft - Suc n) @ anything}" by fact 
       
  1740   have less: "Suc y \<le> x" by fact
       
  1741   have stp1: 
       
  1742     "\<exists> stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (?ft - Suc (Suc n)) @ Suc y # anything)
       
  1743     ?C stp  = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything)"
       
  1744     using assms less
       
  1745     by(rule_tac  pr_loop, auto)
       
  1746   then obtain stp1 where a:
       
  1747     "abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (?ft - Suc (Suc n)) @ Suc y # anything)
       
  1748    ?C stp1 = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything)" ..
       
  1749   moreover have 
       
  1750     "\<exists> stp. abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything)
       
  1751     ?C stp = (length ?C, xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (?ft - Suc n) @ anything)"
       
  1752     using ind less
       
  1753     by(auto simp: abc_Hoare_halt_def, case_tac "abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g) 
       
  1754       (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything) ?C na", rule_tac x = na in exI, simp)
       
  1755   then obtain stp2 where b:
       
  1756     "abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything)
       
  1757     ?C stp2 = (length ?C, xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (?ft - Suc n) @ anything)" ..
       
  1758   from a b show "?case"
       
  1759     by(simp add: abc_Hoare_halt_def, rule_tac x = "stp1 + stp2" in exI, simp add: abc_steps_add)
       
  1760 qed
       
  1761 
       
  1762 lemma compile_pr_correct':
       
  1763   assumes termi_g: "\<forall> y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])"
       
  1764   and g_ind: 
       
  1765   "\<forall> y<x. (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap
       
  1766   {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \<up> (gft - Suc gar) @ anything})"
       
  1767   and termi_f: "terminate f xs"
       
  1768   and f_ind: "\<And> anything. {\<lambda>nl. nl = xs @ 0 \<up> (fft - far) @ anything} fap {\<lambda>nl. nl = xs @ rec_exec f xs # 0 \<up> (fft - Suc far) @ anything}"
       
  1769   and len: "length xs = n"
       
  1770   and compile1: "rec_ci f = (fap, far, fft)"
       
  1771   and compile2: "rec_ci g = (gap, gar, gft)"
       
  1772   shows 
       
  1773   "{\<lambda>nl. nl = xs @ x # 0 \<up> (max (n + 3) (max fft gft) - n) @ anything}
       
  1774   mv_box n (max (n + 3) (max fft gft)) [+]
       
  1775   (fap [+] (mv_box n (Suc n) [+]
       
  1776   ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @
       
  1777   [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)])))
       
  1778   {\<lambda>nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ anything}"
       
  1779 proof -
       
  1780   let ?ft = "max (n+3) (max fft gft)"
       
  1781   let ?A = "mv_box n ?ft"
       
  1782   let ?B = "fap"
       
  1783   let ?C = "mv_box n (Suc n)"
       
  1784   let ?D = "[Dec ?ft (length gap + 7)]"
       
  1785   let ?E = "gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]"
       
  1786   let ?F = "[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]"
       
  1787   let ?P = "\<lambda>nl. nl = xs @ x # 0 \<up> (?ft - n) @ anything"
       
  1788   let ?S = "\<lambda>nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (?ft - Suc n) @ anything"
       
  1789   let ?Q1 = "\<lambda>nl. nl = xs @ 0 \<up> (?ft - n) @  x # anything"
       
  1790   show "{?P} (?A [+] (?B [+] (?C [+] (?D [+] ?E @ ?F)))) {?S}"
       
  1791   proof(rule_tac abc_Hoare_plus_halt)
       
  1792     show "{?P} ?A {?Q1}"
       
  1793       using len by simp
       
  1794   next
       
  1795     let ?Q2 = "\<lambda>nl. nl = xs @ rec_exec f xs # 0 \<up> (?ft - Suc n) @  x # anything"
       
  1796     have a: "?ft \<ge> fft"
       
  1797       by arith
       
  1798     have b: "far = n"
       
  1799       using compile1 termi_f len
       
  1800       by(drule_tac param_pattern, auto)
       
  1801     have c: "fft > far"
       
  1802       using compile1
       
  1803       by(simp add: footprint_ge)
       
  1804     show "{?Q1} (?B [+] (?C [+] (?D [+] ?E @ ?F))) {?S}"
       
  1805     proof(rule_tac abc_Hoare_plus_halt)
       
  1806       have "{\<lambda>nl. nl = xs @ 0 \<up> (fft - far) @ 0\<up>(?ft - fft) @ x # anything} fap 
       
  1807             {\<lambda>nl. nl = xs @ rec_exec f xs # 0 \<up> (fft - Suc far) @ 0\<up>(?ft - fft) @ x # anything}"
       
  1808         by(rule_tac f_ind)
       
  1809       moreover have "fft - far + ?ft - fft = ?ft - far"
       
  1810         using a b c by arith
       
  1811       moreover have "fft - Suc n + ?ft - fft = ?ft - Suc n"
       
  1812         using a b c by arith
       
  1813       ultimately show "{?Q1} ?B {?Q2}"
       
  1814         using b
       
  1815         by(simp add: replicate_merge_anywhere)
  4213     next
  1816     next
  4214       show "rec_ci (Mn n f) = (aprog, rs_pos, a_md)" using h by simp
  1817       let ?Q3 = "\<lambda> nl. nl = xs @ 0 # rec_exec f xs # 0\<up>(?ft - Suc (Suc n)) @ x # anything"
  4215     next
  1818       show "{?Q2} (?C [+] (?D [+] ?E @ ?F)) {?S}"
  4216       show "rec_calc_rel f (lm @ [stp]) rs" using g3 by simp
  1819       proof(rule_tac abc_Hoare_plus_halt)
  4217     next
  1820         show "{?Q2} (?C) {?Q3}"
  4218       show "0 < rs" using g3 by simp
  1821           using mv_box_correct[of n "Suc n" "xs @ rec_exec f xs # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ x # anything"]
  4219     next
  1822           using len
  4220       show "Suc rs_pos < a_md"
  1823           by(auto)
  4221         using g3 h
  1824       next
  4222         apply(auto)
  1825         show "{?Q3} (?D [+] ?E @ ?F) {?S}"
  4223         apply(frule_tac f = f in para_pattern, simp, simp)
  1826           using pr_loop_correct[of x x xs n g  gap gar gft f fft anything] assms
  4224         apply(simp add: rec_ci.simps, auto)
  1827           by(simp add: rec_exec_pr_0_simps)
  4225         apply(subgoal_tac "Suc (length lm) < a_md'")
  1828       qed
  4226         apply(arith)
       
  4227         apply(simp add: ci_ad_ge_paras)
       
  4228         done
       
  4229     next
       
  4230       show "rs_pos' = Suc rs_pos"
       
  4231         using g3 h
       
  4232         apply(auto)
       
  4233         apply(frule_tac f = f in para_pattern, simp, simp)
       
  4234         apply(simp add: rec_ci.simps)
       
  4235         done
       
  4236     qed
  1829     qed
  4237     thus "\<exists>stpa\<ge>Suc stp. abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @
       
  4238                  suf_lm) aprog stpa 
       
  4239       = (0, lm @ Suc stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)"
       
  4240       using g2
       
  4241       apply(erule_tac exE)
       
  4242       apply(case_tac "stpb = 0", simp add: abc_steps_l.simps)
       
  4243       apply(rule_tac x = "stpa + stpb" in exI, simp add:
       
  4244         abc_steps_add)
       
  4245       using g1
       
  4246       apply(arith)
       
  4247       done
       
  4248   qed
  1830   qed
  4249   from this obtain stpa where 
  1831 qed 
  4250     "stp \<le> stpa \<and> abc_steps_l (0, lm @ 0 # 0\<up>(a_md - Suc rs_pos) @ suf_lm)
  1832 
  4251          aprog stpa = (0, lm @ stp # 0\<up>(a_md - Suc rs_pos) @ suf_lm)" ..
  1833 lemma compile_pr_correct:
  4252   thus "case abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog stp
  1834   assumes g_ind: "\<forall>y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) \<and>
  4253     of (ss, e) \<Rightarrow> ss < length aprog"
  1835   (\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow>
  4254     apply(case_tac "abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suf_lm) aprog
  1836   (\<forall>xc. {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \<up> (xb - xa) @ xc} x
  4255       stp", simp, case_tac "a \<ge> length aprog", 
  1837   {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \<up> (xb - Suc xa) @ xc}))"
  4256         simp, simp)
  1838   and termi_f: "terminate f xs"
  4257     apply(subgoal_tac "\<exists> d. stpa = stp + d", erule_tac exE)
  1839   and f_ind:
  4258     apply(subgoal_tac "lm @ 0\<up>(a_md - rs_pos) @ suf_lm = lm @ 0 # 
  1840   "\<And>ap arity fp anything.
  4259              0\<up>(a_md - Suc rs_pos) @ suf_lm", simp add: abc_steps_add)
  1841   rec_ci f = (ap, arity, fp) \<Longrightarrow> {\<lambda>nl. nl = xs @ 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ rec_exec f xs # 0 \<up> (fp - Suc arity) @ anything}"
  4260     apply(frule_tac as = a and lm = b and stp = d in abc_state_keep, 
  1842   and len: "length xs = n"
  4261           simp)
  1843   and compile: "rec_ci (Pr n f g) = (ap, arity, fp)"
  4262     using h  
  1844   shows "{\<lambda>nl. nl = xs @ x # 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (fp - Suc arity) @ anything}"
  4263     apply(simp add: rec_ci.simps, simp, 
  1845 proof(case_tac "rec_ci f", case_tac "rec_ci g")
  4264               simp only: replicate_Suc[THEN sym])
  1846   fix fap far fft gap gar gft
  4265     apply(case_tac rs_pos, simp, simp)
  1847   assume h: "rec_ci f = (fap, far, fft)" "rec_ci g = (gap, gar, gft)"
  4266     apply(rule_tac x = "stpa - stp" in exI, simp, simp)
  1848   have g: 
  4267     done
  1849     "\<forall>y<x. (terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) \<and>
  4268 qed   
  1850      (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap
  4269 
  1851     {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \<up> (gft - Suc gar) @ anything}))"
  4270 lemma abc_append_cons_eq[intro!]: 
  1852     using g_ind h
  4271   "\<lbrakk>ap = bp; cp = dp\<rbrakk> \<Longrightarrow> ap [+] cp = bp [+] dp"
  1853     by(auto)
  4272 by simp 
  1854   hence termi_g: "\<forall> y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])"
  4273 
  1855     by simp
  4274 lemma cn_merge_gs_split: 
  1856   from g have g_newind: 
  4275   "\<lbrakk>i < length gs; rec_ci (gs!i) = (ga, gb, gc)\<rbrakk> \<Longrightarrow> 
  1857     "\<forall> y<x. (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap
  4276      cn_merge_gs (map rec_ci gs) p = 
  1858     {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \<up> (gft - Suc gar) @ anything})"
  4277         cn_merge_gs (map rec_ci (take i gs)) p [+] ga [+] 
  1859     by auto
  4278        mv_box gb (p + i) [+] 
  1860   have f_newind: "\<And> anything. {\<lambda>nl. nl = xs @ 0 \<up> (fft - far) @ anything} fap {\<lambda>nl. nl = xs @ rec_exec f xs # 0 \<up> (fft - Suc far) @ anything}"
  4279       cn_merge_gs (map rec_ci (drop (Suc i) gs)) (p + Suc i)"
  1861     using h
  4280 apply(induct i arbitrary: gs p, case_tac gs, simp, simp)
  1862     by(rule_tac f_ind, simp)
  4281 apply(case_tac gs, simp, case_tac "rec_ci a", 
  1863   show "?thesis"
  4282        simp add: abc_append_commute[THEN sym])
  1864     using termi_f termi_g h compile
  4283 done
  1865     apply(simp add: rec_ci.simps abc_comp_commute, auto)
  4284 
  1866     using g_newind f_newind len
  4285 text {*
  1867     by(rule_tac compile_pr_correct', simp_all)
  4286   Correctness of the complier (non-terminating case for Mn). There are many cases when a 
  1868 qed
  4287   recursive function does not terminate. For the purpose of Uiversal Turing Machine, we only 
  1869 
  4288   need to prove the case for @{text "Mn"} and @{text "Cn"}.
  1870 fun mn_ind_inv ::
  4289   This lemma is for @{text "Cn"}. For @{text "Cn f g1 g2 \<dots>gi, gi+1, \<dots> gn"}, this lemma describes what 
  1871   "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> nat list \<Rightarrow> bool"
  4290   happens when every one of @{text "g1, g2, \<dots> gi"} terminates, but 
  1872   where
  4291   @{text "gi+1"} does not terminate, so that whole function @{text "Cn f g1 g2 \<dots>gi, gi+1, \<dots> gn"}
  1873   "mn_ind_inv (as, lm') ss x rsx suf_lm lm = 
  4292   does not terminate.
  1874            (if as = ss then lm' = lm @ x # rsx # suf_lm
  4293   *}
  1875             else if as = ss + 1 then 
  4294 
  1876                  \<exists>y. (lm' = lm @ x # y # suf_lm) \<and> y \<le> rsx
  4295 lemma cn_gi_uhalt: 
  1877             else if as = ss + 2 then 
  4296   assumes cn_recf: "rf = Cn n f gs"
  1878                  \<exists>y. (lm' = lm @ x # y # suf_lm) \<and> y \<le> rsx
  4297   and compiled_cn_recf: "rec_ci rf = (aprog, rs_pos, a_md)"
  1879             else if as = ss + 3 then lm' = lm @ x # 0 # suf_lm
  4298   and args_length: "length lm = n"
  1880             else if as = ss + 4 then lm' = lm @ Suc x # 0 # suf_lm
  4299   and exist_unhalt_recf: "i < length gs" "gi = gs ! i"
  1881             else if as = 0 then lm' = lm @ Suc x # 0 # suf_lm
  4300   and complied_unhalt_recf: "rec_ci gi = (ga, gb, gc)"  "gb = n"
  1882             else False
  4301   and all_halt_before_gi: "\<forall> j < i. (\<exists> rs. rec_calc_rel (gs!j) lm rs)" 
  1883 )"
  4302   and unhalt_condition: "\<And> slm. \<forall> stp. case abc_steps_l (0, lm @ 0\<up>(gc - gb) @ slm) 
  1884 
  4303      ga stp of (se, e) \<Rightarrow> se < length ga"
  1885 fun mn_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
  4304   shows " \<forall> stp. case abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suflm) aprog
  1886   where
  4305   stp of (ss, e) \<Rightarrow> ss < length aprog"
  1887   "mn_stage1 (as, lm) ss n = 
  4306   using cn_recf compiled_cn_recf args_length exist_unhalt_recf complied_unhalt_recf
  1888             (if as = 0 then 0 
  4307         all_halt_before_gi unhalt_condition
  1889              else if as = ss + 4 then 1
  4308 proof(case_tac "rec_ci f", simp)
  1890              else if as = ss + 3 then 2
  4309   fix a b c
  1891              else if as = ss + 2 \<or> as = ss + 1 then 3
  4310   assume h1: "rf = Cn n f gs" 
  1892              else if as = ss then 4
  4311     "rec_ci (Cn n f gs) = (aprog, rs_pos, a_md)" 
  1893              else 0
  4312     "length lm = n" 
  1894 )"
  4313     "gi = gs ! i" 
  1895 
  4314     "rec_ci (gs!i) = (ga, n, gc)" 
  1896 fun mn_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
  4315     "gb = n" "rec_ci f = (a, b, c)"
  1897   where
  4316     and h2: "\<forall>j<i. \<exists>rs. rec_calc_rel (gs ! j) lm rs"
  1898   "mn_stage2 (as, lm) ss n = 
  4317     "i < length gs"
  1899             (if as = ss + 1 \<or> as = ss + 2 then (lm ! (Suc n))
  4318   and ind:
  1900              else 0)"
  4319     "\<And> slm. \<forall> stp. case abc_steps_l (0, lm @ 0\<up>(gc - n) @ slm) ga stp of (se, e) \<Rightarrow> se < length ga"
  1901 
  4320   have h3: "rs_pos = n"
  1902 fun mn_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
  4321     using h1
  1903   where
  4322     by(rule_tac ci_cn_para_eq, simp)
  1904   "mn_stage3 (as, lm) ss n = (if as = ss + 2 then 1 else 0)"
  4323   let ?ggs = "take i gs"
  1905 
  4324   have "\<exists> ys. (length ys = i \<and> 
  1906  
  4325     (\<forall> k < i. rec_calc_rel (?ggs ! k) lm (ys ! k)))"
  1907 fun mn_measure :: "((nat \<times> nat list) \<times> nat \<times> nat) \<Rightarrow>
  4326     using h2
  1908                                                 (nat \<times> nat \<times> nat)"
  4327     apply(induct i, simp, simp)
  1909   where
  4328     apply(erule_tac exE)
  1910   "mn_measure ((as, lm), ss, n) = 
  4329     apply(erule_tac x = ia in allE, simp)
  1911      (mn_stage1 (as, lm) ss n, mn_stage2 (as, lm) ss n,
  4330     apply(erule_tac exE)
  1912                                        mn_stage3 (as, lm) ss n)"
  4331     apply(rule_tac x = "ys @ [x]" in exI, simp add: nth_append, auto)
  1913 
  4332     apply(subgoal_tac "k = length ys", simp, simp)
  1914 definition mn_LE :: "(((nat \<times> nat list) \<times> nat \<times> nat) \<times>
  4333     done
  1915                      ((nat \<times> nat list) \<times> nat \<times> nat)) set"
  4334   from this obtain ys where g1:
  1916   where "mn_LE \<equiv> (inv_image lex_triple mn_measure)"
  4335     "(length ys = i \<and> (\<forall> k < i. rec_calc_rel (?ggs ! k)
  1917 
  4336                         lm (ys ! k)))" ..
  1918 lemma wf_mn_le[intro]: "wf mn_LE"
  4337   let ?pstr = "Max (set (Suc n # c # map (\<lambda>(aprog, p, n). n)
  1919 by(auto intro:wf_inv_image wf_lex_triple simp: mn_LE_def)
  4338     (map rec_ci (gs))))"
  1920 
  4339   have "\<exists>stp. abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suflm) 
  1921 declare mn_ind_inv.simps[simp del]
  4340     (cn_merge_gs (map rec_ci ?ggs) ?pstr) stp =
  1922 
  4341     (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) ?ggs) +
  1923 lemma [simp]: 
  4342     3 * length ?ggs, lm @ 0\<up>(?pstr - n) @ ys @ 0\<up>(a_md -(?pstr + length ?ggs)) @
  1924   "0 < rsx \<Longrightarrow> 
  4343     suflm) "
  1925  \<exists>y. (xs @ x # rsx # anything)[Suc (length xs) := rsx - Suc 0] = xs @ x # y # anything \<and> y \<le> rsx"
  4344     apply(rule_tac  cn_merge_gs_ex)
  1926 apply(rule_tac x = "rsx - 1" in exI)
  4345     apply(rule_tac  recursive_compile_correct, simp, simp)
  1927 apply(simp add: list_update_append list_update.simps)
  4346     using h1
  1928 done
  4347     apply(simp add: rec_ci.simps, auto)
  1929 
  4348     using g1
  1930 lemma [simp]: 
  4349     apply(simp)
  1931   "\<lbrakk>y \<le> rsx; 0 < y\<rbrakk>
  4350     using h2 g1
  1932             \<Longrightarrow> \<exists>ya. (xs @ x # y # anything)[Suc (length xs) := y - Suc 0] = xs @ x # ya # anything \<and> ya \<le> rsx"
  4351     apply(simp)
  1933 apply(rule_tac x = "y - 1" in exI)
  4352     apply(rule_tac min_max.le_supI2)
  1934 apply(simp add: list_update_append list_update.simps)
  4353     apply(rule_tac Max_ge, simp, simp, rule_tac disjI2)
  1935 done
  4354     apply(subgoal_tac "aa \<in> set gs", simp)
  1936 
  4355     using h2
  1937 lemma abc_comp_null[simp]: "(A [+] B = []) = (A = [] \<and> B = [])"
  4356     apply(rule_tac A = "set (take i gs)" in subsetD, 
  1938 by(auto simp: abc_comp.simps abc_shift.simps)
  4357       simp add: set_take_subset, simp)
  1939 
  4358     done
  1940 lemma rec_ci_not_null[simp]: "(rec_ci f \<noteq> ([], a, b))"
  4359   from this obtain stpa where g2: 
  1941 apply(case_tac f, auto simp: rec_ci_z_def rec_ci_s_def rec_ci.simps addition.simps rec_ci_id.simps)
  4360     "abc_steps_l (0, lm @ 0\<up>(a_md - n) @ suflm) 
  1942 apply(case_tac "rec_ci recf", auto simp: mv_box.simps)
  4361     (cn_merge_gs (map rec_ci ?ggs) ?pstr) stpa =
  1943 apply(case_tac "rec_ci recf1", case_tac "rec_ci recf2", simp)
  4362     (listsum (map ((\<lambda>(ap, pos, n). length ap) \<circ> rec_ci) ?ggs) +
  1944 apply(case_tac "rec_ci recf", simp)
  4363     3 * length ?ggs, lm @ 0\<up>(?pstr - n) @ ys @ 0\<up>(a_md -(?pstr + length ?ggs)) @
  1945 done
  4364     suflm)" ..
  1946 
       
  1947 lemma mn_correct:
       
  1948   assumes compile: "rec_ci rf = (fap, far, fft)"
       
  1949   and ge: "0 < rsx"
       
  1950   and len: "length xs = arity"
       
  1951   and B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]"
       
  1952   and f: "f = (\<lambda> stp. (abc_steps_l (length fap, xs @ x # rsx # anything) (fap @ B) stp, (length fap), arity)) "
       
  1953   and P: "P =(\<lambda> ((as, lm), ss, arity). as = 0)"
       
  1954   and Q: "Q = (\<lambda> ((as, lm), ss, arity). mn_ind_inv (as, lm) (length fap) x rsx anything xs)"
       
  1955   shows "\<exists> stp. P (f stp) \<and> Q (f stp)"
       
  1956 proof(rule_tac halt_lemma2)
       
  1957   show "wf mn_LE"
       
  1958     using wf_mn_le by simp
       
  1959 next
       
  1960   show "Q (f 0)"
       
  1961     by(auto simp: Q f abc_steps_l.simps mn_ind_inv.simps)
       
  1962 next
       
  1963   have "fap \<noteq> []"
       
  1964     using compile by auto
       
  1965   thus "\<not> P (f 0)"
       
  1966     by(auto simp: f P abc_steps_l.simps)
       
  1967 next
       
  1968   have "fap \<noteq> []"
       
  1969     using compile by auto
       
  1970   then have "\<And> stp. \<lbrakk>\<not> P (f stp); Q (f stp)\<rbrakk> \<Longrightarrow> Q (f (Suc stp)) \<and> (f (Suc stp), f stp) \<in> mn_LE"
       
  1971     using ge len
       
  1972     apply(case_tac "(abc_steps_l (length fap, xs @ x # rsx # anything) (fap @ B) stp)")
       
  1973     apply(simp add: abc_step_red2  B f P Q)
       
  1974     apply(auto split:if_splits simp add:abc_steps_l.simps  mn_ind_inv.simps abc_steps_zero B abc_fetch.simps nth_append)
       
  1975     by(auto simp: mn_LE_def lex_triple_def lex_pair_def 
       
  1976                abc_step_l.simps abc_steps_l.simps mn_ind_inv.simps
       
  1977                abc_lm_v.simps abc_lm_s.simps nth_append abc_fetch.simps
       
  1978                 split: if_splits)    
       
  1979   thus "\<forall>stp. \<not> P (f stp) \<and> Q (f stp) \<longrightarrow> Q (f (Suc stp)) \<and> (f (Suc stp), f stp) \<in> mn_LE"
       
  1980     by(auto)
       
  1981 qed
       
  1982 
       
  1983 lemma abc_Hoare_haltE:
       
  1984   "{\<lambda> nl. nl = lm1} p {\<lambda> nl. nl = lm2}
       
  1985     \<Longrightarrow> \<exists> stp. abc_steps_l (0, lm1) p stp = (length p, lm2)"
       
  1986 apply(auto simp: abc_Hoare_halt_def)
       
  1987 apply(rule_tac x = n in exI)
       
  1988 apply(case_tac "abc_steps_l (0, lm1) p n", auto)
       
  1989 done
       
  1990 
       
  1991 lemma mn_loop:
       
  1992   assumes B:  "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]"
       
  1993   and ft: "ft = max (Suc arity) fft"
       
  1994   and len: "length xs = arity"
       
  1995   and far: "far = Suc arity"
       
  1996   and ind: " (\<forall>xc. ({\<lambda>nl. nl = xs @ x # 0 \<up> (fft - far) @ xc} fap
       
  1997     {\<lambda>nl. nl = xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (fft - Suc far) @ xc}))"
       
  1998   and exec_less: "rec_exec f (xs @ [x]) > 0"
       
  1999   and compile: "rec_ci f = (fap, far, fft)"
       
  2000   shows "\<exists> stp > 0. abc_steps_l (0, xs @ x # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp =
       
  2001     (0, xs @ Suc x # 0 \<up> (ft - Suc arity) @ anything)"
       
  2002 proof -
       
  2003   have "\<exists> stp. abc_steps_l (0, xs @ x # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp =
       
  2004     (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)"
       
  2005   proof -
       
  2006     have "\<exists> stp. abc_steps_l (0, xs @ x # 0 \<up> (ft - Suc arity) @ anything) fap stp =
       
  2007       (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)"
       
  2008     proof -
       
  2009       have "{\<lambda>nl. nl = xs @ x # 0 \<up> (fft - far) @ 0\<up>(ft - fft) @ anything} fap 
       
  2010             {\<lambda>nl. nl = xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (fft - Suc far) @ 0\<up>(ft - fft) @ anything}"
       
  2011         using ind by simp
       
  2012       moreover have "fft > far"
       
  2013         using compile
       
  2014         by(erule_tac footprint_ge)
       
  2015       ultimately show "?thesis"
       
  2016         using ft far
       
  2017         apply(drule_tac abc_Hoare_haltE)
       
  2018         by(simp add: replicate_merge_anywhere max_absorb2)
       
  2019     qed
       
  2020     then obtain stp where "abc_steps_l (0, xs @ x # 0 \<up> (ft - Suc arity) @ anything) fap stp =
       
  2021       (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" ..
       
  2022     thus "?thesis"
       
  2023       by(erule_tac abc_append_frist_steps_halt_eq)
       
  2024   qed
  4365   moreover have 
  2025   moreover have 
  4366     "\<exists> cp. aprog = (cn_merge_gs
  2026     "\<exists> stp > 0. abc_steps_l (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything) (fap @ B) stp =
  4367     (map rec_ci ?ggs) ?pstr) [+] ga [+] cp"
  2027     (0, xs @ Suc x # 0 # 0 \<up> (ft - Suc (Suc arity)) @ anything)"
  4368     using h1
  2028     using mn_correct[of f fap far fft "rec_exec f (xs @ [x])" xs arity B
  4369     apply(simp add: rec_ci.simps)
  2029       "(\<lambda>stp. (abc_steps_l (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything) (fap @ B) stp, length fap, arity))"     
  4370     apply(rule_tac x = "mv_box n (?pstr + i) [+] 
  2030       x "0 \<up> (ft - Suc (Suc arity)) @ anything" "(\<lambda>((as, lm), ss, arity). as = 0)" 
  4371       (cn_merge_gs (map rec_ci (drop (Suc i) gs)) (?pstr + Suc i))
  2031       "(\<lambda>((as, lm), ss, aritya). mn_ind_inv (as, lm) (length fap) x (rec_exec f (xs @ [x])) (0 \<up> (ft - Suc (Suc arity)) @ anything) xs) "]  
  4372       [+]mv_boxes 0 (Suc (max (Suc n) (Max (insert c 
  2032       B compile  exec_less len
  4373      (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) +
  2033     apply(subgoal_tac "fap \<noteq> []", auto)
  4374       length gs)) n [+] mv_boxes (max (Suc n) (Max (insert c 
  2034     apply(rule_tac x = stp in exI, auto simp: mn_ind_inv.simps)
  4375       (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) 0 (length gs) [+]
  2035     by(case_tac "stp = 0", simp_all add: abc_steps_l.simps)
  4376       a [+] mv_box b (max (Suc n) 
  2036   moreover have "fft > far"
  4377       (Max (insert c (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) [+]
  2037     using compile
  4378      empty_boxes (length gs) [+] mv_box (max (Suc n) 
  2038     by(erule_tac footprint_ge)
  4379       (Max (insert c (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs)))) n [+]
  2039   ultimately show "?thesis"
  4380       mv_boxes (Suc (max (Suc n) (Max (insert c 
  2040     using ft far
  4381     (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))) + length gs)) 0 n" in exI)
       
  4382     apply(simp add: abc_append_commute [THEN sym])
       
  4383     apply(auto)
  2041     apply(auto)
  4384     using cn_merge_gs_split[of i gs ga "length lm" gc 
  2042     by(rule_tac x = "stp + stpa" in exI, 
  4385       "(max (Suc (length lm))
  2043       simp add: abc_steps_add replicate_Suc[THEN sym] diff_Suc_Suc del: replicate_Suc)
  4386        (Max (insert c (((\<lambda>(aprog, p, n). n) \<circ> rec_ci) ` set gs))))"] 
  2044 qed
  4387       h2
  2045 
  4388     apply(simp)
  2046 lemma mn_loop_correct': 
  4389     done
  2047   assumes B:  "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]"
  4390   from this obtain cp where g3: 
  2048   and ft: "ft = max (Suc arity) fft"
  4391     "aprog = (cn_merge_gs (map rec_ci ?ggs) ?pstr) [+] ga [+] cp" ..
  2049   and len: "length xs = arity"
  4392   show "\<forall> stp. case abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suflm) 
  2050   and ind_all: "\<forall>i\<le>x. (\<forall>xc. ({\<lambda>nl. nl = xs @ i # 0 \<up> (fft - far) @ xc} fap
  4393     aprog stp of (ss, e) \<Rightarrow> ss < length aprog"
  2051     {\<lambda>nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \<up> (fft - Suc far) @ xc}))"
  4394   proof(rule_tac abc_append_unhalt2)
  2052   and exec_ge: "\<forall> i\<le>x. rec_exec f (xs @ [i]) > 0"
  4395     show "abc_steps_l (0, lm @ 0\<up>(a_md - rs_pos) @ suflm) (
  2053   and compile: "rec_ci f = (fap, far, fft)"
  4396       cn_merge_gs (map rec_ci ?ggs) ?pstr) stpa =
  2054   and far: "far = Suc arity"
  4397          (length ((cn_merge_gs (map rec_ci ?ggs) ?pstr)),  
  2055   shows "\<exists> stp > x. abc_steps_l (0, xs @ 0 # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp = 
  4398           lm @ 0\<up>(?pstr - n) @ ys @ 0\<up>(a_md -(?pstr + length ?ggs)) @ suflm)"
  2056                (0, xs @ Suc x # 0 \<up> (ft - Suc arity) @ anything)"
  4399       using h3 g2
  2057 using ind_all exec_ge
  4400       apply(simp add: cn_merge_gs_length)
  2058 proof(induct x)
  4401       done
  2059   case 0
  4402   next
  2060   thus "?case"
  4403     show "ga \<noteq> []"
  2061     using assms
  4404       using h1
  2062     by(rule_tac mn_loop, simp_all)
  4405       apply(simp add: rec_ci_not_null)
  2063 next
  4406       done
  2064   case (Suc x)
  4407   next
  2065   have ind': "\<lbrakk>\<forall>i\<le>x. \<forall>xc. {\<lambda>nl. nl = xs @ i # 0 \<up> (fft - far) @ xc} fap {\<lambda>nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \<up> (fft - Suc far) @ xc};
  4408     show "\<forall>stp. case abc_steps_l (0, lm @ 0\<up>(?pstr - n) @ ys
  2066                \<forall>i\<le>x. 0 < rec_exec f (xs @ [i])\<rbrakk> \<Longrightarrow> 
  4409       @ 0\<up>(a_md - (?pstr + length (take i gs))) @ suflm) ga  stp of
  2067             \<exists>stp > x. abc_steps_l (0, xs @ 0 # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \<up> (ft - Suc arity) @ anything)" by fact
  4410           (ss, e) \<Rightarrow> ss < length ga"
  2068   have exec_ge: "\<forall>i\<le>Suc x. 0 < rec_exec f (xs @ [i])" by fact
  4411       using ind[of "0\<up>(?pstr - gc) @ ys @ 0\<up>(a_md - (?pstr + length (take i gs)))
  2069   have ind_all: "\<forall>i\<le>Suc x. \<forall>xc. {\<lambda>nl. nl = xs @ i # 0 \<up> (fft - far) @ xc} fap 
  4412         @ suflm"]
  2070     {\<lambda>nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \<up> (fft - Suc far) @ xc}" by fact
  4413       apply(subgoal_tac "lm @ 0\<up>(?pstr - n) @ ys
  2071   have ind: "\<exists>stp > x. abc_steps_l (0, xs @ 0 # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp =
  4414         @ 0\<up>(a_md - (?pstr + length (take i gs))) @ suflm
  2072     (0, xs @ Suc x # 0 \<up> (ft - Suc arity) @ anything)" using ind' exec_ge ind_all by simp
  4415                        = lm @ 0\<up>(gc - n) @ 
  2073   have stp: "\<exists> stp > 0. abc_steps_l (0, xs @ Suc x # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp =
  4416         0\<up>(?pstr - gc) @ ys @ 0\<up>(a_md - (?pstr + length (take i gs))) @ suflm", simp)
  2074     (0, xs @ Suc (Suc x) # 0 \<up> (ft - Suc arity) @ anything)"
  4417       apply(simp add: replicate_add[THEN sym])
  2075     using ind_all exec_ge B ft len far compile
  4418       apply(subgoal_tac "gc > n \<and> ?pstr \<ge> gc")
  2076     by(rule_tac mn_loop, simp_all)
  4419       apply(erule_tac conjE)
  2077   from ind stp show "?case"
  4420       apply(simp add: h1)
  2078     apply(auto)
  4421       using h1
  2079     by(rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add)
  4422       apply(auto)
  2080 qed
  4423       apply(rule_tac min_max.le_supI2)
  2081 
  4424       apply(rule_tac Max_ge, simp, simp)
  2082 lemma mn_loop_correct: 
  4425       apply(rule_tac disjI2)
  2083   assumes B:  "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]"
  4426       using h2
  2084   and ft: "ft = max (Suc arity) fft"
  4427       apply(rule_tac x = "gs!i" in rev_image_eqI, simp, simp)
  2085   and len: "length xs = arity"
  4428       done
  2086   and ind_all: "\<forall>i\<le>x. (\<forall>xc. ({\<lambda>nl. nl = xs @ i # 0 \<up> (fft - far) @ xc} fap
  4429   next
  2087     {\<lambda>nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \<up> (fft - Suc far) @ xc}))"
  4430     show "aprog = cn_merge_gs (map rec_ci (take i gs)) 
  2088   and exec_ge: "\<forall> i\<le>x. rec_exec f (xs @ [i]) > 0"
  4431               ?pstr [+] ga [+] cp"
  2089   and compile: "rec_ci f = (fap, far, fft)"
  4432       using g3 by simp
  2090   and far: "far = Suc arity"
       
  2091   shows "\<exists> stp. abc_steps_l (0, xs @ 0 # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp = 
       
  2092                (0, xs @ Suc x # 0 \<up> (ft - Suc arity) @ anything)"
       
  2093 proof -
       
  2094   have "\<exists>stp>x. abc_steps_l (0, xs @ 0 # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \<up> (ft - Suc arity) @ anything)"
       
  2095     using assms
       
  2096     by(rule_tac mn_loop_correct', simp_all)
       
  2097   thus "?thesis"
       
  2098     by(auto)
       
  2099 qed
       
  2100 
       
  2101 lemma compile_mn_correct': 
       
  2102   assumes B:  "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]"
       
  2103   and ft: "ft = max (Suc arity) fft"
       
  2104   and len: "length xs = arity"
       
  2105   and termi_f: "terminate f (xs @ [r])"
       
  2106   and f_ind: "\<And>anything. {\<lambda>nl. nl = xs @ r # 0 \<up> (fft - far) @ anything} fap 
       
  2107         {\<lambda>nl. nl = xs @ r # 0 # 0 \<up> (fft - Suc far) @ anything}"
       
  2108   and ind_all: "\<forall>i < r. (\<forall>xc. ({\<lambda>nl. nl = xs @ i # 0 \<up> (fft - far) @ xc} fap
       
  2109     {\<lambda>nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \<up> (fft - Suc far) @ xc}))"
       
  2110   and exec_less: "\<forall> i<r. rec_exec f (xs @ [i]) > 0"
       
  2111   and exec: "rec_exec f (xs @ [r]) = 0"
       
  2112   and compile: "rec_ci f = (fap, far, fft)"
       
  2113   shows "{\<lambda>nl. nl = xs @ 0 \<up> (max (Suc arity) fft - arity) @ anything}
       
  2114     fap @ B
       
  2115     {\<lambda>nl. nl = xs @ rec_exec (Mn arity f) xs # 0 \<up> (max (Suc arity) fft - Suc arity) @ anything}"
       
  2116 proof -
       
  2117   have a: "far = Suc arity"
       
  2118     using len compile termi_f
       
  2119     by(drule_tac param_pattern, auto)
       
  2120   have b: "fft > far"
       
  2121     using compile
       
  2122     by(erule_tac footprint_ge)
       
  2123   have "\<exists> stp. abc_steps_l (0, xs @ 0 # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp = 
       
  2124     (0, xs @ r # 0 \<up> (ft - Suc arity) @ anything)"
       
  2125     using assms a
       
  2126     apply(case_tac r, rule_tac x = 0 in exI, simp add: abc_steps_l.simps, simp)
       
  2127     by(rule_tac mn_loop_correct, auto)  
       
  2128   moreover have 
       
  2129     "\<exists> stp. abc_steps_l (0, xs @ r # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp = 
       
  2130     (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)"
       
  2131   proof -
       
  2132     have "\<exists> stp. abc_steps_l (0, xs @ r # 0 \<up> (ft - Suc arity) @ anything) fap stp =
       
  2133       (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)"
       
  2134     proof -
       
  2135       have "{\<lambda>nl. nl = xs @ r # 0 \<up> (fft - far) @ 0\<up>(ft - fft) @ anything} fap 
       
  2136             {\<lambda>nl. nl = xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (fft - Suc far) @ 0\<up>(ft - fft) @ anything}"
       
  2137         using f_ind exec by simp
       
  2138       thus "?thesis"
       
  2139         using ft a b
       
  2140         apply(drule_tac abc_Hoare_haltE)
       
  2141         by(simp add: replicate_merge_anywhere max_absorb2)
       
  2142     qed
       
  2143     then obtain stp where "abc_steps_l (0, xs @ r # 0 \<up> (ft - Suc arity) @ anything) fap stp =
       
  2144       (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" ..
       
  2145     thus "?thesis"
       
  2146       by(erule_tac abc_append_frist_steps_halt_eq)
  4433   qed
  2147   qed
  4434 qed
  2148   moreover have 
  4435 
  2149     "\<exists> stp. abc_steps_l (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything) (fap @ B) stp = 
  4436 lemma recursive_compile_correct_spec: 
  2150              (length fap + 5, xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)"
  4437   "\<lbrakk>rec_ci re = (ap, ary, fp); 
  2151     using ft a b len B exec
  4438     rec_calc_rel re args r\<rbrakk>
  2152     apply(rule_tac x = 1 in exI, auto)
  4439   \<Longrightarrow> (\<exists> stp. (abc_steps_l (0, args @ 0\<up>(fp - ary)) ap stp) = 
  2153     by(auto simp: abc_steps_l.simps B abc_step_l.simps 
  4440                      (length ap, args@[r]@0\<up>(fp - ary - 1)))"
  2154       abc_fetch.simps nth_append max_absorb2 abc_lm_v.simps abc_lm_s.simps)
  4441 using recursive_compile_correct[of re ap ary fp args r "[]"]
  2155   moreover have "rec_exec (Mn (length xs) f) xs = r"
  4442 by simp
  2156     using exec exec_less
       
  2157     apply(auto simp: rec_exec.simps Least_def)
       
  2158     thm the_equality
       
  2159     apply(rule_tac the_equality, auto)
       
  2160     apply(metis exec_less less_not_refl3 linorder_not_less)
       
  2161     by (metis le_neq_implies_less less_not_refl3)
       
  2162   ultimately show "?thesis"
       
  2163     using ft a b len B exec
       
  2164     apply(auto simp: abc_Hoare_halt_def)
       
  2165     apply(rule_tac x = "stp + stpa + stpb"  in exI)
       
  2166     by(simp add: abc_steps_add replicate_Suc_iff_anywhere max_absorb2 Suc_diff_Suc del: replicate_Suc)
       
  2167 qed
       
  2168 
       
  2169 lemma compile_mn_correct:
       
  2170   assumes len: "length xs = n"
       
  2171   and termi_f: "terminate f (xs @ [r])"
       
  2172   and f_ind: "\<And>ap arity fp anything. rec_ci f = (ap, arity, fp) \<Longrightarrow> 
       
  2173   {\<lambda>nl. nl = xs @ r # 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ r # 0 # 0 \<up> (fp - Suc arity) @ anything}"
       
  2174   and exec: "rec_exec f (xs @ [r]) = 0"
       
  2175   and ind_all: 
       
  2176   "\<forall>i<r. terminate f (xs @ [i]) \<and>
       
  2177   (\<forall>x xa xb. rec_ci f = (x, xa, xb) \<longrightarrow> 
       
  2178   (\<forall>xc. {\<lambda>nl. nl = xs @ i # 0 \<up> (xb - xa) @ xc} x {\<lambda>nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \<up> (xb - Suc xa) @ xc})) \<and>
       
  2179   0 < rec_exec f (xs @ [i])"
       
  2180   and compile: "rec_ci (Mn n f) = (ap, arity, fp)"
       
  2181   shows "{\<lambda>nl. nl = xs @ 0 \<up> (fp - arity) @ anything} ap 
       
  2182   {\<lambda>nl. nl = xs @ rec_exec (Mn n f) xs # 0 \<up> (fp - Suc arity) @ anything}"
       
  2183 proof(case_tac "rec_ci f")
       
  2184   fix fap far fft
       
  2185   assume h: "rec_ci f = (fap, far, fft)"
       
  2186   hence f_newind: 
       
  2187     "\<And>anything. {\<lambda>nl. nl = xs @ r # 0 \<up> (fft - far) @ anything} fap 
       
  2188         {\<lambda>nl. nl = xs @ r # 0 # 0 \<up> (fft - Suc far) @ anything}"
       
  2189     by(rule_tac f_ind, simp)
       
  2190   have newind_all: 
       
  2191     "\<forall>i < r. (\<forall>xc. ({\<lambda>nl. nl = xs @ i # 0 \<up> (fft - far) @ xc} fap
       
  2192     {\<lambda>nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \<up> (fft - Suc far) @ xc}))"
       
  2193     using ind_all h
       
  2194     by(auto)
       
  2195   have all_less: "\<forall> i<r. rec_exec f (xs @ [i]) > 0"
       
  2196     using ind_all by auto
       
  2197   show "?thesis"
       
  2198     using h compile f_newind newind_all all_less len termi_f exec
       
  2199     apply(auto simp: rec_ci.simps)
       
  2200     by(rule_tac compile_mn_correct', auto)
       
  2201 qed
       
  2202     
       
  2203 lemma recursive_compile_correct:
       
  2204    "\<lbrakk>terminate recf args; rec_ci recf = (ap, arity, fp)\<rbrakk>
       
  2205   \<Longrightarrow> {\<lambda> nl. nl = args @ 0\<up>(fp - arity) @ anything} ap 
       
  2206          {\<lambda> nl. nl = args@ rec_exec recf args # 0\<up>(fp - Suc arity) @ anything}"
       
  2207 apply(induct arbitrary: ap arity fp anything r rule: terminate.induct)
       
  2208 apply(simp_all add: compile_s_correct compile_z_correct compile_id_correct 
       
  2209                     compile_cn_correct compile_pr_correct compile_mn_correct)
       
  2210 done
  4443 
  2211 
  4444 definition dummy_abc :: "nat \<Rightarrow> abc_inst list"
  2212 definition dummy_abc :: "nat \<Rightarrow> abc_inst list"
  4445 where
  2213 where
  4446 "dummy_abc k = [Inc k, Dec k 0, Goto 3]"
  2214 "dummy_abc k = [Inc k, Dec k 0, Goto 3]"
  4447 
  2215 
  4448 definition abc_list_crsp:: "nat list \<Rightarrow> nat list \<Rightarrow> bool"
  2216 definition abc_list_crsp:: "nat list \<Rightarrow> nat list \<Rightarrow> bool"
  4449   where
  2217   where
  4450   "abc_list_crsp xs ys = (\<exists> n. xs = ys @ 0\<up>n \<or> ys = xs @ 0\<up>n)"
  2218   "abc_list_crsp xs ys = (\<exists> n. xs = ys @ 0\<up>n \<or> ys = xs @ 0\<up>n)"
  4451 
  2219 
  4452 lemma [intro]: "abc_list_crsp (lm @ 0\<up>m) lm"
  2220 lemma abc_list_crsp_simp1[intro]: "abc_list_crsp (lm @ 0\<up>m) lm"
  4453 apply(auto simp: abc_list_crsp_def)
  2221 by(auto simp: abc_list_crsp_def)
  4454 done
  2222 
  4455 
  2223 
  4456 lemma abc_list_crsp_lm_v: 
  2224 lemma abc_list_crsp_lm_v: 
  4457   "abc_list_crsp lma lmb \<Longrightarrow> abc_lm_v lma n = abc_lm_v lmb n"
  2225   "abc_list_crsp lma lmb \<Longrightarrow> abc_lm_v lma n = abc_lm_v lmb n"
  4458 apply(auto simp: abc_list_crsp_def abc_lm_v.simps 
  2226 by(auto simp: abc_list_crsp_def abc_lm_v.simps 
  4459                  nth_append)
  2227                  nth_append)
  4460 done
  2228 
  4461 
  2229 
  4462 lemma  rep_app_cons_iff: 
  2230 lemma abc_list_crsp_elim: 
  4463   "k < n \<Longrightarrow> replicate n a[k:=b] = 
  2231   "\<lbrakk>abc_list_crsp lma lmb; \<exists> n. lma = lmb @ 0\<up>n \<or> lmb = lma @ 0\<up>n \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
  4464           replicate k a @ b # replicate (n - k - 1) a"
  2232 by(auto simp: abc_list_crsp_def)
  4465 apply(induct n arbitrary: k, simp)
  2233 
  4466 apply(simp split:nat.splits)
  2234 lemma [simp]: 
  4467 done
  2235   "\<lbrakk>abc_list_crsp lma lmb; m < length lma; m < length lmb\<rbrakk> \<Longrightarrow>
       
  2236           abc_list_crsp (lma[m := n]) (lmb[m := n])"
       
  2237 by(auto simp: abc_list_crsp_def list_update_append)
       
  2238 
       
  2239 lemma [simp]: 
       
  2240   "\<lbrakk>abc_list_crsp lma lmb; m < length lma; \<not> m < length lmb\<rbrakk> \<Longrightarrow> 
       
  2241   abc_list_crsp (lma[m := n]) (lmb @ 0 \<up> (m - length lmb) @ [n])"
       
  2242 apply(auto simp: abc_list_crsp_def list_update_append)
       
  2243 apply(rule_tac x = "na + length lmb - Suc m" in exI)
       
  2244 apply(rule_tac disjI1)
       
  2245 apply(simp add: upd_conv_take_nth_drop min_absorb1)
       
  2246 done
       
  2247 
       
  2248 lemma [simp]:
       
  2249   "\<lbrakk>abc_list_crsp lma lmb; \<not> m < length lma; m < length lmb\<rbrakk> \<Longrightarrow> 
       
  2250   abc_list_crsp (lma @ 0 \<up> (m - length lma) @ [n]) (lmb[m := n])"
       
  2251 apply(auto simp: abc_list_crsp_def list_update_append)
       
  2252 apply(rule_tac x = "na + length lma - Suc m" in exI)
       
  2253 apply(rule_tac disjI2)
       
  2254 apply(simp add: upd_conv_take_nth_drop min_absorb1)
       
  2255 done
       
  2256 
       
  2257 lemma [simp]: "\<lbrakk>abc_list_crsp lma lmb; \<not> m < length lma; \<not> m < length lmb\<rbrakk> \<Longrightarrow> 
       
  2258   abc_list_crsp (lma @ 0 \<up> (m - length lma) @ [n]) (lmb @ 0 \<up> (m - length lmb) @ [n])"
       
  2259   by(auto simp: abc_list_crsp_def list_update_append replicate_merge_anywhere)
  4468 
  2260 
  4469 lemma abc_list_crsp_lm_s: 
  2261 lemma abc_list_crsp_lm_s: 
  4470   "abc_list_crsp lma lmb \<Longrightarrow> 
  2262   "abc_list_crsp lma lmb \<Longrightarrow> 
  4471       abc_list_crsp (abc_lm_s lma m n) (abc_lm_s lmb m n)"
  2263       abc_list_crsp (abc_lm_s lma m n) (abc_lm_s lmb m n)"
  4472 apply(auto simp: abc_list_crsp_def abc_lm_v.simps abc_lm_s.simps)
  2264 by(auto simp: abc_lm_s.simps)
  4473 apply(simp_all add: list_update_append, auto)
       
  4474 proof -
       
  4475   fix na
       
  4476   assume h: "m < length lmb + na" " \<not> m < length lmb"
       
  4477   hence "m - length lmb < na" by simp
       
  4478   hence "replicate na 0[(m- length lmb):= n] = 
       
  4479            replicate (m - length lmb) 0 @ n # 
       
  4480               replicate (na - (m - length lmb) - 1) 0"
       
  4481     apply(erule_tac rep_app_cons_iff)
       
  4482     done
       
  4483   thus "\<exists>nb. replicate na 0[m - length lmb := n] =
       
  4484                  replicate (m - length lmb) 0 @ n # replicate nb 0 \<or>
       
  4485                  replicate (m - length lmb) 0 @ [n] =
       
  4486                  replicate na 0[m - length lmb := n] @ replicate nb 0"
       
  4487     apply(auto)
       
  4488     done
       
  4489 next
       
  4490   fix na
       
  4491   assume h: "\<not> m < length lmb + na"
       
  4492   show 
       
  4493     "\<exists>nb. replicate na 0 @ replicate (m - (length lmb + na)) 0 @ [n] =
       
  4494            replicate (m - length lmb) 0 @ n # replicate nb 0 \<or>
       
  4495           replicate (m - length lmb) 0 @ [n] =
       
  4496             replicate na 0 @
       
  4497             replicate (m - (length lmb + na)) 0 @ n # replicate nb 0"
       
  4498     apply(rule_tac x = 0 in exI, simp, auto)
       
  4499     using h
       
  4500     apply(simp add: replicate_add[THEN sym])
       
  4501     done
       
  4502 next
       
  4503   fix na
       
  4504   assume h: "\<not> m < length lma" "m < length lma + na"
       
  4505   hence "m - length lma < na" by simp
       
  4506   hence 
       
  4507     "replicate na 0[(m- length lma):= n] = replicate (m - length lma) 
       
  4508                   0 @ n # replicate (na - (m - length lma) - 1) 0"
       
  4509     apply(erule_tac rep_app_cons_iff)
       
  4510     done
       
  4511   thus "\<exists>nb. replicate (m - length lma) 0 @ [n] =
       
  4512                  replicate na 0[m - length lma := n] @ replicate nb 0 
       
  4513            \<or> replicate na 0[m - length lma := n] =
       
  4514                  replicate (m - length lma) 0 @ n # replicate nb 0"
       
  4515     apply(auto)
       
  4516     done
       
  4517 next
       
  4518   fix na
       
  4519   assume "\<not> m < length lma + na"
       
  4520   thus " \<exists>nb. replicate (m - length lma) 0 @ [n] =
       
  4521             replicate na 0 @
       
  4522             replicate (m - (length lma + na)) 0 @ n # replicate nb 0 
       
  4523         \<or>   replicate na 0 @ 
       
  4524                replicate (m - (length lma + na)) 0 @ [n] =
       
  4525             replicate (m - length lma) 0 @ n # replicate nb 0"
       
  4526     apply(rule_tac x = 0 in exI, simp, auto)
       
  4527     apply(simp add: replicate_add[THEN sym])
       
  4528     done
       
  4529 qed
       
  4530 
  2265 
  4531 lemma abc_list_crsp_step: 
  2266 lemma abc_list_crsp_step: 
  4532   "\<lbrakk>abc_list_crsp lma lmb; abc_step_l (aa, lma) i = (a, lma'); 
  2267   "\<lbrakk>abc_list_crsp lma lmb; abc_step_l (aa, lma) i = (a, lma'); 
  4533     abc_step_l (aa, lmb) i = (a', lmb')\<rbrakk>
  2268     abc_step_l (aa, lmb) i = (a', lmb')\<rbrakk>
  4534     \<Longrightarrow> a' = a \<and> abc_list_crsp lma' lmb'"
  2269     \<Longrightarrow> a' = a \<and> abc_list_crsp lma' lmb'"
  4575           case_tac "abc_step_l (aa, lma) (abc_fetch aa aprog)", simp)
  2310           case_tac "abc_step_l (aa, lma) (abc_fetch aa aprog)", simp)
  4576     apply(rule_tac abc_list_crsp_step, auto)
  2311     apply(rule_tac abc_list_crsp_step, auto)
  4577     done
  2312     done
  4578 qed
  2313 qed
  4579 
  2314 
  4580 lemma recursive_compile_correct_norm: 
  2315 lemma list_crsp_simp2: "abc_list_crsp (lm1 @ 0\<up>n) lm2 \<Longrightarrow> abc_list_crsp lm1 lm2"
  4581   "\<lbrakk>rec_ci re = (aprog, rs_pos, a_md);  
  2316 proof(induct n)
  4582    rec_calc_rel re lm rs\<rbrakk>
  2317   case 0
  4583   \<Longrightarrow> (\<exists> stp lm' m. (abc_steps_l (0, lm) aprog stp) = 
  2318   thus "?case"
  4584   (length aprog, lm') \<and> abc_list_crsp lm' (lm @ rs # 0\<up>m))"
  2319     by(auto simp: abc_list_crsp_def)
  4585 apply(frule_tac recursive_compile_correct_spec, auto)
  2320 next
  4586 apply(drule_tac abc_list_crsp_steps)
  2321   case (Suc n)
  4587 apply(rule_tac rec_ci_not_null, simp)
  2322   have ind: "abc_list_crsp (lm1 @ 0 \<up> n) lm2 \<Longrightarrow> abc_list_crsp lm1 lm2" by fact
  4588 apply(erule_tac exE, rule_tac x = stp in exI, 
  2323   have h: "abc_list_crsp (lm1 @ 0 \<up> Suc n) lm2" by fact
  4589   auto simp: abc_list_crsp_def)
  2324   then have "abc_list_crsp (lm1 @ 0 \<up> n) lm2"
  4590 done
  2325     apply(auto simp: exp_suc abc_list_crsp_def del: replicate_Suc)
  4591 
  2326     apply(case_tac n, simp_all add: exp_suc replicate_Suc[THEN sym] del: replicate_Suc, auto)
  4592 lemma [simp]: "length (dummy_abc (length lm)) = 3"
  2327     apply(rule_tac x = 1 in exI, simp)
  4593 apply(simp add: dummy_abc_def)
  2328     by(rule_tac x = "Suc n" in exI, simp,  simp add: exp_suc replicate_Suc[THEN sym] del: replicate_Suc)
  4594 done
  2329   thus "?case"
  4595 
  2330     using ind
  4596 lemma [simp]: "dummy_abc (length lm) \<noteq> []"
  2331     by auto
  4597 apply(simp add: dummy_abc_def)
  2332 qed
  4598 done
  2333 
  4599 
  2334 lemma recursive_compile_correct_norm': 
  4600 lemma dummy_abc_steps_ex: 
  2335   "\<lbrakk>rec_ci f = (ap, arity, ft);  
  4601   "\<exists>bstp. abc_steps_l (0, lm') (dummy_abc (length lm)) bstp = 
  2336     terminate f args\<rbrakk>
  4602   ((Suc (Suc (Suc 0))), abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)))"
  2337   \<Longrightarrow> \<exists> stp rl. (abc_steps_l (0, args) ap stp) = (length ap, rl) \<and> abc_list_crsp (args @ [rec_exec f args]) rl"
  4603 apply(rule_tac x = "Suc (Suc (Suc 0))" in exI)
  2338   using recursive_compile_correct[of f args ap arity ft "[]"]
  4604 apply(auto simp: abc_steps_l.simps abc_step_l.simps 
  2339 apply(auto simp: abc_Hoare_halt_def)
  4605   dummy_abc_def abc_fetch.simps)
  2340 apply(rule_tac x = n in exI)
  4606 apply(auto simp: abc_lm_s.simps abc_lm_v.simps nth_append)
  2341 apply(case_tac "abc_steps_l (0, args @ 0 \<up> (ft - arity)) ap n", auto)
  4607 apply(simp add: butlast_append)
  2342 apply(drule_tac abc_list_crsp_steps, auto)
       
  2343 apply(rule_tac list_crsp_simp2, auto)
       
  2344 done
       
  2345 
       
  2346 lemma [simp]:
       
  2347   assumes a: "args @ [rec_exec f args] = lm @ 0 \<up> n"
       
  2348   and b: "length args < length lm"
       
  2349   shows "\<exists>m. lm = args @ rec_exec f args # 0 \<up> m"
       
  2350 using assms
       
  2351 apply(case_tac n, simp)
       
  2352 apply(rule_tac x = 0 in exI, simp)
       
  2353 apply(drule_tac length_equal, simp)
  4608 done
  2354 done
  4609 
  2355 
  4610 lemma [simp]: 
  2356 lemma [simp]: 
  4611   "\<lbrakk>Suc (length lm) - length lm' \<le> n; \<not> length lm < length lm'; lm @ rs # 0 \<up> m = lm' @ 0 \<up> n\<rbrakk> 
  2357 "\<lbrakk>args @ [rec_exec f args] = lm @ 0 \<up> n; \<not> length args < length lm\<rbrakk>
  4612   \<Longrightarrow> lm' @ 0 \<up> Suc (length lm - length lm') = lm @ [rs]"
  2358   \<Longrightarrow> \<exists>m. (lm @ 0 \<up> (length args - length lm) @ [Suc 0])[length args := 0] =
  4613 apply(subgoal_tac "n > m")
  2359   args @ rec_exec f args # 0 \<up> m"
  4614 apply(subgoal_tac "\<exists> d. n = d + m", erule_tac exE)
  2360 apply(case_tac n, simp_all add: exp_suc list_update_append list_update.simps del: replicate_Suc)
  4615 apply(simp add: replicate_add)
  2361 apply(subgoal_tac "length args = Suc (length lm)", simp)
  4616 apply(drule_tac length_equal, simp)
  2362 apply(rule_tac x = "Suc (Suc 0)" in exI, simp)
  4617 apply(simp add: replicate_Suc[THEN sym] del: replicate_Suc)
  2363 apply(drule_tac length_equal, simp, auto)
  4618 apply(rule_tac x = "n - m" in exI, simp)
  2364 done
  4619 apply(drule_tac length_equal, simp)
  2365 
  4620 done
  2366 lemma compile_append_dummy_correct: 
  4621 
  2367   assumes compile: "rec_ci f = (ap, ary, fp)"
  4622 lemma [elim]: 
  2368   and termi: "terminate f args"
  4623   "lm @ rs # 0\<up>m = lm' @ 0\<up>n \<Longrightarrow> 
  2369   shows "{\<lambda> nl. nl = args} (ap [+] dummy_abc (length args)) {\<lambda> nl. (\<exists> m. nl = args @ rec_exec f args # 0\<up>m)}"
  4624   \<exists>m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) = 
  2370 proof(rule_tac abc_Hoare_plus_halt)
  4625                             lm @ rs # 0\<up>m"
  2371   show "{\<lambda>nl. nl = args} ap {\<lambda> nl. abc_list_crsp (args @ [rec_exec f args]) nl}"
  4626 proof(cases "length lm' > length lm")
  2372     using compile termi recursive_compile_correct_norm'[of f ap ary fp args]
  4627   case True 
  2373     apply(auto simp: abc_Hoare_halt_def)
  4628   assume h: "lm @ rs # 0\<up>m = lm' @ 0\<up>n" "length lm < length lm'"
  2374     by(rule_tac x = stp in exI, simp)
  4629   hence "m \<ge> n"
       
  4630     apply(drule_tac length_equal)
       
  4631     apply(simp)
       
  4632     done
       
  4633   hence "\<exists> d. m = d + n"
       
  4634     apply(rule_tac x = "m - n" in exI, simp)
       
  4635     done
       
  4636   from this obtain d where "m = d + n" ..
       
  4637   from h and this show "?thesis"
       
  4638     apply(auto simp: abc_lm_s.simps abc_lm_v.simps 
       
  4639                      replicate_add)
       
  4640     done
       
  4641 next
  2375 next
  4642   case False
  2376   show "{abc_list_crsp (args @ [rec_exec f args])} dummy_abc (length args) 
  4643   assume h:"lm @ rs # 0\<up>m = lm' @ 0\<up>n" 
  2377     {\<lambda>nl. \<exists>m. nl = args @ rec_exec f args # 0 \<up> m}"
  4644     and    g: "\<not> length lm < length lm'"
  2378     apply(auto simp: dummy_abc_def abc_Hoare_halt_def)
  4645   have "take (Suc (length lm)) (lm @ rs # 0\<up>m) = 
  2379     apply(rule_tac x = 3 in exI)
  4646                         take (Suc (length lm)) (lm' @ 0\<up>n)"
  2380     by(auto simp: abc_steps_l.simps abc_list_crsp_def abc_step_l.simps numeral_3_eq_3 abc_fetch.simps
  4647     using h by simp
  2381                      abc_lm_v.simps nth_append abc_lm_s.simps)
  4648   moreover have "n \<ge> (Suc (length lm) - length lm')"
  2382 qed
  4649     using h g
  2383 
  4650     apply(drule_tac length_equal)
  2384 lemma cn_merge_gs_split: 
  4651     apply(simp)
  2385   "\<lbrakk>i < length gs; rec_ci (gs!i) = (ga, gb, gc)\<rbrakk> \<Longrightarrow> 
  4652     done
  2386   cn_merge_gs (map rec_ci gs) p =  cn_merge_gs (map rec_ci (take i gs)) p [+] (ga [+] 
  4653   ultimately show 
  2387        mv_box gb (p + i)) [+]  cn_merge_gs (map rec_ci (drop (Suc i) gs)) (p + Suc i)"
  4654     "\<exists>m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) =
  2388 apply(induct i arbitrary: gs p, case_tac gs, simp, simp)
  4655                                                        lm @ rs # 0\<up>m"
  2389 apply(case_tac gs, simp, case_tac "rec_ci a", 
  4656     using g h
  2390        simp add: abc_comp_commute[THEN sym])
  4657     apply(simp add: abc_lm_s.simps abc_lm_v.simps min_def)
  2391 done
  4658     apply(rule_tac x = 0 in exI, 
  2392 
  4659       simp add:replicate_append_same replicate_Suc[THEN sym]
  2393 lemma cn_unhalt_case:
  4660                                       del:replicate_Suc)
  2394   assumes compile1: "rec_ci (Cn n f gs) = (ap, ar, ft) \<and> length args = ar"
  4661     done    
  2395   and g: "i < length gs"
  4662 qed
  2396   and compile2: "rec_ci (gs!i) = (gap, gar, gft) \<and> gar = length args"
  4663 
  2397   and g_unhalt: "\<And> anything. {\<lambda> nl. nl = args @ 0\<up>(gft - gar) @ anything} gap \<up>"
  4664 lemma [elim]: 
  2398   and g_ind: "\<And> apj arj ftj j anything. \<lbrakk>j < i; rec_ci (gs!j) = (apj, arj, ftj)\<rbrakk> 
  4665   "abc_list_crsp lm' (lm @ rs # 0\<up>m)
  2399   \<Longrightarrow> {\<lambda> nl. nl = args @ 0\<up>(ftj - arj) @ anything} apj {\<lambda> nl. nl = args @ rec_exec (gs!j) args # 0\<up>(ftj - Suc arj) @ anything}"
  4666   \<Longrightarrow> \<exists>m. abc_lm_s lm' (length lm) (abc_lm_v lm' (length lm)) 
  2400   and all_termi: "\<forall> j<i. terminate (gs!j) args"
  4667              = lm @ rs # 0\<up>m"
  2401   shows "{\<lambda> nl. nl = args @ 0\<up>(ft - ar) @ anything} ap \<up>"
  4668 apply(auto simp: abc_list_crsp_def)
  2402 using compile1
  4669 apply(simp add: abc_lm_v.simps abc_lm_s.simps)
  2403 apply(case_tac "rec_ci f", auto simp: rec_ci.simps abc_comp_commute)
  4670 apply(rule_tac x =  "m + n" in exI, 
  2404 proof(rule_tac abc_Hoare_plus_unhalt1)
  4671       simp add: replicate_add)
  2405   fix fap far fft
  4672 done
  2406   let ?ft = "max (Suc (length args)) (Max (insert fft ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
  4673 
  2407   let ?Q = "\<lambda>nl. nl = args @ 0\<up> (?ft - length args) @ map (\<lambda>i. rec_exec i args) (take i gs) @ 
  4674 lemma abc_append_dummy_complie:
  2408     0\<up>(length gs - i) @ 0\<up> Suc (length args) @ anything"
  4675   "\<lbrakk>rec_ci recf = (ap, ary, fp);  
  2409   have "cn_merge_gs (map rec_ci gs) ?ft = 
  4676     rec_calc_rel recf args r; 
  2410     cn_merge_gs (map rec_ci (take i gs)) ?ft [+] (gap [+] 
  4677     length args = k\<rbrakk>
  2411     mv_box gar (?ft + i)) [+]  cn_merge_gs (map rec_ci (drop (Suc i) gs)) (?ft + Suc i)"
  4678   \<Longrightarrow> (\<exists> stp m. (abc_steps_l (0, args) (ap [+] dummy_abc k) stp) = 
  2412     using g compile2 cn_merge_gs_split by simp
  4679                   (length ap + 3, args @ r # 0\<up>m))"
  2413   thus "{\<lambda>nl. nl = args @ 0 # 0 \<up> (?ft + length gs) @ anything} (cn_merge_gs (map rec_ci gs) ?ft) \<up>"
  4680 apply(drule_tac recursive_compile_correct_norm, auto simp: numeral_3_eq_3)
  2414   proof(simp, rule_tac abc_Hoare_plus_unhalt1, rule_tac abc_Hoare_plus_unhalt2, 
  4681 proof -
  2415               rule_tac abc_Hoare_plus_unhalt1)
  4682   fix stp lm' m
  2416     let ?Q_tmp = "\<lambda>nl. nl = args @ 0\<up> (gft - gar) @ 0\<up>(?ft - (length args) - (gft -gar)) @ map (\<lambda>i. rec_exec i args) (take i gs) @ 
  4683   assume h: "rec_calc_rel recf args r"  
  2417       0\<up>(length gs - i) @ 0\<up> Suc (length args) @ anything"
  4684     "abc_steps_l (0, args) ap stp = (length ap, lm')" 
  2418     have a: "{?Q_tmp} gap \<up>"
  4685     "abc_list_crsp lm' (args @ r # 0\<up>m)"
  2419       using g_unhalt[of "0 \<up> (?ft - (length args) - (gft - gar)) @
  4686   have "\<exists>stp. abc_steps_l (0, args) (ap [+] 
  2420         map (\<lambda>i. rec_exec i args) (take i gs) @ 0 \<up> (length gs - i) @ 0 \<up> Suc (length args) @ anything"]
  4687     (dummy_abc (length args))) stp = (length ap + 3, 
  2421       by simp
  4688     abc_lm_s lm' (length args) (abc_lm_v lm' (length args)))"
  2422     moreover have "?ft \<ge> gft"
  4689     using h
  2423       using g compile2
  4690     apply(rule_tac bm = lm' in abc_append_exc2,
  2424       apply(rule_tac min_max.le_supI2, rule_tac Max_ge, simp, rule_tac insertI2)
  4691           auto intro: dummy_abc_steps_ex simp: numeral_3_eq_3)
  2425       apply(rule_tac  x = "rec_ci (gs ! i)" in image_eqI, simp)
  4692     done
  2426       by(rule_tac x = "gs!i"  in image_eqI, simp, simp)
  4693   thus "\<exists>stp m. abc_steps_l (0, args) (ap [+] 
  2427     then have b:"?Q_tmp = ?Q"
  4694     dummy_abc (length args)) stp = (Suc (Suc (Suc (length ap))), args @ r # 0\<up>m)"
  2428       using compile2
  4695     using h
  2429       apply(rule_tac arg_cong)
  4696     apply(erule_tac exE)
  2430       by(simp add: replicate_merge_anywhere)
  4697     apply(rule_tac x = stpa in exI, auto)
  2431     thus "{?Q} gap \<up>"
  4698     done
  2432       using a by simp
  4699 qed
  2433   next
  4700 
  2434     show "{\<lambda>nl. nl = args @ 0 # 0 \<up> (?ft + length gs) @ anything} 
  4701 lemma [simp]: "length (dummy_abc k) = 3"
  2435       cn_merge_gs (map rec_ci (take i gs)) ?ft
  4702 apply(simp add: dummy_abc_def)
  2436        {\<lambda>nl. nl = args @ 0 \<up> (?ft - length args) @
  4703 done
  2437       map (\<lambda>i. rec_exec i args) (take i gs) @ 0 \<up> (length gs - i) @ 0 \<up> Suc (length args) @ anything}"
  4704 
  2438       using all_termi
  4705 lemma [simp]: "length args = k \<Longrightarrow> abc_lm_v (args @ r # 0\<up>m) k = r "
  2439       apply(rule_tac compile_cn_gs_correct', auto simp: set_conv_nth)
  4706 apply(simp add: abc_lm_v.simps nth_append)
  2440       by(drule_tac apj = x and arj = xa and  ftj = xb and j = ia and anything = xc in g_ind, auto)
  4707 done
  2441   qed
  4708 
  2442 qed
  4709 lemma [simp]: "crsp (layout_of (ap [+] dummy_abc k)) (0, args)
  2443 
  4710   (Suc 0, Bk # Bk # ires, <args> @ Bk \<up> rn) ires"
  2444 
  4711 apply(auto simp: crsp.simps start_of.simps)
  2445 
  4712 done
  2446 lemma mn_unhalt_case':
  4713 
  2447   assumes compile: "rec_ci f = (a, b, c)"
  4714 (* cccc *)
  2448   and all_termi: "\<forall>i. terminate f (args @ [i]) \<and> 0 < rec_exec f (args @ [i])"
       
  2449   and B: "B = [Dec (Suc (length args)) (length a + 5), Dec (Suc (length args)) (length a + 3), 
       
  2450   Goto (Suc (length a)), Inc (length args), Goto 0]"
       
  2451   shows "{\<lambda>nl. nl = args @ 0 \<up> (max (Suc (length args)) c - length args) @ anything}
       
  2452   a @ B \<up>"
       
  2453 proof(rule_tac abc_Hoare_unhaltI, auto)
       
  2454   fix n
       
  2455   have a:  "b = Suc (length args)"
       
  2456     using all_termi compile
       
  2457     apply(erule_tac x = 0 in allE)
       
  2458     by(auto, drule_tac param_pattern,auto)
       
  2459   moreover have b: "c > b"
       
  2460     using compile by(elim footprint_ge)
       
  2461   ultimately have c: "max (Suc (length args)) c = c" by arith
       
  2462   have "\<exists> stp > n. abc_steps_l (0, args @ 0 # 0\<up>(c - Suc (length args)) @ anything) (a @ B) stp
       
  2463          = (0, args @ Suc n # 0\<up>(c - Suc (length args)) @ anything)"
       
  2464     using assms a b c
       
  2465   proof(rule_tac mn_loop_correct', auto)
       
  2466     fix i xc
       
  2467     show "{\<lambda>nl. nl = args @ i # 0 \<up> (c - Suc (length args)) @ xc} a 
       
  2468       {\<lambda>nl. nl = args @ i # rec_exec f (args @ [i]) # 0 \<up> (c - Suc (Suc (length args))) @ xc}"
       
  2469       using all_termi recursive_compile_correct[of f "args @ [i]" a b c xc] compile a
       
  2470       by(simp)
       
  2471   qed
       
  2472   then obtain stp where d: "stp > n \<and> abc_steps_l (0, args @ 0 # 0\<up>(c - Suc (length args)) @ anything) (a @ B) stp
       
  2473          = (0, args @ Suc n # 0\<up>(c - Suc (length args)) @ anything)" ..
       
  2474   then obtain d where e: "stp = n + Suc d"
       
  2475     by (metis add_Suc_right less_iff_Suc_add)
       
  2476   obtain s nl where f: "abc_steps_l (0, args @ 0 # 0\<up>(c - Suc (length args)) @ anything) (a @ B) n = (s, nl)"
       
  2477     by (metis prod.exhaust)
       
  2478   have g: "s < length (a @ B)"
       
  2479     using d e f
       
  2480     apply(rule_tac classical, simp only: abc_steps_add)
       
  2481     by(simp add: halt_steps2 leI)
       
  2482   from f g show "abc_notfinal (abc_steps_l (0, args @ 0 \<up> 
       
  2483     (max (Suc (length args)) c - length args) @ anything) (a @ B) n) (a @ B)"
       
  2484     using c b a
       
  2485     by(simp add: replicate_Suc_iff_anywhere Suc_diff_Suc del: replicate_Suc)
       
  2486 qed
       
  2487     
       
  2488 lemma mn_unhalt_case: 
       
  2489   assumes compile: "rec_ci (Mn n f) = (ap, ar, ft) \<and> length args = ar"
       
  2490   and all_term: "\<forall> i. terminate f (args @ [i]) \<and> rec_exec f (args @ [i]) > 0"
       
  2491   shows "{\<lambda> nl. nl = args @ 0\<up>(ft - ar) @ anything} ap \<up> "
       
  2492   using assms
       
  2493 apply(case_tac "rec_ci f", auto simp: rec_ci.simps abc_comp_commute)
       
  2494 by(rule_tac mn_unhalt_case', simp_all)
  4715 
  2495 
  4716 fun tm_of_rec :: "recf \<Rightarrow> instr list"
  2496 fun tm_of_rec :: "recf \<Rightarrow> instr list"
  4717 where "tm_of_rec recf = (let (ap, k, fp) = rec_ci recf in
  2497 where "tm_of_rec recf = (let (ap, k, fp) = rec_ci recf in
  4718                          let tp = tm_of (ap [+] dummy_abc k) in 
  2498                          let tp = tm_of (ap [+] dummy_abc k) in 
  4719                              tp @ (shift (mopup k) (length tp div 2)))"
  2499                            tp @ (shift (mopup k) (length tp div 2)))"
  4720 
  2500 
  4721 lemma recursive_compile_to_tm_correct: 
  2501 lemma recursive_compile_to_tm_correct1: 
  4722   "\<lbrakk>rec_ci recf = (ap, ary, fp); 
  2502   assumes  compile: "rec_ci recf = (ap, ary, fp)"
  4723     rec_calc_rel recf args r;
  2503   and termi: " terminate recf args"
  4724     length args = k;
  2504   and tp: "tp = tm_of (ap [+] dummy_abc (length args))"
  4725     ly = layout_of (ap [+] dummy_abc k);
  2505   shows "\<exists> stp m l. steps0 (Suc 0, Bk # Bk # ires, <args> @ Bk\<up>rn)
  4726     tp = tm_of (ap [+] dummy_abc k)\<rbrakk>
  2506   (tp @ shift (mopup (length args)) (length tp div 2)) stp = (0, Bk\<up>m @ Bk # Bk # ires, Oc\<up>Suc (rec_exec recf args) @ Bk\<up>l)"
  4727   \<Longrightarrow> \<exists> stp m l. steps0 (Suc 0, Bk # Bk # ires, <args> @ Bk\<up>rn)
  2507 proof -
  4728   (tp @ shift (mopup k) (length tp div 2)) stp
  2508   have "{\<lambda>nl. nl = args} ap [+] dummy_abc (length args) {\<lambda>nl. \<exists>m. nl = args @ rec_exec recf args # 0 \<up> m}"
  4729   = (0, Bk\<up>m @ Bk # Bk # ires, Oc\<up>Suc r @ Bk\<up>l)"
  2509     using compile termi compile
  4730   using abc_append_dummy_complie[of recf ap ary fp args r k]
  2510     by(rule_tac compile_append_dummy_correct, auto)
  4731 apply(simp)
  2511   then obtain stp m where h: "abc_steps_l (0, args) (ap [+] dummy_abc (length args)) stp = 
  4732 apply(erule_tac exE)+
  2512     (length (ap [+] dummy_abc (length args)), args @ rec_exec recf args # 0\<up>m) "
  4733 apply(frule_tac tp = tp and n = k 
  2513     apply(simp add: abc_Hoare_halt_def, auto)
  4734                and ires = ires in compile_correct_halt, simp_all add: length_append)
  2514     by(case_tac "abc_steps_l (0, args) (ap [+] dummy_abc (length args)) n", auto)
  4735 apply(simp_all add: length_append)
  2515   thus "?thesis"
  4736 done
  2516     using assms tp
       
  2517     by(rule_tac  lm = args and stp = stp and am = "args @ rec_exec recf args # 0 \<up> m"
       
  2518       in compile_correct_halt, auto simp: crsp.simps start_of.simps length_abc_comp abc_lm_v.simps)
       
  2519 qed
  4737 
  2520 
  4738 lemma recursive_compile_to_tm_correct2: 
  2521 lemma recursive_compile_to_tm_correct2: 
  4739   assumes "rec_ci recf = (ap, ary, fp)" 
  2522   assumes  compile: "rec_ci recf = (ap, ary, fp)"
  4740   and     "rec_calc_rel recf args r"
  2523   and termi: " terminate recf args"
  4741   and     "length args = k"
  2524   shows "\<exists> stp m l. steps0 (Suc 0, [Bk, Bk], <args>) (tm_of_rec recf) stp = 
  4742   and     "tp = tm_of (ap [+] dummy_abc k)"
  2525                      (0, Bk\<up>Suc (Suc m), Oc\<up>Suc (rec_exec recf args) @ Bk\<up>l)"
  4743   shows "\<exists> m n. {\<lambda>tp. tp = ([Bk, Bk], <args>)}
  2526 using recursive_compile_to_tm_correct1[of recf ap ary fp args "tm_of (ap [+] dummy_abc (length args))" "[]" 0]
  4744              (tp @ (shift (mopup k) (length tp div 2)))
  2527 assms param_pattern[of recf args ap ary fp]
  4745              {\<lambda>tp. tp = (Bk \<up> m, Oc \<up> (Suc r) @ Bk \<up> n)}"
  2528 by(simp add: exp_suc replicate_Suc[THEN sym] del: replicate_Suc tm_of_rec_def)
  4746 using recursive_compile_to_tm_correct[where ires="[]" and rn="0", OF assms(1-3) _ assms(4)]
       
  4747 apply(simp add: Hoare_halt_def)
       
  4748 apply(drule_tac x="layout_of (ap [+] dummy_abc k)" in meta_spec)
       
  4749 apply(auto)
       
  4750 apply(rule_tac x="m + 2" in exI)
       
  4751 apply(rule_tac x="l" in exI)
       
  4752 apply(rule_tac x="stp" in exI)
       
  4753 apply(auto)
       
  4754 by (metis append_Nil2 replicate_app_Cons_same)
       
  4755 
  2529 
  4756 lemma recursive_compile_to_tm_correct3: 
  2530 lemma recursive_compile_to_tm_correct3: 
  4757   assumes "rec_calc_rel recf args r"
  2531   assumes compile: "rec_ci recf = (ap, ary, fp)"
  4758   shows "{\<lambda>tp. tp = ([Bk, Bk], <args>)} tm_of_rec recf {\<lambda>tp. \<exists>k l. tp = (Bk \<up> k, <r> @ Bk \<up> l)}"
  2532   and termi: "terminate recf args"
  4759 using recursive_compile_to_tm_correct2[OF _ assms] 
  2533   shows "{\<lambda> (l, r). l = [Bk, Bk] \<and> r = <args>} (tm_of_rec recf) 
  4760 apply(auto)
  2534          {\<lambda> (l, r). \<exists> m i. l = Bk\<up>(Suc (Suc m)) \<and> r = Oc\<up>Suc (rec_exec recf args) @ Bk \<up> i}"
  4761 apply(case_tac "rec_ci recf")
  2535 using recursive_compile_to_tm_correct2[of recf ap ary fp args] assms
  4762 apply(auto)
  2536 apply(simp add: Hoare_halt_def, auto)
  4763 apply(drule_tac x="a" in meta_spec)
  2537 apply(rule_tac x = stp in exI, simp)
  4764 apply(drule_tac x="b" in meta_spec)
  2538 done
  4765 apply(drule_tac x="c" in meta_spec)
       
  4766 apply(drule_tac x="length args" in meta_spec)
       
  4767 apply(drule_tac x="tm_of (a [+] dummy_abc (length args))" in meta_spec)
       
  4768 apply(auto)
       
  4769 apply(simp add: tape_of_nat_abv)
       
  4770 apply(subgoal_tac "b = length args")
       
  4771 apply(simp add: Hoare_halt_def)
       
  4772 apply(auto)[1]
       
  4773 apply(rule_tac x="na" in exI)
       
  4774 apply(auto)[1]
       
  4775 apply(case_tac "steps0 (Suc 0, [Bk, Bk], <args>)
       
  4776                                    (tm_of (a [+] dummy_abc (length args)) @
       
  4777                                     shift (mopup (length args))
       
  4778                                      (listsum
       
  4779  (layout_of (a [+] dummy_abc (length args)))))
       
  4780                                    na")
       
  4781 apply(simp)
       
  4782 by (metis assms para_pattern)
       
  4783 
       
  4784 
  2539 
  4785 lemma [simp]:
  2540 lemma [simp]:
  4786   "list_all (\<lambda>(acn, s). s \<le> Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))) xs \<Longrightarrow>
  2541   "list_all (\<lambda>(acn, s). s \<le> Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))) xs \<Longrightarrow>
  4787   list_all (\<lambda>(acn, s). s \<le> Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))))) xs"
  2542   list_all (\<lambda>(acn, s). s \<le> Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))))) xs"
  4788 apply(induct xs, simp, simp)
  2543 apply(induct xs, simp, simp)
  4906 apply(rule_tac start_of_all_le)
  2661 apply(rule_tac start_of_all_le)
  4907 apply(rule_tac dec_state_all_le, simp_all)
  2662 apply(rule_tac dec_state_all_le, simp_all)
  4908 apply(rule_tac start_of_all_le)
  2663 apply(rule_tac start_of_all_le)
  4909 done
  2664 done
  4910 
  2665 
  4911 lemma concat_in: "i < length (concat xs) \<Longrightarrow> \<exists>k < length xs. concat xs ! i \<in> set (xs ! k)"
  2666 lemma concat_in: "i < length (concat xs) \<Longrightarrow> 
  4912 apply(induct xs rule: list_tl_induct, simp, simp)
  2667   \<exists>k < length xs. concat xs ! i \<in> set (xs ! k)"
  4913 apply(case_tac "i < length (concat list)", simp)
  2668 apply(induct xs rule: rev_induct, simp, simp)
       
  2669 apply(case_tac "i < length (concat xs)", simp)
  4914 apply(erule_tac exE, rule_tac x = k in exI)
  2670 apply(erule_tac exE, rule_tac x = k in exI)
  4915 apply(simp add: nth_append)
  2671 apply(simp add: nth_append)
  4916 apply(rule_tac x = "length list" in exI, simp)
  2672 apply(rule_tac x = "length xs" in exI, simp)
  4917 apply(simp add: nth_append)
  2673 apply(simp add: nth_append)
  4918 done 
  2674 done 
  4919 
  2675 
  4920 lemma [simp]: "length (tms_of ap) = length ap"
  2676 lemma [simp]: "length (tms_of ap) = length ap"
  4921 apply(simp add: tms_of.simps tpairs_of.simps)
  2677 apply(simp add: tms_of.simps tpairs_of.simps)
  4948 apply(case_tac i, auto simp: ci.simps length_findnth
  2704 apply(case_tac i, auto simp: ci.simps length_findnth
  4949   tinc_b_def adjust.simps tdec_b_def)
  2705   tinc_b_def adjust.simps tdec_b_def)
  4950 done
  2706 done
  4951 
  2707 
  4952 lemma [intro]: "listsum (map (length \<circ> (\<lambda>(x, y). ci ly x y)) zs) mod 2 = 0"
  2708 lemma [intro]: "listsum (map (length \<circ> (\<lambda>(x, y). ci ly x y)) zs) mod 2 = 0"
  4953 apply(induct zs rule: list_tl_induct, simp)
  2709 apply(induct zs rule: rev_induct, simp)
  4954 apply(case_tac a, simp)
  2710 apply(case_tac x, simp)
  4955 apply(subgoal_tac "length (ci ly aa b) mod 2 = 0")
  2711 apply(subgoal_tac "length (ci ly a b) mod 2 = 0")
  4956 apply(auto)
  2712 apply(auto)
  4957 done
  2713 done
  4958 
  2714 
  4959 lemma zip_pre:
  2715 lemma zip_pre:
  4960   "(length ys) \<le> length ap \<Longrightarrow>
  2716   "(length ys) \<le> length ap \<Longrightarrow>
  4982 done
  2738 done
  4983 
  2739 
  4984 lemma [simp]: "length mopup_b = 12"
  2740 lemma [simp]: "length mopup_b = 12"
  4985 apply(simp add: mopup_b_def)
  2741 apply(simp add: mopup_b_def)
  4986 done
  2742 done
  4987 (*
       
  4988 lemma [elim]: "\<lbrakk>na < 4 * n; tshift (mop_bef n) q ! na = (a, b)\<rbrakk> \<Longrightarrow> 
       
  4989   b \<le> q + (2 * n + 6)"
       
  4990 apply(induct n, simp, simp add: mop_bef.simps nth_append tshift_append shift_length)
       
  4991 apply(case_tac "na < 4*n", simp, simp)
       
  4992 apply(subgoal_tac "na = 4*n \<or> na = 1 + 4*n \<or> na = 2 + 4*n \<or> na = 3 + 4*n",
       
  4993   auto simp: shift_length)
       
  4994 apply(simp_all add: tshift.simps)
       
  4995 done
       
  4996 *)
       
  4997 
  2743 
  4998 lemma mp_up_all_le: "list_all  (\<lambda>(acn, s). s \<le> q + (2 * n + 6)) 
  2744 lemma mp_up_all_le: "list_all  (\<lambda>(acn, s). s \<le> q + (2 * n + 6)) 
  4999   [(R, Suc (Suc (2 * n + q))), (R, Suc (2 * n + q)), 
  2745   [(R, Suc (Suc (2 * n + q))), (R, Suc (2 * n + q)), 
  5000   (L, 5 + 2 * n + q), (W0, Suc (Suc (Suc (2 * n + q)))), (R, 4 + 2 * n + q),
  2746   (L, 5 + 2 * n + q), (W0, Suc (Suc (Suc (2 * n + q)))), (R, 4 + 2 * n + q),
  5001   (W0, Suc (Suc (Suc (2 * n + q)))), (R, Suc (Suc (2 * n + q))),
  2747   (W0, Suc (Suc (Suc (2 * n + q)))), (R, Suc (Suc (2 * n + q))),
  5026 lemma [simp]: "b \<le> Suc x
  2772 lemma [simp]: "b \<le> Suc x
  5027           \<Longrightarrow> b \<le> (2 * x + length (mopup n)) div 2"
  2773           \<Longrightarrow> b \<le> (2 * x + length (mopup n)) div 2"
  5028 apply(auto simp: mopup.simps)
  2774 apply(auto simp: mopup.simps)
  5029 done
  2775 done
  5030 
  2776 
  5031 lemma t_compiled_correct: 
  2777 lemma wf_tm_from_abacus: 
  5032   "\<lbrakk>tp = tm_of ap; ly = layout_of ap; mop_ss = start_of ly (length ap)\<rbrakk> \<Longrightarrow> 
  2778   "tp = tm_of ap \<Longrightarrow> 
  5033     tm_wf (tp @ shift( mopup n) (length tp div 2), 0)"
  2779     tm_wf (tp @ shift( mopup n) (length tp div 2), 0)"
  5034   using length_start_of_tm[of ap] all_le_start_of[of ap]
  2780   using length_start_of_tm[of ap] all_le_start_of[of ap]
  5035 apply(auto simp: tm_wf.simps List.list_all_iff)
  2781 apply(auto simp: tm_wf.simps List.list_all_iff)
  5036 done
  2782 done
  5037 
  2783 
       
  2784 lemma wf_tm_from_recf:
       
  2785   assumes compile: "tp = tm_of_rec recf"
       
  2786   shows "tm_wf0 tp"
       
  2787 proof -
       
  2788   obtain a b c where "rec_ci recf = (a, b, c)"
       
  2789     by (metis prod_cases3)
       
  2790   thus "?thesis"
       
  2791     using compile
       
  2792     using wf_tm_from_abacus[of "tm_of (a [+] dummy_abc b)" "(a [+] dummy_abc b)" b]
       
  2793     by simp
       
  2794 qed
       
  2795  
  5038 end
  2796 end
  5039 
       
  5040     
       
  5041   
       
  5042 
       
  5043 
       
  5044   
       
  5045