327 {\<lambda>nl. nl = n # Suc n # 0 # anything}" |
327 {\<lambda>nl. nl = n # Suc n # 0 # anything}" |
328 by(rule_tac abc_Hoare_haltI, rule_tac x = 1 in exI, auto simp: addition_inv.simps |
328 by(rule_tac abc_Hoare_haltI, rule_tac x = 1 in exI, auto simp: addition_inv.simps |
329 abc_steps_l.simps abc_step_l.simps abc_fetch.simps numeral_2_eq_2 abc_lm_s.simps abc_lm_v.simps) |
329 abc_steps_l.simps abc_step_l.simps abc_fetch.simps numeral_2_eq_2 abc_lm_s.simps abc_lm_v.simps) |
330 qed |
330 qed |
331 |
331 |
332 declare rec_exec.simps[simp del] |
332 declare rec_eval.simps[simp del] |
333 |
333 |
334 lemma abc_comp_commute: "(A [+] B) [+] C = A [+] (B [+] C)" |
334 lemma abc_comp_commute: "(A [+] B) [+] C = A [+] (B [+] C)" |
335 apply(auto simp: abc_comp.simps abc_shift.simps) |
335 apply(auto simp: abc_comp.simps abc_shift.simps) |
336 apply(case_tac x, auto) |
336 apply(case_tac x, auto) |
337 done |
337 done |
338 |
338 |
339 |
339 |
340 |
340 |
341 lemma compile_z_correct: |
341 lemma compile_z_correct: |
342 "\<lbrakk>rec_ci z = (ap, arity, fp); rec_exec z [n] = r\<rbrakk> \<Longrightarrow> |
342 "\<lbrakk>rec_ci z = (ap, arity, fp); rec_eval z [n] = r\<rbrakk> \<Longrightarrow> |
343 {\<lambda>nl. nl = n # 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = n # r # 0 \<up> (fp - Suc arity) @ anything}" |
343 {\<lambda>nl. nl = n # 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = n # r # 0 \<up> (fp - Suc arity) @ anything}" |
344 apply(rule_tac abc_Hoare_haltI) |
344 apply(rule_tac abc_Hoare_haltI) |
345 apply(rule_tac x = 1 in exI) |
345 apply(rule_tac x = 1 in exI) |
346 apply(auto simp: abc_steps_l.simps rec_ci.simps rec_ci_z_def |
346 apply(auto simp: abc_steps_l.simps rec_ci.simps rec_ci_z_def |
347 numeral_2_eq_2 abc_fetch.simps abc_step_l.simps rec_exec.simps) |
347 numeral_2_eq_2 abc_fetch.simps abc_step_l.simps rec_eval.simps) |
348 done |
348 done |
349 |
349 |
350 lemma compile_s_correct: |
350 lemma compile_s_correct: |
351 "\<lbrakk>rec_ci s = (ap, arity, fp); rec_exec s [n] = r\<rbrakk> \<Longrightarrow> |
351 "\<lbrakk>rec_ci s = (ap, arity, fp); rec_eval s [n] = r\<rbrakk> \<Longrightarrow> |
352 {\<lambda>nl. nl = n # 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = n # r # 0 \<up> (fp - Suc arity) @ anything}" |
352 {\<lambda>nl. nl = n # 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = n # r # 0 \<up> (fp - Suc arity) @ anything}" |
353 apply(auto simp: rec_ci.simps rec_ci_s_def compile_s_correct' rec_exec.simps) |
353 apply(auto simp: rec_ci.simps rec_ci_s_def compile_s_correct' rec_eval.simps) |
354 done |
354 done |
355 |
355 |
356 lemma compile_id_correct': |
356 lemma compile_id_correct': |
357 assumes "n < length args" |
357 assumes "n < length args" |
358 shows "{\<lambda>nl. nl = args @ 0 \<up> 2 @ anything} addition n (length args) (Suc (length args)) |
358 shows "{\<lambda>nl. nl = args @ 0 \<up> 2 @ anything} addition n (length args) (Suc (length args)) |
359 {\<lambda>nl. nl = args @ rec_exec (recf.id (length args) n) args # 0 # anything}" |
359 {\<lambda>nl. nl = args @ rec_eval (recf.id (length args) n) args # 0 # anything}" |
360 proof - |
360 proof - |
361 have "{\<lambda>nl. nl = args @ 0 \<up> 2 @ anything} addition n (length args) (Suc (length args)) |
361 have "{\<lambda>nl. nl = args @ 0 \<up> 2 @ anything} addition n (length args) (Suc (length args)) |
362 {\<lambda>nl. addition_inv (7, nl) n (length args) (Suc (length args)) (args @ 0 \<up> 2 @ anything)}" |
362 {\<lambda>nl. addition_inv (7, nl) n (length args) (Suc (length args)) (args @ 0 \<up> 2 @ anything)}" |
363 using assms |
363 using assms |
364 by(rule_tac addition_correct, auto simp: numeral_2_eq_2 nth_append) |
364 by(rule_tac addition_correct, auto simp: numeral_2_eq_2 nth_append) |
365 thus "?thesis" |
365 thus "?thesis" |
366 using assms |
366 using assms |
367 by(simp add: addition_inv.simps rec_exec.simps |
367 by(simp add: addition_inv.simps rec_eval.simps |
368 nth_append numeral_2_eq_2 list_update_append) |
368 nth_append numeral_2_eq_2 list_update_append) |
369 qed |
369 qed |
370 |
370 |
371 lemma compile_id_correct: |
371 lemma compile_id_correct: |
372 "\<lbrakk>n < m; length xs = m; rec_ci (recf.id m n) = (ap, arity, fp); rec_exec (recf.id m n) xs = r\<rbrakk> |
372 "\<lbrakk>n < m; length xs = m; rec_ci (recf.id m n) = (ap, arity, fp); rec_eval (recf.id m n) xs = r\<rbrakk> |
373 \<Longrightarrow> {\<lambda>nl. nl = xs @ 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ r # 0 \<up> (fp - Suc arity) @ anything}" |
373 \<Longrightarrow> {\<lambda>nl. nl = xs @ 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ r # 0 \<up> (fp - Suc arity) @ anything}" |
374 apply(auto simp: rec_ci.simps rec_ci_id.simps compile_id_correct') |
374 apply(auto simp: rec_ci.simps rec_ci_id.simps compile_id_correct') |
375 done |
375 done |
376 |
376 |
377 lemma cn_merge_gs_tl_app: |
377 lemma cn_merge_gs_tl_app: |
670 by(simp, rule_tac min_max.le_supI2, rule_tac Max_ge, simp, |
670 by(simp, rule_tac min_max.le_supI2, rule_tac Max_ge, simp, |
671 rule_tac insertI2, |
671 rule_tac insertI2, |
672 rule_tac f = "(\<lambda>(aprog, p, n). n)" and x = "rec_ci (gs!n)" in image_eqI, simp, |
672 rule_tac f = "(\<lambda>(aprog, p, n). n)" and x = "rec_ci (gs!n)" in image_eqI, simp, |
673 rule_tac x = "gs!n" in image_eqI, simp, simp) |
673 rule_tac x = "gs!n" in image_eqI, simp, simp) |
674 show "{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ |
674 show "{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ |
675 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) |
675 map (\<lambda>i. rec_eval i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything} gp [+] mv_box ga (ft + n) |
676 {\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) |
676 {\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_eval i xs) |
677 (take n gs) @ rec_exec (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything}" |
677 (take n gs) @ rec_eval (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything}" |
678 proof(rule_tac abc_Hoare_plus_halt) |
678 proof(rule_tac abc_Hoare_plus_halt) |
679 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 |
679 show "{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_eval i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything} gp |
680 {\<lambda>nl. nl = xs @ (rec_exec (gs!n) xs) # 0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_exec i xs) |
680 {\<lambda>nl. nl = xs @ (rec_eval (gs!n) xs) # 0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_eval i xs) |
681 (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything}" |
681 (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything}" |
682 proof - |
682 proof - |
683 have |
683 have |
684 "({\<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} |
684 "({\<lambda>nl. nl = xs @ 0 \<up> (gf - ga) @ 0\<up>(ft - gf)@map (\<lambda>i. rec_eval i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything} |
685 gp {\<lambda>nl. nl = xs @ (rec_exec (gs!n) xs) # 0 \<up> (gf - Suc ga) @ |
685 gp {\<lambda>nl. nl = xs @ (rec_eval (gs!n) xs) # 0 \<up> (gf - Suc ga) @ |
686 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})" |
686 0\<up>(ft - gf)@map (\<lambda>i. rec_eval i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything})" |
687 using a g_newcond h x |
687 using a g_newcond h x |
688 apply(erule_tac x = "gs!n" in ballE) |
688 apply(erule_tac x = "gs!n" in ballE) |
689 apply(simp, simp) |
689 apply(simp, simp) |
690 done |
690 done |
691 thus "?thesis" |
691 thus "?thesis" |
692 using a b c d e |
692 using a b c d e |
693 by(simp add: replicate_merge_anywhere) |
693 by(simp add: replicate_merge_anywhere) |
694 qed |
694 qed |
695 next |
695 next |
696 show |
696 show |
697 "{\<lambda>nl. nl = xs @ rec_exec (gs ! n) xs # |
697 "{\<lambda>nl. nl = xs @ rec_eval (gs ! n) xs # |
698 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} |
698 0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_eval i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything} |
699 mv_box ga (ft + n) |
699 mv_box ga (ft + n) |
700 {\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ |
700 {\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_eval i xs) (take n gs) @ |
701 rec_exec (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything}" |
701 rec_eval (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything}" |
702 proof - |
702 proof - |
703 have "{\<lambda>nl. nl = xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ |
703 have "{\<lambda>nl. nl = xs @ rec_eval (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ |
704 map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything} |
704 map (\<lambda>i. rec_eval i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything} |
705 mv_box ga (ft + n) {\<lambda>nl. nl = (xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ |
705 mv_box ga (ft + n) {\<lambda>nl. nl = (xs @ rec_eval (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ |
706 map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) |
706 map (\<lambda>i. rec_eval i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) |
707 [ft + n := (xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ |
707 [ft + n := (xs @ rec_eval (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_eval i xs) (take n gs) @ |
708 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) ! ga + |
708 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) ! ga + |
709 (xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ |
709 (xs @ rec_eval (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ |
710 map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) ! |
710 map (\<lambda>i. rec_eval i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) ! |
711 (ft + n), ga := 0]}" |
711 (ft + n), ga := 0]}" |
712 using a c d e h |
712 using a c d e h |
713 apply(rule_tac mv_box_correct) |
713 apply(rule_tac mv_box_correct) |
714 apply(simp, arith, arith) |
714 apply(simp, arith, arith) |
715 done |
715 done |
716 moreover have "(xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ |
716 moreover have "(xs @ rec_eval (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ |
717 map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) |
717 map (\<lambda>i. rec_eval i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) |
718 [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 [ft + n := (xs @ rec_eval (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_eval i xs) (take n gs) @ |
719 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) ! ga + |
719 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) ! ga + |
720 (xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ |
720 (xs @ rec_eval (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ |
721 map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) ! |
721 map (\<lambda>i. rec_eval i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) ! |
722 (ft + n), ga := 0]= |
722 (ft + n), ga := 0]= |
723 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" |
723 xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_eval i xs) (take n gs) @ rec_eval (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything" |
724 using a c d e h |
724 using a c d e h |
725 by(simp add: list_update_append nth_append length_replicate split: if_splits del: list_update.simps(2), auto) |
725 by(simp add: list_update_append nth_append length_replicate split: if_splits del: list_update.simps(2), auto) |
726 ultimately show "?thesis" |
726 ultimately show "?thesis" |
727 by(simp) |
727 by(simp) |
728 qed |
728 qed |
1005 |
1005 |
1006 lemma save_rs: |
1006 lemma save_rs: |
1007 "\<lbrakk>far = length gs; |
1007 "\<lbrakk>far = length gs; |
1008 ffp \<le> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))); |
1008 ffp \<le> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))); |
1009 far < ffp\<rbrakk> |
1009 far < ffp\<rbrakk> |
1010 \<Longrightarrow> {\<lambda>nl. nl = map (\<lambda>i. rec_exec i xs) gs @ |
1010 \<Longrightarrow> {\<lambda>nl. nl = map (\<lambda>i. rec_eval i xs) gs @ |
1011 rec_exec (Cn (length xs) f gs) xs # 0 \<up> max (Suc (length xs)) |
1011 rec_eval (Cn (length xs) f gs) xs # 0 \<up> max (Suc (length xs)) |
1012 (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @ xs @ anything} |
1012 (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @ xs @ anything} |
1013 mv_box far (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) |
1013 mv_box far (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) |
1014 {\<lambda>nl. nl = map (\<lambda>i. rec_exec i xs) gs @ |
1014 {\<lambda>nl. nl = map (\<lambda>i. rec_eval i xs) gs @ |
1015 0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - length gs) @ |
1015 0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - length gs) @ |
1016 rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything}" |
1016 rec_eval (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything}" |
1017 proof - |
1017 proof - |
1018 let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))" |
1018 let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))" |
1019 thm mv_box_correct |
1019 thm mv_box_correct |
1020 let ?lm= " map (\<lambda>i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> ?ft @ xs @ anything" |
1020 let ?lm= " map (\<lambda>i. rec_eval i xs) gs @ rec_eval (Cn (length xs) f gs) xs # 0 \<up> ?ft @ xs @ anything" |
1021 assume h: "far = length gs" "ffp \<le> ?ft" "far < ffp" |
1021 assume h: "far = length gs" "ffp \<le> ?ft" "far < ffp" |
1022 hence "{\<lambda> nl. nl = ?lm} mv_box far ?ft {\<lambda> nl. nl = ?lm[?ft := ?lm!far + ?lm!?ft, far := 0]}" |
1022 hence "{\<lambda> nl. nl = ?lm} mv_box far ?ft {\<lambda> nl. nl = ?lm[?ft := ?lm!far + ?lm!?ft, far := 0]}" |
1023 apply(rule_tac mv_box_correct) |
1023 apply(rule_tac mv_box_correct) |
1024 by(case_tac "rec_ci a", auto) |
1024 by(case_tac "rec_ci a", auto) |
1025 moreover have "?lm[?ft := ?lm!far + ?lm!?ft, far := 0] |
1025 moreover have "?lm[?ft := ?lm!far + ?lm!?ft, far := 0] |
1026 = map (\<lambda>i. rec_exec i xs) gs @ |
1026 = map (\<lambda>i. rec_eval i xs) gs @ |
1027 0 \<up> (?ft - length gs) @ |
1027 0 \<up> (?ft - length gs) @ |
1028 rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything" |
1028 rec_eval (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything" |
1029 using h |
1029 using h |
1030 apply(simp add: nth_append) |
1030 apply(simp add: nth_append) |
1031 using list_update_length[of "map (\<lambda>i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # |
1031 using list_update_length[of "map (\<lambda>i. rec_eval i xs) gs @ rec_eval (Cn (length xs) f gs) xs # |
1032 0 \<up> (?ft - Suc (length gs))" 0 "0 \<up> length gs @ xs @ anything" "rec_exec (Cn (length xs) f gs) xs"] |
1032 0 \<up> (?ft - Suc (length gs))" 0 "0 \<up> length gs @ xs @ anything" "rec_eval (Cn (length xs) f gs) xs"] |
1033 apply(simp add: replicate_merge_anywhere replicate_Suc_iff_anywhere del: replicate_Suc) |
1033 apply(simp add: replicate_merge_anywhere replicate_Suc_iff_anywhere del: replicate_Suc) |
1034 by(simp add: list_update_append list_update.simps replicate_Suc_iff_anywhere del: replicate_Suc) |
1034 by(simp add: list_update_append list_update.simps replicate_Suc_iff_anywhere del: replicate_Suc) |
1035 ultimately show "?thesis" |
1035 ultimately show "?thesis" |
1036 by(simp) |
1036 by(simp) |
1037 qed |
1037 qed |
1106 done |
1106 done |
1107 |
1107 |
1108 |
1108 |
1109 lemma clean_paras: |
1109 lemma clean_paras: |
1110 "ffp \<ge> length gs \<Longrightarrow> |
1110 "ffp \<ge> length gs \<Longrightarrow> |
1111 {\<lambda>nl. nl = map (\<lambda>i. rec_exec i xs) gs @ |
1111 {\<lambda>nl. nl = map (\<lambda>i. rec_eval i xs) gs @ |
1112 0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - length gs) @ |
1112 0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - length gs) @ |
1113 rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything} |
1113 rec_eval (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything} |
1114 empty_boxes (length gs) |
1114 empty_boxes (length gs) |
1115 {\<lambda>nl. nl = 0 \<up> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @ |
1115 {\<lambda>nl. nl = 0 \<up> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @ |
1116 rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything}" |
1116 rec_eval (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything}" |
1117 proof- |
1117 proof- |
1118 let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))" |
1118 let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))" |
1119 assume h: "length gs \<le> ffp" |
1119 assume h: "length gs \<le> ffp" |
1120 let ?lm = "map (\<lambda>i. rec_exec i xs) gs @ 0 \<up> (?ft - length gs) @ |
1120 let ?lm = "map (\<lambda>i. rec_eval i xs) gs @ 0 \<up> (?ft - length gs) @ |
1121 rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything" |
1121 rec_eval (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything" |
1122 have "{\<lambda> nl. nl = ?lm} empty_boxes (length gs) {\<lambda> nl. nl = 0\<up>length gs @ drop (length gs) ?lm}" |
1122 have "{\<lambda> nl. nl = ?lm} empty_boxes (length gs) {\<lambda> nl. nl = 0\<up>length gs @ drop (length gs) ?lm}" |
1123 by(rule_tac empty_boxes_correct, simp) |
1123 by(rule_tac empty_boxes_correct, simp) |
1124 moreover have "0\<up>length gs @ drop (length gs) ?lm |
1124 moreover have "0\<up>length gs @ drop (length gs) ?lm |
1125 = 0 \<up> ?ft @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything" |
1125 = 0 \<up> ?ft @ rec_eval (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything" |
1126 using h |
1126 using h |
1127 by(simp add: replicate_merge_anywhere) |
1127 by(simp add: replicate_merge_anywhere) |
1128 ultimately show "?thesis" |
1128 ultimately show "?thesis" |
1129 by metis |
1129 by metis |
1130 qed |
1130 qed |
1131 |
1131 |
1132 |
1132 |
1133 lemma restore_rs: |
1133 lemma restore_rs: |
1134 "{\<lambda>nl. nl = 0 \<up> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @ |
1134 "{\<lambda>nl. nl = 0 \<up> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @ |
1135 rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything} |
1135 rec_eval (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything} |
1136 mv_box (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) (length xs) |
1136 mv_box (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) (length xs) |
1137 {\<lambda>nl. nl = 0 \<up> length xs @ |
1137 {\<lambda>nl. nl = 0 \<up> length xs @ |
1138 rec_exec (Cn (length xs) f gs) xs # |
1138 rec_eval (Cn (length xs) f gs) xs # |
1139 0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - (length xs)) @ |
1139 0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - (length xs)) @ |
1140 0 \<up> length gs @ xs @ anything}" |
1140 0 \<up> length gs @ xs @ anything}" |
1141 proof - |
1141 proof - |
1142 let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))" |
1142 let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))" |
1143 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" |
1143 let ?lm = "0\<up>(length xs) @ 0\<up>(?ft - (length xs)) @ rec_eval (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything" |
1144 thm mv_box_correct |
1144 thm mv_box_correct |
1145 have "{\<lambda> nl. nl = ?lm} mv_box ?ft (length xs) {\<lambda> nl. nl = ?lm[length xs := ?lm!?ft + ?lm!(length xs), ?ft := 0]}" |
1145 have "{\<lambda> nl. nl = ?lm} mv_box ?ft (length xs) {\<lambda> nl. nl = ?lm[length xs := ?lm!?ft + ?lm!(length xs), ?ft := 0]}" |
1146 by(rule_tac mv_box_correct, simp, simp) |
1146 by(rule_tac mv_box_correct, simp, simp) |
1147 moreover have "?lm[length xs := ?lm!?ft + ?lm!(length xs), ?ft := 0] |
1147 moreover have "?lm[length xs := ?lm!?ft + ?lm!(length xs), ?ft := 0] |
1148 = 0 \<up> length xs @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> (?ft - (length xs)) @ 0 \<up> length gs @ xs @ anything" |
1148 = 0 \<up> length xs @ rec_eval (Cn (length xs) f gs) xs # 0 \<up> (?ft - (length xs)) @ 0 \<up> length gs @ xs @ anything" |
1149 apply(auto simp: list_update_append nth_append) |
1149 apply(auto simp: list_update_append nth_append) |
1150 apply(case_tac ?ft, simp_all add: Suc_diff_le list_update.simps) |
1150 apply(case_tac ?ft, simp_all add: Suc_diff_le list_update.simps) |
1151 apply(simp add: exp_suc replicate_Suc[THEN sym] del: replicate_Suc) |
1151 apply(simp add: exp_suc replicate_Suc[THEN sym] del: replicate_Suc) |
1152 done |
1152 done |
1153 ultimately show "?thesis" |
1153 ultimately show "?thesis" |
1154 by(simp add: replicate_merge_anywhere) |
1154 by(simp add: replicate_merge_anywhere) |
1155 qed |
1155 qed |
1156 |
1156 |
1157 lemma restore_orgin_paras: |
1157 lemma restore_orgin_paras: |
1158 "{\<lambda>nl. nl = 0 \<up> length xs @ |
1158 "{\<lambda>nl. nl = 0 \<up> length xs @ |
1159 rec_exec (Cn (length xs) f gs) xs # |
1159 rec_eval (Cn (length xs) f gs) xs # |
1160 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} |
1160 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} |
1161 mv_boxes (Suc (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) 0 (length xs) |
1161 mv_boxes (Suc (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) 0 (length xs) |
1162 {\<lambda>nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> |
1162 {\<lambda>nl. nl = xs @ rec_eval (Cn (length xs) f gs) xs # 0 \<up> |
1163 (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything}" |
1163 (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything}" |
1164 proof - |
1164 proof - |
1165 let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))" |
1165 let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))" |
1166 thm mv_boxes_correct2 |
1166 thm mv_boxes_correct2 |
1167 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} |
1167 have "{\<lambda> nl. nl = [] @ 0\<up>(length xs) @ (rec_eval (Cn (length xs) f gs) xs # 0 \<up> (?ft - length xs) @ 0 \<up> length gs) @ xs @ anything} |
1168 mv_boxes (Suc ?ft + length gs) 0 (length xs) |
1168 mv_boxes (Suc ?ft + length gs) 0 (length xs) |
1169 {\<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}" |
1169 {\<lambda> nl. nl = [] @ xs @ (rec_eval (Cn (length xs) f gs) xs # 0 \<up> (?ft - length xs) @ 0 \<up> length gs) @ 0\<up>length xs @ anything}" |
1170 by(rule_tac mv_boxes_correct2, auto) |
1170 by(rule_tac mv_boxes_correct2, auto) |
1171 thus "?thesis" |
1171 thus "?thesis" |
1172 by(simp add: replicate_merge_anywhere) |
1172 by(simp add: replicate_merge_anywhere) |
1173 qed |
1173 qed |
1174 |
1174 |
1175 lemma compile_cn_correct': |
1175 lemma compile_cn_correct': |
1176 assumes f_ind: |
1176 assumes f_ind: |
1177 "\<And> anything r. rec_exec f (map (\<lambda>g. rec_exec g xs) gs) = rec_exec (Cn (length xs) f gs) xs \<Longrightarrow> |
1177 "\<And> anything r. rec_eval f (map (\<lambda>g. rec_eval g xs) gs) = rec_eval (Cn (length xs) f gs) xs \<Longrightarrow> |
1178 {\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ 0 \<up> (ffp - far) @ anything} fap |
1178 {\<lambda>nl. nl = map (\<lambda>g. rec_eval g xs) gs @ 0 \<up> (ffp - far) @ anything} fap |
1179 {\<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}" |
1179 {\<lambda>nl. nl = map (\<lambda>g. rec_eval g xs) gs @ rec_eval (Cn (length xs) f gs) xs # 0 \<up> (ffp - Suc far) @ anything}" |
1180 and compile: "rec_ci f = (fap, far, ffp)" |
1180 and compile: "rec_ci f = (fap, far, ffp)" |
1181 and term_f: "terminate f (map (\<lambda>g. rec_exec g xs) gs)" |
1181 and term_f: "terminates f (map (\<lambda>g. rec_eval g xs) gs)" |
1182 and g_cond: "\<forall>g\<in>set gs. terminate g xs \<and> |
1182 and g_cond: "\<forall>g\<in>set gs. terminates g xs \<and> |
1183 (\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow> |
1183 (\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow> |
1184 (\<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}))" |
1184 (\<forall>xc. {\<lambda>nl. nl = xs @ 0 \<up> (xb - xa) @ xc} x {\<lambda>nl. nl = xs @ rec_eval g xs # 0 \<up> (xb - Suc xa) @ xc}))" |
1185 shows |
1185 shows |
1186 "{\<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} |
1186 "{\<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} |
1187 cn_merge_gs (map rec_ci gs) (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) [+] |
1187 cn_merge_gs (map rec_ci gs) (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) [+] |
1188 (mv_boxes 0 (Suc (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) (length xs) [+] |
1188 (mv_boxes 0 (Suc (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) (length xs) [+] |
1189 (mv_boxes (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) 0 (length gs) [+] |
1189 (mv_boxes (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) 0 (length gs) [+] |
1190 (fap [+] (mv_box far (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) [+] |
1190 (fap [+] (mv_box far (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) [+] |
1191 (empty_boxes (length gs) [+] |
1191 (empty_boxes (length gs) [+] |
1192 (mv_box (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) (length xs) [+] |
1192 (mv_box (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) (length xs) [+] |
1193 mv_boxes (Suc (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) 0 (length xs))))))) |
1193 mv_boxes (Suc (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) 0 (length xs))))))) |
1194 {\<lambda>nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # |
1194 {\<lambda>nl. nl = xs @ rec_eval (Cn (length xs) f gs) xs # |
1195 0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything}" |
1195 0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything}" |
1196 proof - |
1196 proof - |
1197 let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))" |
1197 let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))" |
1198 let ?A = "cn_merge_gs (map rec_ci gs) ?ft" |
1198 let ?A = "cn_merge_gs (map rec_ci gs) ?ft" |
1199 let ?B = "mv_boxes 0 (Suc (?ft+length gs)) (length xs)" |
1199 let ?B = "mv_boxes 0 (Suc (?ft+length gs)) (length xs)" |
1202 let ?E = "mv_box far ?ft" |
1202 let ?E = "mv_box far ?ft" |
1203 let ?F = "empty_boxes (length gs)" |
1203 let ?F = "empty_boxes (length gs)" |
1204 let ?G = "mv_box ?ft (length xs)" |
1204 let ?G = "mv_box ?ft (length xs)" |
1205 let ?H = "mv_boxes (Suc (?ft + length gs)) 0 (length xs)" |
1205 let ?H = "mv_boxes (Suc (?ft + length gs)) 0 (length xs)" |
1206 let ?P1 = "\<lambda>nl. nl = xs @ 0 # 0 \<up> (?ft + length gs) @ anything" |
1206 let ?P1 = "\<lambda>nl. nl = xs @ 0 # 0 \<up> (?ft + length gs) @ anything" |
1207 let ?S = "\<lambda>nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> (?ft + length gs) @ anything" |
1207 let ?S = "\<lambda>nl. nl = xs @ rec_eval (Cn (length xs) f gs) xs # 0 \<up> (?ft + length gs) @ anything" |
1208 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" |
1208 let ?Q1 = "\<lambda> nl. nl = xs @ 0\<up>(?ft - length xs) @ map (\<lambda> i. rec_eval i xs) gs @ 0\<up>(Suc (length xs)) @ anything" |
1209 show "{?P1} (?A [+] (?B [+] (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H))))))) {?S}" |
1209 show "{?P1} (?A [+] (?B [+] (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H))))))) {?S}" |
1210 proof(rule_tac abc_Hoare_plus_halt) |
1210 proof(rule_tac abc_Hoare_plus_halt) |
1211 show "{?P1} ?A {?Q1}" |
1211 show "{?P1} ?A {?Q1}" |
1212 using g_cond |
1212 using g_cond |
1213 by(rule_tac compile_cn_gs_correct, auto) |
1213 by(rule_tac compile_cn_gs_correct, auto) |
1214 next |
1214 next |
1215 let ?Q2 = "\<lambda>nl. nl = 0 \<up> ?ft @ |
1215 let ?Q2 = "\<lambda>nl. nl = 0 \<up> ?ft @ |
1216 map (\<lambda>i. rec_exec i xs) gs @ 0 # xs @ anything" |
1216 map (\<lambda>i. rec_eval i xs) gs @ 0 # xs @ anything" |
1217 show "{?Q1} (?B [+] (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H)))))) {?S}" |
1217 show "{?Q1} (?B [+] (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H)))))) {?S}" |
1218 proof(rule_tac abc_Hoare_plus_halt) |
1218 proof(rule_tac abc_Hoare_plus_halt) |
1219 show "{?Q1} ?B {?Q2}" |
1219 show "{?Q1} ?B {?Q2}" |
1220 by(rule_tac save_paras) |
1220 by(rule_tac save_paras) |
1221 next |
1221 next |
1222 let ?Q3 = "\<lambda> nl. nl = map (\<lambda>i. rec_exec i xs) gs @ 0\<up>?ft @ 0 # xs @ anything" |
1222 let ?Q3 = "\<lambda> nl. nl = map (\<lambda>i. rec_eval i xs) gs @ 0\<up>?ft @ 0 # xs @ anything" |
1223 show "{?Q2} (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H))))) {?S}" |
1223 show "{?Q2} (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H))))) {?S}" |
1224 proof(rule_tac abc_Hoare_plus_halt) |
1224 proof(rule_tac abc_Hoare_plus_halt) |
1225 have "ffp \<ge> length gs" |
1225 have "ffp \<ge> length gs" |
1226 using compile term_f |
1226 using compile term_f |
1227 apply(subgoal_tac "length gs = far") |
1227 apply(subgoal_tac "length gs = far") |
1228 apply(drule_tac footprint_ge, simp) |
1228 apply(drule_tac footprint_ge, simp) |
1229 by(drule_tac param_pattern, auto) |
1229 by(drule_tac param_pattern, auto) |
1230 thus "{?Q2} ?C {?Q3}" |
1230 thus "{?Q2} ?C {?Q3}" |
1231 by(erule_tac restore_new_paras) |
1231 by(erule_tac restore_new_paras) |
1232 next |
1232 next |
1233 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" |
1233 let ?Q4 = "\<lambda> nl. nl = map (\<lambda>i. rec_eval i xs) gs @ rec_eval (Cn (length xs) f gs) xs # 0\<up>?ft @ xs @ anything" |
1234 have a: "far = length gs" |
1234 have a: "far = length gs" |
1235 using compile term_f |
1235 using compile term_f |
1236 by(drule_tac param_pattern, auto) |
1236 by(drule_tac param_pattern, auto) |
1237 have b:"?ft \<ge> ffp" |
1237 have b:"?ft \<ge> ffp" |
1238 by auto |
1238 by auto |
1239 have c: "ffp > far" |
1239 have c: "ffp > far" |
1240 using compile |
1240 using compile |
1241 by(erule_tac footprint_ge) |
1241 by(erule_tac footprint_ge) |
1242 show "{?Q3} (?D [+] (?E [+] (?F [+] (?G [+] ?H)))) {?S}" |
1242 show "{?Q3} (?D [+] (?E [+] (?F [+] (?G [+] ?H)))) {?S}" |
1243 proof(rule_tac abc_Hoare_plus_halt) |
1243 proof(rule_tac abc_Hoare_plus_halt) |
1244 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 |
1244 have "{\<lambda>nl. nl = map (\<lambda>g. rec_eval g xs) gs @ 0 \<up> (ffp - far) @ 0\<up>(?ft - ffp + far) @ 0 # xs @ anything} fap |
1245 {\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ rec_exec (Cn (length xs) f gs) xs # |
1245 {\<lambda>nl. nl = map (\<lambda>g. rec_eval g xs) gs @ rec_eval (Cn (length xs) f gs) xs # |
1246 0 \<up> (ffp - Suc far) @ 0\<up>(?ft - ffp + far) @ 0 # xs @ anything}" |
1246 0 \<up> (ffp - Suc far) @ 0\<up>(?ft - ffp + far) @ 0 # xs @ anything}" |
1247 by(rule_tac f_ind, simp add: rec_exec.simps) |
1247 by(rule_tac f_ind, simp add: rec_eval.simps) |
1248 thus "{?Q3} fap {?Q4}" |
1248 thus "{?Q3} fap {?Q4}" |
1249 using a b c |
1249 using a b c |
1250 by(simp add: replicate_merge_anywhere, |
1250 by(simp add: replicate_merge_anywhere, |
1251 case_tac "?ft", simp_all add: exp_suc del: replicate_Suc) |
1251 case_tac "?ft", simp_all add: exp_suc del: replicate_Suc) |
1252 next |
1252 next |
1253 let ?Q5 = "\<lambda>nl. nl = map (\<lambda>i. rec_exec i xs) gs @ |
1253 let ?Q5 = "\<lambda>nl. nl = map (\<lambda>i. rec_eval i xs) gs @ |
1254 0\<up>(?ft - length gs) @ rec_exec (Cn (length xs) f gs) xs # 0\<up>(length gs)@ xs @ anything" |
1254 0\<up>(?ft - length gs) @ rec_eval (Cn (length xs) f gs) xs # 0\<up>(length gs)@ xs @ anything" |
1255 show "{?Q4} (?E [+] (?F [+] (?G [+] ?H))) {?S}" |
1255 show "{?Q4} (?E [+] (?F [+] (?G [+] ?H))) {?S}" |
1256 proof(rule_tac abc_Hoare_plus_halt) |
1256 proof(rule_tac abc_Hoare_plus_halt) |
1257 from a b c show "{?Q4} ?E {?Q5}" |
1257 from a b c show "{?Q4} ?E {?Q5}" |
1258 by(erule_tac save_rs, simp_all) |
1258 by(erule_tac save_rs, simp_all) |
1259 next |
1259 next |
1260 let ?Q6 = "\<lambda>nl. nl = 0\<up>?ft @ rec_exec (Cn (length xs) f gs) xs # 0\<up>(length gs)@ xs @ anything" |
1260 let ?Q6 = "\<lambda>nl. nl = 0\<up>?ft @ rec_eval (Cn (length xs) f gs) xs # 0\<up>(length gs)@ xs @ anything" |
1261 show "{?Q5} (?F [+] (?G [+] ?H)) {?S}" |
1261 show "{?Q5} (?F [+] (?G [+] ?H)) {?S}" |
1262 proof(rule_tac abc_Hoare_plus_halt) |
1262 proof(rule_tac abc_Hoare_plus_halt) |
1263 have "length gs \<le> ffp" using a b c |
1263 have "length gs \<le> ffp" using a b c |
1264 by simp |
1264 by simp |
1265 thus "{?Q5} ?F {?Q6}" |
1265 thus "{?Q5} ?F {?Q6}" |
1266 by(erule_tac clean_paras) |
1266 by(erule_tac clean_paras) |
1267 next |
1267 next |
1268 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" |
1268 let ?Q7 = "\<lambda>nl. nl = 0\<up>length xs @ rec_eval (Cn (length xs) f gs) xs # 0\<up>(?ft - (length xs)) @ 0\<up>(length gs)@ xs @ anything" |
1269 show "{?Q6} (?G [+] ?H) {?S}" |
1269 show "{?Q6} (?G [+] ?H) {?S}" |
1270 proof(rule_tac abc_Hoare_plus_halt) |
1270 proof(rule_tac abc_Hoare_plus_halt) |
1271 show "{?Q6} ?G {?Q7}" |
1271 show "{?Q6} ?G {?Q7}" |
1272 by(rule_tac restore_rs) |
1272 by(rule_tac restore_rs) |
1273 next |
1273 next |
1561 apply(rule_tac x = "stp + 1" in exI) |
1561 apply(rule_tac x = "stp + 1" in exI) |
1562 apply(simp only: abc_steps_add, simp) |
1562 apply(simp only: abc_steps_add, simp) |
1563 apply(simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_lm_v.simps abc_lm_s.simps) |
1563 apply(simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_lm_v.simps abc_lm_s.simps) |
1564 done |
1564 done |
1565 |
1565 |
1566 lemma rec_exec_pr_0_simps: "rec_exec (Pr n f g) (xs @ [0]) = rec_exec f xs" |
1566 lemma rec_eval_pr_0_simps: "rec_eval (Pr n f g) (xs @ [0]) = rec_eval f xs" |
1567 by(simp add: rec_exec.simps) |
1567 by(simp add: rec_eval.simps) |
1568 |
1568 |
1569 lemma rec_exec_pr_Suc_simps: "rec_exec (Pr n f g) (xs @ [Suc y]) |
1569 lemma rec_eval_pr_Suc_simps: "rec_eval (Pr n f g) (xs @ [Suc y]) |
1570 = rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])" |
1570 = rec_eval g (xs @ [y, rec_eval (Pr n f g) (xs @ [y])])" |
1571 apply(induct y) |
1571 apply(induct y) |
1572 apply(simp add: rec_exec.simps) |
1572 apply(simp add: rec_eval.simps) |
1573 apply(simp add: rec_exec.simps) |
1573 apply(simp add: rec_eval.simps) |
1574 done |
1574 done |
1575 |
1575 |
1576 lemma [simp]: "Suc (max (n + 3) (max fft gft) - Suc (Suc (Suc n))) = max (n + 3) (max fft gft) - Suc (Suc n)" |
1576 lemma [simp]: "Suc (max (n + 3) (max fft gft) - Suc (Suc (Suc n))) = max (n + 3) (max fft gft) - Suc (Suc n)" |
1577 by arith |
1577 by arith |
1578 |
1578 |
1579 lemma pr_loop: |
1579 lemma pr_loop: |
1580 assumes code: "code = ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @ |
1580 assumes code: "code = ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @ |
1581 [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" |
1581 [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" |
1582 and len: "length xs = n" |
1582 and len: "length xs = n" |
1583 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 |
1583 and g_ind: "\<forall> y<x. (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_eval (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap |
1584 {\<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})" |
1584 {\<lambda>nl. nl = xs @ y # rec_eval (Pr n f g) (xs @ [y]) # rec_eval g (xs @ [y, rec_eval (Pr n f g) (xs @ [y])]) # 0 \<up> (gft - Suc gar) @ anything})" |
1585 and compile_g: "rec_ci g = (gap, gar, gft)" |
1585 and compile_g: "rec_ci g = (gap, gar, gft)" |
1586 and termi_g: "\<forall> y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])" |
1586 and termi_g: "\<forall> y<x. terminates g (xs @ [y, rec_eval (Pr n f g) (xs @ [y])])" |
1587 and ft: "ft = max (n + 3) (max fft gft)" |
1587 and ft: "ft = max (n + 3) (max fft gft)" |
1588 and less: "Suc y \<le> x" |
1588 and less: "Suc y \<le> x" |
1589 shows |
1589 shows |
1590 "\<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) |
1590 "\<exists>stp. abc_steps_l (0, xs @ (x - Suc y) # rec_eval (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything) |
1591 code stp = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (ft - Suc (Suc n)) @ y # anything)" |
1591 code stp = (0, xs @ (x - y) # rec_eval (Pr n f g) (xs @ [x - y]) # 0 \<up> (ft - Suc (Suc n)) @ y # anything)" |
1592 proof - |
1592 proof - |
1593 let ?A = "[Dec ft (length gap + 7)]" |
1593 let ?A = "[Dec ft (length gap + 7)]" |
1594 let ?B = "gap" |
1594 let ?B = "gap" |
1595 let ?C = "[Inc n, Dec (Suc n) 3, Goto (Suc 0)]" |
1595 let ?C = "[Inc n, Dec (Suc n) 3, Goto (Suc 0)]" |
1596 let ?D = "[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" |
1596 let ?D = "[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" |
1597 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) |
1597 have "\<exists> stp. abc_steps_l (0, xs @ (x - Suc y) # rec_eval (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything) |
1598 ((?A [+] (?B [+] ?C)) @ ?D) stp = (length (?A [+] (?B [+] ?C)), |
1598 ((?A [+] (?B [+] ?C)) @ ?D) stp = (length (?A [+] (?B [+] ?C)), |
1599 xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) |
1599 xs @ (x - y) # 0 # rec_eval g (xs @ [x - Suc y, rec_eval (Pr n f g) (xs @ [x - Suc y])]) |
1600 # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)" |
1600 # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)" |
1601 proof - |
1601 proof - |
1602 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) |
1602 have "\<exists> stp. abc_steps_l (0, xs @ (x - Suc y) # rec_eval (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything) |
1603 ((?A [+] (?B [+] ?C))) stp = (length (?A [+] (?B [+] ?C)), xs @ (x - y) # 0 # |
1603 ((?A [+] (?B [+] ?C))) stp = (length (?A [+] (?B [+] ?C)), xs @ (x - y) # 0 # |
1604 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)" |
1604 rec_eval g (xs @ [x - Suc y, rec_eval (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)" |
1605 proof - |
1605 proof - |
1606 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} |
1606 have "{\<lambda> nl. nl = xs @ (x - Suc y) # rec_eval (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything} |
1607 (?A [+] (?B [+] ?C)) |
1607 (?A [+] (?B [+] ?C)) |
1608 {\<lambda> nl. nl = xs @ (x - y) # 0 # |
1608 {\<lambda> nl. nl = xs @ (x - y) # 0 # |
1609 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}" |
1609 rec_eval g (xs @ [x - Suc y, rec_eval (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything}" |
1610 proof(rule_tac abc_Hoare_plus_halt) |
1610 proof(rule_tac abc_Hoare_plus_halt) |
1611 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} |
1611 show "{\<lambda>nl. nl = xs @ (x - Suc y) # rec_eval (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything} |
1612 [Dec ft (length gap + 7)] |
1612 [Dec ft (length gap + 7)] |
1613 {\<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}" |
1613 {\<lambda>nl. nl = xs @ (x - Suc y) # rec_eval (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ y # anything}" |
1614 using ft len |
1614 using ft len |
1615 by(simp) |
1615 by(simp) |
1616 next |
1616 next |
1617 show |
1617 show |
1618 "{\<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} |
1618 "{\<lambda>nl. nl = xs @ (x - Suc y) # rec_eval (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ y # anything} |
1619 ?B [+] ?C |
1619 ?B [+] ?C |
1620 {\<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}" |
1620 {\<lambda>nl. nl = xs @ (x - y) # 0 # rec_eval g (xs @ [x - Suc y, rec_eval (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything}" |
1621 proof(rule_tac abc_Hoare_plus_halt) |
1621 proof(rule_tac abc_Hoare_plus_halt) |
1622 have a: "gar = Suc (Suc n)" |
1622 have a: "gar = Suc (Suc n)" |
1623 using compile_g termi_g len less |
1623 using compile_g termi_g len less |
1624 by simp |
1624 by simp |
1625 have b: "gft > gar" |
1625 have b: "gft > gar" |
1626 using compile_g |
1626 using compile_g |
1627 by(erule_tac footprint_ge) |
1627 by(erule_tac footprint_ge) |
1628 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 |
1628 show "{\<lambda>nl. nl = xs @ (x - Suc y) # rec_eval (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ y # anything} gap |
1629 {\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # |
1629 {\<lambda>nl. nl = xs @ (x - Suc y) # rec_eval (Pr n f g) (xs @ [x - Suc y]) # |
1630 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}" |
1630 rec_eval g (xs @ [x - Suc y, rec_eval (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything}" |
1631 proof - |
1631 proof - |
1632 have |
1632 have |
1633 "{\<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 |
1633 "{\<lambda>nl. nl = xs @ (x - Suc y) # rec_eval (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (gft - gar) @ 0\<up>(ft - gft) @ y # anything} gap |
1634 {\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # |
1634 {\<lambda>nl. nl = xs @ (x - Suc y) # rec_eval (Pr n f g) (xs @ [x - Suc y]) # |
1635 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}" |
1635 rec_eval g (xs @ [(x - Suc y), rec_eval (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (gft - Suc gar) @ 0\<up>(ft - gft) @ y # anything}" |
1636 using g_ind less by simp |
1636 using g_ind less by simp |
1637 thus "?thesis" |
1637 thus "?thesis" |
1638 using a b ft |
1638 using a b ft |
1639 by(simp add: replicate_merge_anywhere numeral_3_eq_3) |
1639 by(simp add: replicate_merge_anywhere numeral_3_eq_3) |
1640 qed |
1640 qed |
1641 next |
1641 next |
1642 show "{\<lambda>nl. nl = xs @ (x - Suc y) # |
1642 show "{\<lambda>nl. nl = xs @ (x - Suc y) # |
1643 rec_exec (Pr n f g) (xs @ [x - Suc y]) # |
1643 rec_eval (Pr n f g) (xs @ [x - Suc y]) # |
1644 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} |
1644 rec_eval g (xs @ [x - Suc y, rec_eval (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything} |
1645 [Inc n, Dec (Suc n) 3, Goto (Suc 0)] |
1645 [Inc n, Dec (Suc n) 3, Goto (Suc 0)] |
1646 {\<lambda>nl. nl = xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) |
1646 {\<lambda>nl. nl = xs @ (x - y) # 0 # rec_eval g (xs @ [x - Suc y, rec_eval (Pr n f g) |
1647 (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything}" |
1647 (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything}" |
1648 using len less |
1648 using len less |
1649 using adjust_paras[of xs n "x - Suc y" " rec_exec (Pr n f g) (xs @ [x - Suc y])" |
1649 using adjust_paras[of xs n "x - Suc y" " rec_eval (Pr n f g) (xs @ [x - Suc y])" |
1650 " rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # |
1650 " rec_eval g (xs @ [x - Suc y, rec_eval (Pr n f g) (xs @ [x - Suc y])]) # |
1651 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything"] |
1651 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything"] |
1652 by(simp add: Suc_diff_Suc) |
1652 by(simp add: Suc_diff_Suc) |
1653 qed |
1653 qed |
1654 qed |
1654 qed |
1655 thus "?thesis" |
1655 thus "?thesis" |
1656 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]) # |
1656 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_eval (Pr n f g) (xs @ [x - Suc y]) # |
1657 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything) |
1657 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything) |
1658 ([Dec ft (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) na", simp) |
1658 ([Dec ft (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) na", simp) |
1659 qed |
1659 qed |
1660 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) |
1660 then obtain stpa where "abc_steps_l (0, xs @ (x - Suc y) # rec_eval (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything) |
1661 ((?A [+] (?B [+] ?C))) stpa = (length (?A [+] (?B [+] ?C)), |
1661 ((?A [+] (?B [+] ?C))) stpa = (length (?A [+] (?B [+] ?C)), |
1662 xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) |
1662 xs @ (x - y) # 0 # rec_eval g (xs @ [x - Suc y, rec_eval (Pr n f g) (xs @ [x - Suc y])]) |
1663 # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)" .. |
1663 # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)" .. |
1664 thus "?thesis" |
1664 thus "?thesis" |
1665 by(erule_tac abc_append_frist_steps_halt_eq) |
1665 by(erule_tac abc_append_frist_steps_halt_eq) |
1666 qed |
1666 qed |
1667 moreover have |
1667 moreover have |
1668 "\<exists> stp. abc_steps_l (length (?A [+] (?B [+] ?C)), |
1668 "\<exists> stp. abc_steps_l (length (?A [+] (?B [+] ?C)), |
1669 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) |
1669 xs @ (x - y) # 0 # rec_eval g (xs @ [x - Suc y, rec_eval (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything) |
1670 ((?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])]) # |
1670 ((?A [+] (?B [+] ?C)) @ ?D) stp = (0, xs @ (x - y) # rec_eval g (xs @ [x - Suc y, rec_eval (Pr n f g) (xs @ [x - Suc y])]) # |
1671 0 # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)" |
1671 0 # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)" |
1672 using len |
1672 using len |
1673 by(rule_tac loop_back, simp_all) |
1673 by(rule_tac loop_back, simp_all) |
1674 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])" |
1674 moreover have "rec_eval g (xs @ [x - Suc y, rec_eval (Pr n f g) (xs @ [x - Suc y])]) = rec_eval (Pr n f g) (xs @ [x - y])" |
1675 using less |
1675 using less |
1676 thm rec_exec.simps |
1676 thm rec_eval.simps |
1677 apply(case_tac "x - y", simp_all add: rec_exec_pr_Suc_simps) |
1677 apply(case_tac "x - y", simp_all add: rec_eval_pr_Suc_simps) |
1678 by(subgoal_tac "nat = x - Suc y", simp, arith) |
1678 by(subgoal_tac "nat = x - Suc y", simp, arith) |
1679 ultimately show "?thesis" |
1679 ultimately show "?thesis" |
1680 using code ft |
1680 using code ft |
1681 by(auto, rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add replicate_Suc_iff_anywhere del: replicate_Suc) |
1681 by(auto, rule_tac x = "stp + stpa" in exI, simp add: abc_steps_add replicate_Suc_iff_anywhere del: replicate_Suc) |
1682 qed |
1682 qed |
1683 |
1683 |
1684 lemma [simp]: |
1684 lemma [simp]: |
1685 "length xs = n \<Longrightarrow> abc_lm_s (xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) |
1685 "length xs = n \<Longrightarrow> abc_lm_s (xs @ x # rec_eval (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) |
1686 (max fft gft) - Suc (Suc n)) @ 0 # anything) (max (n + 3) (max fft gft)) 0 = |
1686 (max fft gft) - Suc (Suc n)) @ 0 # anything) (max (n + 3) (max fft gft)) 0 = |
1687 xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ anything" |
1687 xs @ x # rec_eval (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ anything" |
1688 apply(simp add: abc_lm_s.simps) |
1688 apply(simp add: abc_lm_s.simps) |
1689 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))" |
1689 using list_update_length[of "xs @ x # rec_eval (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) (max fft gft) - Suc (Suc n))" |
1690 0 anything 0] |
1690 0 anything 0] |
1691 apply(auto simp: Suc_diff_Suc) |
1691 apply(auto simp: Suc_diff_Suc) |
1692 apply(simp add: exp_suc[THEN sym] Suc_diff_Suc del: replicate_Suc) |
1692 apply(simp add: exp_suc[THEN sym] Suc_diff_Suc del: replicate_Suc) |
1693 done |
1693 done |
1694 |
1694 |
1695 lemma [simp]: |
1695 lemma [simp]: |
1696 "(xs @ x # rec_exec (Pr (length xs) f g) (xs @ [x]) # 0 \<up> (max (length xs + 3) |
1696 "(xs @ x # rec_eval (Pr (length xs) f g) (xs @ [x]) # 0 \<up> (max (length xs + 3) |
1697 (max fft gft) - Suc (Suc (length xs))) @ 0 # anything) ! |
1697 (max fft gft) - Suc (Suc (length xs))) @ 0 # anything) ! |
1698 max (length xs + 3) (max fft gft) = 0" |
1698 max (length xs + 3) (max fft gft) = 0" |
1699 using nth_append_length[of "xs @ x # rec_exec (Pr (length xs) f g) (xs @ [x]) # |
1699 using nth_append_length[of "xs @ x # rec_eval (Pr (length xs) f g) (xs @ [x]) # |
1700 0 \<up> (max (length xs + 3) (max fft gft) - Suc (Suc (length xs)))" 0 anything] |
1700 0 \<up> (max (length xs + 3) (max fft gft) - Suc (Suc (length xs)))" 0 anything] |
1701 by(simp) |
1701 by(simp) |
1702 |
1702 |
1703 lemma pr_loop_correct: |
1703 lemma pr_loop_correct: |
1704 assumes less: "y \<le> x" |
1704 assumes less: "y \<le> x" |
1705 and len: "length xs = n" |
1705 and len: "length xs = n" |
1706 and compile_g: "rec_ci g = (gap, gar, gft)" |
1706 and compile_g: "rec_ci g = (gap, gar, gft)" |
1707 and termi_g: "\<forall> y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])" |
1707 and termi_g: "\<forall> y<x. terminates g (xs @ [y, rec_eval (Pr n f g) (xs @ [y])])" |
1708 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 |
1708 and g_ind: "\<forall> y<x. (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_eval (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap |
1709 {\<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})" |
1709 {\<lambda>nl. nl = xs @ y # rec_eval (Pr n f g) (xs @ [y]) # rec_eval g (xs @ [y, rec_eval (Pr n f g) (xs @ [y])]) # 0 \<up> (gft - Suc gar) @ anything})" |
1710 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} |
1710 shows "{\<lambda>nl. nl = xs @ (x - y) # rec_eval (Pr n f g) (xs @ [x - y]) # 0 \<up> (max (n + 3) (max fft gft) - Suc (Suc n)) @ y # anything} |
1711 ([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)] |
1711 ([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)] |
1712 {\<lambda>nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ anything}" |
1712 {\<lambda>nl. nl = xs @ x # rec_eval (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ anything}" |
1713 using less |
1713 using less |
1714 proof(induct y) |
1714 proof(induct y) |
1715 case 0 |
1715 case 0 |
1716 thus "?case" |
1716 thus "?case" |
1717 using len |
1717 using len |
1723 case (Suc y) |
1723 case (Suc y) |
1724 let ?ft = "max (n + 3) (max fft gft)" |
1724 let ?ft = "max (n + 3) (max fft gft)" |
1725 let ?C = "[Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] |
1725 let ?C = "[Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] |
1726 [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" |
1726 [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" |
1727 have ind: "y \<le> x \<Longrightarrow> |
1727 have ind: "y \<le> x \<Longrightarrow> |
1728 {\<lambda>nl. nl = xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything} |
1728 {\<lambda>nl. nl = xs @ (x - y) # rec_eval (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything} |
1729 ?C {\<lambda>nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (?ft - Suc n) @ anything}" by fact |
1729 ?C {\<lambda>nl. nl = xs @ x # rec_eval (Pr n f g) (xs @ [x]) # 0 \<up> (?ft - Suc n) @ anything}" by fact |
1730 have less: "Suc y \<le> x" by fact |
1730 have less: "Suc y \<le> x" by fact |
1731 have stp1: |
1731 have stp1: |
1732 "\<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) |
1732 "\<exists> stp. abc_steps_l (0, xs @ (x - Suc y) # rec_eval (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (?ft - Suc (Suc n)) @ Suc y # anything) |
1733 ?C stp = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything)" |
1733 ?C stp = (0, xs @ (x - y) # rec_eval (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything)" |
1734 using assms less |
1734 using assms less |
1735 by(rule_tac pr_loop, auto) |
1735 by(rule_tac pr_loop, auto) |
1736 then obtain stp1 where a: |
1736 then obtain stp1 where a: |
1737 "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) |
1737 "abc_steps_l (0, xs @ (x - Suc y) # rec_eval (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (?ft - Suc (Suc n)) @ Suc y # anything) |
1738 ?C stp1 = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything)" .. |
1738 ?C stp1 = (0, xs @ (x - y) # rec_eval (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything)" .. |
1739 moreover have |
1739 moreover have |
1740 "\<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) |
1740 "\<exists> stp. abc_steps_l (0, xs @ (x - y) # rec_eval (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything) |
1741 ?C stp = (length ?C, xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (?ft - Suc n) @ anything)" |
1741 ?C stp = (length ?C, xs @ x # rec_eval (Pr n f g) (xs @ [x]) # 0 \<up> (?ft - Suc n) @ anything)" |
1742 using ind less |
1742 using ind less |
1743 by(auto simp: abc_Hoare_halt_def, case_tac "abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g) |
1743 by(auto simp: abc_Hoare_halt_def, case_tac "abc_steps_l (0, xs @ (x - y) # rec_eval (Pr n f g) |
1744 (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything) ?C na", rule_tac x = na in exI, simp) |
1744 (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything) ?C na", rule_tac x = na in exI, simp) |
1745 then obtain stp2 where b: |
1745 then obtain stp2 where b: |
1746 "abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything) |
1746 "abc_steps_l (0, xs @ (x - y) # rec_eval (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything) |
1747 ?C stp2 = (length ?C, xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (?ft - Suc n) @ anything)" .. |
1747 ?C stp2 = (length ?C, xs @ x # rec_eval (Pr n f g) (xs @ [x]) # 0 \<up> (?ft - Suc n) @ anything)" .. |
1748 from a b show "?case" |
1748 from a b show "?case" |
1749 by(simp add: abc_Hoare_halt_def, rule_tac x = "stp1 + stp2" in exI, simp add: abc_steps_add) |
1749 by(simp add: abc_Hoare_halt_def, rule_tac x = "stp1 + stp2" in exI, simp add: abc_steps_add) |
1750 qed |
1750 qed |
1751 |
1751 |
1752 lemma compile_pr_correct': |
1752 lemma compile_pr_correct': |
1753 assumes termi_g: "\<forall> y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])" |
1753 assumes termi_g: "\<forall> y<x. terminates g (xs @ [y, rec_eval (Pr n f g) (xs @ [y])])" |
1754 and g_ind: |
1754 and g_ind: |
1755 "\<forall> y<x. (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap |
1755 "\<forall> y<x. (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_eval (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap |
1756 {\<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})" |
1756 {\<lambda>nl. nl = xs @ y # rec_eval (Pr n f g) (xs @ [y]) # rec_eval g (xs @ [y, rec_eval (Pr n f g) (xs @ [y])]) # 0 \<up> (gft - Suc gar) @ anything})" |
1757 and termi_f: "terminate f xs" |
1757 and termi_f: "terminates f xs" |
1758 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}" |
1758 and f_ind: "\<And> anything. {\<lambda>nl. nl = xs @ 0 \<up> (fft - far) @ anything} fap {\<lambda>nl. nl = xs @ rec_eval f xs # 0 \<up> (fft - Suc far) @ anything}" |
1759 and len: "length xs = n" |
1759 and len: "length xs = n" |
1760 and compile1: "rec_ci f = (fap, far, fft)" |
1760 and compile1: "rec_ci f = (fap, far, fft)" |
1761 and compile2: "rec_ci g = (gap, gar, gft)" |
1761 and compile2: "rec_ci g = (gap, gar, gft)" |
1762 shows |
1762 shows |
1763 "{\<lambda>nl. nl = xs @ x # 0 \<up> (max (n + 3) (max fft gft) - n) @ anything} |
1763 "{\<lambda>nl. nl = xs @ x # 0 \<up> (max (n + 3) (max fft gft) - n) @ anything} |
1764 mv_box n (max (n + 3) (max fft gft)) [+] |
1764 mv_box n (max (n + 3) (max fft gft)) [+] |
1765 (fap [+] (mv_box n (Suc n) [+] |
1765 (fap [+] (mv_box n (Suc n) [+] |
1766 ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ |
1766 ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ |
1767 [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]))) |
1767 [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]))) |
1768 {\<lambda>nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ anything}" |
1768 {\<lambda>nl. nl = xs @ x # rec_eval (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ anything}" |
1769 proof - |
1769 proof - |
1770 let ?ft = "max (n+3) (max fft gft)" |
1770 let ?ft = "max (n+3) (max fft gft)" |
1771 let ?A = "mv_box n ?ft" |
1771 let ?A = "mv_box n ?ft" |
1772 let ?B = "fap" |
1772 let ?B = "fap" |
1773 let ?C = "mv_box n (Suc n)" |
1773 let ?C = "mv_box n (Suc n)" |
1774 let ?D = "[Dec ?ft (length gap + 7)]" |
1774 let ?D = "[Dec ?ft (length gap + 7)]" |
1775 let ?E = "gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]" |
1775 let ?E = "gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]" |
1776 let ?F = "[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" |
1776 let ?F = "[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]" |
1777 let ?P = "\<lambda>nl. nl = xs @ x # 0 \<up> (?ft - n) @ anything" |
1777 let ?P = "\<lambda>nl. nl = xs @ x # 0 \<up> (?ft - n) @ anything" |
1778 let ?S = "\<lambda>nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (?ft - Suc n) @ anything" |
1778 let ?S = "\<lambda>nl. nl = xs @ x # rec_eval (Pr n f g) (xs @ [x]) # 0 \<up> (?ft - Suc n) @ anything" |
1779 let ?Q1 = "\<lambda>nl. nl = xs @ 0 \<up> (?ft - n) @ x # anything" |
1779 let ?Q1 = "\<lambda>nl. nl = xs @ 0 \<up> (?ft - n) @ x # anything" |
1780 show "{?P} (?A [+] (?B [+] (?C [+] (?D [+] ?E @ ?F)))) {?S}" |
1780 show "{?P} (?A [+] (?B [+] (?C [+] (?D [+] ?E @ ?F)))) {?S}" |
1781 proof(rule_tac abc_Hoare_plus_halt) |
1781 proof(rule_tac abc_Hoare_plus_halt) |
1782 show "{?P} ?A {?Q1}" |
1782 show "{?P} ?A {?Q1}" |
1783 using len by simp |
1783 using len by simp |
1784 next |
1784 next |
1785 let ?Q2 = "\<lambda>nl. nl = xs @ rec_exec f xs # 0 \<up> (?ft - Suc n) @ x # anything" |
1785 let ?Q2 = "\<lambda>nl. nl = xs @ rec_eval f xs # 0 \<up> (?ft - Suc n) @ x # anything" |
1786 have a: "?ft \<ge> fft" |
1786 have a: "?ft \<ge> fft" |
1787 by arith |
1787 by arith |
1788 have b: "far = n" |
1788 have b: "far = n" |
1789 using compile1 termi_f len |
1789 using compile1 termi_f len |
1790 by(drule_tac param_pattern, auto) |
1790 by(drule_tac param_pattern, auto) |
1792 using compile1 |
1792 using compile1 |
1793 by(simp add: footprint_ge) |
1793 by(simp add: footprint_ge) |
1794 show "{?Q1} (?B [+] (?C [+] (?D [+] ?E @ ?F))) {?S}" |
1794 show "{?Q1} (?B [+] (?C [+] (?D [+] ?E @ ?F))) {?S}" |
1795 proof(rule_tac abc_Hoare_plus_halt) |
1795 proof(rule_tac abc_Hoare_plus_halt) |
1796 have "{\<lambda>nl. nl = xs @ 0 \<up> (fft - far) @ 0\<up>(?ft - fft) @ x # anything} fap |
1796 have "{\<lambda>nl. nl = xs @ 0 \<up> (fft - far) @ 0\<up>(?ft - fft) @ x # anything} fap |
1797 {\<lambda>nl. nl = xs @ rec_exec f xs # 0 \<up> (fft - Suc far) @ 0\<up>(?ft - fft) @ x # anything}" |
1797 {\<lambda>nl. nl = xs @ rec_eval f xs # 0 \<up> (fft - Suc far) @ 0\<up>(?ft - fft) @ x # anything}" |
1798 by(rule_tac f_ind) |
1798 by(rule_tac f_ind) |
1799 moreover have "fft - far + ?ft - fft = ?ft - far" |
1799 moreover have "fft - far + ?ft - fft = ?ft - far" |
1800 using a b c by arith |
1800 using a b c by arith |
1801 moreover have "fft - Suc n + ?ft - fft = ?ft - Suc n" |
1801 moreover have "fft - Suc n + ?ft - fft = ?ft - Suc n" |
1802 using a b c by arith |
1802 using a b c by arith |
1803 ultimately show "{?Q1} ?B {?Q2}" |
1803 ultimately show "{?Q1} ?B {?Q2}" |
1804 using b |
1804 using b |
1805 by(simp add: replicate_merge_anywhere) |
1805 by(simp add: replicate_merge_anywhere) |
1806 next |
1806 next |
1807 let ?Q3 = "\<lambda> nl. nl = xs @ 0 # rec_exec f xs # 0\<up>(?ft - Suc (Suc n)) @ x # anything" |
1807 let ?Q3 = "\<lambda> nl. nl = xs @ 0 # rec_eval f xs # 0\<up>(?ft - Suc (Suc n)) @ x # anything" |
1808 show "{?Q2} (?C [+] (?D [+] ?E @ ?F)) {?S}" |
1808 show "{?Q2} (?C [+] (?D [+] ?E @ ?F)) {?S}" |
1809 proof(rule_tac abc_Hoare_plus_halt) |
1809 proof(rule_tac abc_Hoare_plus_halt) |
1810 show "{?Q2} (?C) {?Q3}" |
1810 show "{?Q2} (?C) {?Q3}" |
1811 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"] |
1811 using mv_box_correct[of n "Suc n" "xs @ rec_eval f xs # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ x # anything"] |
1812 using len |
1812 using len |
1813 by(auto) |
1813 by(auto) |
1814 next |
1814 next |
1815 show "{?Q3} (?D [+] ?E @ ?F) {?S}" |
1815 show "{?Q3} (?D [+] ?E @ ?F) {?S}" |
1816 using pr_loop_correct[of x x xs n g gap gar gft f fft anything] assms |
1816 using pr_loop_correct[of x x xs n g gap gar gft f fft anything] assms |
1817 by(simp add: rec_exec_pr_0_simps) |
1817 by(simp add: rec_eval_pr_0_simps) |
1818 qed |
1818 qed |
1819 qed |
1819 qed |
1820 qed |
1820 qed |
1821 qed |
1821 qed |
1822 |
1822 |
1823 lemma compile_pr_correct: |
1823 lemma compile_pr_correct: |
1824 assumes g_ind: "\<forall>y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) \<and> |
1824 assumes g_ind: "\<forall>y<x. terminates g (xs @ [y, rec_eval (Pr n f g) (xs @ [y])]) \<and> |
1825 (\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow> |
1825 (\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow> |
1826 (\<forall>xc. {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \<up> (xb - xa) @ xc} x |
1826 (\<forall>xc. {\<lambda>nl. nl = xs @ y # rec_eval (Pr n f g) (xs @ [y]) # 0 \<up> (xb - xa) @ xc} x |
1827 {\<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}))" |
1827 {\<lambda>nl. nl = xs @ y # rec_eval (Pr n f g) (xs @ [y]) # rec_eval g (xs @ [y, rec_eval (Pr n f g) (xs @ [y])]) # 0 \<up> (xb - Suc xa) @ xc}))" |
1828 and termi_f: "terminate f xs" |
1828 and termi_f: "terminates f xs" |
1829 and f_ind: |
1829 and f_ind: |
1830 "\<And>ap arity fp anything. |
1830 "\<And>ap arity fp anything. |
1831 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}" |
1831 rec_ci f = (ap, arity, fp) \<Longrightarrow> {\<lambda>nl. nl = xs @ 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ rec_eval f xs # 0 \<up> (fp - Suc arity) @ anything}" |
1832 and len: "length xs = n" |
1832 and len: "length xs = n" |
1833 and compile: "rec_ci (Pr n f g) = (ap, arity, fp)" |
1833 and compile: "rec_ci (Pr n f g) = (ap, arity, fp)" |
1834 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}" |
1834 shows "{\<lambda>nl. nl = xs @ x # 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ x # rec_eval (Pr n f g) (xs @ [x]) # 0 \<up> (fp - Suc arity) @ anything}" |
1835 proof(case_tac "rec_ci f", case_tac "rec_ci g") |
1835 proof(case_tac "rec_ci f", case_tac "rec_ci g") |
1836 fix fap far fft gap gar gft |
1836 fix fap far fft gap gar gft |
1837 assume h: "rec_ci f = (fap, far, fft)" "rec_ci g = (gap, gar, gft)" |
1837 assume h: "rec_ci f = (fap, far, fft)" "rec_ci g = (gap, gar, gft)" |
1838 have g: |
1838 have g: |
1839 "\<forall>y<x. (terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) \<and> |
1839 "\<forall>y<x. (terminates g (xs @ [y, rec_eval (Pr n f g) (xs @ [y])]) \<and> |
1840 (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap |
1840 (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_eval (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap |
1841 {\<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}))" |
1841 {\<lambda>nl. nl = xs @ y # rec_eval (Pr n f g) (xs @ [y]) # rec_eval g (xs @ [y, rec_eval (Pr n f g) (xs @ [y])]) # 0 \<up> (gft - Suc gar) @ anything}))" |
1842 using g_ind h |
1842 using g_ind h |
1843 by(auto) |
1843 by(auto) |
1844 hence termi_g: "\<forall> y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])" |
1844 hence termi_g: "\<forall> y<x. terminates g (xs @ [y, rec_eval (Pr n f g) (xs @ [y])])" |
1845 by simp |
1845 by simp |
1846 from g have g_newind: |
1846 from g have g_newind: |
1847 "\<forall> y<x. (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap |
1847 "\<forall> y<x. (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_eval (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap |
1848 {\<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})" |
1848 {\<lambda>nl. nl = xs @ y # rec_eval (Pr n f g) (xs @ [y]) # rec_eval g (xs @ [y, rec_eval (Pr n f g) (xs @ [y])]) # 0 \<up> (gft - Suc gar) @ anything})" |
1849 by auto |
1849 by auto |
1850 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}" |
1850 have f_newind: "\<And> anything. {\<lambda>nl. nl = xs @ 0 \<up> (fft - far) @ anything} fap {\<lambda>nl. nl = xs @ rec_eval f xs # 0 \<up> (fft - Suc far) @ anything}" |
1851 using h |
1851 using h |
1852 by(rule_tac f_ind, simp) |
1852 by(rule_tac f_ind, simp) |
1853 show "?thesis" |
1853 show "?thesis" |
1854 using termi_f termi_g h compile |
1854 using termi_f termi_g h compile |
1855 apply(simp add: rec_ci.simps abc_comp_commute, auto) |
1855 apply(simp add: rec_ci.simps abc_comp_commute, auto) |
1982 assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" |
1982 assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]" |
1983 and ft: "ft = max (Suc arity) fft" |
1983 and ft: "ft = max (Suc arity) fft" |
1984 and len: "length xs = arity" |
1984 and len: "length xs = arity" |
1985 and far: "far = Suc arity" |
1985 and far: "far = Suc arity" |
1986 and ind: " (\<forall>xc. ({\<lambda>nl. nl = xs @ x # 0 \<up> (fft - far) @ xc} fap |
1986 and ind: " (\<forall>xc. ({\<lambda>nl. nl = xs @ x # 0 \<up> (fft - far) @ xc} fap |
1987 {\<lambda>nl. nl = xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (fft - Suc far) @ xc}))" |
1987 {\<lambda>nl. nl = xs @ x # rec_eval f (xs @ [x]) # 0 \<up> (fft - Suc far) @ xc}))" |
1988 and exec_less: "rec_exec f (xs @ [x]) > 0" |
1988 and exec_less: "rec_eval f (xs @ [x]) > 0" |
1989 and compile: "rec_ci f = (fap, far, fft)" |
1989 and compile: "rec_ci f = (fap, far, fft)" |
1990 shows "\<exists> stp > 0. abc_steps_l (0, xs @ x # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp = |
1990 shows "\<exists> stp > 0. abc_steps_l (0, xs @ x # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp = |
1991 (0, xs @ Suc x # 0 \<up> (ft - Suc arity) @ anything)" |
1991 (0, xs @ Suc x # 0 \<up> (ft - Suc arity) @ anything)" |
1992 proof - |
1992 proof - |
1993 have "\<exists> stp. abc_steps_l (0, xs @ x # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp = |
1993 have "\<exists> stp. abc_steps_l (0, xs @ x # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp = |
1994 (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" |
1994 (length fap, xs @ x # rec_eval f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" |
1995 proof - |
1995 proof - |
1996 have "\<exists> stp. abc_steps_l (0, xs @ x # 0 \<up> (ft - Suc arity) @ anything) fap stp = |
1996 have "\<exists> stp. abc_steps_l (0, xs @ x # 0 \<up> (ft - Suc arity) @ anything) fap stp = |
1997 (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" |
1997 (length fap, xs @ x # rec_eval f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" |
1998 proof - |
1998 proof - |
1999 have "{\<lambda>nl. nl = xs @ x # 0 \<up> (fft - far) @ 0\<up>(ft - fft) @ anything} fap |
1999 have "{\<lambda>nl. nl = xs @ x # 0 \<up> (fft - far) @ 0\<up>(ft - fft) @ anything} fap |
2000 {\<lambda>nl. nl = xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (fft - Suc far) @ 0\<up>(ft - fft) @ anything}" |
2000 {\<lambda>nl. nl = xs @ x # rec_eval f (xs @ [x]) # 0 \<up> (fft - Suc far) @ 0\<up>(ft - fft) @ anything}" |
2001 using ind by simp |
2001 using ind by simp |
2002 moreover have "fft > far" |
2002 moreover have "fft > far" |
2003 using compile |
2003 using compile |
2004 by(erule_tac footprint_ge) |
2004 by(erule_tac footprint_ge) |
2005 ultimately show "?thesis" |
2005 ultimately show "?thesis" |
2006 using ft far |
2006 using ft far |
2007 apply(drule_tac abc_Hoare_haltE) |
2007 apply(drule_tac abc_Hoare_haltE) |
2008 by(simp add: replicate_merge_anywhere max_absorb2) |
2008 by(simp add: replicate_merge_anywhere max_absorb2) |
2009 qed |
2009 qed |
2010 then obtain stp where "abc_steps_l (0, xs @ x # 0 \<up> (ft - Suc arity) @ anything) fap stp = |
2010 then obtain stp where "abc_steps_l (0, xs @ x # 0 \<up> (ft - Suc arity) @ anything) fap stp = |
2011 (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" .. |
2011 (length fap, xs @ x # rec_eval f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" .. |
2012 thus "?thesis" |
2012 thus "?thesis" |
2013 by(erule_tac abc_append_frist_steps_halt_eq) |
2013 by(erule_tac abc_append_frist_steps_halt_eq) |
2014 qed |
2014 qed |
2015 moreover have |
2015 moreover have |
2016 "\<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 = |
2016 "\<exists> stp > 0. abc_steps_l (length fap, xs @ x # rec_eval f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything) (fap @ B) stp = |
2017 (0, xs @ Suc x # 0 # 0 \<up> (ft - Suc (Suc arity)) @ anything)" |
2017 (0, xs @ Suc x # 0 # 0 \<up> (ft - Suc (Suc arity)) @ anything)" |
2018 using mn_correct[of f fap far fft "rec_exec f (xs @ [x])" xs arity B |
2018 using mn_correct[of f fap far fft "rec_eval f (xs @ [x])" xs arity B |
2019 "(\<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))" |
2019 "(\<lambda>stp. (abc_steps_l (length fap, xs @ x # rec_eval f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything) (fap @ B) stp, length fap, arity))" |
2020 x "0 \<up> (ft - Suc (Suc arity)) @ anything" "(\<lambda>((as, lm), ss, arity). as = 0)" |
2020 x "0 \<up> (ft - Suc (Suc arity)) @ anything" "(\<lambda>((as, lm), ss, arity). as = 0)" |
2021 "(\<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) "] |
2021 "(\<lambda>((as, lm), ss, aritya). mn_ind_inv (as, lm) (length fap) x (rec_eval f (xs @ [x])) (0 \<up> (ft - Suc (Suc arity)) @ anything) xs) "] |
2022 B compile exec_less len |
2022 B compile exec_less len |
2023 apply(subgoal_tac "fap \<noteq> []", auto) |
2023 apply(subgoal_tac "fap \<noteq> []", auto) |
2024 apply(rule_tac x = stp in exI, auto simp: mn_ind_inv.simps) |
2024 apply(rule_tac x = stp in exI, auto simp: mn_ind_inv.simps) |
2025 by(case_tac "stp = 0", simp_all add: abc_steps_l.simps) |
2025 by(case_tac "stp = 0", simp_all add: abc_steps_l.simps) |
2026 moreover have "fft > far" |
2026 moreover have "fft > far" |
2115 using assms a |
2115 using assms a |
2116 apply(case_tac r, rule_tac x = 0 in exI, simp add: abc_steps_l.simps, simp) |
2116 apply(case_tac r, rule_tac x = 0 in exI, simp add: abc_steps_l.simps, simp) |
2117 by(rule_tac mn_loop_correct, auto) |
2117 by(rule_tac mn_loop_correct, auto) |
2118 moreover have |
2118 moreover have |
2119 "\<exists> stp. abc_steps_l (0, xs @ r # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp = |
2119 "\<exists> stp. abc_steps_l (0, xs @ r # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp = |
2120 (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" |
2120 (length fap, xs @ r # rec_eval f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" |
2121 proof - |
2121 proof - |
2122 have "\<exists> stp. abc_steps_l (0, xs @ r # 0 \<up> (ft - Suc arity) @ anything) fap stp = |
2122 have "\<exists> stp. abc_steps_l (0, xs @ r # 0 \<up> (ft - Suc arity) @ anything) fap stp = |
2123 (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" |
2123 (length fap, xs @ r # rec_eval f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" |
2124 proof - |
2124 proof - |
2125 have "{\<lambda>nl. nl = xs @ r # 0 \<up> (fft - far) @ 0\<up>(ft - fft) @ anything} fap |
2125 have "{\<lambda>nl. nl = xs @ r # 0 \<up> (fft - far) @ 0\<up>(ft - fft) @ anything} fap |
2126 {\<lambda>nl. nl = xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (fft - Suc far) @ 0\<up>(ft - fft) @ anything}" |
2126 {\<lambda>nl. nl = xs @ r # rec_eval f (xs @ [r]) # 0 \<up> (fft - Suc far) @ 0\<up>(ft - fft) @ anything}" |
2127 using f_ind exec by simp |
2127 using f_ind exec by simp |
2128 thus "?thesis" |
2128 thus "?thesis" |
2129 using ft a b |
2129 using ft a b |
2130 apply(drule_tac abc_Hoare_haltE) |
2130 apply(drule_tac abc_Hoare_haltE) |
2131 by(simp add: replicate_merge_anywhere max_absorb2) |
2131 by(simp add: replicate_merge_anywhere max_absorb2) |
2132 qed |
2132 qed |
2133 then obtain stp where "abc_steps_l (0, xs @ r # 0 \<up> (ft - Suc arity) @ anything) fap stp = |
2133 then obtain stp where "abc_steps_l (0, xs @ r # 0 \<up> (ft - Suc arity) @ anything) fap stp = |
2134 (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" .. |
2134 (length fap, xs @ r # rec_eval f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" .. |
2135 thus "?thesis" |
2135 thus "?thesis" |
2136 by(erule_tac abc_append_frist_steps_halt_eq) |
2136 by(erule_tac abc_append_frist_steps_halt_eq) |
2137 qed |
2137 qed |
2138 moreover have |
2138 moreover have |
2139 "\<exists> stp. abc_steps_l (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything) (fap @ B) stp = |
2139 "\<exists> stp. abc_steps_l (length fap, xs @ r # rec_eval f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything) (fap @ B) stp = |
2140 (length fap + 5, xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" |
2140 (length fap + 5, xs @ r # rec_eval f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" |
2141 using ft a b len B exec |
2141 using ft a b len B exec |
2142 apply(rule_tac x = 1 in exI, auto) |
2142 apply(rule_tac x = 1 in exI, auto) |
2143 by(auto simp: abc_steps_l.simps B abc_step_l.simps |
2143 by(auto simp: abc_steps_l.simps B abc_step_l.simps |
2144 abc_fetch.simps nth_append max_absorb2 abc_lm_v.simps abc_lm_s.simps) |
2144 abc_fetch.simps nth_append max_absorb2 abc_lm_v.simps abc_lm_s.simps) |
2145 moreover have "rec_exec (Mn (length xs) f) xs = r" |
2145 moreover have "rec_eval (Mn (length xs) f) xs = r" |
2146 using exec exec_less |
2146 using exec exec_less |
2147 apply(auto simp: rec_exec.simps Least_def) |
2147 apply(auto simp: rec_eval.simps Least_def) |
2148 thm the_equality |
2148 thm the_equality |
2149 apply(rule_tac the_equality, auto) |
2149 apply(rule_tac the_equality, auto) |
2150 apply(metis exec_less less_not_refl3 linorder_not_less) |
2150 apply(metis exec_less less_not_refl3 linorder_not_less) |
2151 by (metis le_neq_implies_less less_not_refl3) |
2151 by (metis le_neq_implies_less less_not_refl3) |
2152 ultimately show "?thesis" |
2152 ultimately show "?thesis" |
2156 by(simp add: abc_steps_add replicate_Suc_iff_anywhere max_absorb2 Suc_diff_Suc del: replicate_Suc) |
2156 by(simp add: abc_steps_add replicate_Suc_iff_anywhere max_absorb2 Suc_diff_Suc del: replicate_Suc) |
2157 qed |
2157 qed |
2158 |
2158 |
2159 lemma compile_mn_correct: |
2159 lemma compile_mn_correct: |
2160 assumes len: "length xs = n" |
2160 assumes len: "length xs = n" |
2161 and termi_f: "terminate f (xs @ [r])" |
2161 and termi_f: "terminates f (xs @ [r])" |
2162 and f_ind: "\<And>ap arity fp anything. rec_ci f = (ap, arity, fp) \<Longrightarrow> |
2162 and f_ind: "\<And>ap arity fp anything. rec_ci f = (ap, arity, fp) \<Longrightarrow> |
2163 {\<lambda>nl. nl = xs @ r # 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ r # 0 # 0 \<up> (fp - Suc arity) @ anything}" |
2163 {\<lambda>nl. nl = xs @ r # 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ r # 0 # 0 \<up> (fp - Suc arity) @ anything}" |
2164 and exec: "rec_exec f (xs @ [r]) = 0" |
2164 and exec: "rec_eval f (xs @ [r]) = 0" |
2165 and ind_all: |
2165 and ind_all: |
2166 "\<forall>i<r. terminate f (xs @ [i]) \<and> |
2166 "\<forall>i<r. terminates f (xs @ [i]) \<and> |
2167 (\<forall>x xa xb. rec_ci f = (x, xa, xb) \<longrightarrow> |
2167 (\<forall>x xa xb. rec_ci f = (x, xa, xb) \<longrightarrow> |
2168 (\<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> |
2168 (\<forall>xc. {\<lambda>nl. nl = xs @ i # 0 \<up> (xb - xa) @ xc} x {\<lambda>nl. nl = xs @ i # rec_eval f (xs @ [i]) # 0 \<up> (xb - Suc xa) @ xc})) \<and> |
2169 0 < rec_exec f (xs @ [i])" |
2169 0 < rec_eval f (xs @ [i])" |
2170 and compile: "rec_ci (Mn n f) = (ap, arity, fp)" |
2170 and compile: "rec_ci (Mn n f) = (ap, arity, fp)" |
2171 shows "{\<lambda>nl. nl = xs @ 0 \<up> (fp - arity) @ anything} ap |
2171 shows "{\<lambda>nl. nl = xs @ 0 \<up> (fp - arity) @ anything} ap |
2172 {\<lambda>nl. nl = xs @ rec_exec (Mn n f) xs # 0 \<up> (fp - Suc arity) @ anything}" |
2172 {\<lambda>nl. nl = xs @ rec_eval (Mn n f) xs # 0 \<up> (fp - Suc arity) @ anything}" |
2173 proof(case_tac "rec_ci f") |
2173 proof(case_tac "rec_ci f") |
2174 fix fap far fft |
2174 fix fap far fft |
2175 assume h: "rec_ci f = (fap, far, fft)" |
2175 assume h: "rec_ci f = (fap, far, fft)" |
2176 hence f_newind: |
2176 hence f_newind: |
2177 "\<And>anything. {\<lambda>nl. nl = xs @ r # 0 \<up> (fft - far) @ anything} fap |
2177 "\<And>anything. {\<lambda>nl. nl = xs @ r # 0 \<up> (fft - far) @ anything} fap |
2178 {\<lambda>nl. nl = xs @ r # 0 # 0 \<up> (fft - Suc far) @ anything}" |
2178 {\<lambda>nl. nl = xs @ r # 0 # 0 \<up> (fft - Suc far) @ anything}" |
2179 by(rule_tac f_ind, simp) |
2179 by(rule_tac f_ind, simp) |
2180 have newind_all: |
2180 have newind_all: |
2181 "\<forall>i < r. (\<forall>xc. ({\<lambda>nl. nl = xs @ i # 0 \<up> (fft - far) @ xc} fap |
2181 "\<forall>i < r. (\<forall>xc. ({\<lambda>nl. nl = xs @ i # 0 \<up> (fft - far) @ xc} fap |
2182 {\<lambda>nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \<up> (fft - Suc far) @ xc}))" |
2182 {\<lambda>nl. nl = xs @ i # rec_eval f (xs @ [i]) # 0 \<up> (fft - Suc far) @ xc}))" |
2183 using ind_all h |
2183 using ind_all h |
2184 by(auto) |
2184 by(auto) |
2185 have all_less: "\<forall> i<r. rec_exec f (xs @ [i]) > 0" |
2185 have all_less: "\<forall> i<r. rec_eval f (xs @ [i]) > 0" |
2186 using ind_all by auto |
2186 using ind_all by auto |
2187 show "?thesis" |
2187 show "?thesis" |
2188 using h compile f_newind newind_all all_less len termi_f exec |
2188 using h compile f_newind newind_all all_less len termi_f exec |
2189 apply(auto simp: rec_ci.simps) |
2189 apply(auto simp: rec_ci.simps) |
2190 by(rule_tac compile_mn_correct', auto) |
2190 by(rule_tac compile_mn_correct', auto) |
2191 qed |
2191 qed |
2192 |
2192 |
2193 lemma recursive_compile_correct: |
2193 lemma recursive_compile_correct: |
2194 "\<lbrakk>terminate recf args; rec_ci recf = (ap, arity, fp)\<rbrakk> |
2194 "\<lbrakk>terminates recf args; rec_ci recf = (ap, arity, fp)\<rbrakk> |
2195 \<Longrightarrow> {\<lambda> nl. nl = args @ 0\<up>(fp - arity) @ anything} ap |
2195 \<Longrightarrow> {\<lambda> nl. nl = args @ 0\<up>(fp - arity) @ anything} ap |
2196 {\<lambda> nl. nl = args@ rec_exec recf args # 0\<up>(fp - Suc arity) @ anything}" |
2196 {\<lambda> nl. nl = args@ rec_eval recf args # 0\<up>(fp - Suc arity) @ anything}" |
2197 apply(induct arbitrary: ap arity fp anything r rule: terminate.induct) |
2197 apply(induct arbitrary: ap arity fp anything r rule: terminates.induct) |
2198 apply(simp_all add: compile_s_correct compile_z_correct compile_id_correct |
2198 apply(simp_all add: compile_s_correct compile_z_correct compile_id_correct |
2199 compile_cn_correct compile_pr_correct compile_mn_correct) |
2199 compile_cn_correct compile_pr_correct compile_mn_correct) |
2200 done |
2200 done |
2201 |
2201 |
2202 definition dummy_abc :: "nat \<Rightarrow> abc_inst list" |
2202 definition dummy_abc :: "nat \<Rightarrow> abc_inst list" |
2321 by auto |
2321 by auto |
2322 qed |
2322 qed |
2323 |
2323 |
2324 lemma recursive_compile_correct_norm': |
2324 lemma recursive_compile_correct_norm': |
2325 "\<lbrakk>rec_ci f = (ap, arity, ft); |
2325 "\<lbrakk>rec_ci f = (ap, arity, ft); |
2326 terminate f args\<rbrakk> |
2326 terminates f args\<rbrakk> |
2327 \<Longrightarrow> \<exists> stp rl. (abc_steps_l (0, args) ap stp) = (length ap, rl) \<and> abc_list_crsp (args @ [rec_exec f args]) rl" |
2327 \<Longrightarrow> \<exists> stp rl. (abc_steps_l (0, args) ap stp) = (length ap, rl) \<and> abc_list_crsp (args @ [rec_eval f args]) rl" |
2328 using recursive_compile_correct[of f args ap arity ft "[]"] |
2328 using recursive_compile_correct[of f args ap arity ft "[]"] |
2329 apply(auto simp: abc_Hoare_halt_def) |
2329 apply(auto simp: abc_Hoare_halt_def) |
2330 apply(rule_tac x = n in exI) |
2330 apply(rule_tac x = n in exI) |
2331 apply(case_tac "abc_steps_l (0, args @ 0 \<up> (ft - arity)) ap n", auto) |
2331 apply(case_tac "abc_steps_l (0, args @ 0 \<up> (ft - arity)) ap n", auto) |
2332 apply(drule_tac abc_list_crsp_steps, auto) |
2332 apply(drule_tac abc_list_crsp_steps, auto) |
2333 apply(rule_tac list_crsp_simp2, auto) |
2333 apply(rule_tac list_crsp_simp2, auto) |
2334 done |
2334 done |
2335 |
2335 |
2336 lemma [simp]: |
2336 lemma [simp]: |
2337 assumes a: "args @ [rec_exec f args] = lm @ 0 \<up> n" |
2337 assumes a: "args @ [rec_eval f args] = lm @ 0 \<up> n" |
2338 and b: "length args < length lm" |
2338 and b: "length args < length lm" |
2339 shows "\<exists>m. lm = args @ rec_exec f args # 0 \<up> m" |
2339 shows "\<exists>m. lm = args @ rec_eval f args # 0 \<up> m" |
2340 using assms |
2340 using assms |
2341 apply(case_tac n, simp) |
2341 apply(case_tac n, simp) |
2342 apply(rule_tac x = 0 in exI, simp) |
2342 apply(rule_tac x = 0 in exI, simp) |
2343 apply(drule_tac length_equal, simp) |
2343 apply(drule_tac length_equal, simp) |
2344 done |
2344 done |
2345 |
2345 |
2346 lemma [simp]: |
2346 lemma [simp]: |
2347 "\<lbrakk>args @ [rec_exec f args] = lm @ 0 \<up> n; \<not> length args < length lm\<rbrakk> |
2347 "\<lbrakk>args @ [rec_eval f args] = lm @ 0 \<up> n; \<not> length args < length lm\<rbrakk> |
2348 \<Longrightarrow> \<exists>m. (lm @ 0 \<up> (length args - length lm) @ [Suc 0])[length args := 0] = |
2348 \<Longrightarrow> \<exists>m. (lm @ 0 \<up> (length args - length lm) @ [Suc 0])[length args := 0] = |
2349 args @ rec_exec f args # 0 \<up> m" |
2349 args @ rec_eval f args # 0 \<up> m" |
2350 apply(case_tac n, simp_all add: exp_suc list_update_append list_update.simps del: replicate_Suc) |
2350 apply(case_tac n, simp_all add: exp_suc list_update_append list_update.simps del: replicate_Suc) |
2351 apply(subgoal_tac "length args = Suc (length lm)", simp) |
2351 apply(subgoal_tac "length args = Suc (length lm)", simp) |
2352 apply(rule_tac x = "Suc (Suc 0)" in exI, simp) |
2352 apply(rule_tac x = "Suc (Suc 0)" in exI, simp) |
2353 apply(drule_tac length_equal, simp, auto) |
2353 apply(drule_tac length_equal, simp, auto) |
2354 done |
2354 done |
2355 |
2355 |
2356 lemma compile_append_dummy_correct: |
2356 lemma compile_append_dummy_correct: |
2357 assumes compile: "rec_ci f = (ap, ary, fp)" |
2357 assumes compile: "rec_ci f = (ap, ary, fp)" |
2358 and termi: "terminate f args" |
2358 and termi: "terminates f args" |
2359 shows "{\<lambda> nl. nl = args} (ap [+] dummy_abc (length args)) {\<lambda> nl. (\<exists> m. nl = args @ rec_exec f args # 0\<up>m)}" |
2359 shows "{\<lambda> nl. nl = args} (ap [+] dummy_abc (length args)) {\<lambda> nl. (\<exists> m. nl = args @ rec_eval f args # 0\<up>m)}" |
2360 proof(rule_tac abc_Hoare_plus_halt) |
2360 proof(rule_tac abc_Hoare_plus_halt) |
2361 show "{\<lambda>nl. nl = args} ap {\<lambda> nl. abc_list_crsp (args @ [rec_exec f args]) nl}" |
2361 show "{\<lambda>nl. nl = args} ap {\<lambda> nl. abc_list_crsp (args @ [rec_eval f args]) nl}" |
2362 using compile termi recursive_compile_correct_norm'[of f ap ary fp args] |
2362 using compile termi recursive_compile_correct_norm'[of f ap ary fp args] |
2363 apply(auto simp: abc_Hoare_halt_def) |
2363 apply(auto simp: abc_Hoare_halt_def) |
2364 by(rule_tac x = stp in exI, simp) |
2364 by(rule_tac x = stp in exI, simp) |
2365 next |
2365 next |
2366 show "{abc_list_crsp (args @ [rec_exec f args])} dummy_abc (length args) |
2366 show "{abc_list_crsp (args @ [rec_eval f args])} dummy_abc (length args) |
2367 {\<lambda>nl. \<exists>m. nl = args @ rec_exec f args # 0 \<up> m}" |
2367 {\<lambda>nl. \<exists>m. nl = args @ rec_eval f args # 0 \<up> m}" |
2368 apply(auto simp: dummy_abc_def abc_Hoare_halt_def) |
2368 apply(auto simp: dummy_abc_def abc_Hoare_halt_def) |
2369 apply(rule_tac x = 3 in exI) |
2369 apply(rule_tac x = 3 in exI) |
2370 by(auto simp: abc_steps_l.simps abc_list_crsp_def abc_step_l.simps numeral_3_eq_3 abc_fetch.simps |
2370 by(auto simp: abc_steps_l.simps abc_list_crsp_def abc_step_l.simps numeral_3_eq_3 abc_fetch.simps |
2371 abc_lm_v.simps nth_append abc_lm_s.simps) |
2371 abc_lm_v.simps nth_append abc_lm_s.simps) |
2372 qed |
2372 qed |
2384 assumes compile1: "rec_ci (Cn n f gs) = (ap, ar, ft) \<and> length args = ar" |
2384 assumes compile1: "rec_ci (Cn n f gs) = (ap, ar, ft) \<and> length args = ar" |
2385 and g: "i < length gs" |
2385 and g: "i < length gs" |
2386 and compile2: "rec_ci (gs!i) = (gap, gar, gft) \<and> gar = length args" |
2386 and compile2: "rec_ci (gs!i) = (gap, gar, gft) \<and> gar = length args" |
2387 and g_unhalt: "\<And> anything. {\<lambda> nl. nl = args @ 0\<up>(gft - gar) @ anything} gap \<up>" |
2387 and g_unhalt: "\<And> anything. {\<lambda> nl. nl = args @ 0\<up>(gft - gar) @ anything} gap \<up>" |
2388 and g_ind: "\<And> apj arj ftj j anything. \<lbrakk>j < i; rec_ci (gs!j) = (apj, arj, ftj)\<rbrakk> |
2388 and g_ind: "\<And> apj arj ftj j anything. \<lbrakk>j < i; rec_ci (gs!j) = (apj, arj, ftj)\<rbrakk> |
2389 \<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}" |
2389 \<Longrightarrow> {\<lambda> nl. nl = args @ 0\<up>(ftj - arj) @ anything} apj {\<lambda> nl. nl = args @ rec_eval (gs!j) args # 0\<up>(ftj - Suc arj) @ anything}" |
2390 and all_termi: "\<forall> j<i. terminate (gs!j) args" |
2390 and all_termi: "\<forall> j<i. terminates (gs!j) args" |
2391 shows "{\<lambda> nl. nl = args @ 0\<up>(ft - ar) @ anything} ap \<up>" |
2391 shows "{\<lambda> nl. nl = args @ 0\<up>(ft - ar) @ anything} ap \<up>" |
2392 using compile1 |
2392 using compile1 |
2393 apply(case_tac "rec_ci f", auto simp: rec_ci.simps abc_comp_commute) |
2393 apply(case_tac "rec_ci f", auto simp: rec_ci.simps abc_comp_commute) |
2394 proof(rule_tac abc_Hoare_plus_unhalt1) |
2394 proof(rule_tac abc_Hoare_plus_unhalt1) |
2395 fix fap far fft |
2395 fix fap far fft |
2396 let ?ft = "max (Suc (length args)) (Max (insert fft ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))" |
2396 let ?ft = "max (Suc (length args)) (Max (insert fft ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))" |
2397 let ?Q = "\<lambda>nl. nl = args @ 0\<up> (?ft - length args) @ map (\<lambda>i. rec_exec i args) (take i gs) @ |
2397 let ?Q = "\<lambda>nl. nl = args @ 0\<up> (?ft - length args) @ map (\<lambda>i. rec_eval i args) (take i gs) @ |
2398 0\<up>(length gs - i) @ 0\<up> Suc (length args) @ anything" |
2398 0\<up>(length gs - i) @ 0\<up> Suc (length args) @ anything" |
2399 have "cn_merge_gs (map rec_ci gs) ?ft = |
2399 have "cn_merge_gs (map rec_ci gs) ?ft = |
2400 cn_merge_gs (map rec_ci (take i gs)) ?ft [+] (gap [+] |
2400 cn_merge_gs (map rec_ci (take i gs)) ?ft [+] (gap [+] |
2401 mv_box gar (?ft + i)) [+] cn_merge_gs (map rec_ci (drop (Suc i) gs)) (?ft + Suc i)" |
2401 mv_box gar (?ft + i)) [+] cn_merge_gs (map rec_ci (drop (Suc i) gs)) (?ft + Suc i)" |
2402 using g compile2 cn_merge_gs_split by simp |
2402 using g compile2 cn_merge_gs_split by simp |
2403 thus "{\<lambda>nl. nl = args @ 0 # 0 \<up> (?ft + length gs) @ anything} (cn_merge_gs (map rec_ci gs) ?ft) \<up>" |
2403 thus "{\<lambda>nl. nl = args @ 0 # 0 \<up> (?ft + length gs) @ anything} (cn_merge_gs (map rec_ci gs) ?ft) \<up>" |
2404 proof(simp, rule_tac abc_Hoare_plus_unhalt1, rule_tac abc_Hoare_plus_unhalt2, |
2404 proof(simp, rule_tac abc_Hoare_plus_unhalt1, rule_tac abc_Hoare_plus_unhalt2, |
2405 rule_tac abc_Hoare_plus_unhalt1) |
2405 rule_tac abc_Hoare_plus_unhalt1) |
2406 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) @ |
2406 let ?Q_tmp = "\<lambda>nl. nl = args @ 0\<up> (gft - gar) @ 0\<up>(?ft - (length args) - (gft -gar)) @ map (\<lambda>i. rec_eval i args) (take i gs) @ |
2407 0\<up>(length gs - i) @ 0\<up> Suc (length args) @ anything" |
2407 0\<up>(length gs - i) @ 0\<up> Suc (length args) @ anything" |
2408 have a: "{?Q_tmp} gap \<up>" |
2408 have a: "{?Q_tmp} gap \<up>" |
2409 using g_unhalt[of "0 \<up> (?ft - (length args) - (gft - gar)) @ |
2409 using g_unhalt[of "0 \<up> (?ft - (length args) - (gft - gar)) @ |
2410 map (\<lambda>i. rec_exec i args) (take i gs) @ 0 \<up> (length gs - i) @ 0 \<up> Suc (length args) @ anything"] |
2410 map (\<lambda>i. rec_eval i args) (take i gs) @ 0 \<up> (length gs - i) @ 0 \<up> Suc (length args) @ anything"] |
2411 by simp |
2411 by simp |
2412 moreover have "?ft \<ge> gft" |
2412 moreover have "?ft \<ge> gft" |
2413 using g compile2 |
2413 using g compile2 |
2414 apply(rule_tac min_max.le_supI2, rule_tac Max_ge, simp, rule_tac insertI2) |
2414 apply(rule_tac min_max.le_supI2, rule_tac Max_ge, simp, rule_tac insertI2) |
2415 apply(rule_tac x = "rec_ci (gs ! i)" in image_eqI, simp) |
2415 apply(rule_tac x = "rec_ci (gs ! i)" in image_eqI, simp) |
2488 let tp = tm_of (ap [+] dummy_abc k) in |
2488 let tp = tm_of (ap [+] dummy_abc k) in |
2489 tp @ (shift (mopup k) (length tp div 2)))" |
2489 tp @ (shift (mopup k) (length tp div 2)))" |
2490 |
2490 |
2491 lemma recursive_compile_to_tm_correct1: |
2491 lemma recursive_compile_to_tm_correct1: |
2492 assumes compile: "rec_ci recf = (ap, ary, fp)" |
2492 assumes compile: "rec_ci recf = (ap, ary, fp)" |
2493 and termi: " terminate recf args" |
2493 and termi: " terminates recf args" |
2494 and tp: "tp = tm_of (ap [+] dummy_abc (length args))" |
2494 and tp: "tp = tm_of (ap [+] dummy_abc (length args))" |
2495 shows "\<exists> stp m l. steps0 (Suc 0, Bk # Bk # ires, <args> @ Bk\<up>rn) |
2495 shows "\<exists> stp m l. steps0 (Suc 0, Bk # Bk # ires, <args> @ Bk\<up>rn) |
2496 (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)" |
2496 (tp @ shift (mopup (length args)) (length tp div 2)) stp = (0, Bk\<up>m @ Bk # Bk # ires, Oc\<up>Suc (rec_eval recf args) @ Bk\<up>l)" |
2497 proof - |
2497 proof - |
2498 have "{\<lambda>nl. nl = args} ap [+] dummy_abc (length args) {\<lambda>nl. \<exists>m. nl = args @ rec_exec recf args # 0 \<up> m}" |
2498 have "{\<lambda>nl. nl = args} ap [+] dummy_abc (length args) {\<lambda>nl. \<exists>m. nl = args @ rec_eval recf args # 0 \<up> m}" |
2499 using compile termi compile |
2499 using compile termi compile |
2500 by(rule_tac compile_append_dummy_correct, auto) |
2500 by(rule_tac compile_append_dummy_correct, auto) |
2501 then obtain stp m where h: "abc_steps_l (0, args) (ap [+] dummy_abc (length args)) stp = |
2501 then obtain stp m where h: "abc_steps_l (0, args) (ap [+] dummy_abc (length args)) stp = |
2502 (length (ap [+] dummy_abc (length args)), args @ rec_exec recf args # 0\<up>m) " |
2502 (length (ap [+] dummy_abc (length args)), args @ rec_eval recf args # 0\<up>m) " |
2503 apply(simp add: abc_Hoare_halt_def, auto) |
2503 apply(simp add: abc_Hoare_halt_def, auto) |
2504 by(case_tac "abc_steps_l (0, args) (ap [+] dummy_abc (length args)) n", auto) |
2504 by(case_tac "abc_steps_l (0, args) (ap [+] dummy_abc (length args)) n", auto) |
2505 thus "?thesis" |
2505 thus "?thesis" |
2506 using assms tp |
2506 using assms tp |
2507 by(rule_tac lm = args and stp = stp and am = "args @ rec_exec recf args # 0 \<up> m" |
2507 by(rule_tac lm = args and stp = stp and am = "args @ rec_eval recf args # 0 \<up> m" |
2508 in compile_correct_halt, auto simp: crsp.simps start_of.simps length_abc_comp abc_lm_v.simps) |
2508 in compile_correct_halt, auto simp: crsp.simps start_of.simps length_abc_comp abc_lm_v.simps) |
2509 qed |
2509 qed |
2510 |
2510 |
2511 lemma recursive_compile_to_tm_correct2: |
2511 lemma recursive_compile_to_tm_correct2: |
2512 assumes termi: " terminate recf args" |
2512 assumes termi: " terminates recf args" |
2513 shows "\<exists> stp m l. steps0 (Suc 0, [Bk, Bk], <args>) (tm_of_rec recf) stp = |
2513 shows "\<exists> stp m l. steps0 (Suc 0, [Bk, Bk], <args>) (tm_of_rec recf) stp = |
2514 (0, Bk\<up>Suc (Suc m), Oc\<up>Suc (rec_exec recf args) @ Bk\<up>l)" |
2514 (0, Bk\<up>Suc (Suc m), Oc\<up>Suc (rec_eval recf args) @ Bk\<up>l)" |
2515 proof(case_tac "rec_ci recf", simp add: tm_of_rec.simps) |
2515 proof(case_tac "rec_ci recf", simp add: tm_of_rec.simps) |
2516 fix ap ar fp |
2516 fix ap ar fp |
2517 assume "rec_ci recf = (ap, ar, fp)" |
2517 assume "rec_ci recf = (ap, ar, fp)" |
2518 thus "\<exists>stp m l. steps0 (Suc 0, [Bk, Bk], <args>) |
2518 thus "\<exists>stp m l. steps0 (Suc 0, [Bk, Bk], <args>) |
2519 (tm_of (ap [+] dummy_abc ar) @ shift (mopup ar) (listsum (layout_of (ap [+] dummy_abc ar)))) stp = |
2519 (tm_of (ap [+] dummy_abc ar) @ shift (mopup ar) (listsum (layout_of (ap [+] dummy_abc ar)))) stp = |
2520 (0, Bk # Bk # Bk \<up> m, Oc # Oc \<up> rec_exec recf args @ Bk \<up> l)" |
2520 (0, Bk # Bk # Bk \<up> m, Oc # Oc \<up> rec_eval recf args @ Bk \<up> l)" |
2521 using recursive_compile_to_tm_correct1[of recf ap ar fp args "tm_of (ap [+] dummy_abc (length args))" "[]" 0] |
2521 using recursive_compile_to_tm_correct1[of recf ap ar fp args "tm_of (ap [+] dummy_abc (length args))" "[]" 0] |
2522 assms param_pattern[of recf args ap ar fp] |
2522 assms param_pattern[of recf args ap ar fp] |
2523 by(simp add: replicate_Suc[THEN sym] replicate_Suc_iff_anywhere del: replicate_Suc tm_of_rec_def, |
2523 by(simp add: replicate_Suc[THEN sym] replicate_Suc_iff_anywhere del: replicate_Suc tm_of_rec_def, |
2524 simp add: exp_suc del: replicate_Suc) |
2524 simp add: exp_suc del: replicate_Suc) |
2525 qed |
2525 qed |
2526 |
2526 |
2527 lemma recursive_compile_to_tm_correct3: |
2527 lemma recursive_compile_to_tm_correct3: |
2528 assumes termi: "terminate recf args" |
2528 assumes termi: "terminates recf args" |
2529 shows "{\<lambda> tp. tp =([Bk, Bk], <args>)} (tm_of_rec recf) |
2529 shows "{\<lambda> tp. tp =([Bk, Bk], <args>)} (tm_of_rec recf) |
2530 {\<lambda> tp. \<exists> k l. tp = (Bk\<up> k, <rec_exec recf args> @ Bk \<up> l)}" |
2530 {\<lambda> tp. \<exists> k l. tp = (Bk\<up> k, <rec_eval recf args> @ Bk \<up> l)}" |
2531 using recursive_compile_to_tm_correct2[OF assms] |
2531 using recursive_compile_to_tm_correct2[OF assms] |
2532 apply(auto simp add: Hoare_halt_def) |
2532 apply(auto simp add: Hoare_halt_def) |
2533 apply(rule_tac x = stp in exI) |
2533 apply(rule_tac x = stp in exI) |
2534 apply(auto simp add: tape_of_nat_abv) |
2534 apply(auto simp add: tape_of_nat_abv) |
2535 apply(rule_tac x = "Suc (Suc m)" in exI) |
2535 apply(rule_tac x = "Suc (Suc m)" in exI) |