507 |
507 |
508 definition fourtimes_ly :: "nat list" |
508 definition fourtimes_ly :: "nat list" |
509 where |
509 where |
510 "fourtimes_ly = layout_of abc_fourtimes" |
510 "fourtimes_ly = layout_of abc_fourtimes" |
511 |
511 |
512 definition t_twice :: "tprog" |
512 definition t_twice_compile :: "instr list" |
513 where |
513 where |
514 "t_twice = change_termi_state (tm_of (abc_twice) @ (tMp 1 (start_of twice_ly (length abc_twice) - Suc 0)))" |
514 "t_twice_compile= (tm_of abc_twice @ (shift (mopup 1) (length (tm_of abc_twice) div 2)))" |
515 |
515 |
516 definition t_fourtimes :: "tprog" |
516 definition t_twice :: "instr list" |
517 where |
517 where |
518 "t_fourtimes = change_termi_state (tm_of (abc_fourtimes) @ |
518 "t_twice = adjust t_twice_compile" |
519 (tMp 1 (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)))" |
519 |
520 |
520 definition t_fourtimes_compile :: "instr list" |
|
521 where |
|
522 "t_fourtimes_compile= (tm_of abc_fourtimes @ (shift (mopup 1) (length (tm_of abc_fourtimes) div 2)))" |
|
523 |
|
524 definition t_fourtimes :: "instr list" |
|
525 where |
|
526 "t_fourtimes = adjust t_fourtimes_compile" |
521 |
527 |
522 definition t_twice_len :: "nat" |
528 definition t_twice_len :: "nat" |
523 where |
529 where |
524 "t_twice_len = length t_twice div 2" |
530 "t_twice_len = length t_twice div 2" |
525 |
531 |
526 definition t_wcode_main_first_part:: "tprog" |
532 definition t_wcode_main_first_part:: "instr list" |
527 where |
533 where |
528 "t_wcode_main_first_part \<equiv> |
534 "t_wcode_main_first_part \<equiv> |
529 [(L, 1), (L, 2), (L, 7), (R, 3), |
535 [(L, 1), (L, 2), (L, 7), (R, 3), |
530 (R, 4), (W0, 3), (R, 4), (R, 5), |
536 (R, 4), (W0, 3), (R, 4), (R, 5), |
531 (W1, 6), (R, 5), (R, 13), (L, 6), |
537 (W1, 6), (R, 5), (R, 13), (L, 6), |
532 (R, 0), (R, 8), (R, 9), (Nop, 8), |
538 (R, 0), (R, 8), (R, 9), (Nop, 8), |
533 (R, 10), (W0, 9), (R, 10), (R, 11), |
539 (R, 10), (W0, 9), (R, 10), (R, 11), |
534 (W1, 12), (R, 11), (R, t_twice_len + 14), (L, 12)]" |
540 (W1, 12), (R, 11), (R, t_twice_len + 14), (L, 12)]" |
535 |
541 |
536 definition t_wcode_main :: "tprog" |
542 definition t_wcode_main :: "instr list" |
537 where |
543 where |
538 "t_wcode_main = (t_wcode_main_first_part @ tshift t_twice 12 @ [(L, 1), (L, 1)] |
544 "t_wcode_main = (t_wcode_main_first_part @ shift t_twice 12 @ [(L, 1), (L, 1)] |
539 @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])" |
545 @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])" |
540 |
546 |
541 fun bl_bin :: "block list \<Rightarrow> nat" |
547 fun bl_bin :: "cell list \<Rightarrow> nat" |
542 where |
548 where |
543 "bl_bin [] = 0" |
549 "bl_bin [] = 0" |
544 | "bl_bin (Bk # xs) = 2 * bl_bin xs" |
550 | "bl_bin (Bk # xs) = 2 * bl_bin xs" |
545 | "bl_bin (Oc # xs) = Suc (2 * bl_bin xs)" |
551 | "bl_bin (Oc # xs) = Suc (2 * bl_bin xs)" |
546 |
552 |
547 declare bl_bin.simps[simp del] |
553 declare bl_bin.simps[simp del] |
548 |
554 |
549 type_synonym bin_inv_t = "block list \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
555 type_synonym bin_inv_t = "cell list \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
550 |
556 |
551 fun wcode_before_double :: "bin_inv_t" |
557 fun wcode_before_double :: "bin_inv_t" |
552 where |
558 where |
553 "wcode_before_double ires rs (l, r) = |
559 "wcode_before_double ires rs (l, r) = |
554 (\<exists> ln rn. l = Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> |
560 (\<exists> ln rn. l = Bk # Bk # Bk\<up>(ln) @ Oc # ires \<and> |
555 r = Oc\<^bsup>(Suc (Suc rs))\<^esup> @ Bk\<^bsup>rn \<^esup>)" |
561 r = Oc\<up>((Suc (Suc rs))) @ Bk\<up>(rn ))" |
556 |
562 |
557 declare wcode_before_double.simps[simp del] |
563 declare wcode_before_double.simps[simp del] |
558 |
564 |
559 fun wcode_after_double :: "bin_inv_t" |
565 fun wcode_after_double :: "bin_inv_t" |
560 where |
566 where |
561 "wcode_after_double ires rs (l, r) = |
567 "wcode_after_double ires rs (l, r) = |
562 (\<exists> ln rn. l = Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires \<and> |
568 (\<exists> ln rn. l = Bk # Bk # Bk\<up>(ln) @ Oc # ires \<and> |
563 r = Oc\<^bsup>Suc (Suc (Suc 2*rs))\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
569 r = Oc\<up>(Suc (Suc (Suc 2*rs))) @ Bk\<up>(rn))" |
564 |
570 |
565 declare wcode_after_double.simps[simp del] |
571 declare wcode_after_double.simps[simp del] |
566 |
572 |
567 fun wcode_on_left_moving_1_B :: "bin_inv_t" |
573 fun wcode_on_left_moving_1_B :: "bin_inv_t" |
568 where |
574 where |
569 "wcode_on_left_moving_1_B ires rs (l, r) = |
575 "wcode_on_left_moving_1_B ires rs (l, r) = |
570 (\<exists> ml mr rn. l = Bk\<^bsup>ml\<^esup> @ Oc # Oc # ires \<and> |
576 (\<exists> ml mr rn. l = Bk\<up>(ml) @ Oc # Oc # ires \<and> |
571 r = Bk\<^bsup>mr\<^esup> @ Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> |
577 r = Bk\<up>(mr) @ Oc\<up>(Suc rs) @ Bk\<up>(rn) \<and> |
572 ml + mr > Suc 0 \<and> mr > 0)" |
578 ml + mr > Suc 0 \<and> mr > 0)" |
573 |
579 |
574 declare wcode_on_left_moving_1_B.simps[simp del] |
580 declare wcode_on_left_moving_1_B.simps[simp del] |
575 |
581 |
576 fun wcode_on_left_moving_1_O :: "bin_inv_t" |
582 fun wcode_on_left_moving_1_O :: "bin_inv_t" |
577 where |
583 where |
578 "wcode_on_left_moving_1_O ires rs (l, r) = |
584 "wcode_on_left_moving_1_O ires rs (l, r) = |
579 (\<exists> ln rn. |
585 (\<exists> ln rn. |
580 l = Oc # ires \<and> |
586 l = Oc # ires \<and> |
581 r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
587 r = Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))" |
582 |
588 |
583 declare wcode_on_left_moving_1_O.simps[simp del] |
589 declare wcode_on_left_moving_1_O.simps[simp del] |
584 |
590 |
585 fun wcode_on_left_moving_1 :: "bin_inv_t" |
591 fun wcode_on_left_moving_1 :: "bin_inv_t" |
586 where |
592 where |
1088 apply(rule_tac disjI1, rule_tac conjI) |
1101 apply(rule_tac disjI1, rule_tac conjI) |
1089 apply(rule_tac x = ln in exI, simp, rule_tac x = rn in exI, simp) |
1102 apply(rule_tac x = ln in exI, simp, rule_tac x = rn in exI, simp) |
1090 apply(rule_tac disjI2) |
1103 apply(rule_tac disjI2) |
1091 apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = ln in exI, |
1104 apply(rule_tac x = nat in exI, rule_tac x = "Suc mr" in exI, rule_tac x = ln in exI, |
1092 rule_tac x = rn in exI, simp) |
1105 rule_tac x = rn in exI, simp) |
1093 apply(simp add: exp_ind_def) |
1106 done |
1094 done |
1107 |
1095 |
1108 declare nth_of.simps[simp del] fetch.simps[simp del] |
1096 declare new_tape.simps[simp del] nth_of.simps[simp del] fetch.simps[simp del] |
|
1097 lemma wcode_double_case_first_correctness: |
1109 lemma wcode_double_case_first_correctness: |
1098 "let P = (\<lambda> (st, l, r). st = 13) in |
1110 "let P = (\<lambda> (st, l, r). st = 13) in |
1099 let Q = (\<lambda> (st, l, r). wcode_double_case_inv st ires rs (l, r)) in |
1111 let Q = (\<lambda> (st, l, r). wcode_double_case_inv st ires rs (l, r)) in |
1100 let f = (\<lambda> stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp) in |
1112 let f = (\<lambda> stp. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp) in |
1101 \<exists> n .P (f n) \<and> Q (f (n::nat))" |
1113 \<exists> n .P (f n) \<and> Q (f (n::nat))" |
1102 proof - |
1114 proof - |
1103 let ?P = "(\<lambda> (st, l, r). st = 13)" |
1115 let ?P = "(\<lambda> (st, l, r). st = 13)" |
1104 let ?Q = "(\<lambda> (st, l, r). wcode_double_case_inv st ires rs (l, r))" |
1116 let ?Q = "(\<lambda> (st, l, r). wcode_double_case_inv st ires rs (l, r))" |
1105 let ?f = "(\<lambda> stp. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp)" |
1117 let ?f = "(\<lambda> stp. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp)" |
1106 have "\<exists> n. ?P (?f n) \<and> ?Q (?f (n::nat))" |
1118 have "\<exists> n. ?P (?f n) \<and> ?Q (?f (n::nat))" |
1107 proof(rule_tac halt_lemma2) |
1119 proof(rule_tac halt_lemma2) |
1108 show "wf wcode_double_case_le" |
1120 show "wf wcode_double_case_le" |
1109 by auto |
1121 by auto |
1110 next |
1122 next |
1111 show "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow> |
1123 show "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow> |
1112 ?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_double_case_le" |
1124 ?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_double_case_le" |
1113 proof(rule_tac allI, case_tac "?f na", simp add: tstep_red) |
1125 proof(rule_tac allI, case_tac "?f na", simp add: step_red) |
1114 fix na a b c |
1126 fix na a b c |
1115 show "a \<noteq> 13 \<and> wcode_double_case_inv a ires rs (b, c) \<longrightarrow> |
1127 show "a \<noteq> 13 \<and> wcode_double_case_inv a ires rs (b, c) \<longrightarrow> |
1116 (case tstep (a, b, c) t_wcode_main of (st, x) \<Rightarrow> |
1128 (case step0 (a, b, c) t_wcode_main of (st, x) \<Rightarrow> |
1117 wcode_double_case_inv st ires rs x) \<and> |
1129 wcode_double_case_inv st ires rs x) \<and> |
1118 (tstep (a, b, c) t_wcode_main, a, b, c) \<in> wcode_double_case_le" |
1130 (step0 (a, b, c) t_wcode_main, a, b, c) \<in> wcode_double_case_le" |
1119 apply(rule_tac impI, simp add: wcode_double_case_inv.simps) |
1131 apply(rule_tac impI, simp add: wcode_double_case_inv.simps) |
1120 apply(auto split: if_splits simp: tstep.simps, |
1132 apply(auto split: if_splits simp: step.simps, |
1121 case_tac [!] c, simp_all, case_tac [!] "(c::block list)!0") |
1133 case_tac [!] c, simp_all, case_tac [!] "(c::cell list)!0") |
1122 apply(simp_all add: new_tape.simps wcode_double_case_inv.simps wcode_double_case_le_def |
1134 apply(simp_all add: wcode_double_case_inv.simps wcode_double_case_le_def |
1123 lex_pair_def) |
1135 lex_pair_def) |
1124 apply(auto split: if_splits) |
1136 apply(auto split: if_splits) |
1125 done |
1137 done |
1126 qed |
1138 qed |
1127 next |
1139 next |
1128 show "?Q (?f 0)" |
1140 show "?Q (?f 0)" |
1129 apply(simp add: steps.simps wcode_double_case_inv.simps |
1141 apply(simp add: steps.simps wcode_double_case_inv.simps |
1130 wcode_on_left_moving_1.simps |
1142 wcode_on_left_moving_1.simps |
1131 wcode_on_left_moving_1_B.simps) |
1143 wcode_on_left_moving_1_B.simps) |
1132 apply(rule_tac disjI1) |
1144 apply(rule_tac disjI1) |
1133 apply(rule_tac x = "Suc m" in exI, simp add: exp_ind_def) |
1145 apply(rule_tac x = "Suc m" in exI, simp) |
1134 apply(rule_tac x = "Suc 0" in exI, simp add: exp_ind_def) |
1146 apply(rule_tac x = "Suc 0" in exI, simp) |
1135 apply(auto) |
|
1136 done |
1147 done |
1137 next |
1148 next |
1138 show "\<not> ?P (?f 0)" |
1149 show "\<not> ?P (?f 0)" |
1139 apply(simp add: steps.simps) |
1150 apply(simp add: steps.simps) |
1140 done |
1151 done |
1141 qed |
1152 qed |
1142 thus "let P = \<lambda>(st, l, r). st = 13; |
1153 thus "let P = \<lambda>(st, l, r). st = 13; |
1143 Q = \<lambda>(st, l, r). wcode_double_case_inv st ires rs (l, r); |
1154 Q = \<lambda>(st, l, r). wcode_double_case_inv st ires rs (l, r); |
1144 f = steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main |
1155 f = steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main |
1145 in \<exists>n. P (f n) \<and> Q (f n)" |
1156 in \<exists>n. P (f n) \<and> Q (f n)" |
1146 apply(simp add: Let_def) |
1157 apply(simp add: Let_def) |
1147 done |
1158 done |
1148 qed |
1159 qed |
1149 |
1160 |
1150 lemma [elim]: "t_ncorrect tp |
1161 lemma tm_append_shift_append_steps: |
1151 \<Longrightarrow> t_ncorrect (tshift tp a)" |
1162 "\<lbrakk>steps0 (st, l, r) tp stp = (st', l', r'); |
1152 apply(simp add: t_ncorrect.simps shift_length) |
1163 0 < st'; |
1153 done |
1164 length tp1 mod 2 = 0 |
1154 |
1165 \<rbrakk> |
1155 lemma tshift_fetch: "\<lbrakk> fetch tp a b = (aa, st'); 0 < st'\<rbrakk> |
1166 \<Longrightarrow> steps0 (st + length tp1 div 2, l, r) (tp1 @ shift tp (length tp1 div 2) @ tp2) stp |
1156 \<Longrightarrow> fetch (tshift tp (length tp1 div 2)) a b |
1167 = (st' + length tp1 div 2, l', r')" |
1157 = (aa, st' + length tp1 div 2)" |
|
1158 apply(subgoal_tac "a > 0") |
|
1159 apply(auto simp: fetch.simps nth_of.simps shift_length nth_map |
|
1160 tshift.simps split: block.splits if_splits) |
|
1161 done |
|
1162 |
|
1163 lemma t_steps_steps_eq: "\<lbrakk>steps (st, l, r) tp stp = (st', l', r'); |
|
1164 0 < st'; |
|
1165 0 < st \<and> st \<le> length tp div 2; |
|
1166 t_ncorrect tp1; |
|
1167 t_ncorrect tp\<rbrakk> |
|
1168 \<Longrightarrow> t_steps (st + length tp1 div 2, l, r) (tshift tp (length tp1 div 2), |
|
1169 length tp1 div 2) stp |
|
1170 = (st' + length tp1 div 2, l', r')" |
|
1171 apply(induct stp arbitrary: st' l' r', simp add: steps.simps t_steps.simps, |
|
1172 simp add: tstep_red stepn) |
|
1173 apply(case_tac "(steps (st, l, r) tp stp)", simp) |
|
1174 proof - |
1168 proof - |
1175 fix stp st' l' r' a b c |
1169 assume h: |
1176 assume ind: "\<And>st' l' r'. |
1170 "steps0 (st, l, r) tp stp = (st', l', r')" |
1177 \<lbrakk>a = st' \<and> b = l' \<and> c = r'; 0 < st'\<rbrakk> |
1171 "0 < st'" |
1178 \<Longrightarrow> t_steps (st + length tp1 div 2, l, r) |
1172 "length tp1 mod 2 = 0 " |
1179 (tshift tp (length tp1 div 2), length tp1 div 2) stp = |
1173 from h have |
1180 (st' + length tp1 div 2, l', r')" |
1174 "steps (st + length tp1 div 2, l, r) (tp1 @ shift tp (length tp1 div 2), 0) stp = |
1181 and h: "tstep (a, b, c) tp = (st', l', r')" "0 < st'" "t_ncorrect tp1" "t_ncorrect tp" |
1175 (st' + length tp1 div 2, l', r')" |
1182 have k: "t_steps (st + length tp1 div 2, l, r) (tshift tp (length tp1 div 2), |
1176 by(rule_tac tm_append_second_steps_eq, simp_all) |
1183 length tp1 div 2) stp = (a + length tp1 div 2, b, c)" |
1177 then have "steps (st + length tp1 div 2, l, r) ((tp1 @ shift tp (length tp1 div 2)) @ tp2, 0) stp = |
1184 apply(rule_tac ind, simp) |
1178 (st' + length tp1 div 2, l', r')" |
1185 using h |
1179 using h |
1186 apply(case_tac a, simp_all add: tstep.simps fetch.simps) |
1180 apply(rule_tac tm_append_first_steps_eq, simp_all) |
1187 done |
1181 done |
1188 from h and this show "t_step (t_steps (st + length tp1 div 2, l, r) (tshift tp (length tp1 div 2), length tp1 div 2) stp) |
1182 thus "?thesis" |
1189 (tshift tp (length tp1 div 2), length tp1 div 2) = |
1183 by simp |
1190 (st' + length tp1 div 2, l', r')" |
|
1191 apply(simp add: k) |
|
1192 apply(simp add: tstep.simps t_step.simps) |
|
1193 apply(case_tac "fetch tp a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp) |
|
1194 apply(subgoal_tac "fetch (tshift tp (length tp1 div 2)) a |
|
1195 (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) = (aa, st' + length tp1 div 2)", simp) |
|
1196 apply(simp add: tshift_fetch) |
|
1197 done |
|
1198 qed |
1184 qed |
1199 |
1185 |
1200 lemma t_tshift_lemma: "\<lbrakk> steps (st, l, r) tp stp = (st', l', r'); |
|
1201 st' \<noteq> 0; |
|
1202 stp > 0; |
|
1203 0 < st \<and> st \<le> length tp div 2; |
|
1204 t_ncorrect tp1; |
|
1205 t_ncorrect tp; |
|
1206 t_ncorrect tp2 |
|
1207 \<rbrakk> |
|
1208 \<Longrightarrow> \<exists> stp>0. steps (st + length tp1 div 2, l, r) (tp1 @ tshift tp (length tp1 div 2) @ tp2) stp |
|
1209 = (st' + length tp1 div 2, l', r')" |
|
1210 proof - |
|
1211 assume h: "steps (st, l, r) tp stp = (st', l', r')" |
|
1212 "st' \<noteq> 0" "stp > 0" |
|
1213 "0 < st \<and> st \<le> length tp div 2" |
|
1214 "t_ncorrect tp1" |
|
1215 "t_ncorrect tp" |
|
1216 "t_ncorrect tp2" |
|
1217 from h have |
|
1218 "\<exists>stp>0. t_steps (st + length tp1 div 2, l, r) (tp1 @ tshift tp (length tp1 div 2) @ tp2, 0) stp = |
|
1219 (st' + length tp1 div 2, l', r')" |
|
1220 apply(rule_tac stp = stp in turing_shift, simp_all add: shift_length) |
|
1221 apply(simp add: t_steps_steps_eq) |
|
1222 apply(simp add: t_ncorrect.simps shift_length) |
|
1223 done |
|
1224 thus "\<exists> stp>0. steps (st + length tp1 div 2, l, r) (tp1 @ tshift tp (length tp1 div 2) @ tp2) stp |
|
1225 = (st' + length tp1 div 2, l', r')" |
|
1226 apply(erule_tac exE) |
|
1227 apply(rule_tac x = stp in exI, simp) |
|
1228 apply(subgoal_tac "length (tp1 @ tshift tp (length tp1 div 2) @ tp2) mod 2 = 0") |
|
1229 apply(simp only: steps_eq) |
|
1230 using h |
|
1231 apply(auto simp: t_ncorrect.simps shift_length) |
|
1232 apply arith |
|
1233 done |
|
1234 qed |
|
1235 |
|
1236 |
|
1237 lemma t_twice_len_ge: "Suc 0 \<le> length t_twice div 2" |
1186 lemma t_twice_len_ge: "Suc 0 \<le> length t_twice div 2" |
1238 apply(simp add: t_twice_def tMp.simps shift_length) |
1187 apply(simp add: t_twice_def mopup.simps t_twice_compile_def) |
1239 done |
1188 done |
1240 |
1189 |
1241 lemma [intro]: "rec_calc_rel (recf.id (Suc 0) 0) [rs] rs" |
1190 lemma [intro]: "rec_calc_rel (recf.id (Suc 0) 0) [rs] rs" |
1242 apply(rule_tac calc_id, simp_all) |
1191 apply(rule_tac calc_id, simp_all) |
1243 done |
1192 done |
1249 |
1198 |
1250 lemma [intro]: "rec_calc_rel rec_mult [rs, 2] (2 * rs)" |
1199 lemma [intro]: "rec_calc_rel rec_mult [rs, 2] (2 * rs)" |
1251 using prime_rel_exec_eq[of "rec_mult" "[rs, 2]" "2*rs"] |
1200 using prime_rel_exec_eq[of "rec_mult" "[rs, 2]" "2*rs"] |
1252 apply(subgoal_tac "primerec rec_mult (Suc (Suc 0))", auto) |
1201 apply(subgoal_tac "primerec rec_mult (Suc (Suc 0))", auto) |
1253 done |
1202 done |
1254 lemma t_twice_correct: "\<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) |
1203 |
1255 (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0)) stp = |
1204 declare start_of.simps[simp del] |
1256 (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
1205 |
|
1206 lemma t_twice_correct: |
|
1207 "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) |
|
1208 (tm_of abc_twice @ shift (mopup (Suc 0)) ((length (tm_of abc_twice) div 2))) stp = |
|
1209 (0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))" |
1257 proof(case_tac "rec_ci rec_twice") |
1210 proof(case_tac "rec_ci rec_twice") |
1258 fix a b c |
1211 fix a b c |
1259 assume h: "rec_ci rec_twice = (a, b, c)" |
1212 assume h: "rec_ci rec_twice = (a, b, c)" |
1260 have "\<exists>stp m l. steps (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\<^bsup>n\<^esup>) (tm_of abc_twice @ tMp (Suc 0) |
1213 have "\<exists>stp m l. steps0 (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\<up>(n)) (tm_of abc_twice @ shift (mopup 1) |
1261 (start_of twice_ly (length abc_twice) - 1)) stp = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2*rs)\<^esup> @ Bk\<^bsup>l\<^esup>)" |
1214 (length (tm_of abc_twice) div 2)) stp = (0, Bk\<up>(m) @ Bk # Bk # ires, Oc\<up>(Suc (2*rs)) @ Bk\<up>(l))" |
1262 proof(rule_tac t_compiled_by_rec) |
1215 proof(rule_tac recursive_compile_to_tm_correct) |
1263 show "rec_ci rec_twice = (a, b, c)" by (simp add: h) |
1216 show "rec_ci rec_twice = (a, b, c)" by (simp add: h) |
1264 next |
1217 next |
1265 show "rec_calc_rel rec_twice [rs] (2 * rs)" |
1218 show "rec_calc_rel rec_twice [rs] (2 * rs)" |
1266 apply(simp add: rec_twice_def) |
1219 apply(simp add: rec_twice_def) |
1267 apply(rule_tac rs = "[rs, 2]" in calc_cn, simp_all) |
1220 apply(rule_tac rs = "[rs, 2]" in calc_cn, simp_all) |
1268 apply(rule_tac allI, case_tac k, auto) |
1221 apply(rule_tac allI, case_tac k, auto) |
1269 done |
1222 done |
1270 next |
1223 next |
1271 show "length [rs] = Suc 0" by simp |
1224 show "length [rs] = 1" by simp |
|
1225 next |
|
1226 show "layout_of (a [+] dummy_abc 1) = layout_of (a [+] dummy_abc 1)" by simp |
1272 next |
1227 next |
1273 show "layout_of (a [+] dummy_abc (Suc 0)) = layout_of (a [+] dummy_abc (Suc 0))" |
1228 show "tm_of abc_twice = tm_of (a [+] dummy_abc 1)" |
1274 by simp |
|
1275 next |
|
1276 show "start_of twice_ly (length abc_twice) = |
|
1277 start_of (layout_of (a [+] dummy_abc (Suc 0))) (length (a [+] dummy_abc (Suc 0)))" |
|
1278 using h |
|
1279 apply(simp add: twice_ly_def abc_twice_def) |
|
1280 done |
|
1281 next |
|
1282 show "tm_of abc_twice = tm_of (a [+] dummy_abc (Suc 0))" |
|
1283 using h |
1229 using h |
1284 apply(simp add: abc_twice_def) |
1230 apply(simp add: abc_twice_def) |
1285 done |
1231 done |
1286 qed |
1232 qed |
1287 thus "\<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) |
1233 thus "?thesis" |
1288 (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0)) stp = |
|
1289 (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
|
1290 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) |
1234 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) |
1291 done |
1235 done |
1292 qed |
1236 qed |
1293 |
1237 |
1294 lemma change_termi_state_fetch: "\<lbrakk>fetch ap a b = (aa, st); st > 0\<rbrakk> |
1238 |
1295 \<Longrightarrow> fetch (change_termi_state ap) a b = (aa, st)" |
1239 lemma adjust_fetch0: |
1296 apply(case_tac b, auto simp: fetch.simps nth_of.simps change_termi_state.simps nth_map |
1240 "\<lbrakk>0 < a; a \<le> length ap div 2; fetch ap a b = (aa, 0)\<rbrakk> |
1297 split: if_splits block.splits) |
1241 \<Longrightarrow> fetch (adjust ap) a b = (aa, Suc (length ap div 2))" |
1298 done |
1242 apply(case_tac b, auto simp: fetch.simps nth_of.simps nth_map |
1299 |
1243 split: if_splits) |
1300 lemma change_termi_state_exec_in_range: |
1244 apply(case_tac [!] a, auto simp: fetch.simps nth_of.simps) |
1301 "\<lbrakk>steps (st, l, r) ap stp = (st', l', r'); st' \<noteq> 0\<rbrakk> |
1245 done |
1302 \<Longrightarrow> steps (st, l, r) (change_termi_state ap) stp = (st', l', r')" |
1246 |
1303 proof(induct stp arbitrary: st l r st' l' r', simp add: steps.simps) |
1247 lemma adjust_fetch_norm: |
1304 fix stp st l r st' l' r' |
1248 "\<lbrakk>st > 0; st \<le> length tp div 2; fetch ap st b = (aa, ns); ns \<noteq> 0\<rbrakk> |
1305 assume ind: "\<And>st l r st' l' r'. |
1249 \<Longrightarrow> fetch (turing_basic.adjust ap) st b = (aa, ns)" |
1306 \<lbrakk>steps (st, l, r) ap stp = (st', l', r'); st' \<noteq> 0\<rbrakk> \<Longrightarrow> |
1250 apply(case_tac b, auto simp: fetch.simps nth_of.simps nth_map |
1307 steps (st, l, r) (change_termi_state ap) stp = (st', l', r')" |
1251 split: if_splits) |
1308 and h: "steps (st, l, r) ap (Suc stp) = (st', l', r')" "st' \<noteq> 0" |
1252 apply(case_tac [!] st, auto simp: fetch.simps nth_of.simps) |
1309 from h show "steps (st, l, r) (change_termi_state ap) (Suc stp) = (st', l', r')" |
1253 done |
1310 proof(simp add: tstep_red, case_tac "steps (st, l, r) ap stp", simp) |
1254 |
1311 fix a b c |
1255 lemma adjust_step_eq: |
1312 assume g: "steps (st, l, r) ap stp = (a, b, c)" |
1256 assumes exec: "step0 (st,l,r) ap = (st', l', r')" |
1313 "tstep (a, b, c) ap = (st', l', r')" "0 < st'" |
1257 and wf_tm: "tm_wf (ap, 0)" |
1314 hence "steps (st, l, r) (change_termi_state ap) stp = (a, b, c)" |
1258 and notfinal: "st' > 0" |
1315 apply(rule_tac ind, simp) |
1259 shows "step0 (st, l, r) (adjust ap) = (st', l', r')" |
1316 apply(case_tac a, simp_all add: tstep_0) |
1260 using assms |
1317 done |
1261 proof - |
1318 from g and this show "tstep (steps (st, l, r) (change_termi_state ap) stp) |
1262 have "st > 0" |
1319 (change_termi_state ap) = (st', l', r')" |
1263 using assms |
1320 apply(simp add: tstep.simps) |
1264 by(case_tac st, simp_all add: step.simps fetch.simps) |
1321 apply(case_tac "fetch ap a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp) |
1265 moreover hence "st \<le> (length ap) div 2" |
1322 apply(subgoal_tac "fetch (change_termi_state ap) a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) |
1266 using assms |
1323 = (aa, st')", simp) |
1267 apply(case_tac "st \<le> (length ap) div 2", simp) |
1324 apply(simp add: change_termi_state_fetch) |
1268 apply(case_tac st, auto simp: step.simps fetch.simps) |
1325 done |
1269 apply(case_tac "read r", simp_all add: fetch.simps nth_of.simps) |
1326 qed |
1270 done |
|
1271 ultimately have "fetch (adjust ap) st (read r) = fetch ap st (read r)" |
|
1272 using assms |
|
1273 apply(case_tac "fetch ap st (read r)") |
|
1274 apply(drule_tac adjust_fetch_norm, simp_all) |
|
1275 apply(simp add: step.simps) |
|
1276 done |
|
1277 thus "?thesis" |
|
1278 using exec |
|
1279 by(simp add: step.simps) |
1327 qed |
1280 qed |
1328 |
1281 |
1329 lemma change_termi_state_fetch0: |
1282 declare adjust.simps[simp del] |
1330 "\<lbrakk>0 < a; a \<le> length ap div 2; t_correct ap; fetch ap a b = (aa, 0)\<rbrakk> |
1283 |
1331 \<Longrightarrow> fetch (change_termi_state ap) a b = (aa, Suc (length ap div 2))" |
1284 lemma adjust_steps_eq: |
1332 apply(case_tac b, auto simp: fetch.simps nth_of.simps change_termi_state.simps nth_map |
1285 assumes exec: "steps0 (st,l,r) ap stp = (st', l', r')" |
1333 split: if_splits block.splits) |
1286 and wf_tm: "tm_wf (ap, 0)" |
1334 done |
1287 and notfinal: "st' > 0" |
1335 |
1288 shows "steps0 (st, l, r) (adjust ap) stp = (st', l', r')" |
1336 lemma turing_change_termi_state: |
1289 using exec notfinal |
1337 "\<lbrakk>steps (Suc 0, l, r) ap stp = (0, l', r'); t_correct ap\<rbrakk> |
1290 proof(induct stp arbitrary: st' l' r') |
1338 \<Longrightarrow> \<exists> stp. steps (Suc 0, l, r) (change_termi_state ap) stp = |
1291 case 0 |
|
1292 thus "?case" |
|
1293 by(simp add: steps.simps) |
|
1294 next |
|
1295 case (Suc stp st' l' r') |
|
1296 have ind: "\<And>st' l' r'. \<lbrakk>steps0 (st, l, r) ap stp = (st', l', r'); 0 < st'\<rbrakk> |
|
1297 \<Longrightarrow> steps0 (st, l, r) (turing_basic.adjust ap) stp = (st', l', r')" by fact |
|
1298 have h: "steps0 (st, l, r) ap (Suc stp) = (st', l', r')" by fact |
|
1299 have g: "0 < st'" by fact |
|
1300 obtain st'' l'' r'' where a: "steps0 (st, l, r) ap stp = (st'', l'', r'')" |
|
1301 by (metis prod_cases3) |
|
1302 hence c:"0 < st''" |
|
1303 using h g |
|
1304 apply(simp add: step_red) |
|
1305 apply(case_tac st'', auto) |
|
1306 done |
|
1307 hence b: "steps0 (st, l, r) (turing_basic.adjust ap) stp = (st'', l'', r'')" |
|
1308 using a |
|
1309 by(rule_tac ind, simp_all) |
|
1310 thus "?case" |
|
1311 using assms a b h g |
|
1312 apply(simp add: step_red) |
|
1313 apply(rule_tac adjust_step_eq, simp_all) |
|
1314 done |
|
1315 qed |
|
1316 |
|
1317 lemma adjust_halt_eq: |
|
1318 assumes exec: "steps0 (1, l, r) ap stp = (0, l', r')" |
|
1319 and tm_wf: "tm_wf (ap, 0)" |
|
1320 shows "\<exists> stp. steps0 (Suc 0, l, r) (adjust ap) stp = |
1339 (Suc (length ap div 2), l', r')" |
1321 (Suc (length ap div 2), l', r')" |
1340 apply(drule first_halt_point) |
1322 proof - |
1341 apply(erule_tac exE) |
1323 have "\<exists> stp. \<not> is_final (steps0 (1, l, r) ap stp) \<and> (steps0 (1, l, r) ap (Suc stp) = (0, l', r'))" |
1342 apply(rule_tac x = "Suc stp" in exI, simp add: tstep_red) |
1324 thm before_final using exec |
1343 apply(case_tac "steps (Suc 0, l, r) ap stp") |
1325 by(erule_tac before_final) |
1344 apply(simp add: isS0_def change_termi_state_exec_in_range) |
1326 then obtain stpa where a: |
1345 apply(subgoal_tac "steps (Suc 0, l, r) (change_termi_state ap) stp = (a, b, c)", simp) |
1327 "\<not> is_final (steps0 (1, l, r) ap stpa) \<and> (steps0 (1, l, r) ap (Suc stpa) = (0, l', r'))" .. |
1346 apply(simp add: tstep.simps) |
1328 obtain sa la ra where b:"steps0 (1, l, r) ap stpa = (sa, la, ra)" by (metis prod_cases3) |
1347 apply(case_tac "fetch ap a (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp) |
1329 hence c: "steps0 (Suc 0, l, r) (adjust ap) stpa = (sa, la, ra)" |
1348 apply(subgoal_tac "fetch (change_termi_state ap) a |
1330 using assms a |
1349 (case c of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) = (aa, Suc (length ap div 2))", simp) |
1331 apply(rule_tac adjust_steps_eq, simp_all) |
1350 apply(rule_tac ap = ap in change_termi_state_fetch0, simp_all) |
1332 done |
1351 apply(rule_tac tp = "(l, r)" and l = b and r = c and stp = stp and A = ap in s_keep, simp_all) |
1333 have d: "sa \<le> length ap div 2" |
1352 apply(simp add: change_termi_state_exec_in_range) |
1334 using steps_in_range[of "(l, r)" ap stpa] a tm_wf b |
|
1335 by(simp) |
|
1336 obtain ac ns where e: "fetch ap sa (read ra) = (ac, ns)" |
|
1337 by (metis prod.exhaust) |
|
1338 hence f: "ns = 0" |
|
1339 using b a |
|
1340 apply(simp add: step_red step.simps) |
|
1341 done |
|
1342 have k: "fetch (adjust ap) sa (read ra) = (ac, Suc (length ap div 2))" |
|
1343 using a b c d e f |
|
1344 apply(rule_tac adjust_fetch0, simp_all) |
|
1345 done |
|
1346 from a b e f k and c show "?thesis" |
|
1347 apply(rule_tac x = "Suc stpa" in exI) |
|
1348 apply(simp add: step_red, auto) |
|
1349 apply(simp add: step.simps) |
|
1350 done |
|
1351 qed |
|
1352 |
|
1353 declare tm_wf.simps[simp del] |
|
1354 |
|
1355 lemma [simp]: " tm_wf (t_twice_compile, 0)" |
|
1356 apply(simp only: t_twice_compile_def) |
|
1357 apply(rule_tac t_compiled_correct) |
|
1358 apply(simp_all add: abc_twice_def) |
1353 done |
1359 done |
1354 |
1360 |
1355 lemma t_twice_change_term_state: |
1361 lemma t_twice_change_term_state: |
1356 "\<exists> stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp |
1362 "\<exists> stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) t_twice stp |
1357 = (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
1363 = (Suc t_twice_len, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))" |
1358 using t_twice_correct[of ires rs n] |
1364 proof - |
1359 apply(erule_tac exE) |
1365 have "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) |
1360 apply(erule_tac exE) |
1366 (tm_of abc_twice @ shift (mopup (Suc 0)) ((length (tm_of abc_twice) div 2))) stp = |
1361 apply(erule_tac exE) |
1367 (0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))" |
1362 proof(drule_tac turing_change_termi_state) |
1368 by(rule_tac t_twice_correct) |
1363 fix stp ln rn |
1369 then obtain stp ln rn where " steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) |
1364 show "t_correct (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0))" |
1370 (tm_of abc_twice @ shift (mopup (Suc 0)) ((length (tm_of abc_twice) div 2))) stp = |
1365 apply(rule_tac t_compiled_correct, simp_all) |
1371 (0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))" by blast |
1366 apply(simp add: twice_ly_def) |
1372 hence "\<exists> stp. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) |
|
1373 (adjust t_twice_compile) stp |
|
1374 = (Suc (length t_twice_compile div 2), Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))" |
|
1375 thm adjust_halt_eq |
|
1376 apply(rule_tac stp = stp in adjust_halt_eq) |
|
1377 apply(simp add: t_twice_compile_def, auto) |
1367 done |
1378 done |
1368 next |
1379 then obtain stpb where |
1369 fix stp ln rn |
1380 "steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) |
1370 show "\<exists>stp. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) |
1381 (adjust t_twice_compile) stpb |
1371 (change_termi_state (tm_of abc_twice @ tMp (Suc 0) |
1382 = (Suc (length t_twice_compile div 2), Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))" .. |
1372 (start_of twice_ly (length abc_twice) - Suc 0))) stp = |
1383 thus "?thesis" |
1373 (Suc (length (tm_of abc_twice @ tMp (Suc 0) (start_of twice_ly (length abc_twice) - Suc 0)) div 2), |
1384 apply(simp add: t_twice_def t_twice_len_def) |
1374 Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>) \<Longrightarrow> |
1385 by metis |
1375 \<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp = |
|
1376 (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
|
1377 apply(erule_tac exE) |
|
1378 apply(simp add: t_twice_len_def t_twice_def) |
|
1379 apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI, simp) |
|
1380 done |
|
1381 qed |
1386 qed |
1382 |
1387 |
|
1388 lemma [intro]: "length t_wcode_main_first_part mod 2 = 0" |
|
1389 apply(auto simp: t_wcode_main_first_part_def) |
|
1390 done |
|
1391 |
1383 lemma t_twice_append_pre: |
1392 lemma t_twice_append_pre: |
1384 "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp |
1393 "steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) t_twice stp |
1385 = (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>) |
1394 = (Suc t_twice_len, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn)) |
1386 \<Longrightarrow> \<exists> stp>0. steps (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) |
1395 \<Longrightarrow> steps0 (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) |
1387 (t_wcode_main_first_part @ tshift t_twice (length t_wcode_main_first_part div 2) @ |
1396 (t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @ |
1388 ([(L, 1), (L, 1)] @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp |
1397 ([(L, 1), (L, 1)] @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp |
1389 = (Suc (t_twice_len) + length t_wcode_main_first_part div 2, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
1398 = (Suc (t_twice_len) + length t_wcode_main_first_part div 2, |
1390 proof(rule_tac t_tshift_lemma, simp_all add: t_twice_len_ge) |
1399 Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))" |
1391 assume "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_twice stp = |
1400 by(rule_tac tm_append_shift_append_steps, auto) |
1392 (Suc t_twice_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
1401 |
1393 thus "0 < stp" |
|
1394 apply(case_tac stp, simp add: steps.simps t_twice_len_ge t_twice_len_def) |
|
1395 using t_twice_len_ge |
|
1396 apply(simp, simp) |
|
1397 done |
|
1398 next |
|
1399 show "t_ncorrect t_wcode_main_first_part" |
|
1400 apply(simp add: t_ncorrect.simps t_wcode_main_first_part_def) |
|
1401 done |
|
1402 next |
|
1403 show "t_ncorrect t_twice" |
|
1404 using length_tm_even[of abc_twice] |
|
1405 apply(auto simp: t_ncorrect.simps t_twice_def) |
|
1406 apply(arith) |
|
1407 done |
|
1408 next |
|
1409 show "t_ncorrect ((L, Suc 0) # (L, Suc 0) # |
|
1410 tshift t_fourtimes (t_twice_len + 13) @ [(L, Suc 0), (L, Suc 0)])" |
|
1411 using length_tm_even[of abc_fourtimes] |
|
1412 apply(simp add: t_ncorrect.simps shift_length t_fourtimes_def) |
|
1413 apply arith |
|
1414 done |
|
1415 qed |
|
1416 |
|
1417 lemma t_twice_append: |
1402 lemma t_twice_append: |
1418 "\<exists> stp ln rn. steps (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) |
1403 "\<exists> stp ln rn. steps0 (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) |
1419 (t_wcode_main_first_part @ tshift t_twice (length t_wcode_main_first_part div 2) @ |
1404 (t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @ |
1420 ([(L, 1), (L, 1)] @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp |
1405 ([(L, 1), (L, 1)] @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp |
1421 = (Suc (t_twice_len) + length t_wcode_main_first_part div 2, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
1406 = (Suc (t_twice_len) + length t_wcode_main_first_part div 2, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))" |
1422 using t_twice_change_term_state[of ires rs n] |
1407 using t_twice_change_term_state[of ires rs n] |
1423 apply(erule_tac exE) |
1408 apply(erule_tac exE) |
1424 apply(erule_tac exE) |
1409 apply(erule_tac exE) |
1425 apply(erule_tac exE) |
1410 apply(erule_tac exE) |
1426 apply(drule_tac t_twice_append_pre) |
1411 apply(drule_tac t_twice_append_pre) |
1427 apply(erule_tac exE) |
1412 apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI) |
1428 apply(rule_tac x = stpa in exI, rule_tac x = ln in exI, rule_tac x = rn in exI) |
|
1429 apply(simp) |
1413 apply(simp) |
1430 done |
1414 done |
1431 |
1415 |
|
1416 lemma mopup_mod2: "length (mopup k) mod 2 = 0" |
|
1417 apply(auto simp: mopup.simps) |
|
1418 by arith |
|
1419 |
1432 lemma [simp]: "fetch t_wcode_main (Suc (t_twice_len + length t_wcode_main_first_part div 2)) Oc |
1420 lemma [simp]: "fetch t_wcode_main (Suc (t_twice_len + length t_wcode_main_first_part div 2)) Oc |
1433 = (L, Suc 0)" |
1421 = (L, Suc 0)" |
1434 apply(subgoal_tac "length (t_twice) mod 2 = 0") |
1422 apply(subgoal_tac "length (t_twice) mod 2 = 0") |
1435 apply(simp add: t_wcode_main_def nth_append fetch.simps t_wcode_main_first_part_def |
1423 apply(simp add: t_wcode_main_def nth_append fetch.simps t_wcode_main_first_part_def |
1436 nth_of.simps shift_length t_twice_len_def, auto) |
1424 nth_of.simps t_twice_len_def, auto) |
1437 apply(simp add: t_twice_def) |
1425 apply(simp add: t_twice_def t_twice_compile_def) |
1438 apply(subgoal_tac "length (tm_of abc_twice) mod 2 = 0") |
1426 using mopup_mod2[of 1] |
1439 apply arith |
1427 apply(simp) |
1440 apply(rule_tac tm_even) |
1428 by arith |
1441 done |
|
1442 |
1429 |
1443 lemma wcode_jump1: |
1430 lemma wcode_jump1: |
1444 "\<exists> stp ln rn. steps (Suc (t_twice_len) + length t_wcode_main_first_part div 2, |
1431 "\<exists> stp ln rn. steps0 (Suc (t_twice_len) + length t_wcode_main_first_part div 2, |
1445 Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>n\<^esup>) |
1432 Bk\<up>(m) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(n)) |
1446 t_wcode_main stp |
1433 t_wcode_main stp |
1447 = (Suc 0, Bk\<^bsup>ln\<^esup> @ Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
1434 = (Suc 0, Bk\<up>(ln) @ Bk # ires, Bk # Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))" |
1448 apply(rule_tac x = "Suc 0" in exI, rule_tac x = "m" in exI, rule_tac x = n in exI) |
1435 apply(rule_tac x = "Suc 0" in exI, rule_tac x = "m" in exI, rule_tac x = n in exI) |
1449 apply(simp add: steps.simps tstep.simps exp_ind_def new_tape.simps) |
1436 apply(simp add: steps.simps step.simps exp_ind) |
1450 apply(case_tac m, simp, simp add: exp_ind_def) |
1437 apply(case_tac m, simp_all) |
1451 apply(simp add: exp_ind_def[THEN sym] exp_ind[THEN sym]) |
1438 apply(simp add: exp_ind[THEN sym]) |
1452 done |
1439 done |
1453 |
1440 |
1454 lemma wcode_main_first_part_len: |
1441 lemma wcode_main_first_part_len: |
1455 "length t_wcode_main_first_part = 24" |
1442 "length t_wcode_main_first_part = 24" |
1456 apply(simp add: t_wcode_main_first_part_def) |
1443 apply(simp add: t_wcode_main_first_part_def) |
1457 done |
1444 done |
1458 |
1445 |
1459 lemma wcode_double_case: |
1446 lemma wcode_double_case: |
1460 shows "\<exists>stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
1447 shows "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp = |
1461 (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
1448 (Suc 0, Bk # Bk\<up>(ln) @ Oc # ires, Bk # Oc\<up>(Suc (2 * rs + 2)) @ Bk\<up>(rn))" |
1462 proof - |
1449 proof - |
1463 have "\<exists>stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
1450 have "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp = |
1464 (13, Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
1451 (13, Bk # Bk # Bk\<up>(ln) @ Oc # ires, Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn))" |
1465 using wcode_double_case_first_correctness[of ires rs m n] |
1452 using wcode_double_case_first_correctness[of ires rs m n] |
1466 apply(simp) |
1453 apply(simp) |
1467 apply(erule_tac exE) |
1454 apply(erule_tac exE) |
1468 apply(case_tac "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, |
1455 apply(case_tac "steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, |
1469 Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main na", |
1456 Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main na", |
1470 auto simp: wcode_double_case_inv.simps |
1457 auto simp: wcode_double_case_inv.simps |
1471 wcode_before_double.simps) |
1458 wcode_before_double.simps) |
1472 apply(rule_tac x = na in exI, rule_tac x = ln in exI, rule_tac x = rn in exI) |
1459 apply(rule_tac x = na in exI, rule_tac x = ln in exI, rule_tac x = rn in exI) |
1473 apply(simp) |
1460 apply(simp) |
1474 done |
1461 done |
1475 from this obtain stpa lna rna where stp1: |
1462 from this obtain stpa lna rna where stp1: |
1476 "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa = |
1463 "steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stpa = |
1477 (13, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast |
1464 (13, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rna))" by blast |
1478 have "\<exists> stp ln rn. steps (13, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp = |
1465 have "\<exists> stp ln rn. steps0 (13, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rna)) t_wcode_main stp = |
1479 (13 + t_twice_len, Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
1466 (13 + t_twice_len, Bk # Bk # Bk\<up>(ln) @ Oc # ires, Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rn))" |
1480 using t_twice_append[of "Bk\<^bsup>lna\<^esup> @ Oc # ires" "Suc rs" rna] |
1467 using t_twice_append[of "Bk\<up>(lna) @ Oc # ires" "Suc rs" rna] |
1481 apply(erule_tac exE) |
1468 apply(erule_tac exE) |
1482 apply(erule_tac exE) |
1469 apply(erule_tac exE) |
1483 apply(erule_tac exE) |
1470 apply(erule_tac exE) |
1484 apply(simp add: wcode_main_first_part_len) |
1471 apply(simp add: wcode_main_first_part_len) |
1485 apply(rule_tac x = stp in exI, rule_tac x = "ln + lna" in exI, |
1472 apply(rule_tac x = stp in exI, rule_tac x = "ln + lna" in exI, |
1486 rule_tac x = rn in exI) |
1473 rule_tac x = rn in exI) |
1487 apply(simp add: t_wcode_main_def) |
1474 apply(simp add: t_wcode_main_def) |
1488 apply(simp add: exp_ind_def[THEN sym] exp_add[THEN sym]) |
1475 apply(simp add: replicate_Suc[THEN sym] exp_add[THEN sym] del: replicate_Suc) |
1489 done |
1476 done |
1490 from this obtain stpb lnb rnb where stp2: |
1477 from this obtain stpb lnb rnb where stp2: |
1491 "steps (13, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc rs)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb = |
1478 "steps0 (13, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rna)) t_wcode_main stpb = |
1492 (13 + t_twice_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnb\<^esup>)" by blast |
1479 (13 + t_twice_len, Bk # Bk # Bk\<up>(lnb) @ Oc # ires, Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rnb))" by blast |
1493 have "\<exists>stp ln rn. steps (13 + t_twice_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, |
1480 have "\<exists>stp ln rn. steps0 (13 + t_twice_len, Bk # Bk # Bk\<up>(lnb) @ Oc # ires, |
1494 Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnb\<^esup>) t_wcode_main stp = |
1481 Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rnb)) t_wcode_main stp = |
1495 (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
1482 (Suc 0, Bk # Bk\<up>(ln) @ Oc # ires, Bk # Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rn))" |
1496 using wcode_jump1[of lnb "Oc # ires" "Suc rs" rnb] |
1483 using wcode_jump1[of lnb "Oc # ires" "Suc rs" rnb] |
1497 apply(erule_tac exE) |
1484 apply(erule_tac exE) |
1498 apply(erule_tac exE) |
1485 apply(erule_tac exE) |
1499 apply(erule_tac exE) |
1486 apply(erule_tac exE) |
1500 apply(rule_tac x = stp in exI, |
1487 apply(rule_tac x = stp in exI, |
1501 rule_tac x = ln in exI, |
1488 rule_tac x = ln in exI, |
1502 rule_tac x = rn in exI, simp add:wcode_main_first_part_len t_wcode_main_def) |
1489 rule_tac x = rn in exI, simp add:wcode_main_first_part_len t_wcode_main_def) |
1503 apply(subgoal_tac "Bk\<^bsup>lnb\<^esup> @ Bk # Bk # Oc # ires = Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires", simp) |
1490 apply(subgoal_tac "Bk\<up>(lnb) @ Bk # Bk # Oc # ires = Bk # Bk # Bk\<up>(lnb) @ Oc # ires", simp) |
1504 apply(simp add: exp_ind_def[THEN sym] exp_ind[THEN sym]) |
1491 apply(simp add: replicate_Suc[THEN sym] exp_ind[THEN sym] del: replicate_Suc) |
1505 apply(simp) |
1492 apply(simp) |
1506 apply(case_tac lnb, simp, simp add: exp_ind_def[THEN sym] exp_ind) |
1493 apply(simp add: replicate_Suc[THEN sym] exp_ind del: replicate_Suc) |
1507 done |
1494 done |
1508 from this obtain stpc lnc rnc where stp3: |
1495 from this obtain stpc lnc rnc where stp3: |
1509 "steps (13 + t_twice_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, |
1496 "steps0 (13 + t_twice_len, Bk # Bk # Bk\<up>(lnb) @ Oc # ires, |
1510 Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnb\<^esup>) t_wcode_main stpc = |
1497 Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rnb)) t_wcode_main stpc = |
1511 (Suc 0, Bk # Bk\<^bsup>lnc\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (Suc (Suc (2 *rs)))\<^esup> @ Bk\<^bsup>rnc\<^esup>)" |
1498 (Suc 0, Bk # Bk\<up>(lnc) @ Oc # ires, Bk # Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rnc))" |
1512 by blast |
1499 by blast |
1513 from stp1 stp2 stp3 show "?thesis" |
1500 from stp1 stp2 stp3 show "?thesis" |
1514 apply(rule_tac x = "stpa + stpb + stpc" in exI, rule_tac x = lnc in exI, |
1501 apply(rule_tac x = "stpa + stpb + stpc" in exI, rule_tac x = lnc in exI, |
1515 rule_tac x = rnc in exI) |
1502 rule_tac x = rnc in exI) |
1516 apply(simp add: steps_add) |
1503 apply(simp add: steps_add) |
1979 definition t_fourtimes_len :: "nat" |
1982 definition t_fourtimes_len :: "nat" |
1980 where |
1983 where |
1981 "t_fourtimes_len = (length t_fourtimes div 2)" |
1984 "t_fourtimes_len = (length t_fourtimes div 2)" |
1982 |
1985 |
1983 lemma t_fourtimes_len_gr: "t_fourtimes_len > 0" |
1986 lemma t_fourtimes_len_gr: "t_fourtimes_len > 0" |
1984 apply(simp add: t_fourtimes_len_def t_fourtimes_def) |
1987 apply(simp add: t_fourtimes_len_def t_fourtimes_def mopup.simps t_fourtimes_compile_def) |
|
1988 done |
|
1989 |
|
1990 lemma [intro]: "rec_calc_rel (constn 4) [rs] 4" |
|
1991 using prime_rel_exec_eq[of "constn 4" "[rs]" 4] |
|
1992 apply(subgoal_tac "primerec (constn 4) 1", auto) |
|
1993 done |
|
1994 |
|
1995 lemma [intro]: "rec_calc_rel rec_mult [rs, 4] (4 * rs)" |
|
1996 using prime_rel_exec_eq[of "rec_mult" "[rs, 4]" "4*rs"] |
|
1997 apply(subgoal_tac "primerec rec_mult 2", auto simp: numeral_2_eq_2) |
1985 done |
1998 done |
1986 |
1999 |
1987 lemma t_fourtimes_correct: |
2000 lemma t_fourtimes_correct: |
1988 "\<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) |
2001 "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) |
1989 (tm_of abc_fourtimes @ tMp (Suc 0) (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) stp = |
2002 (tm_of abc_fourtimes @ shift (mopup 1) (length (tm_of abc_fourtimes) div 2)) stp = |
1990 (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2003 (0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))" |
1991 proof(case_tac "rec_ci rec_fourtimes") |
2004 proof(case_tac "rec_ci rec_fourtimes") |
1992 fix a b c |
2005 fix a b c |
1993 assume h: "rec_ci rec_fourtimes = (a, b, c)" |
2006 assume h: "rec_ci rec_fourtimes = (a, b, c)" |
1994 have "\<exists>stp m l. steps (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\<^bsup>n\<^esup>) (tm_of abc_fourtimes @ tMp (Suc 0) |
2007 have "\<exists>stp m l. steps0 (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\<up>(n)) (tm_of abc_fourtimes @ shift (mopup 1) |
1995 (start_of fourtimes_ly (length abc_fourtimes) - 1)) stp = (0, Bk\<^bsup>m\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4*rs)\<^esup> @ Bk\<^bsup>l\<^esup>)" |
2008 (length (tm_of abc_fourtimes) div 2)) stp = (0, Bk\<up>(m) @ Bk # Bk # ires, Oc\<up>(Suc (4*rs)) @ Bk\<up>(l))" |
1996 proof(rule_tac t_compiled_by_rec) |
2009 proof(rule_tac recursive_compile_to_tm_correct) |
1997 show "rec_ci rec_fourtimes = (a, b, c)" by (simp add: h) |
2010 show "rec_ci rec_fourtimes = (a, b, c)" by (simp add: h) |
1998 next |
2011 next |
1999 show "rec_calc_rel rec_fourtimes [rs] (4 * rs)" |
2012 show "rec_calc_rel rec_fourtimes [rs] (4 * rs)" |
2000 using prime_rel_exec_eq [of rec_fourtimes "[rs]" "4 * rs"] |
2013 apply(simp add: rec_fourtimes_def) |
2001 apply(subgoal_tac "primerec rec_fourtimes (length [rs])") |
2014 apply(rule_tac rs = "[rs, 4]" in calc_cn, simp_all) |
2002 apply(simp_all add: rec_fourtimes_def rec_exec.simps) |
2015 apply(rule_tac allI, case_tac k, auto simp: mult_lemma) |
2003 apply(auto) |
|
2004 apply(simp only: Nat.One_nat_def[THEN sym], auto) |
|
2005 done |
2016 done |
2006 next |
2017 next |
2007 show "length [rs] = Suc 0" by simp |
2018 show "length [rs] = 1" by simp |
|
2019 next |
|
2020 show "layout_of (a [+] dummy_abc 1) = layout_of (a [+] dummy_abc 1)" by simp |
2008 next |
2021 next |
2009 show "layout_of (a [+] dummy_abc (Suc 0)) = layout_of (a [+] dummy_abc (Suc 0))" |
2022 show "tm_of abc_fourtimes = tm_of (a [+] dummy_abc 1)" |
2010 by simp |
|
2011 next |
|
2012 show "start_of fourtimes_ly (length abc_fourtimes) = |
|
2013 start_of (layout_of (a [+] dummy_abc (Suc 0))) (length (a [+] dummy_abc (Suc 0)))" |
|
2014 using h |
|
2015 apply(simp add: fourtimes_ly_def abc_fourtimes_def) |
|
2016 done |
|
2017 next |
|
2018 show "tm_of abc_fourtimes = tm_of (a [+] dummy_abc (Suc 0))" |
|
2019 using h |
2023 using h |
2020 apply(simp add: abc_fourtimes_def) |
2024 apply(simp add: abc_fourtimes_def) |
2021 done |
2025 done |
2022 qed |
2026 qed |
2023 thus "\<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) |
2027 thus "?thesis" |
2024 (tm_of abc_fourtimes @ tMp (Suc 0) (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) stp = |
|
2025 (0, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
|
2026 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) |
2028 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) |
2027 done |
2029 done |
2028 qed |
2030 qed |
2029 |
2031 |
|
2032 lemma wf_fourtimes[intro]: "tm_wf (t_fourtimes_compile, 0)" |
|
2033 apply(simp only: t_fourtimes_compile_def) |
|
2034 apply(rule_tac t_compiled_correct) |
|
2035 apply(simp_all add: abc_twice_def) |
|
2036 done |
|
2037 |
2030 lemma t_fourtimes_change_term_state: |
2038 lemma t_fourtimes_change_term_state: |
2031 "\<exists> stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp |
2039 "\<exists> stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) t_fourtimes stp |
2032 = (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2040 = (Suc t_fourtimes_len, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))" |
2033 using t_fourtimes_correct[of ires rs n] |
2041 proof - |
2034 apply(erule_tac exE) |
2042 have "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) |
2035 apply(erule_tac exE) |
2043 (tm_of abc_fourtimes @ shift (mopup 1) ((length (tm_of abc_fourtimes) div 2))) stp = |
2036 apply(erule_tac exE) |
2044 (0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))" |
2037 proof(drule_tac turing_change_termi_state) |
2045 by(rule_tac t_fourtimes_correct) |
2038 fix stp ln rn |
2046 then obtain stp ln rn where |
2039 show "t_correct (tm_of abc_fourtimes @ tMp (Suc 0) |
2047 "steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) |
2040 (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))" |
2048 (tm_of abc_fourtimes @ shift (mopup 1) ((length (tm_of abc_fourtimes) div 2))) stp = |
2041 apply(rule_tac t_compiled_correct, auto simp: fourtimes_ly_def) |
2049 (0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))" by blast |
|
2050 hence "\<exists> stp. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) |
|
2051 (adjust t_fourtimes_compile) stp |
|
2052 = (Suc (length t_fourtimes_compile div 2), Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))" |
|
2053 apply(rule_tac stp = stp in adjust_halt_eq) |
|
2054 apply(simp add: t_fourtimes_compile_def, auto) |
2042 done |
2055 done |
2043 next |
2056 then obtain stpb where |
2044 fix stp ln rn |
2057 "steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) |
2045 show "\<exists>stp. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) |
2058 (adjust t_fourtimes_compile) stpb |
2046 (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) |
2059 = (Suc (length t_fourtimes_compile div 2), Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))" .. |
2047 (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) stp = |
2060 thus "?thesis" |
2048 (Suc (length (tm_of abc_fourtimes @ tMp (Suc 0) (start_of fourtimes_ly |
2061 apply(simp add: t_fourtimes_def t_fourtimes_len_def) |
2049 (length abc_fourtimes) - Suc 0)) div 2), Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>) \<Longrightarrow> |
2062 by metis |
2050 \<exists>stp ln rn. steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp = |
|
2051 (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
|
2052 apply(erule_tac exE) |
|
2053 apply(simp add: t_fourtimes_len_def t_fourtimes_def) |
|
2054 apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI, simp) |
|
2055 done |
|
2056 qed |
2063 qed |
2057 |
2064 |
|
2065 lemma [intro]: "length t_twice mod 2 = 0" |
|
2066 apply(auto simp: t_twice_def t_twice_compile_def) |
|
2067 done |
|
2068 |
2058 lemma t_fourtimes_append_pre: |
2069 lemma t_fourtimes_append_pre: |
2059 "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp |
2070 "steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) t_fourtimes stp |
2060 = (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>) |
2071 = (Suc t_fourtimes_len, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn)) |
2061 \<Longrightarrow> \<exists> stp>0. steps (Suc 0 + length (t_wcode_main_first_part @ |
2072 \<Longrightarrow> steps0 (Suc 0 + length (t_wcode_main_first_part @ |
2062 tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, |
2073 shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, |
2063 Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) |
2074 Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) |
2064 ((t_wcode_main_first_part @ |
2075 ((t_wcode_main_first_part @ |
2065 tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) @ |
2076 shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) @ |
2066 tshift t_fourtimes (length (t_wcode_main_first_part @ |
2077 shift t_fourtimes (length (t_wcode_main_first_part @ |
2067 tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2) @ ([(L, 1), (L, 1)])) stp |
2078 shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2) @ ([(L, 1), (L, 1)])) stp |
2068 = (Suc t_fourtimes_len + length (t_wcode_main_first_part @ |
2079 = ((Suc t_fourtimes_len) + length (t_wcode_main_first_part @ |
2069 tshift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, |
2080 shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, |
2070 Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2081 Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))" |
2071 proof(rule_tac t_tshift_lemma, auto) |
2082 apply(rule_tac tm_append_shift_append_steps, simp_all) |
2072 assume "steps (Suc 0, Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_fourtimes stp = |
2083 apply(auto simp: t_wcode_main_first_part_def) |
2073 (Suc t_fourtimes_len, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2084 done |
2074 thus "0 < stp" |
2085 |
2075 using t_fourtimes_len_gr |
|
2076 apply(case_tac stp, simp_all add: steps.simps) |
|
2077 done |
|
2078 next |
|
2079 show "Suc 0 \<le> length t_fourtimes div 2" |
|
2080 apply(simp add: t_fourtimes_def shift_length tMp.simps) |
|
2081 done |
|
2082 next |
|
2083 show "t_ncorrect (t_wcode_main_first_part @ |
|
2084 tshift t_twice (length t_wcode_main_first_part div 2) @ |
|
2085 [(L, Suc 0), (L, Suc 0)])" |
|
2086 apply(simp add: t_ncorrect.simps t_wcode_main_first_part_def shift_length |
|
2087 t_twice_def) |
|
2088 using tm_even[of abc_twice] |
|
2089 by arith |
|
2090 next |
|
2091 show "t_ncorrect t_fourtimes" |
|
2092 apply(simp add: t_fourtimes_def steps.simps t_ncorrect.simps) |
|
2093 using tm_even[of abc_fourtimes] |
|
2094 by arith |
|
2095 next |
|
2096 show "t_ncorrect [(L, Suc 0), (L, Suc 0)]" |
|
2097 apply(simp add: t_ncorrect.simps) |
|
2098 done |
|
2099 qed |
|
2100 |
2086 |
2101 lemma [simp]: "length t_wcode_main_first_part = 24" |
2087 lemma [simp]: "length t_wcode_main_first_part = 24" |
2102 apply(simp add: t_wcode_main_first_part_def) |
2088 apply(simp add: t_wcode_main_first_part_def) |
2103 done |
2089 done |
2104 |
2090 |
2105 lemma [simp]: "(26 + length t_twice) div 2 = (length t_twice) div 2 + 13" |
2091 lemma [simp]: "(26 + length t_twice) div 2 = (length t_twice) div 2 + 13" |
2106 using tm_even[of abc_twice] |
2092 apply(simp add: t_twice_def t_twice_def) |
2107 apply(simp add: t_twice_def) |
2093 done |
2108 done |
2094 |
2109 |
2095 lemma [simp]: "((26 + length (shift t_twice 12)) div 2) |
2110 lemma [simp]: "((26 + length (tshift t_twice 12)) div 2) |
2096 = (length (shift t_twice 12) div 2 + 13)" |
2111 = (length (tshift t_twice 12) div 2 + 13)" |
|
2112 using tm_even[of abc_twice] |
|
2113 apply(simp add: t_twice_def) |
2097 apply(simp add: t_twice_def) |
2114 done |
2098 done |
2115 |
2099 |
2116 lemma [simp]: "t_twice_len + 14 = 14 + length (tshift t_twice 12) div 2" |
2100 lemma [simp]: "t_twice_len + 14 = 14 + length (shift t_twice 12) div 2" |
2117 using tm_even[of abc_twice] |
2101 apply(simp add: t_twice_def t_twice_len_def) |
2118 apply(simp add: t_twice_def t_twice_len_def shift_length) |
|
2119 done |
2102 done |
2120 |
2103 |
2121 lemma t_fourtimes_append: |
2104 lemma t_fourtimes_append: |
2122 "\<exists> stp ln rn. |
2105 "\<exists> stp ln rn. |
2123 steps (Suc 0 + length (t_wcode_main_first_part @ tshift t_twice |
2106 steps0 (Suc 0 + length (t_wcode_main_first_part @ shift t_twice |
2124 (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, |
2107 (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, |
2125 Bk # Bk # ires, Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) |
2108 Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) |
2126 ((t_wcode_main_first_part @ tshift t_twice (length t_wcode_main_first_part div 2) @ |
2109 ((t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @ |
2127 [(L, 1), (L, 1)]) @ tshift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)]) stp |
2110 [(L, 1), (L, 1)]) @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)]) stp |
2128 = (Suc t_fourtimes_len + length (t_wcode_main_first_part @ tshift t_twice |
2111 = (Suc t_fourtimes_len + length (t_wcode_main_first_part @ shift t_twice |
2129 (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, Bk\<^bsup>ln\<^esup> @ Bk # Bk # ires, |
2112 (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, Bk\<up>(ln) @ Bk # Bk # ires, |
2130 Oc\<^bsup>Suc (4 * rs)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2113 Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))" |
2131 using t_fourtimes_change_term_state[of ires rs n] |
2114 using t_fourtimes_change_term_state[of ires rs n] |
2132 apply(erule_tac exE) |
2115 apply(erule_tac exE) |
2133 apply(erule_tac exE) |
2116 apply(erule_tac exE) |
2134 apply(erule_tac exE) |
2117 apply(erule_tac exE) |
2135 apply(drule_tac t_fourtimes_append_pre) |
2118 apply(drule_tac t_fourtimes_append_pre) |
2136 apply(erule_tac exE) |
2119 apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI) |
2137 apply(rule_tac x = stpa in exI, rule_tac x = ln in exI, rule_tac x = rn in exI) |
2120 apply(simp add: t_twice_len_def) |
2138 apply(simp add: t_twice_len_def shift_length) |
|
2139 done |
2121 done |
2140 |
2122 |
2141 lemma t_wcode_main_len: "length t_wcode_main = length t_twice + length t_fourtimes + 28" |
2123 lemma t_wcode_main_len: "length t_wcode_main = length t_twice + length t_fourtimes + 28" |
2142 apply(simp add: t_wcode_main_def shift_length) |
2124 apply(simp add: t_wcode_main_def) |
2143 done |
2125 done |
2144 |
2126 |
|
2127 lemma even_twice_len: "length t_twice mod 2 = 0" |
|
2128 apply(auto simp: t_twice_def t_twice_compile_def) |
|
2129 done |
|
2130 |
|
2131 lemma even_fourtimes_len: "length t_fourtimes mod 2 = 0" |
|
2132 apply(auto simp: t_fourtimes_def t_fourtimes_compile_def) |
|
2133 done |
|
2134 |
|
2135 lemma [simp]: "2 * (length t_twice div 2) = length t_twice" |
|
2136 using even_twice_len |
|
2137 by arith |
|
2138 |
|
2139 lemma [simp]: "2 * (length t_fourtimes div 2) = length t_fourtimes" |
|
2140 using even_fourtimes_len |
|
2141 by arith |
|
2142 |
|
2143 lemma [simp]: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) Oc |
|
2144 = (L, Suc 0)" |
|
2145 apply(subgoal_tac "14 = Suc 13") |
|
2146 apply(simp only: fetch.simps add_Suc nth_of.simps t_wcode_main_def) |
|
2147 apply(simp add:length_append length_shift Parity.two_times_even_div_two even_twice_len t_fourtimes_len_def) |
|
2148 by arith |
|
2149 |
|
2150 lemma [simp]: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) Bk |
|
2151 = (L, Suc 0)" |
|
2152 apply(subgoal_tac "14 = Suc 13") |
|
2153 apply(simp only: fetch.simps add_Suc nth_of.simps t_wcode_main_def) |
|
2154 apply(simp add:length_append length_shift Parity.two_times_even_div_two even_twice_len t_fourtimes_len_def nth_append) |
|
2155 by arith |
|
2156 |
2145 lemma [simp]: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) b |
2157 lemma [simp]: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) b |
2146 = (L, Suc 0)" |
2158 = (L, Suc 0)" |
2147 using tm_even[of "abc_twice"] tm_even[of "abc_fourtimes"] |
2159 apply(case_tac b, simp_all) |
2148 apply(case_tac b) |
|
2149 apply(simp_all only: fetch.simps) |
|
2150 apply(auto simp: nth_of.simps t_wcode_main_len t_twice_len_def |
|
2151 t_fourtimes_def t_twice_def t_fourtimes_def t_fourtimes_len_def) |
|
2152 apply(auto simp: t_wcode_main_def t_wcode_main_first_part_def shift_length t_twice_def nth_append |
|
2153 t_fourtimes_def) |
|
2154 done |
2160 done |
2155 |
2161 |
2156 lemma wcode_jump2: |
2162 lemma wcode_jump2: |
2157 "\<exists> stp ln rn. steps (t_twice_len + 14 + t_fourtimes_len |
2163 "\<exists> stp ln rn. steps0 (t_twice_len + 14 + t_fourtimes_len |
2158 , Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, Oc\<^bsup>Suc (4 * rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>) t_wcode_main stp = |
2164 , Bk # Bk # Bk\<up>(lnb) @ Oc # ires, Oc\<up>(Suc (4 * rs + 4)) @ Bk\<up>(rnb)) t_wcode_main stp = |
2159 (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4 * rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2165 (Suc 0, Bk # Bk\<up>(ln) @ Oc # ires, Bk # Oc\<up>(Suc (4 * rs + 4)) @ Bk\<up>(rn))" |
2160 apply(rule_tac x = "Suc 0" in exI) |
2166 apply(rule_tac x = "Suc 0" in exI) |
2161 apply(simp add: steps.simps shift_length) |
2167 apply(simp add: steps.simps) |
2162 apply(rule_tac x = lnb in exI, rule_tac x = rnb in exI) |
2168 apply(rule_tac x = lnb in exI, rule_tac x = rnb in exI) |
2163 apply(simp add: tstep.simps new_tape.simps) |
2169 apply(simp add: step.simps) |
2164 done |
2170 done |
2165 |
2171 |
2166 lemma wcode_fourtimes_case: |
2172 lemma wcode_fourtimes_case: |
2167 shows "\<exists>stp ln rn. |
2173 shows "\<exists>stp ln rn. |
2168 steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
2174 steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp = |
2169 (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2175 (Suc 0, Bk # Bk\<up>(ln) @ Oc # ires, Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rn))" |
2170 proof - |
2176 proof - |
2171 have "\<exists>stp ln rn. |
2177 have "\<exists>stp ln rn. |
2172 steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
2178 steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp = |
2173 (t_twice_len + 14, Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2179 (t_twice_len + 14, Bk # Bk # Bk\<up>(ln) @ Oc # ires, Oc\<up>(Suc (rs + 1)) @ Bk\<up>(rn))" |
2174 using wcode_fourtimes_case_first_correctness[of ires rs m n] |
2180 using wcode_fourtimes_case_first_correctness[of ires rs m n] |
2175 apply(simp add: wcode_fourtimes_case_inv.simps, auto) |
2181 apply(simp add: wcode_fourtimes_case_inv.simps, auto) |
2176 apply(rule_tac x = na in exI, rule_tac x = ln in exI, |
2182 apply(rule_tac x = na in exI, rule_tac x = ln in exI, |
2177 rule_tac x = rn in exI) |
2183 rule_tac x = rn in exI) |
2178 apply(simp) |
2184 apply(simp) |
2179 done |
2185 done |
2180 from this obtain stpa lna rna where stp1: |
2186 from this obtain stpa lna rna where stp1: |
2181 "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Oc # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa = |
2187 "steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stpa = |
2182 (t_twice_len + 14, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast |
2188 (t_twice_len + 14, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (rs + 1)) @ Bk\<up>(rna))" by blast |
2183 have "\<exists>stp ln rn. steps (t_twice_len + 14, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rna\<^esup>) |
2189 have "\<exists>stp ln rn. steps0 (t_twice_len + 14, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (rs + 1)) @ Bk\<up>(rna)) |
2184 t_wcode_main stp = |
2190 t_wcode_main stp = |
2185 (t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2191 (t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\<up>(ln) @ Oc # ires, Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rn))" |
2186 using t_fourtimes_append[of " Bk\<^bsup>lna\<^esup> @ Oc # ires" "rs + 1" rna] |
2192 using t_fourtimes_append[of " Bk\<up>(lna) @ Oc # ires" "rs + 1" rna] |
2187 apply(erule_tac exE) |
2193 apply(erule_tac exE) |
2188 apply(erule_tac exE) |
2194 apply(erule_tac exE) |
2189 apply(erule_tac exE) |
2195 apply(erule_tac exE) |
2190 apply(simp add: t_wcode_main_def) |
2196 apply(simp add: t_wcode_main_def) |
2191 apply(rule_tac x = stp in exI, |
2197 apply(rule_tac x = stp in exI, |
2192 rule_tac x = "ln + lna" in exI, |
2198 rule_tac x = "ln + lna" in exI, |
2193 rule_tac x = rn in exI, simp) |
2199 rule_tac x = rn in exI, simp) |
2194 apply(simp add: exp_ind_def[THEN sym] exp_add[THEN sym]) |
2200 apply(simp add: replicate_Suc[THEN sym] exp_add[THEN sym] del: replicate_Suc) |
2195 done |
2201 done |
2196 from this obtain stpb lnb rnb where stp2: |
2202 from this obtain stpb lnb rnb where stp2: |
2197 "steps (t_twice_len + 14, Bk # Bk # Bk\<^bsup>lna\<^esup> @ Oc # ires, Oc\<^bsup>Suc (rs + 1)\<^esup> @ Bk\<^bsup>rna\<^esup>) |
2203 "steps0 (t_twice_len + 14, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (rs + 1)) @ Bk\<up>(rna)) |
2198 t_wcode_main stpb = |
2204 t_wcode_main stpb = |
2199 (t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>)" |
2205 (t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\<up>(lnb) @ Oc # ires, Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rnb))" |
2200 by blast |
2206 by blast |
2201 have "\<exists>stp ln rn. steps (t_twice_len + 14 + t_fourtimes_len, |
2207 have "\<exists>stp ln rn. steps0 (t_twice_len + 14 + t_fourtimes_len, |
2202 Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>) |
2208 Bk # Bk # Bk\<up>(lnb) @ Oc # ires, Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rnb)) |
2203 t_wcode_main stp = |
2209 t_wcode_main stp = |
2204 (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2210 (Suc 0, Bk # Bk\<up>(ln) @ Oc # ires, Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rn))" |
2205 apply(rule wcode_jump2) |
2211 apply(rule wcode_jump2) |
2206 done |
2212 done |
2207 from this obtain stpc lnc rnc where stp3: |
2213 from this obtain stpc lnc rnc where stp3: |
2208 "steps (t_twice_len + 14 + t_fourtimes_len, |
2214 "steps0 (t_twice_len + 14 + t_fourtimes_len, |
2209 Bk # Bk # Bk\<^bsup>lnb\<^esup> @ Oc # ires, Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnb\<^esup>) |
2215 Bk # Bk # Bk\<up>(lnb) @ Oc # ires, Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rnb)) |
2210 t_wcode_main stpc = |
2216 t_wcode_main stpc = |
2211 (Suc 0, Bk # Bk\<^bsup>lnc\<^esup> @ Oc # ires, Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rnc\<^esup>)" |
2217 (Suc 0, Bk # Bk\<up>(lnc) @ Oc # ires, Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rnc))" |
2212 by blast |
2218 by blast |
2213 from stp1 stp2 stp3 show "?thesis" |
2219 from stp1 stp2 stp3 show "?thesis" |
2214 apply(rule_tac x = "stpa + stpb + stpc" in exI, |
2220 apply(rule_tac x = "stpa + stpb + stpc" in exI, |
2215 rule_tac x = lnc in exI, rule_tac x = rnc in exI) |
2221 rule_tac x = lnc in exI, rule_tac x = rnc in exI) |
2216 apply(simp add: steps_add) |
2222 apply(simp add: steps_add) |
2405 apply(auto) |
2413 apply(auto) |
2406 done |
2414 done |
2407 qed |
2415 qed |
2408 |
2416 |
2409 declare wcode_halt_case_inv.simps[simp del] |
2417 declare wcode_halt_case_inv.simps[simp del] |
2410 lemma [intro]: "\<exists> xs. (<rev list @ [aa::nat]> :: block list) = Oc # xs" |
2418 lemma [intro]: "\<exists> xs. (<rev list @ [aa::nat]> :: cell list) = Oc # xs" |
2411 apply(case_tac "rev list", simp) |
2419 apply(case_tac "rev list", simp) |
2412 apply(simp add: tape_of_nat_abv tape_of_nat_list.simps exp_ind_def) |
2420 apply(simp add: tape_of_nl_cons) |
2413 apply(case_tac list, simp, simp) |
|
2414 done |
2421 done |
2415 |
2422 |
2416 lemma wcode_halt_case: |
2423 lemma wcode_halt_case: |
2417 "\<exists>stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) |
2424 "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) |
2418 t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2425 t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))" |
2419 using t_halt_case_correctness[of ires rs m n] |
2426 using t_halt_case_correctness[of ires rs m n] |
2420 apply(simp) |
2427 apply(simp) |
2421 apply(erule_tac exE) |
2428 apply(erule_tac exE) |
2422 apply(case_tac "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires, |
2429 apply(case_tac "steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Bk # ires, |
2423 Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main na") |
2430 Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main na") |
2424 apply(auto simp: wcode_halt_case_inv.simps wcode_stop.simps) |
2431 apply(auto simp: wcode_halt_case_inv.simps wcode_stop.simps) |
2425 apply(rule_tac x = na in exI, rule_tac x = ln in exI, |
2432 apply(rule_tac x = na in exI, rule_tac x = ln in exI, |
2426 rule_tac x = rn in exI, simp) |
2433 rule_tac x = rn in exI, simp) |
2427 done |
2434 done |
2428 |
2435 |
2429 lemma bl_bin_one: "bl_bin [Oc] = Suc 0" |
2436 lemma bl_bin_one: "bl_bin [Oc] = Suc 0" |
2430 apply(simp add: bl_bin.simps) |
2437 apply(simp add: bl_bin.simps) |
2431 done |
2438 done |
2432 |
2439 |
|
2440 lemma [simp]: "bl_bin [Oc] = 1" |
|
2441 apply(simp add: bl_bin.simps) |
|
2442 done |
|
2443 |
|
2444 lemma [intro]: "2 * 2 ^ a = Suc (Suc (2 * bl_bin (Oc \<up> a)))" |
|
2445 apply(induct a, auto simp: bl_bin.simps) |
|
2446 done |
|
2447 declare replicate_Suc[simp del] |
|
2448 |
2433 lemma t_wcode_main_lemma_pre: |
2449 lemma t_wcode_main_lemma_pre: |
2434 "\<lbrakk>args \<noteq> []; lm = <args::nat list>\<rbrakk> \<Longrightarrow> |
2450 "\<lbrakk>args \<noteq> []; lm = <args::nat list>\<rbrakk> \<Longrightarrow> |
2435 \<exists> stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev lm @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main |
2451 \<exists> stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main |
2436 stp |
2452 stp |
2437 = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin lm + rs * 2^(length lm - 1) \<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2453 = (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin lm + rs * 2^(length lm - 1) ) @ Bk\<up>(rn))" |
2438 proof(induct "length args" arbitrary: args lm rs m n, simp) |
2454 proof(induct "length args" arbitrary: args lm rs m n, simp) |
2439 fix x args lm rs m n |
2455 fix x args lm rs m n |
2440 assume ind: |
2456 assume ind: |
2441 "\<And>args lm rs m n. |
2457 "\<And>args lm rs m n. |
2442 \<lbrakk>x = length args; (args::nat list) \<noteq> []; lm = <args>\<rbrakk> |
2458 \<lbrakk>x = length args; (args::nat list) \<noteq> []; lm = <args>\<rbrakk> |
2443 \<Longrightarrow> \<exists>stp ln rn. |
2459 \<Longrightarrow> \<exists>stp ln rn. |
2444 steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev lm @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
2460 steps0 (Suc 0, Bk # Bk\<up>(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp = |
2445 (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin lm + rs * 2 ^ (length lm - 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2461 (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin lm + rs * 2 ^ (length lm - 1)) @ Bk\<up>(rn))" |
2446 |
|
2447 and h: "Suc x = length args" "(args::nat list) \<noteq> []" "lm = <args>" |
2462 and h: "Suc x = length args" "(args::nat list) \<noteq> []" "lm = <args>" |
2448 from h have "\<exists> (a::nat) xs. args = xs @ [a]" |
2463 from h have "\<exists> (a::nat) xs. args = xs @ [a]" |
2449 apply(rule_tac x = "last args" in exI) |
2464 apply(rule_tac x = "last args" in exI) |
2450 apply(rule_tac x = "butlast args" in exI, auto) |
2465 apply(rule_tac x = "butlast args" in exI, auto) |
2451 done |
2466 done |
2452 from this obtain a xs where "args = xs @ [a]" by blast |
2467 from this obtain a xs where "args = xs @ [a]" by blast |
2453 from h and this show |
2468 from h and this show |
2454 "\<exists>stp ln rn. |
2469 "\<exists>stp ln rn. |
2455 steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev lm @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
2470 steps0 (Suc 0, Bk # Bk\<up>(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp = |
2456 (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin lm + rs * 2 ^ (length lm - 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2471 (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin lm + rs * 2 ^ (length lm - 1)) @ Bk\<up>(rn))" |
2457 proof(case_tac "xs::nat list", simp) |
2472 proof(case_tac "xs::nat list", simp) |
2458 show "\<exists>stp ln rn. |
2473 show "\<exists>stp ln rn. |
2459 steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
2474 steps0 (Suc 0, Bk # Bk \<up> m @ Oc \<up> Suc a @ Bk # Bk # ires, Bk # Oc \<up> Suc rs @ Bk \<up> n) t_wcode_main stp = |
2460 (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<a>) + rs * 2 ^ a\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2475 (0, Bk # ires, Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> (bl_bin (Oc \<up> Suc a) + rs * 2 ^ a) @ Bk \<up> rn)" |
2461 proof(induct "a" arbitrary: m n rs ires, simp) |
2476 proof(induct "a" arbitrary: m n rs ires, simp) |
2462 fix m n rs ires |
2477 fix m n rs ires |
2463 show "\<exists>stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) |
2478 show "\<exists>stp ln rn. |
2464 t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin [Oc] + rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2479 steps0 (Suc 0, Bk # Bk \<up> m @ Oc # Bk # Bk # ires, Bk # Oc \<up> Suc rs @ Bk \<up> n) t_wcode_main stp = |
2465 apply(simp add: bl_bin_one) |
2480 (0, Bk # ires, Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> Suc rs @ Bk \<up> rn)" |
2466 apply(rule_tac wcode_halt_case) |
2481 apply(rule_tac wcode_halt_case) |
2467 done |
2482 done |
2468 next |
2483 next |
2469 fix a m n rs ires |
2484 fix a m n rs ires |
2470 assume ind2: |
2485 assume ind2: |
2471 "\<And>m n rs ires. |
2486 "\<And>m n rs ires. |
2472 \<exists>stp ln rn. |
2487 \<exists>stp ln rn. |
2473 steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
2488 steps0 (Suc 0, Bk # Bk \<up> m @ Oc \<up> Suc a @ Bk # Bk # ires, Bk # Oc \<up> Suc rs @ Bk \<up> n) t_wcode_main stp = |
2474 (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<a>) + rs * 2 ^ a\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2489 (0, Bk # ires, Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> (bl_bin (Oc \<up> Suc a) + rs * 2 ^ a) @ Bk \<up> rn)" |
2475 show "\<exists>stp ln rn. |
2490 show " \<exists>stp ln rn. |
2476 steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev (<Suc a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
2491 steps0 (Suc 0, Bk # Bk \<up> m @ Oc \<up> Suc (Suc a) @ Bk # Bk # ires, Bk # Oc \<up> Suc rs @ Bk \<up> n) t_wcode_main stp = |
2477 (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<Suc a>) + rs * 2 ^ Suc a\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2492 (0, Bk # ires, Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> (bl_bin (Oc \<up> Suc (Suc a)) + rs * 2 ^ Suc a) @ Bk \<up> rn)" |
2478 proof - |
2493 proof - |
2479 have "\<exists>stp ln rn. |
2494 have "\<exists>stp ln rn. |
2480 steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev (<Suc a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
2495 steps0 (Suc 0, Bk # Bk\<up>(m) @ rev (<Suc a>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp = |
2481 (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2496 (Suc 0, Bk # Bk\<up>(ln) @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc (2 * rs + 2)) @ Bk\<up>(rn))" |
2482 apply(simp add: tape_of_nat) |
2497 apply(simp add: tape_of_nat) |
2483 using wcode_double_case[of m "Oc\<^bsup>a\<^esup> @ Bk # Bk # ires" rs n] |
2498 using wcode_double_case[of m "Oc\<up>(a) @ Bk # Bk # ires" rs n] |
2484 apply(simp add: exp_ind_def) |
2499 apply(simp add: replicate_Suc) |
2485 done |
2500 done |
2486 from this obtain stpa lna rna where stp1: |
2501 from this obtain stpa lna rna where stp1: |
2487 "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev (<Suc a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa = |
2502 "steps0 (Suc 0, Bk # Bk\<up>(m) @ rev (<Suc a>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stpa = |
2488 (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast |
2503 (Suc 0, Bk # Bk\<up>(lna) @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc (2 * rs + 2)) @ Bk\<up>(rna))" by blast |
2489 moreover have |
2504 moreover have |
2490 "\<exists>stp ln rn. |
2505 "\<exists>stp ln rn. |
2491 steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp = |
2506 steps0 (Suc 0, Bk # Bk\<up>(lna) @ rev (<a::nat>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc (2 * rs + 2)) @ Bk\<up>(rna)) t_wcode_main stp = |
2492 (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<a>) + (2*rs + 2) * 2 ^ a\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2507 (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin (<a>) + (2*rs + 2) * 2 ^ a) @ Bk\<up>(rn))" |
2493 using ind2[of lna ires "2*rs + 2" rna] by simp |
2508 using ind2[of lna ires "2*rs + 2" rna] by(simp add: tape_of_nl_abv tape_of_nat_abv) |
2494 from this obtain stpb lnb rnb where stp2: |
2509 from this obtain stpb lnb rnb where stp2: |
2495 "steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc (2 * rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb = |
2510 "steps0 (Suc 0, Bk # Bk\<up>(lna) @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc (2 * rs + 2)) @ Bk\<up>(rna)) t_wcode_main stpb = |
2496 (0, Bk # ires, Bk # Oc # Bk\<^bsup>lnb\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<a>) + (2*rs + 2) * 2 ^ a\<^esup> @ Bk\<^bsup>rnb\<^esup>)" |
2511 (0, Bk # ires, Bk # Oc # Bk\<up>(lnb) @ Bk # Bk # Oc\<up>(bl_bin (<a>) + (2*rs + 2) * 2 ^ a) @ Bk\<up>(rnb))" |
2497 by blast |
2512 by blast |
2498 from stp1 and stp2 show "?thesis" |
2513 from stp1 and stp2 show "?thesis" |
2499 apply(rule_tac x = "stpa + stpb" in exI, |
2514 apply(rule_tac x = "stpa + stpb" in exI, |
2500 rule_tac x = lnb in exI, rule_tac x = rnb in exI, simp) |
2515 rule_tac x = lnb in exI, rule_tac x = rnb in exI, simp add: tape_of_nat_abv) |
2501 apply(simp add: steps_add bl_bin_nat_Suc exponent_def) |
2516 apply(simp add: bl_bin.simps replicate_Suc) |
|
2517 apply(auto) |
2502 done |
2518 done |
2503 qed |
2519 qed |
2504 qed |
2520 qed |
2505 next |
2521 next |
2506 fix aa list |
2522 fix aa list |
2507 assume g: "Suc x = length args" "args \<noteq> []" "lm = <args>" "args = xs @ [a::nat]" "xs = (aa::nat) # list" |
2523 assume g: "Suc x = length args" "args \<noteq> []" "lm = <args>" "args = xs @ [a::nat]" "xs = (aa::nat) # list" |
2508 thus "\<exists>stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ rev lm @ Bk # Bk # ires, Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
2524 thus "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp = |
2509 (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin lm + rs * 2 ^ (length lm - 1)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2525 (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin lm + rs * 2 ^ (length lm - 1)) @ Bk\<up>(rn))" |
2510 proof(induct a arbitrary: m n rs args lm, simp_all add: tape_of_nl_rev, |
2526 proof(induct a arbitrary: m n rs args lm, simp_all add: tape_of_nl_rev, |
2511 simp only: tape_of_nl_cons_app1, simp) |
2527 simp only: tape_of_nl_cons_app1, simp) |
2512 fix m n rs args lm |
2528 fix m n rs args lm |
2513 have "\<exists>stp ln rn. |
2529 have "\<exists>stp ln rn. |
2514 steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # rev (<(aa::nat) # list>) @ Bk # Bk # ires, |
2530 steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # rev (<(aa::nat) # list>) @ Bk # Bk # ires, |
2515 Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
2531 Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp = |
2516 (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ rev (<aa # list>) @ Bk # Bk # ires, |
2532 (Suc 0, Bk # Bk\<up>(ln) @ rev (<aa # list>) @ Bk # Bk # ires, |
2517 Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2533 Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rn))" |
2518 proof(simp add: tape_of_nl_rev) |
2534 proof(simp add: tape_of_nl_rev) |
2519 have "\<exists> xs. (<rev list @ [aa]>) = Oc # xs" by auto |
2535 have "\<exists> xs. (<rev list @ [aa]>) = Oc # xs" by auto |
2520 from this obtain xs where "(<rev list @ [aa]>) = Oc # xs" .. |
2536 from this obtain xs where "(<rev list @ [aa]>) = Oc # xs" .. |
2521 thus "\<exists>stp ln rn. |
2537 thus "\<exists>stp ln rn. |
2522 steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # <rev list @ [aa]> @ Bk # Bk # ires, |
2538 steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # <rev list @ [aa]> @ Bk # Bk # ires, |
2523 Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
2539 Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp = |
2524 (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ <rev list @ [aa]> @ Bk # Bk # ires, Bk # Oc\<^bsup>5 + 4 * rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2540 (Suc 0, Bk # Bk\<up>(ln) @ <rev list @ [aa]> @ Bk # Bk # ires, Bk # Oc\<up>(5 + 4 * rs) @ Bk\<up>(rn))" |
2525 apply(simp) |
2541 apply(simp) |
2526 using wcode_fourtimes_case[of m "xs @ Bk # Bk # ires" rs n] |
2542 using wcode_fourtimes_case[of m "xs @ Bk # Bk # ires" rs n] |
2527 apply(simp) |
2543 apply(simp) |
2528 done |
2544 done |
2529 qed |
2545 qed |
2530 from this obtain stpa lna rna where stp1: |
2546 from this obtain stpa lna rna where stp1: |
2531 "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # rev (<aa # list>) @ Bk # Bk # ires, |
2547 "steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # rev (<aa # list>) @ Bk # Bk # ires, |
2532 Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa = |
2548 Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stpa = |
2533 (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<aa # list>) @ Bk # Bk # ires, |
2549 (Suc 0, Bk # Bk\<up>(lna) @ rev (<aa # list>) @ Bk # Bk # ires, |
2534 Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast |
2550 Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rna))" by blast |
2535 from g have |
2551 from g have |
2536 "\<exists> stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, |
2552 "\<exists> stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(lna) @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, |
2537 Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp = (0, Bk # ires, |
2553 Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rna)) t_wcode_main stp = (0, Bk # ires, |
2538 Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<aa#list>)+ (4*rs + 4) * 2^(length (<aa#list>) - 1) \<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2554 Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin (<aa#list>)+ (4*rs + 4) * 2^(length (<aa#list>) - 1) ) @ Bk\<up>(rn))" |
2539 apply(rule_tac args = "(aa::nat)#list" in ind, simp_all) |
2555 apply(rule_tac args = "(aa::nat)#list" in ind, simp_all) |
2540 done |
2556 done |
2541 from this obtain stpb lnb rnb where stp2: |
2557 from this obtain stpb lnb rnb where stp2: |
2542 "steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, |
2558 "steps0 (Suc 0, Bk # Bk\<up>(lna) @ rev (<(aa::nat) # list>) @ Bk # Bk # ires, |
2543 Bk # Oc\<^bsup>Suc (4*rs + 4)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb = (0, Bk # ires, |
2559 Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rna)) t_wcode_main stpb = (0, Bk # ires, |
2544 Bk # Oc # Bk\<^bsup>lnb\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<aa#list>)+ (4*rs + 4) * 2^(length (<aa#list>) - 1) \<^esup> @ Bk\<^bsup>rnb\<^esup>)" |
2560 Bk # Oc # Bk\<up>(lnb) @ Bk # Bk # Oc\<up>(bl_bin (<aa#list>)+ (4*rs + 4) * 2^(length (<aa#list>) - 1) ) @ Bk\<up>(rnb))" |
2545 by blast |
2561 by blast |
2546 from stp1 and stp2 and h |
2562 from stp1 and stp2 and h |
2547 show "\<exists>stp ln rn. |
2563 show "\<exists>stp ln rn. |
2548 steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc # Bk # <rev list @ [aa]> @ Bk # Bk # ires, |
2564 steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # <rev list @ [aa]> @ Bk # Bk # ires, |
2549 Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
2565 Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp = |
2550 (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # |
2566 (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # |
2551 Bk # Oc\<^bsup>bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # <list @ [0]>) + rs * (2 * 2 ^ (aa + length (<list @ [0]>)))\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2567 Bk # Oc\<up>(bl_bin (Oc\<up>(Suc aa) @ Bk # <list @ [0]>) + rs * (2 * 2 ^ (aa + length (<list @ [0]>)))) @ Bk\<up>(rn))" |
2552 apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI, |
2568 apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI, |
2553 rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_rev) |
2569 rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_rev) |
2554 done |
2570 done |
2555 next |
2571 next |
2556 fix ab m n rs args lm |
2572 fix ab m n rs args lm |
2557 assume ind2: |
2573 assume ind2: |
2558 "\<And> m n rs args lm. |
2574 "\<And> m n rs args lm. |
2559 \<lbrakk>lm = <aa # list @ [ab]>; args = aa # list @ [ab]\<rbrakk> |
2575 \<lbrakk>lm = <aa # list @ [ab]>; args = aa # list @ [ab]\<rbrakk> |
2560 \<Longrightarrow> \<exists>stp ln rn. |
2576 \<Longrightarrow> \<exists>stp ln rn. |
2561 steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ <ab # rev list @ [aa]> @ Bk # Bk # ires, |
2577 steps0 (Suc 0, Bk # Bk\<up>(m) @ <ab # rev list @ [aa]> @ Bk # Bk # ires, |
2562 Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
2578 Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp = |
2563 (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # |
2579 (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # |
2564 Bk # Oc\<^bsup>bl_bin (<aa # list @ [ab]>) + rs * 2 ^ (length (<aa # list @ [ab]>) - Suc 0)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2580 Bk # Oc\<up>(bl_bin (<aa # list @ [ab]>) + rs * 2 ^ (length (<aa # list @ [ab]>) - Suc 0)) @ Bk\<up>(rn))" |
2565 and k: "args = aa # list @ [Suc ab]" "lm = <aa # list @ [Suc ab]>" |
2581 and k: "args = aa # list @ [Suc ab]" "lm = <aa # list @ [Suc ab]>" |
2566 show "\<exists>stp ln rn. |
2582 show "\<exists>stp ln rn. |
2567 steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ <Suc ab # rev list @ [aa]> @ Bk # Bk # ires, |
2583 steps0 (Suc 0, Bk # Bk\<up>(m) @ <Suc ab # rev list @ [aa]> @ Bk # Bk # ires, |
2568 Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
2584 Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp = |
2569 (0, Bk # ires,Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # |
2585 (0, Bk # ires,Bk # Oc # Bk\<up>(ln) @ Bk # |
2570 Bk # Oc\<^bsup>bl_bin (<aa # list @ [Suc ab]>) + rs * 2 ^ (length (<aa # list @ [Suc ab]>) - Suc 0)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2586 Bk # Oc\<up>(bl_bin (<aa # list @ [Suc ab]>) + rs * 2 ^ (length (<aa # list @ [Suc ab]>) - Suc 0)) @ Bk\<up>(rn))" |
2571 proof(simp add: tape_of_nl_cons_app1) |
2587 proof(simp add: tape_of_nl_cons_app1) |
2572 have "\<exists>stp ln rn. |
2588 have "\<exists>stp ln rn. |
2573 steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc\<^bsup>Suc (Suc ab)\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, |
2589 steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc\<up>(Suc (Suc ab)) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, |
2574 Bk # Oc # Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp |
2590 Bk # Oc # Oc\<up>(rs) @ Bk\<up>(n)) t_wcode_main stp |
2575 = (Suc 0, Bk # Bk\<^bsup>ln\<^esup> @ Oc\<^bsup>Suc ab\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, |
2591 = (Suc 0, Bk # Bk\<up>(ln) @ Oc\<up>(Suc ab) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, |
2576 Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2592 Bk # Oc\<up>(Suc (2*rs + 2)) @ Bk\<up>(rn))" |
2577 using wcode_double_case[of m "Oc\<^bsup>ab\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires" |
2593 using wcode_double_case[of m "Oc\<up>(ab) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires" |
2578 rs n] |
2594 rs n] |
2579 apply(simp add: exp_ind_def) |
2595 apply(simp add: replicate_Suc) |
2580 done |
2596 done |
2581 from this obtain stpa lna rna where stp1: |
2597 from this obtain stpa lna rna where stp1: |
2582 "steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc\<^bsup>Suc (Suc ab)\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, |
2598 "steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc\<up>(Suc (Suc ab)) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, |
2583 Bk # Oc # Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stpa |
2599 Bk # Oc # Oc\<up>(rs) @ Bk\<up>(n)) t_wcode_main stpa |
2584 = (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ Oc\<^bsup>Suc ab\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, |
2600 = (Suc 0, Bk # Bk\<up>(lna) @ Oc\<up>(Suc ab) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, |
2585 Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>)" by blast |
2601 Bk # Oc\<up>(Suc (2*rs + 2)) @ Bk\<up>(rna))" by blast |
2586 from k have |
2602 from k have |
2587 "\<exists> stp ln rn. steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ <ab # rev list @ [aa]> @ Bk # Bk # ires, |
2603 "\<exists> stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(lna) @ <ab # rev list @ [aa]> @ Bk # Bk # ires, |
2588 Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stp |
2604 Bk # Oc\<up>(Suc (2*rs + 2)) @ Bk\<up>(rna)) t_wcode_main stp |
2589 = (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # |
2605 = (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # |
2590 Bk # Oc\<^bsup>bl_bin (<aa # list @ [ab]> ) + (2*rs + 2)* 2^(length (<aa # list @ [ab]>) - Suc 0)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2606 Bk # Oc\<up>(bl_bin (<aa # list @ [ab]> ) + (2*rs + 2)* 2^(length (<aa # list @ [ab]>) - Suc 0)) @ Bk\<up>(rn))" |
2591 apply(rule_tac ind2, simp_all) |
2607 apply(rule_tac ind2, simp_all) |
2592 done |
2608 done |
2593 from this obtain stpb lnb rnb where stp2: |
2609 from this obtain stpb lnb rnb where stp2: |
2594 "steps (Suc 0, Bk # Bk\<^bsup>lna\<^esup> @ <ab # rev list @ [aa]> @ Bk # Bk # ires, |
2610 "steps0 (Suc 0, Bk # Bk\<up>(lna) @ <ab # rev list @ [aa]> @ Bk # Bk # ires, |
2595 Bk # Oc\<^bsup>Suc (2*rs + 2)\<^esup> @ Bk\<^bsup>rna\<^esup>) t_wcode_main stpb |
2611 Bk # Oc\<up>(Suc (2*rs + 2)) @ Bk\<up>(rna)) t_wcode_main stpb |
2596 = (0, Bk # ires, Bk # Oc # Bk\<^bsup>lnb\<^esup> @ Bk # |
2612 = (0, Bk # ires, Bk # Oc # Bk\<up>(lnb) @ Bk # |
2597 Bk # Oc\<^bsup>bl_bin (<aa # list @ [ab]> ) + (2*rs + 2)* 2^(length (<aa # list @ [ab]>) - Suc 0)\<^esup> @ Bk\<^bsup>rnb\<^esup>)" |
2613 Bk # Oc\<up>(bl_bin (<aa # list @ [ab]> ) + (2*rs + 2)* 2^(length (<aa # list @ [ab]>) - Suc 0)) @ Bk\<up>(rnb))" |
2598 by blast |
2614 by blast |
2599 from stp1 and stp2 show |
2615 from stp1 and stp2 show |
2600 "\<exists>stp ln rn. |
2616 "\<exists>stp ln rn. |
2601 steps (Suc 0, Bk # Bk\<^bsup>m\<^esup> @ Oc\<^bsup>Suc (Suc ab)\<^esup> @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, |
2617 steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc\<up>(Suc (Suc ab)) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires, |
2602 Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>) t_wcode_main stp = |
2618 Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp = |
2603 (0, Bk # ires, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # |
2619 (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # |
2604 Oc\<^bsup>bl_bin (Oc\<^bsup>Suc aa\<^esup> @ Bk # <list @ [Suc ab]>) + rs * (2 * 2 ^ (aa + length (<list @ [Suc ab]>)))\<^esup> |
2620 Oc\<up>(bl_bin (Oc\<up>(Suc aa) @ Bk # <list @ [Suc ab]>) + rs * (2 * 2 ^ (aa + length (<list @ [Suc ab]>)))) |
2605 @ Bk\<^bsup>rn\<^esup>)" |
2621 @ Bk\<up>(rn))" |
2606 apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI, |
2622 apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI, |
2607 rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_cons_app1 exp_ind_def) |
2623 rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_cons_app1 replicate_Suc) |
2608 done |
2624 done |
2609 qed |
2625 qed |
2610 qed |
2626 qed |
2611 qed |
2627 qed |
2612 qed |
2628 qed |
2613 |
2629 |
2614 |
2630 |
2615 |
2631 definition t_wcode_prepare :: "instr list" |
2616 (* turing_shift can be used*) |
|
2617 term t_wcode_main |
|
2618 definition t_wcode_prepare :: "tprog" |
|
2619 where |
2632 where |
2620 "t_wcode_prepare \<equiv> |
2633 "t_wcode_prepare \<equiv> |
2621 [(W1, 2), (L, 1), (L, 3), (R, 2), (R, 4), (W0, 3), |
2634 [(W1, 2), (L, 1), (L, 3), (R, 2), (R, 4), (W0, 3), |
2622 (R, 4), (R, 5), (R, 6), (R, 5), (R, 7), (R, 5), |
2635 (R, 4), (R, 5), (R, 6), (R, 5), (R, 7), (R, 5), |
2623 (W1, 7), (L, 0)]" |
2636 (W1, 7), (L, 0)]" |
2624 |
2637 |
2625 fun wprepare_add_one :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2638 fun wprepare_add_one :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2626 where |
2639 where |
2627 "wprepare_add_one m lm (l, r) = |
2640 "wprepare_add_one m lm (l, r) = |
2628 (\<exists> rn. l = [] \<and> |
2641 (\<exists> rn. l = [] \<and> |
2629 (r = <m # lm> @ Bk\<^bsup>rn\<^esup> \<or> |
2642 (r = <m # lm> @ Bk\<up>(rn) \<or> |
2630 r = Bk # <m # lm> @ Bk\<^bsup>rn\<^esup>))" |
2643 r = Bk # <m # lm> @ Bk\<up>(rn)))" |
2631 |
2644 |
2632 fun wprepare_goto_first_end :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2645 fun wprepare_goto_first_end :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2633 where |
2646 where |
2634 "wprepare_goto_first_end m lm (l, r) = |
2647 "wprepare_goto_first_end m lm (l, r) = |
2635 (\<exists> ml mr rn. l = Oc\<^bsup>ml\<^esup> \<and> |
2648 (\<exists> ml mr rn. l = Oc\<up>(ml) \<and> |
2636 r = Oc\<^bsup>mr\<^esup> @ Bk # <lm> @ Bk\<^bsup>rn\<^esup> \<and> |
2649 r = Oc\<up>(mr) @ Bk # <lm> @ Bk\<up>(rn) \<and> |
2637 ml + mr = Suc (Suc m))" |
2650 ml + mr = Suc (Suc m))" |
2638 |
2651 |
2639 fun wprepare_erase :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2652 fun wprepare_erase :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2640 where |
2653 where |
2641 "wprepare_erase m lm (l, r) = |
2654 "wprepare_erase m lm (l, r) = |
2642 (\<exists> rn. l = Oc\<^bsup>Suc m\<^esup> \<and> |
2655 (\<exists> rn. l = Oc\<up>(Suc m) \<and> |
2643 tl r = Bk # <lm> @ Bk\<^bsup>rn\<^esup>)" |
2656 tl r = Bk # <lm> @ Bk\<up>(rn))" |
2644 |
2657 |
2645 fun wprepare_goto_start_pos_B :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2658 fun wprepare_goto_start_pos_B :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2646 where |
2659 where |
2647 "wprepare_goto_start_pos_B m lm (l, r) = |
2660 "wprepare_goto_start_pos_B m lm (l, r) = |
2648 (\<exists> rn. l = Bk # Oc\<^bsup>Suc m\<^esup> \<and> |
2661 (\<exists> rn. l = Bk # Oc\<up>(Suc m) \<and> |
2649 r = Bk # <lm> @ Bk\<^bsup>rn\<^esup>)" |
2662 r = Bk # <lm> @ Bk\<up>(rn))" |
2650 |
2663 |
2651 fun wprepare_goto_start_pos_O :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2664 fun wprepare_goto_start_pos_O :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2652 where |
2665 where |
2653 "wprepare_goto_start_pos_O m lm (l, r) = |
2666 "wprepare_goto_start_pos_O m lm (l, r) = |
2654 (\<exists> rn. l = Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and> |
2667 (\<exists> rn. l = Bk # Bk # Oc\<up>(Suc m) \<and> |
2655 r = <lm> @ Bk\<^bsup>rn\<^esup>)" |
2668 r = <lm> @ Bk\<up>(rn))" |
2656 |
2669 |
2657 fun wprepare_goto_start_pos :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2670 fun wprepare_goto_start_pos :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2658 where |
2671 where |
2659 "wprepare_goto_start_pos m lm (l, r) = |
2672 "wprepare_goto_start_pos m lm (l, r) = |
2660 (wprepare_goto_start_pos_B m lm (l, r) \<or> |
2673 (wprepare_goto_start_pos_B m lm (l, r) \<or> |
2661 wprepare_goto_start_pos_O m lm (l, r))" |
2674 wprepare_goto_start_pos_O m lm (l, r))" |
2662 |
2675 |
2663 fun wprepare_loop_start_on_rightmost :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2676 fun wprepare_loop_start_on_rightmost :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2664 where |
2677 where |
2665 "wprepare_loop_start_on_rightmost m lm (l, r) = |
2678 "wprepare_loop_start_on_rightmost m lm (l, r) = |
2666 (\<exists> rn mr. rev l @ r = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm> @ Bk\<^bsup>rn\<^esup> \<and> l \<noteq> [] \<and> |
2679 (\<exists> rn mr. rev l @ r = Oc\<up>(Suc m) @ Bk # Bk # <lm> @ Bk\<up>(rn) \<and> l \<noteq> [] \<and> |
2667 r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
2680 r = Oc\<up>(mr) @ Bk\<up>(rn))" |
2668 |
2681 |
2669 fun wprepare_loop_start_in_middle :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2682 fun wprepare_loop_start_in_middle :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2670 where |
2683 where |
2671 "wprepare_loop_start_in_middle m lm (l, r) = |
2684 "wprepare_loop_start_in_middle m lm (l, r) = |
2672 (\<exists> rn (mr:: nat) (lm1::nat list). |
2685 (\<exists> rn (mr:: nat) (lm1::nat list). |
2673 rev l @ r = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm> @ Bk\<^bsup>rn\<^esup> \<and> l \<noteq> [] \<and> |
2686 rev l @ r = Oc\<up>(Suc m) @ Bk # Bk # <lm> @ Bk\<up>(rn) \<and> l \<noteq> [] \<and> |
2674 r = Oc\<^bsup>mr\<^esup> @ Bk # <lm1> @ Bk\<^bsup>rn\<^esup> \<and> lm1 \<noteq> [])" |
2687 r = Oc\<up>(mr) @ Bk # <lm1> @ Bk\<up>(rn) \<and> lm1 \<noteq> [])" |
2675 |
2688 |
2676 fun wprepare_loop_start :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2689 fun wprepare_loop_start :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2677 where |
2690 where |
2678 "wprepare_loop_start m lm (l, r) = (wprepare_loop_start_on_rightmost m lm (l, r) \<or> |
2691 "wprepare_loop_start m lm (l, r) = (wprepare_loop_start_on_rightmost m lm (l, r) \<or> |
2679 wprepare_loop_start_in_middle m lm (l, r))" |
2692 wprepare_loop_start_in_middle m lm (l, r))" |
2680 |
2693 |
2681 fun wprepare_loop_goon_on_rightmost :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2694 fun wprepare_loop_goon_on_rightmost :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2682 where |
2695 where |
2683 "wprepare_loop_goon_on_rightmost m lm (l, r) = |
2696 "wprepare_loop_goon_on_rightmost m lm (l, r) = |
2684 (\<exists> rn. l = Bk # <rev lm> @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and> |
2697 (\<exists> rn. l = Bk # <rev lm> @ Bk # Bk # Oc\<up>(Suc m) \<and> |
2685 r = Bk\<^bsup>rn\<^esup>)" |
2698 r = Bk\<up>(rn))" |
2686 |
2699 |
2687 fun wprepare_loop_goon_in_middle :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2700 fun wprepare_loop_goon_in_middle :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2688 where |
2701 where |
2689 "wprepare_loop_goon_in_middle m lm (l, r) = |
2702 "wprepare_loop_goon_in_middle m lm (l, r) = |
2690 (\<exists> rn (mr:: nat) (lm1::nat list). |
2703 (\<exists> rn (mr:: nat) (lm1::nat list). |
2691 rev l @ r = Oc\<^bsup>Suc m\<^esup> @ Bk # Bk # <lm> @ Bk\<^bsup>rn\<^esup> \<and> l \<noteq> [] \<and> |
2704 rev l @ r = Oc\<up>(Suc m) @ Bk # Bk # <lm> @ Bk\<up>(rn) \<and> l \<noteq> [] \<and> |
2692 (if lm1 = [] then r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> |
2705 (if lm1 = [] then r = Oc\<up>(mr) @ Bk\<up>(rn) |
2693 else r = Oc\<^bsup>mr\<^esup> @ Bk # <lm1> @ Bk\<^bsup>rn\<^esup>) \<and> mr > 0)" |
2706 else r = Oc\<up>(mr) @ Bk # <lm1> @ Bk\<up>(rn)) \<and> mr > 0)" |
2694 |
2707 |
2695 fun wprepare_loop_goon :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2708 fun wprepare_loop_goon :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2696 where |
2709 where |
2697 "wprepare_loop_goon m lm (l, r) = |
2710 "wprepare_loop_goon m lm (l, r) = |
2698 (wprepare_loop_goon_in_middle m lm (l, r) \<or> |
2711 (wprepare_loop_goon_in_middle m lm (l, r) \<or> |
2699 wprepare_loop_goon_on_rightmost m lm (l, r))" |
2712 wprepare_loop_goon_on_rightmost m lm (l, r))" |
2700 |
2713 |
2701 fun wprepare_add_one2 :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2714 fun wprepare_add_one2 :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2702 where |
2715 where |
2703 "wprepare_add_one2 m lm (l, r) = |
2716 "wprepare_add_one2 m lm (l, r) = |
2704 (\<exists> rn. l = Bk # Bk # <rev lm> @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and> |
2717 (\<exists> rn. l = Bk # Bk # <rev lm> @ Bk # Bk # Oc\<up>(Suc m) \<and> |
2705 (r = [] \<or> tl r = Bk\<^bsup>rn\<^esup>))" |
2718 (r = [] \<or> tl r = Bk\<up>(rn)))" |
2706 |
2719 |
2707 fun wprepare_stop :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2720 fun wprepare_stop :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2708 where |
2721 where |
2709 "wprepare_stop m lm (l, r) = |
2722 "wprepare_stop m lm (l, r) = |
2710 (\<exists> rn. l = Bk # <rev lm> @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and> |
2723 (\<exists> rn. l = Bk # <rev lm> @ Bk # Bk # Oc\<up>(Suc m) \<and> |
2711 r = Bk # Oc # Bk\<^bsup>rn\<^esup>)" |
2724 r = Bk # Oc # Bk\<up>(rn))" |
2712 |
2725 |
2713 fun wprepare_inv :: "nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2726 fun wprepare_inv :: "nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool" |
2714 where |
2727 where |
2715 "wprepare_inv st m lm (l, r) = |
2728 "wprepare_inv st m lm (l, r) = |
2716 (if st = 0 then wprepare_stop m lm (l, r) |
2729 (if st = 0 then wprepare_stop m lm (l, r) |
2806 lemma [simp]: "fetch t_wcode_prepare (Suc (Suc (Suc 0))) Oc = (W0, 3)" |
2819 lemma [simp]: "fetch t_wcode_prepare (Suc (Suc (Suc 0))) Oc = (W0, 3)" |
2807 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) |
2820 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) |
2808 done |
2821 done |
2809 |
2822 |
2810 lemma [simp]: "fetch t_wcode_prepare 4 Bk = (R, 4)" |
2823 lemma [simp]: "fetch t_wcode_prepare 4 Bk = (R, 4)" |
2811 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) |
2824 apply(subgoal_tac "4 = Suc 3") |
|
2825 apply(simp_all only: fetch.simps t_wcode_prepare_def nth_of.simps) |
|
2826 apply(auto) |
2812 done |
2827 done |
2813 |
2828 |
2814 lemma [simp]: "fetch t_wcode_prepare 4 Oc = (R, 5)" |
2829 lemma [simp]: "fetch t_wcode_prepare 4 Oc = (R, 5)" |
2815 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) |
2830 apply(subgoal_tac "4 = Suc 3") |
2816 done |
2831 apply(simp_all only: fetch.simps t_wcode_prepare_def nth_of.simps) |
|
2832 apply(auto) |
|
2833 done |
|
2834 |
2817 |
2835 |
2818 lemma [simp]: "fetch t_wcode_prepare 5 Oc = (R, 5)" |
2836 lemma [simp]: "fetch t_wcode_prepare 5 Oc = (R, 5)" |
2819 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) |
2837 apply(subgoal_tac "5 = Suc 4") |
|
2838 apply(simp_all only: fetch.simps t_wcode_prepare_def nth_of.simps) |
|
2839 apply(auto) |
2820 done |
2840 done |
2821 |
2841 |
2822 lemma [simp]: "fetch t_wcode_prepare 5 Bk = (R, 6)" |
2842 lemma [simp]: "fetch t_wcode_prepare 5 Bk = (R, 6)" |
2823 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) |
2843 apply(subgoal_tac "5 = Suc 4") |
|
2844 apply(simp_all only: fetch.simps t_wcode_prepare_def nth_of.simps) |
|
2845 apply(auto) |
2824 done |
2846 done |
2825 |
2847 |
2826 lemma [simp]: "fetch t_wcode_prepare 6 Oc = (R, 5)" |
2848 lemma [simp]: "fetch t_wcode_prepare 6 Oc = (R, 5)" |
2827 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) |
2849 apply(subgoal_tac "6 = Suc 5") |
|
2850 apply(simp_all only: fetch.simps t_wcode_prepare_def nth_of.simps) |
|
2851 apply(auto) |
2828 done |
2852 done |
2829 |
2853 |
2830 lemma [simp]: "fetch t_wcode_prepare 6 Bk = (R, 7)" |
2854 lemma [simp]: "fetch t_wcode_prepare 6 Bk = (R, 7)" |
2831 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) |
2855 apply(subgoal_tac "6 = Suc 5") |
|
2856 apply(simp_all only: fetch.simps t_wcode_prepare_def nth_of.simps) |
|
2857 apply(auto) |
2832 done |
2858 done |
2833 |
2859 |
2834 lemma [simp]: "fetch t_wcode_prepare 7 Oc = (L, 0)" |
2860 lemma [simp]: "fetch t_wcode_prepare 7 Oc = (L, 0)" |
2835 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) |
2861 apply(subgoal_tac "7 = Suc 6") |
|
2862 apply(simp_all only: fetch.simps t_wcode_prepare_def nth_of.simps) |
|
2863 apply(auto) |
2836 done |
2864 done |
2837 |
2865 |
2838 lemma [simp]: "fetch t_wcode_prepare 7 Bk = (W1, 7)" |
2866 lemma [simp]: "fetch t_wcode_prepare 7 Bk = (W1, 7)" |
2839 apply(simp add: fetch.simps t_wcode_prepare_def nth_of.simps) |
2867 apply(subgoal_tac "7 = Suc 6") |
2840 done |
2868 apply(simp_all only: fetch.simps t_wcode_prepare_def nth_of.simps) |
2841 |
2869 apply(auto) |
2842 lemma tape_of_nl_not_null: "lm \<noteq> [] \<Longrightarrow> <lm::nat list> \<noteq> []" |
|
2843 apply(case_tac lm, auto) |
|
2844 apply(case_tac list, auto simp: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def) |
|
2845 done |
2870 done |
2846 |
2871 |
2847 lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_add_one m lm (b, []) = False" |
2872 lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_add_one m lm (b, []) = False" |
2848 apply(simp add: wprepare_invs) |
2873 apply(simp add: wprepare_invs) |
2849 apply(simp add: tape_of_nl_not_null) |
|
2850 done |
2874 done |
2851 |
2875 |
2852 lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_goto_first_end m lm (b, []) = False" |
2876 lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_goto_first_end m lm (b, []) = False" |
2853 apply(simp add: wprepare_invs) |
2877 apply(simp add: wprepare_invs) |
2854 done |
2878 done |
2855 |
2879 |
2856 lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_erase m lm (b, []) = False" |
2880 lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_erase m lm (b, []) = False" |
2857 apply(simp add: wprepare_invs) |
2881 apply(simp add: wprepare_invs) |
2858 done |
2882 done |
2859 |
2883 |
2860 |
|
2861 |
|
2862 lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_goto_start_pos m lm (b, []) = False" |
2884 lemma [simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_goto_start_pos m lm (b, []) = False" |
2863 apply(simp add: wprepare_invs tape_of_nl_not_null) |
2885 apply(simp add: wprepare_invs) |
2864 done |
2886 done |
2865 |
2887 |
2866 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [])\<rbrakk> \<Longrightarrow> b \<noteq> []" |
2888 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [])\<rbrakk> \<Longrightarrow> b \<noteq> []" |
2867 apply(simp add: wprepare_invs tape_of_nl_not_null, auto) |
2889 apply(simp add: wprepare_invs) |
2868 done |
2890 done |
|
2891 |
|
2892 lemma rev_eq: "rev xs = rev ys \<Longrightarrow> xs = ys" |
|
2893 by auto |
2869 |
2894 |
2870 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [])\<rbrakk> \<Longrightarrow> |
2895 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [])\<rbrakk> \<Longrightarrow> |
2871 wprepare_loop_goon m lm (Bk # b, [])" |
2896 wprepare_loop_goon m lm (Bk # b, [])" |
2872 apply(simp only: wprepare_invs tape_of_nl_not_null) |
2897 apply(simp only: wprepare_invs) |
2873 apply(erule_tac disjE) |
2898 apply(erule_tac disjE) |
2874 apply(rule_tac disjI2) |
2899 apply(rule_tac disjI2) |
2875 apply(simp add: wprepare_loop_start_on_rightmost.simps |
2900 apply(simp add: wprepare_loop_start_on_rightmost.simps |
2876 wprepare_loop_goon_on_rightmost.simps, auto) |
2901 wprepare_loop_goon_on_rightmost.simps, auto) |
2877 apply(rule_tac rev_eq, simp add: tape_of_nl_rev) |
2902 apply(rule_tac rev_eq, simp add: tape_of_nl_rev) |
2878 done |
2903 done |
2879 |
2904 |
2880 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, [])\<rbrakk> \<Longrightarrow> b \<noteq> []" |
2905 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, [])\<rbrakk> \<Longrightarrow> b \<noteq> []" |
2881 apply(simp only: wprepare_invs tape_of_nl_not_null, auto) |
2906 apply(simp only: wprepare_invs, auto) |
2882 done |
2907 done |
2883 |
2908 |
2884 lemma [simp]:"\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, [])\<rbrakk> \<Longrightarrow> |
2909 lemma [simp]:"\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, [])\<rbrakk> \<Longrightarrow> |
2885 wprepare_add_one2 m lm (Bk # b, [])" |
2910 wprepare_add_one2 m lm (Bk # b, [])" |
2886 apply(simp only: wprepare_invs tape_of_nl_not_null, auto split: if_splits) |
2911 apply(simp only: wprepare_invs, auto split: if_splits) |
2887 apply(case_tac mr, simp, simp add: exp_ind_def) |
|
2888 done |
2912 done |
2889 |
2913 |
2890 lemma [simp]: "wprepare_add_one2 m lm (b, []) \<Longrightarrow> b \<noteq> []" |
2914 lemma [simp]: "wprepare_add_one2 m lm (b, []) \<Longrightarrow> b \<noteq> []" |
2891 apply(simp only: wprepare_invs tape_of_nl_not_null, auto) |
2915 apply(simp only: wprepare_invs, auto) |
2892 done |
2916 done |
2893 |
2917 |
2894 lemma [simp]: "wprepare_add_one2 m lm (b, []) \<Longrightarrow> wprepare_add_one2 m lm (b, [Oc])" |
2918 lemma [simp]: "wprepare_add_one2 m lm (b, []) \<Longrightarrow> wprepare_add_one2 m lm (b, [Oc])" |
2895 apply(simp only: wprepare_invs tape_of_nl_not_null, auto) |
2919 apply(simp only: wprepare_invs, auto) |
2896 done |
2920 done |
2897 |
2921 |
2898 lemma [simp]: "Bk # list = <(m::nat) # lm> @ ys = False" |
2922 lemma [simp]: "Bk # list = <(m::nat) # lm> @ ys = False" |
2899 apply(case_tac lm, auto simp: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def) |
2923 apply(case_tac lm, auto simp: tape_of_nl_cons replicate_Suc) |
2900 done |
2924 done |
2901 |
2925 |
2902 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_add_one m lm (b, Bk # list)\<rbrakk> |
2926 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_add_one m lm (b, Bk # list)\<rbrakk> |
2903 \<Longrightarrow> (b = [] \<longrightarrow> wprepare_goto_first_end m lm ([], Oc # list)) \<and> |
2927 \<Longrightarrow> (b = [] \<longrightarrow> wprepare_goto_first_end m lm ([], Oc # list)) \<and> |
2904 (b \<noteq> [] \<longrightarrow> wprepare_goto_first_end m lm (b, Oc # list))" |
2928 (b \<noteq> [] \<longrightarrow> wprepare_goto_first_end m lm (b, Oc # list))" |
2905 apply(simp only: wprepare_invs, auto) |
2929 apply(simp only: wprepare_invs) |
2906 apply(rule_tac x = 0 in exI, simp add: exp_ind_def) |
2930 apply(auto simp: tape_of_nl_cons split: if_splits) |
2907 apply(case_tac lm, simp, simp add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def) |
2931 apply(rule_tac x = 0 in exI, simp add: replicate_Suc) |
2908 apply(rule_tac x = rn in exI, simp) |
2932 apply(case_tac lm, simp, simp add: tape_of_nl_abv tape_of_nat_list.simps replicate_Suc) |
2909 done |
2933 done |
2910 |
2934 |
2911 lemma [simp]: "wprepare_goto_first_end m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []" |
2935 lemma [simp]: "wprepare_goto_first_end m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []" |
2912 apply(simp only: wprepare_invs tape_of_nl_not_null, auto) |
2936 apply(simp only: wprepare_invs , auto simp: replicate_Suc) |
2913 apply(case_tac mr, simp_all add: exp_ind_def) |
2937 done |
2914 done |
2938 |
|
2939 declare replicate_Suc[simp] |
2915 |
2940 |
2916 lemma [simp]: "wprepare_goto_first_end m lm (b, Bk # list) \<Longrightarrow> |
2941 lemma [simp]: "wprepare_goto_first_end m lm (b, Bk # list) \<Longrightarrow> |
2917 wprepare_erase m lm (tl b, hd b # Bk # list)" |
2942 wprepare_erase m lm (tl b, hd b # Bk # list)" |
2918 apply(simp only: wprepare_invs tape_of_nl_not_null, auto) |
2943 apply(simp only: wprepare_invs, auto) |
2919 apply(case_tac mr, simp_all add: exp_ind_def) |
2944 apply(case_tac mr, simp_all) |
2920 apply(case_tac mr, auto simp: exp_ind_def) |
2945 apply(case_tac mr, auto) |
2921 done |
2946 done |
2922 |
2947 |
2923 lemma [simp]: "wprepare_erase m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []" |
2948 lemma [simp]: "wprepare_erase m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []" |
2924 apply(simp only: wprepare_invs exp_ind_def, auto) |
2949 apply(simp only: wprepare_invs, auto) |
2925 done |
2950 done |
2926 |
2951 |
2927 lemma [simp]: "wprepare_erase m lm (b, Bk # list) \<Longrightarrow> |
2952 lemma [simp]: "wprepare_erase m lm (b, Bk # list) \<Longrightarrow> |
2928 wprepare_goto_start_pos m lm (Bk # b, list)" |
2953 wprepare_goto_start_pos m lm (Bk # b, list)" |
2929 apply(simp only: wprepare_invs, auto) |
2954 apply(simp only: wprepare_invs, auto) |
2930 done |
2955 done |
2931 |
2956 |
2932 lemma [simp]: "\<lbrakk>wprepare_add_one m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []" |
2957 lemma [simp]: "\<lbrakk>wprepare_add_one m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []" |
2933 apply(simp only: wprepare_invs) |
2958 apply(simp only: wprepare_invs) |
2934 apply(case_tac lm, simp_all add: tape_of_nl_abv |
2959 apply(case_tac lm, simp_all add: tape_of_nl_abv |
2935 tape_of_nat_list.simps exp_ind_def, auto) |
2960 tape_of_nat_list.simps, auto) |
2936 done |
2961 done |
2937 |
2962 |
2938 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_first_end m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []" |
2963 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_first_end m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []" |
2939 apply(simp only: wprepare_invs, auto) |
2964 apply(simp only: wprepare_invs, auto) |
2940 apply(case_tac mr, simp_all add: exp_ind_def) |
2965 apply(case_tac mr, simp_all) |
2941 apply(simp add: tape_of_nl_not_null) |
|
2942 done |
2966 done |
2943 |
2967 |
2944 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_first_end m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []" |
2968 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_first_end m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []" |
2945 apply(simp only: wprepare_invs, auto) |
2969 apply(simp only: wprepare_invs, auto) |
2946 apply(case_tac mr, simp_all add: exp_ind_def tape_of_nl_not_null) |
|
2947 done |
2970 done |
2948 |
2971 |
2949 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_erase m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []" |
2972 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_erase m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []" |
2950 apply(simp only: wprepare_invs, auto) |
2973 apply(simp only: wprepare_invs, auto) |
2951 done |
2974 done |
2952 |
2975 |
2953 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_erase m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []" |
2976 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_erase m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []" |
2954 apply(simp only: wprepare_invs, auto simp: exp_ind_def) |
2977 apply(simp only: wprepare_invs, auto) |
2955 done |
2978 done |
2956 |
2979 |
2957 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []" |
2980 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []" |
2958 apply(simp only: wprepare_invs, auto) |
2981 apply(simp only: wprepare_invs, auto) |
2959 apply(simp add: tape_of_nl_not_null) |
|
2960 apply(case_tac lm, simp, case_tac list) |
2982 apply(case_tac lm, simp, case_tac list) |
2961 apply(simp_all add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def) |
2983 apply(simp_all add: tape_of_nl_abv tape_of_nat_list.simps) |
2962 done |
2984 done |
2963 |
2985 |
2964 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []" |
2986 lemma [simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []" |
2965 apply(simp only: wprepare_invs) |
2987 apply(simp only: wprepare_invs) |
2966 apply(auto) |
2988 apply(auto) |
3323 lemma [intro]: "t_correct (tm_of abc_fourtimes @ tMp (Suc 0) |
3332 lemma [intro]: "t_correct (tm_of abc_fourtimes @ tMp (Suc 0) |
3324 (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))" |
3333 (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))" |
3325 apply(rule_tac t_compiled_correct, simp_all) |
3334 apply(rule_tac t_compiled_correct, simp_all) |
3326 apply(simp add: fourtimes_ly_def) |
3335 apply(simp add: fourtimes_ly_def) |
3327 done |
3336 done |
3328 |
3337 *) |
3329 |
3338 |
3330 lemma [intro]: "t_correct t_wcode_main" |
3339 lemma [intro]: "(28 + (length t_twice_compile + length t_fourtimes_compile)) mod 2 = 0" |
3331 apply(auto simp: t_wcode_main_def t_correct.simps shift_length |
3340 apply(auto simp: t_twice_compile_def t_fourtimes_compile_def) |
3332 t_twice_def t_fourtimes_def) |
3341 by arith |
|
3342 |
|
3343 lemma [elim]: "(a, b) \<in> set t_wcode_main_first_part \<Longrightarrow> |
|
3344 b \<le> (28 + (length t_twice_compile + length t_fourtimes_compile)) div 2" |
|
3345 apply(auto simp: t_wcode_main_first_part_def t_twice_def) |
|
3346 done |
|
3347 |
|
3348 |
|
3349 |
|
3350 lemma tm_wf_change_termi: "tm_wf (tp, 0) \<Longrightarrow> |
|
3351 list_all (\<lambda>(acn, st). (st \<le> Suc (length tp div 2))) (adjust tp)" |
|
3352 apply(auto simp: tm_wf.simps List.list_all_length) |
|
3353 apply(case_tac "tp!n", auto simp: adjust.simps split: if_splits) |
|
3354 apply(erule_tac x = "(a, b)" in ballE, auto) |
|
3355 by (metis in_set_conv_nth) |
|
3356 |
|
3357 lemma tm_wf_shift: |
|
3358 "list_all (\<lambda>(acn, st). (st \<le> y)) tp \<Longrightarrow> |
|
3359 list_all (\<lambda>(acn, st). (st \<le> y + off)) (shift tp off) " |
|
3360 apply(auto simp: tm_wf.simps List.list_all_length) |
|
3361 apply(erule_tac x = n in allE, simp add: length_shift) |
|
3362 apply(case_tac "tp!n", auto simp: shift.simps) |
|
3363 done |
|
3364 |
|
3365 declare length_tp'[simp del] |
|
3366 |
|
3367 lemma [simp]: "length (mopup (Suc 0)) = 16" |
|
3368 apply(auto simp: mopup.simps) |
|
3369 done |
|
3370 |
|
3371 lemma [elim]: "(a, b) \<in> set (shift (turing_basic.adjust t_twice_compile) 12) \<Longrightarrow> |
|
3372 b \<le> (28 + (length t_twice_compile + length t_fourtimes_compile)) div 2" |
|
3373 apply(simp add: t_twice_compile_def t_fourtimes_compile_def) |
3333 proof - |
3374 proof - |
3334 show "iseven (60 + (length (tm_of abc_twice) + |
3375 assume g: "(a, b) \<in> set (shift (turing_basic.adjust (tm_of abc_twice @ shift (mopup (Suc 0)) (length (tm_of abc_twice) div 2))) 12)" |
3335 length (tm_of abc_fourtimes)))" |
3376 moreover have "length (tm_of abc_twice) mod 2 = 0" by auto |
3336 using twice_len_even fourtimes_len_even |
3377 moreover have "length (tm_of abc_fourtimes) mod 2 = 0" by auto |
3337 apply(auto simp: iseven_def) |
3378 ultimately have "list_all (\<lambda>(acn, st). (st \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2)) |
3338 apply(rule_tac x = "30 + q + qa" in exI, simp) |
3379 (shift (turing_basic.adjust t_twice_compile) 12)" |
|
3380 proof(auto simp: mod_ex1) |
|
3381 fix q qa |
|
3382 assume h: "length (tm_of abc_twice) = 2 * q" "length (tm_of abc_fourtimes) = 2 * qa" |
|
3383 hence "list_all (\<lambda>(acn, st). st \<le> (18 + (q + qa)) + 12) (shift (turing_basic.adjust t_twice_compile) 12)" |
|
3384 proof(rule_tac tm_wf_shift t_twice_compile_def) |
|
3385 have "list_all (\<lambda>(acn, st). st \<le> Suc (length t_twice_compile div 2)) (adjust t_twice_compile)" |
|
3386 by(rule_tac tm_wf_change_termi, auto) |
|
3387 thus "list_all (\<lambda>(acn, st). st \<le> 18 + (q + qa)) (turing_basic.adjust t_twice_compile)" |
|
3388 using h |
|
3389 apply(simp add: t_twice_compile_def, auto simp: List.list_all_length) |
|
3390 done |
|
3391 qed |
|
3392 thus "list_all (\<lambda>(acn, st). st \<le> 30 + (q + qa)) (shift (turing_basic.adjust t_twice_compile) 12)" |
|
3393 by simp |
|
3394 qed |
|
3395 thus "b \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2" |
|
3396 using g |
|
3397 apply(auto simp:t_twice_compile_def) |
|
3398 apply(simp add: Ball_set[THEN sym]) |
|
3399 apply(erule_tac x = "(a, b)" in ballE, simp, simp) |
3339 done |
3400 done |
3340 next |
3401 qed |
3341 show " list_all (\<lambda>(acn, s). s \<le> (60 + (length (tm_of abc_twice) + |
3402 |
3342 length (tm_of abc_fourtimes))) div 2) t_wcode_main_first_part" |
3403 lemma [elim]: "(a, b) \<in> set (shift (turing_basic.adjust t_fourtimes_compile) (t_twice_len + 13)) |
3343 apply(auto simp: t_wcode_main_first_part_def shift_length t_twice_def) |
3404 \<Longrightarrow> b \<le> (28 + (length t_twice_compile + length t_fourtimes_compile)) div 2" |
3344 done |
3405 apply(simp add: t_twice_compile_def t_fourtimes_compile_def t_twice_len_def) |
3345 next |
3406 proof - |
3346 have "list_all (\<lambda>(acn, s). s \<le> Suc (length (tm_of abc_twice @ tMp (Suc 0) |
3407 assume g: "(a, b) \<in> set (shift (turing_basic.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) |
3347 (start_of twice_ly (length abc_twice) - Suc 0)) div 2)) |
3408 (length (tm_of abc_fourtimes) div 2))) (length t_twice div 2 + 13))" |
3348 (change_termi_state (tm_of abc_twice @ tMp (Suc 0) |
3409 moreover have "length (tm_of abc_twice) mod 2 = 0" by auto |
3349 (start_of twice_ly (length abc_twice) - Suc 0)))" |
3410 moreover have "length (tm_of abc_fourtimes) mod 2 = 0" by auto |
3350 apply(rule_tac t_correct_termi, auto) |
3411 ultimately have "list_all (\<lambda>(acn, st). (st \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2)) |
3351 done |
3412 (shift (turing_basic.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) |
3352 hence "list_all (\<lambda>(acn, s). s \<le> Suc (length (tm_of abc_twice @ tMp (Suc 0) |
3413 (length (tm_of abc_fourtimes) div 2))) (length t_twice div 2 + 13))" |
3353 (start_of twice_ly (length abc_twice) - Suc 0)) div 2) + 12) |
3414 proof(auto simp: mod_ex1 t_twice_def t_twice_compile_def) |
3354 (tshift (change_termi_state (tm_of abc_twice @ tMp (Suc 0) |
3415 fix q qa |
3355 (start_of twice_ly (length abc_twice) - Suc 0))) 12)" |
3416 assume h: "length (tm_of abc_twice) = 2 * q" "length (tm_of abc_fourtimes) = 2 * qa" |
3356 apply(rule_tac t_correct_shift, simp) |
3417 hence "list_all (\<lambda>(acn, st). st \<le> (9 + qa + (21 + q))) |
3357 done |
3418 (shift (turing_basic.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa)) (21 + q))" |
3358 thus "list_all (\<lambda>(acn, s). s \<le> |
3419 proof(rule_tac tm_wf_shift t_twice_compile_def) |
3359 (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2) |
3420 have "list_all (\<lambda>(acn, st). st \<le> Suc (length (tm_of abc_fourtimes @ shift |
3360 (tshift (change_termi_state (tm_of abc_twice @ tMp (Suc 0) |
3421 (mopup (Suc 0)) qa) div 2)) (adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa))" |
3361 (start_of twice_ly (length abc_twice) - Suc 0))) 12)" |
3422 apply(rule_tac tm_wf_change_termi) |
3362 apply(simp) |
3423 using wf_fourtimes h |
3363 apply(simp add: list_all_length, auto) |
3424 apply(simp add: t_fourtimes_compile_def) |
3364 done |
3425 done |
3365 next |
3426 thus "list_all (\<lambda>(acn, st). st \<le> 9 + qa) ((turing_basic.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa)))" |
3366 have "list_all (\<lambda>(acn, s). s \<le> Suc (length (tm_of abc_fourtimes @ tMp (Suc 0) |
3427 using h |
3367 (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) div 2)) |
3428 apply(simp) |
3368 (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) |
3429 done |
3369 (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) " |
3430 qed |
3370 apply(rule_tac t_correct_termi, auto) |
3431 thus "list_all (\<lambda>(acn, st). st \<le> 30 + (q + qa)) (shift (turing_basic.adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa)) (21 + q))" |
3371 done |
3432 apply(subgoal_tac "qa + q = q + qa") |
3372 hence "list_all (\<lambda>(acn, s). s \<le> Suc (length (tm_of abc_fourtimes @ tMp (Suc 0) |
3433 apply(simp, simp) |
3373 (start_of fourtimes_ly (length abc_fourtimes) - Suc 0)) div 2) + (t_twice_len + 13)) |
3434 done |
3374 (tshift (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) |
3435 qed |
3375 (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) (t_twice_len + 13))" |
3436 thus "b \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2" |
3376 apply(rule_tac t_correct_shift, simp) |
3437 using g |
3377 done |
3438 apply(simp add: Ball_set[THEN sym]) |
3378 thus "list_all (\<lambda>(acn, s). s \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2) |
3439 apply(erule_tac x = "(a, b)" in ballE, simp, simp) |
3379 (tshift (change_termi_state (tm_of abc_fourtimes @ tMp (Suc 0) |
|
3380 (start_of fourtimes_ly (length abc_fourtimes) - Suc 0))) (t_twice_len + 13))" |
|
3381 apply(simp add: t_twice_len_def t_twice_def) |
|
3382 using twice_len_even fourtimes_len_even |
|
3383 apply(auto simp: list_all_length) |
|
3384 done |
3440 done |
3385 qed |
3441 qed |
3386 |
3442 |
3387 lemma [intro]: "t_correct (t_wcode_prepare |+| t_wcode_main)" |
3443 lemma [intro]: "tm_wf (t_wcode_main, 0)" |
3388 apply(auto intro: t_correct_add) |
3444 apply(auto simp: t_wcode_main_def tm_wf.simps |
|
3445 t_twice_def t_fourtimes_def del: List.list_all_iff) |
|
3446 done |
|
3447 |
|
3448 declare tm_comp.simps[simp del] |
|
3449 lemma tm_wf_comp: "\<lbrakk>tm_wf (A, 0); tm_wf (B, 0)\<rbrakk> \<Longrightarrow> tm_wf (A |+| B, 0)" |
|
3450 apply(auto simp: tm_wf.simps shift.simps adjust.simps tm_comp_length |
|
3451 tm_comp.simps) |
|
3452 done |
|
3453 |
|
3454 lemma [intro]: "tm_wf (t_wcode_prepare |+| t_wcode_main, 0)" |
|
3455 apply(rule_tac tm_wf_comp, auto) |
3389 done |
3456 done |
3390 |
3457 |
3391 lemma prepare_mainpart_lemma: |
3458 lemma prepare_mainpart_lemma: |
3392 "args \<noteq> [] \<Longrightarrow> |
3459 "args \<noteq> [] \<Longrightarrow> |
3393 \<exists> stp ln rn. steps (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) stp |
3460 \<exists> stp ln rn. steps0 (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) stp |
3394 = (0, Bk # Oc\<^bsup>Suc m\<^esup>, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
3461 = (0, Bk # Oc\<up>(Suc m), Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin (<args>)) @ Bk\<up>(rn))" |
3395 proof - |
3462 proof - |
3396 let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <m # args>" |
3463 let ?P1 = "(\<lambda> (l, r). (l::cell list) = [] \<and> r = <m # args>)" |
3397 let ?Q1 = "\<lambda> (l, r). wprepare_stop m args (l, r)" |
3464 let ?Q1 = "(\<lambda> (l, r). wprepare_stop m args (l, r))" |
3398 let ?P2 = ?Q1 |
3465 let ?P2 = ?Q1 |
3399 let ?Q2 = "\<lambda> (l, r). (\<exists> ln rn. l = Bk # Oc\<^bsup>Suc m\<^esup> \<and> |
3466 let ?Q2 = "(\<lambda> (l, r). (\<exists> ln rn. l = Bk # Oc\<up>(Suc m) \<and> |
3400 r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
3467 r = Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin (<args>)) @ Bk\<up>(rn)))" |
3401 let ?P3 = "\<lambda> tp. False" |
3468 let ?P3 = "\<lambda> tp. False" |
3402 assume h: "args \<noteq> []" |
3469 assume h: "args \<noteq> []" |
3403 have "?P1 \<turnstile>-> \<lambda> tp. (\<exists> stp tp'. steps (Suc 0, tp) |
3470 have "{?P1} t_wcode_prepare |+| t_wcode_main {?Q2}" |
3404 (t_wcode_prepare |+| t_wcode_main) stp = (0, tp') \<and> ?Q2 tp')" |
3471 proof(rule_tac Hoare_plus_halt) |
3405 proof(rule_tac turing_merge.t_merge_halt[of t_wcode_prepare t_wcode_main ?P1 ?P2 ?P3 ?P3 ?Q1 ?Q2], |
3472 show "?Q1 \<mapsto> ?P2" |
3406 auto simp: turing_merge_def) |
3473 by(simp add: assert_imp_def) |
3407 show "\<exists>stp. case steps (Suc 0, [], <m # args>) t_wcode_prepare stp of (st, tp') |
|
3408 \<Rightarrow> st = 0 \<and> wprepare_stop m args tp'" |
|
3409 using wprepare_correctness[of args m] h |
|
3410 apply(simp, auto) |
|
3411 apply(rule_tac x = n in exI, simp add: wprepare_inv.simps) |
|
3412 done |
|
3413 next |
3474 next |
3414 fix a b |
3475 show "tm_wf (t_wcode_prepare, 0)" |
3415 assume "wprepare_stop m args (a, b)" |
3476 by auto |
3416 thus "\<exists>stp. case steps (Suc 0, a, b) t_wcode_main stp of |
3477 next |
3417 (st, tp') \<Rightarrow> (st = 0) \<and> (case tp' of (l, r) \<Rightarrow> l = Bk # Oc\<^bsup>Suc m\<^esup> \<and> |
3478 show "{?P1} t_wcode_prepare {?Q1}" |
3418 (\<exists>ln rn. r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>))" |
3479 proof(rule_tac HoareI, auto) |
3419 proof(simp only: wprepare_stop.simps, erule_tac exE) |
3480 show "\<exists>n. is_final (steps0 (Suc 0, [], <m # args>) t_wcode_prepare n) \<and> |
|
3481 wprepare_stop m args holds_for steps0 (Suc 0, [], <m # args>) t_wcode_prepare n" |
|
3482 using wprepare_correctness[of args m] h |
|
3483 apply(auto) |
|
3484 apply(rule_tac x = n in exI, simp add: wprepare_inv.simps) |
|
3485 done |
|
3486 qed |
|
3487 next |
|
3488 show "{?P2} t_wcode_main {?Q2}" |
|
3489 proof(rule_tac HoareI, auto) |
|
3490 fix l r |
|
3491 assume "wprepare_stop m args (l, r)" |
|
3492 thus "\<exists>n. is_final (steps0 (Suc 0, l, r) t_wcode_main n) \<and> |
|
3493 (\<lambda>(l, r). l = Bk # Oc # Oc \<up> m \<and> (\<exists>ln rn. r = Bk # Oc # Bk \<up> ln @ |
|
3494 Bk # Bk # Oc \<up> bl_bin (<args>) @ Bk \<up> rn)) holds_for steps0 (Suc 0, l, r) t_wcode_main n" |
|
3495 proof(auto simp: wprepare_stop.simps) |
3420 fix rn |
3496 fix rn |
3421 assume "a = Bk # <rev args> @ Bk # Bk # Oc\<^bsup>Suc m\<^esup> \<and> |
3497 show " \<exists>n. is_final (steps0 (Suc 0, Bk # <rev args> @ Bk # Bk # Oc # Oc \<up> m, Bk # Oc # Bk \<up> rn) t_wcode_main n) \<and> |
3422 b = Bk # Oc # Bk\<^bsup>rn\<^esup>" |
3498 (\<lambda>(l, r). l = Bk # Oc # Oc \<up> m \<and> |
3423 thus "?thesis" |
3499 (\<exists>ln rn. r = Bk # Oc # Bk \<up> ln @ |
3424 using t_wcode_main_lemma_pre[of "args" "<args>" 0 "Oc\<^bsup>Suc m\<^esup>" 0 rn] h |
3500 Bk # Bk # Oc \<up> bl_bin (<args>) @ |
3425 apply(simp) |
3501 Bk \<up> rn)) holds_for steps0 (Suc 0, Bk # <rev args> @ Bk # Bk # Oc # Oc \<up> m, Bk # Oc # Bk \<up> rn) t_wcode_main n" |
3426 apply(erule_tac exE)+ |
3502 using t_wcode_main_lemma_pre[of "args" "<args>" 0 "Oc\<up>(Suc m)" 0 rn] h |
3427 apply(rule_tac x = stp in exI, simp add: tape_of_nl_rev, auto) |
3503 apply(auto simp: tape_of_nl_rev) |
|
3504 apply(rule_tac x = stp in exI, auto) |
3428 done |
3505 done |
3429 qed |
3506 qed |
3430 next |
3507 qed |
3431 show "wprepare_stop m args \<turnstile>-> wprepare_stop m args" |
|
3432 by(simp add: t_imply_def) |
|
3433 qed |
3508 qed |
3434 thus "\<exists> stp ln rn. steps (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) stp |
3509 thus "?thesis" |
3435 = (0, Bk # Oc\<^bsup>Suc m\<^esup>, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
3510 apply(auto simp: Hoare_def) |
3436 apply(simp add: t_imply_def) |
3511 apply(rule_tac x = n in exI) |
3437 apply(erule_tac exE)+ |
3512 apply(case_tac "(steps0 (Suc 0, [], <m # args>) |
3438 apply(auto) |
3513 (turing_basic.adjust t_wcode_prepare @ shift t_wcode_main (length t_wcode_prepare div 2)) n)") |
|
3514 apply(auto simp: tm_comp.simps) |
3439 done |
3515 done |
3440 qed |
3516 qed |
3441 |
3517 |
3442 |
|
3443 lemma [simp]: "tinres r r' \<Longrightarrow> |
3518 lemma [simp]: "tinres r r' \<Longrightarrow> |
3444 fetch t ss (case r of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) = |
3519 fetch t ss (read r) = |
3445 fetch t ss (case r' of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)" |
3520 fetch t ss (read r')" |
3446 apply(simp add: fetch.simps, auto split: if_splits simp: tinres_def) |
3521 apply(simp add: fetch.simps, auto split: if_splits simp: tinres_def) |
3447 apply(case_tac [!] r', simp_all) |
3522 apply(case_tac [!] n, simp_all) |
3448 apply(case_tac [!] n, simp_all add: exp_ind_def) |
3523 done |
3449 apply(case_tac [!] r, simp_all) |
3524 |
3450 done |
3525 lemma [intro]: "\<exists> n. (a::cell)\<up>(n) = []" |
3451 |
|
3452 lemma [intro]: "\<exists> n. (a::block)\<^bsup>n\<^esup> = []" |
|
3453 by auto |
3526 by auto |
3454 |
3527 |
3455 lemma [simp]: "\<lbrakk>tinres r r'; r \<noteq> []; r' \<noteq> []\<rbrakk> \<Longrightarrow> hd r = hd r'" |
3528 lemma [simp]: "\<lbrakk>tinres r r'; r \<noteq> []; r' \<noteq> []\<rbrakk> \<Longrightarrow> hd r = hd r'" |
3456 apply(auto simp: tinres_def) |
3529 apply(auto simp: tinres_def) |
3457 done |
3530 done |
3458 |
3531 |
3459 lemma [intro]: "hd (Bk\<^bsup>Suc n\<^esup>) = Bk" |
3532 lemma [intro]: "hd (Bk\<up>(Suc n)) = Bk" |
3460 apply(simp add: exp_ind_def) |
3533 apply(simp add: ) |
3461 done |
3534 done |
3462 |
3535 |
3463 lemma [simp]: "\<lbrakk>tinres r []; r \<noteq> []\<rbrakk> \<Longrightarrow> hd r = Bk" |
3536 lemma [simp]: "\<lbrakk>tinres r []; r \<noteq> []\<rbrakk> \<Longrightarrow> hd r = Bk" |
3464 apply(auto simp: tinres_def) |
3537 apply(auto simp: tinres_def) |
3465 apply(case_tac n, auto) |
|
3466 done |
3538 done |
3467 |
3539 |
3468 lemma [simp]: "\<lbrakk>tinres [] r'; r' \<noteq> []\<rbrakk> \<Longrightarrow> hd r' = Bk" |
3540 lemma [simp]: "\<lbrakk>tinres [] r'; r' \<noteq> []\<rbrakk> \<Longrightarrow> hd r' = Bk" |
3469 apply(auto simp: tinres_def) |
3541 apply(auto simp: tinres_def) |
3470 done |
3542 done |
3471 |
3543 |
3472 lemma [intro]: "\<exists>na. tl r = tl (r @ Bk\<^bsup>n\<^esup>) @ Bk\<^bsup>na\<^esup> \<or> tl (r @ Bk\<^bsup>n\<^esup>) = tl r @ Bk\<^bsup>na\<^esup>" |
3544 lemma [intro]: "\<exists>na. tl r = tl (r @ Bk\<up>(n)) @ Bk\<up>(na) \<or> tl (r @ Bk\<up>(n)) = tl r @ Bk\<up>(na)" |
3473 apply(case_tac r, simp) |
3545 apply(case_tac r, simp) |
3474 apply(case_tac n, simp) |
3546 apply(case_tac n, simp, simp) |
3475 apply(rule_tac x = 0 in exI, simp) |
3547 apply(rule_tac x = nat in exI, simp) |
3476 apply(rule_tac x = nat in exI, simp add: exp_ind_def) |
|
3477 apply(simp) |
|
3478 apply(rule_tac x = n in exI, simp) |
3548 apply(rule_tac x = n in exI, simp) |
3479 done |
3549 done |
3480 |
3550 |
3481 lemma [simp]: "tinres r r' \<Longrightarrow> tinres (tl r) (tl r')" |
3551 lemma [simp]: "tinres r r' \<Longrightarrow> tinres (tl r) (tl r')" |
3482 apply(auto simp: tinres_def) |
3552 apply(auto simp: tinres_def) |
3483 apply(case_tac r', simp_all) |
3553 apply(case_tac r', simp) |
3484 apply(case_tac n, simp_all add: exp_ind_def) |
3554 apply(case_tac n, simp_all) |
3485 apply(rule_tac x = 0 in exI, simp) |
3555 apply(rule_tac x = nat in exI, simp) |
3486 apply(rule_tac x = nat in exI, simp_all) |
|
3487 apply(rule_tac x = n in exI, simp) |
3556 apply(rule_tac x = n in exI, simp) |
3488 done |
3557 done |
3489 |
3558 |
3490 lemma [simp]: "\<lbrakk>tinres r []; r \<noteq> []\<rbrakk> \<Longrightarrow> tinres (tl r) []" |
3559 lemma [simp]: "\<lbrakk>tinres r []; r \<noteq> []\<rbrakk> \<Longrightarrow> tinres (tl r) []" |
3491 apply(case_tac r, auto simp: tinres_def) |
3560 apply(case_tac r, auto simp: tinres_def) |
3492 apply(case_tac n, simp_all add: exp_ind_def) |
3561 apply(case_tac n, simp_all add: ) |
3493 apply(rule_tac x = nat in exI, simp) |
3562 apply(rule_tac x = nat in exI, simp) |
3494 done |
3563 done |
3495 |
3564 |
3496 lemma [simp]: "\<lbrakk>tinres [] r'\<rbrakk> \<Longrightarrow> tinres [] (tl r')" |
3565 lemma [simp]: "\<lbrakk>tinres [] r'\<rbrakk> \<Longrightarrow> tinres [] (tl r')" |
3497 apply(case_tac r', auto simp: tinres_def) |
3566 apply(case_tac r', auto simp: tinres_def) |
3498 apply(case_tac n, simp_all add: exp_ind_def) |
3567 apply(case_tac n, simp_all add: ) |
3499 apply(rule_tac x = nat in exI, simp) |
3568 apply(rule_tac x = nat in exI, simp) |
3500 done |
3569 done |
3501 |
3570 |
3502 lemma [simp]: "tinres r r' \<Longrightarrow> tinres (b # r) (b # r')" |
3571 lemma [simp]: "tinres r r' \<Longrightarrow> tinres (b # r) (b # r')" |
3503 apply(auto simp: tinres_def) |
3572 apply(auto simp: tinres_def) |
3504 done |
3573 done |
3505 |
3574 |
|
3575 lemma [simp]: "tinres r [] \<Longrightarrow> tinres (Bk # tl r) [Bk]" |
|
3576 apply(auto simp: tinres_def) |
|
3577 done |
|
3578 |
|
3579 lemma [simp]: "tinres r [] \<Longrightarrow> tinres (Oc # tl r) [Oc]" |
|
3580 apply(auto simp: tinres_def) |
|
3581 done |
|
3582 |
3506 lemma tinres_step2: |
3583 lemma tinres_step2: |
3507 "\<lbrakk>tinres r r'; tstep (ss, l, r) t = (sa, la, ra); tstep (ss, l, r') t = (sb, lb, rb)\<rbrakk> |
3584 "\<lbrakk>tinres r r'; step0 (ss, l, r) t = (sa, la, ra); step0 (ss, l, r') t = (sb, lb, rb)\<rbrakk> |
3508 \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb" |
3585 \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb" |
3509 apply(case_tac "ss = 0", simp add: tstep_0) |
3586 apply(case_tac "ss = 0", simp add: step_0) |
3510 apply(simp add: tstep.simps [simp del]) |
3587 apply(simp add: step.simps [simp del], auto) |
3511 apply(case_tac "fetch t ss (case r of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x)", simp) |
3588 apply(case_tac [!] "fetch t ss (read r')", simp) |
3512 apply(auto simp: new_tape.simps) |
3589 apply(auto simp: update.simps) |
3513 apply(simp_all split: taction.splits if_splits) |
3590 apply(case_tac [!] a, auto split: if_splits) |
3514 apply(auto) |
3591 done |
3515 done |
|
3516 |
|
3517 |
3592 |
3518 lemma tinres_steps2: |
3593 lemma tinres_steps2: |
3519 "\<lbrakk>tinres r r'; steps (ss, l, r) t stp = (sa, la, ra); steps (ss, l, r') t stp = (sb, lb, rb)\<rbrakk> |
3594 "\<lbrakk>tinres r r'; steps0 (ss, l, r) t stp = (sa, la, ra); steps0 (ss, l, r') t stp = (sb, lb, rb)\<rbrakk> |
3520 \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb" |
3595 \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb" |
3521 apply(induct stp arbitrary: sa la ra sb lb rb, simp add: steps.simps) |
3596 apply(induct stp arbitrary: sa la ra sb lb rb, simp add: steps.simps) |
3522 apply(simp add: tstep_red) |
3597 apply(simp add: step_red) |
3523 apply(case_tac "(steps (ss, l, r) t stp)") |
3598 apply(case_tac "(steps0 (ss, l, r) t stp)") |
3524 apply(case_tac "(steps (ss, l, r') t stp)") |
3599 apply(case_tac "(steps0 (ss, l, r') t stp)") |
3525 proof - |
3600 proof - |
3526 fix stp sa la ra sb lb rb a b c aa ba ca |
3601 fix stp sa la ra sb lb rb a b c aa ba ca |
3527 assume ind: "\<And>sa la ra sb lb rb. \<lbrakk>steps (ss, l, r) t stp = (sa, la, ra); |
3602 assume ind: "\<And>sa la ra sb lb rb. \<lbrakk>steps0 (ss, l, r) t stp = (sa, la, ra); |
3528 steps (ss, l, r') t stp = (sb, lb, rb)\<rbrakk> \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb" |
3603 steps0 (ss, l, r') t stp = (sb, lb, rb)\<rbrakk> \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb" |
3529 and h: " tinres r r'" "tstep (steps (ss, l, r) t stp) t = (sa, la, ra)" |
3604 and h: " tinres r r'" "step0 (steps0 (ss, l, r) t stp) t = (sa, la, ra)" |
3530 "tstep (steps (ss, l, r') t stp) t = (sb, lb, rb)" "steps (ss, l, r) t stp = (a, b, c)" |
3605 "step0 (steps0 (ss, l, r') t stp) t = (sb, lb, rb)" "steps0 (ss, l, r) t stp = (a, b, c)" |
3531 "steps (ss, l, r') t stp = (aa, ba, ca)" |
3606 "steps0 (ss, l, r') t stp = (aa, ba, ca)" |
3532 have "b = ba \<and> tinres c ca \<and> a = aa" |
3607 have "b = ba \<and> tinres c ca \<and> a = aa" |
3533 apply(rule_tac ind, simp_all add: h) |
3608 apply(rule_tac ind, simp_all add: h) |
3534 done |
3609 done |
3535 thus "la = lb \<and> tinres ra rb \<and> sa = sb" |
3610 thus "la = lb \<and> tinres ra rb \<and> sa = sb" |
3536 apply(rule_tac l = b and r = c and ss = a and r' = ca |
3611 apply(rule_tac l = b and r = c and ss = a and r' = ca |
3537 and t = t in tinres_step2) |
3612 and t = t in tinres_step2) |
3538 using h |
3613 using h |
3539 apply(simp, simp, simp) |
3614 apply(simp, simp, simp) |
3540 done |
3615 done |
3541 qed |
3616 qed |
3542 |
3617 |
3543 definition t_wcode_adjust :: "tprog" |
3618 definition t_wcode_adjust :: "instr list" |
3544 where |
3619 where |
3545 "t_wcode_adjust = [(W1, 1), (R, 2), (Nop, 2), (R, 3), (R, 3), (R, 4), |
3620 "t_wcode_adjust = [(W1, 1), (R, 2), (Nop, 2), (R, 3), (R, 3), (R, 4), |
3546 (L, 8), (L, 5), (L, 6), (W0, 5), (L, 6), (R, 7), |
3621 (L, 8), (L, 5), (L, 6), (W0, 5), (L, 6), (R, 7), |
3547 (W1, 2), (Nop, 7), (L, 9), (W0, 8), (L, 9), (L, 10), |
3622 (W1, 2), (Nop, 7), (L, 9), (W0, 8), (L, 9), (L, 10), |
3548 (L, 11), (L, 10), (R, 0), (L, 11)]" |
3623 (L, 11), (L, 10), (R, 0), (L, 11)]" |
3564 done |
3639 done |
3565 |
3640 |
3566 lemma [simp]: "fetch t_wcode_adjust (Suc (Suc (Suc 0))) Bk = (R, 3)" |
3641 lemma [simp]: "fetch t_wcode_adjust (Suc (Suc (Suc 0))) Bk = (R, 3)" |
3567 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3642 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3568 done |
3643 done |
3569 |
3644 |
3570 lemma [simp]: "fetch t_wcode_adjust 4 Bk = (L, 8)" |
3645 lemma [simp]: "fetch t_wcode_adjust (Suc (Suc (Suc (Suc 0)))) Bk = (L, 8)" |
3571 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3646 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_4_eq_4) |
3572 done |
3647 done |
3573 |
3648 |
3574 lemma [simp]: "fetch t_wcode_adjust 4 Oc = (L, 5)" |
3649 lemma [simp]: "fetch t_wcode_adjust (Suc (Suc (Suc (Suc 0)))) Oc = (L, 5)" |
3575 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3650 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_4_eq_4) |
3576 done |
3651 done |
|
3652 |
|
3653 thm numeral_5_eq_5 |
3577 |
3654 |
3578 lemma [simp]: "fetch t_wcode_adjust 5 Oc = (W0, 5)" |
3655 lemma [simp]: "fetch t_wcode_adjust 5 Oc = (W0, 5)" |
3579 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3656 apply(simp only: fetch.simps t_wcode_adjust_def nth_of.simps numeral_5_eq_5, simp) |
3580 done |
3657 done |
3581 |
3658 |
3582 lemma [simp]: "fetch t_wcode_adjust 5 Bk = (L, 6)" |
3659 lemma [simp]: "fetch t_wcode_adjust 5 Bk = (L, 6)" |
3583 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3660 apply(simp only: fetch.simps t_wcode_adjust_def nth_of.simps numeral_5_eq_5, auto) |
3584 done |
3661 done |
3585 |
3662 |
|
3663 thm numeral_6_eq_6 |
3586 lemma [simp]: "fetch t_wcode_adjust 6 Oc = (R, 7)" |
3664 lemma [simp]: "fetch t_wcode_adjust 6 Oc = (R, 7)" |
3587 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3665 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_6_eq_6) |
3588 done |
3666 done |
3589 |
3667 |
3590 lemma [simp]: "fetch t_wcode_adjust 6 Bk = (L, 6)" |
3668 lemma [simp]: "fetch t_wcode_adjust 6 Bk = (L, 6)" |
3591 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3669 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_6_eq_6) |
3592 done |
3670 done |
3593 |
3671 |
3594 lemma [simp]: "fetch t_wcode_adjust 7 Bk = (W1, 2)" |
3672 lemma [simp]: "fetch t_wcode_adjust 7 Bk = (W1, 2)" |
3595 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3673 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_7_eq_7) |
3596 done |
3674 done |
3597 |
3675 |
3598 lemma [simp]: "fetch t_wcode_adjust 8 Bk = (L, 9)" |
3676 lemma [simp]: "fetch t_wcode_adjust 8 Bk = (L, 9)" |
3599 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3677 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_8_eq_8) |
3600 done |
3678 done |
3601 |
3679 |
3602 lemma [simp]: "fetch t_wcode_adjust 8 Oc = (W0, 8)" |
3680 lemma [simp]: "fetch t_wcode_adjust 8 Oc = (W0, 8)" |
3603 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3681 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_8_eq_8) |
3604 done |
3682 done |
3605 |
3683 |
3606 lemma [simp]: "fetch t_wcode_adjust 9 Oc = (L, 10)" |
3684 lemma [simp]: "fetch t_wcode_adjust 9 Oc = (L, 10)" |
3607 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3685 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_9_eq_9) |
3608 done |
3686 done |
3609 |
3687 |
3610 lemma [simp]: "fetch t_wcode_adjust 9 Bk = (L, 9)" |
3688 lemma [simp]: "fetch t_wcode_adjust 9 Bk = (L, 9)" |
3611 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3689 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_9_eq_9) |
3612 done |
3690 done |
3613 |
3691 |
3614 lemma [simp]: "fetch t_wcode_adjust 10 Bk = (L, 11)" |
3692 lemma [simp]: "fetch t_wcode_adjust 10 Bk = (L, 11)" |
3615 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3693 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps numeral_10_eq_10) |
3616 done |
3694 done |
3617 |
3695 |
3618 lemma [simp]: "fetch t_wcode_adjust 10 Oc = (L, 10)" |
3696 lemma [simp]: "fetch t_wcode_adjust 10 Oc = (L, 10)" |
3619 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3697 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps eval_nat_numeral) |
3620 done |
3698 done |
3621 |
3699 |
3622 lemma [simp]: "fetch t_wcode_adjust 11 Oc = (L, 11)" |
3700 lemma [simp]: "fetch t_wcode_adjust 11 Oc = (L, 11)" |
3623 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3701 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps eval_nat_numeral) |
3624 done |
3702 done |
3625 |
3703 |
3626 lemma [simp]: "fetch t_wcode_adjust 11 Bk = (R, 0)" |
3704 lemma [simp]: "fetch t_wcode_adjust 11 Bk = (R, 0)" |
3627 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps) |
3705 apply(simp add: fetch.simps t_wcode_adjust_def nth_of.simps eval_nat_numeral) |
3628 done |
3706 done |
3629 |
3707 |
3630 fun wadjust_start :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
3708 fun wadjust_start :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
3631 where |
3709 where |
3632 "wadjust_start m rs (l, r) = |
3710 "wadjust_start m rs (l, r) = |
3633 (\<exists> ln rn. l = Bk # Oc\<^bsup>Suc m\<^esup> \<and> |
3711 (\<exists> ln rn. l = Bk # Oc\<up>(Suc m) \<and> |
3634 tl r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
3712 tl r = Oc # Bk\<up>(ln) @ Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))" |
3635 |
3713 |
3636 fun wadjust_loop_start :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
3714 fun wadjust_loop_start :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
3637 where |
3715 where |
3638 "wadjust_loop_start m rs (l, r) = |
3716 "wadjust_loop_start m rs (l, r) = |
3639 (\<exists> ln rn ml mr. l = Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and> |
3717 (\<exists> ln rn ml mr. l = Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and> |
3640 r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> |
3718 r = Oc # Bk\<up>(ln) @ Bk # Oc\<up>(mr) @ Bk\<up>(rn) \<and> |
3641 ml + mr = Suc (Suc rs) \<and> mr > 0)" |
3719 ml + mr = Suc (Suc rs) \<and> mr > 0)" |
3642 |
3720 |
3643 fun wadjust_loop_right_move :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
3721 fun wadjust_loop_right_move :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
3644 where |
3722 where |
3645 "wadjust_loop_right_move m rs (l, r) = |
3723 "wadjust_loop_right_move m rs (l, r) = |
3646 (\<exists> ml mr nl nr rn. l = Bk\<^bsup>nl\<^esup> @ Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and> |
3724 (\<exists> ml mr nl nr rn. l = Bk\<up>(nl) @ Oc # Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and> |
3647 r = Bk\<^bsup>nr\<^esup> @ Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> |
3725 r = Bk\<up>(nr) @ Oc\<up>(mr) @ Bk\<up>(rn) \<and> |
3648 ml + mr = Suc (Suc rs) \<and> mr > 0 \<and> |
3726 ml + mr = Suc (Suc rs) \<and> mr > 0 \<and> |
3649 nl + nr > 0)" |
3727 nl + nr > 0)" |
3650 |
3728 |
3651 fun wadjust_loop_check :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
3729 fun wadjust_loop_check :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
3652 where |
3730 where |
3653 "wadjust_loop_check m rs (l, r) = |
3731 "wadjust_loop_check m rs (l, r) = |
3654 (\<exists> ml mr ln rn. l = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and> |
3732 (\<exists> ml mr ln rn. l = Oc # Bk\<up>(ln) @ Bk # Oc # Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and> |
3655 r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> ml + mr = (Suc rs))" |
3733 r = Oc\<up>(mr) @ Bk\<up>(rn) \<and> ml + mr = (Suc rs))" |
3656 |
3734 |
3657 fun wadjust_loop_erase :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
3735 fun wadjust_loop_erase :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
3658 where |
3736 where |
3659 "wadjust_loop_erase m rs (l, r) = |
3737 "wadjust_loop_erase m rs (l, r) = |
3660 (\<exists> ml mr ln rn. l = Bk\<^bsup>ln\<^esup> @ Bk # Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and> |
3738 (\<exists> ml mr ln rn. l = Bk\<up>(ln) @ Bk # Oc # Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and> |
3661 tl r = Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> ml + mr = (Suc rs) \<and> mr > 0)" |
3739 tl r = Oc\<up>(mr) @ Bk\<up>(rn) \<and> ml + mr = (Suc rs) \<and> mr > 0)" |
3662 |
3740 |
3663 fun wadjust_loop_on_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
3741 fun wadjust_loop_on_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
3664 where |
3742 where |
3665 "wadjust_loop_on_left_moving_O m rs (l, r) = |
3743 "wadjust_loop_on_left_moving_O m rs (l, r) = |
3666 (\<exists> ml mr ln rn. l = Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m \<^esup>\<and> |
3744 (\<exists> ml mr ln rn. l = Oc\<up>(ml) @ Bk # Oc\<up>(Suc m )\<and> |
3667 r = Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> |
3745 r = Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(mr) @ Bk\<up>(rn) \<and> |
3668 ml + mr = Suc rs \<and> mr > 0)" |
3746 ml + mr = Suc rs \<and> mr > 0)" |
3669 |
3747 |
3670 fun wadjust_loop_on_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
3748 fun wadjust_loop_on_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
3671 where |
3749 where |
3672 "wadjust_loop_on_left_moving_B m rs (l, r) = |
3750 "wadjust_loop_on_left_moving_B m rs (l, r) = |
3673 (\<exists> ml mr nl nr rn. l = Bk\<^bsup>nl\<^esup> @ Oc # Oc\<^bsup>ml\<^esup> @ Bk # Oc\<^bsup>Suc m\<^esup> \<and> |
3751 (\<exists> ml mr nl nr rn. l = Bk\<up>(nl) @ Oc # Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and> |
3674 r = Bk\<^bsup>nr\<^esup> @ Bk # Bk # Oc\<^bsup>mr\<^esup> @ Bk\<^bsup>rn\<^esup> \<and> |
3752 r = Bk\<up>(nr) @ Bk # Bk # Oc\<up>(mr) @ Bk\<up>(rn) \<and> |
3675 ml + mr = Suc rs \<and> mr > 0)" |
3753 ml + mr = Suc rs \<and> mr > 0)" |
3676 |
3754 |
3677 fun wadjust_loop_on_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
3755 fun wadjust_loop_on_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool" |
3678 where |
3756 where |
3679 "wadjust_loop_on_left_moving m rs (l, r) = |
3757 "wadjust_loop_on_left_moving m rs (l, r) = |
4409 declare numeral_2_eq_2[simp del] |
4468 declare numeral_2_eq_2[simp del] |
4410 |
4469 |
4411 lemma wadjust_correctness: |
4470 lemma wadjust_correctness: |
4412 shows "let P = (\<lambda> (len, st, l, r). st = 0) in |
4471 shows "let P = (\<lambda> (len, st, l, r). st = 0) in |
4413 let Q = (\<lambda> (len, st, l, r). wadjust_inv st m rs (l, r)) in |
4472 let Q = (\<lambda> (len, st, l, r). wadjust_inv st m rs (l, r)) in |
4414 let f = (\<lambda> stp. (Suc (Suc rs), steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, |
4473 let f = (\<lambda> stp. (Suc (Suc rs), steps0 (Suc 0, Bk # Oc\<up>(Suc m), |
4415 Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust stp)) in |
4474 Bk # Oc # Bk\<up>(ln) @ Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn)) t_wcode_adjust stp)) in |
4416 \<exists> n .P (f n) \<and> Q (f n)" |
4475 \<exists> n .P (f n) \<and> Q (f n)" |
4417 proof - |
4476 proof - |
4418 let ?P = "(\<lambda> (len, st, l, r). st = 0)" |
4477 let ?P = "(\<lambda> (len, st, l, r). st = 0)" |
4419 let ?Q = "\<lambda> (len, st, l, r). wadjust_inv st m rs (l, r)" |
4478 let ?Q = "\<lambda> (len, st, l, r). wadjust_inv st m rs (l, r)" |
4420 let ?f = "\<lambda> stp. (Suc (Suc rs), steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, |
4479 let ?f = "\<lambda> stp. (Suc (Suc rs), steps0 (Suc 0, Bk # Oc\<up>(Suc m), |
4421 Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust stp)" |
4480 Bk # Oc # Bk\<up>(ln) @ Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn)) t_wcode_adjust stp)" |
4422 have "\<exists> n. ?P (?f n) \<and> ?Q (?f n)" |
4481 have "\<exists> n. ?P (?f n) \<and> ?Q (?f n)" |
4423 proof(rule_tac halt_lemma2) |
4482 proof(rule_tac halt_lemma2) |
4424 show "wf wadjust_le" by auto |
4483 show "wf wadjust_le" by auto |
4425 next |
4484 next |
4426 show "\<forall> n. \<not> ?P (?f n) \<and> ?Q (?f n) \<longrightarrow> |
4485 show "\<forall> n. \<not> ?P (?f n) \<and> ?Q (?f n) \<longrightarrow> |
4427 ?Q (?f (Suc n)) \<and> (?f (Suc n), ?f n) \<in> wadjust_le" |
4486 ?Q (?f (Suc n)) \<and> (?f (Suc n), ?f n) \<in> wadjust_le" |
4428 proof(rule_tac allI, rule_tac impI, case_tac "?f n", |
4487 apply(rule_tac allI, rule_tac impI, case_tac "?f n", simp) |
4429 simp add: tstep_red tstep.simps, rule_tac conjI, erule_tac conjE, |
4488 apply(simp add: step.simps) |
4430 erule_tac conjE) |
4489 apply(case_tac d, case_tac [2] aa, simp_all) |
4431 fix n a b c d |
4490 apply(simp_all add: wadjust_inv.simps wadjust_le_def |
4432 assume "0 < b" "wadjust_inv b m rs (c, d)" "Suc (Suc rs) = a" |
4491 abacus.lex_triple_def abacus.lex_pair_def lex_square_def numeral_4_eq_4 |
4433 thus "case case fetch t_wcode_adjust b (case d of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) |
|
4434 of (ac, ns) \<Rightarrow> (ns, new_tape ac (c, d)) of (st, x) \<Rightarrow> wadjust_inv st m rs x" |
|
4435 apply(case_tac d, simp, case_tac [2] aa) |
|
4436 apply(simp_all add: wadjust_inv.simps wadjust_le_def new_tape.simps |
|
4437 abacus.lex_triple_def abacus.lex_pair_def lex_square_def |
|
4438 split: if_splits) |
|
4439 done |
|
4440 next |
|
4441 fix n a b c d |
|
4442 assume "0 < b \<and> wadjust_inv b m rs (c, d)" |
|
4443 "Suc (Suc rs) = a \<and> steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, |
|
4444 Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust n = (b, c, d)" |
|
4445 thus "((a, case fetch t_wcode_adjust b (case d of [] \<Rightarrow> Bk | x # xs \<Rightarrow> x) |
|
4446 of (ac, ns) \<Rightarrow> (ns, new_tape ac (c, d))), a, b, c, d) \<in> wadjust_le" |
|
4447 proof(erule_tac conjE, erule_tac conjE, erule_tac conjE) |
|
4448 assume "0 < b" "wadjust_inv b m rs (c, d)" "Suc (Suc rs) = a" |
|
4449 thus "?thesis" |
|
4450 apply(case_tac d, case_tac [2] aa) |
|
4451 apply(simp_all add: wadjust_inv.simps wadjust_le_def new_tape.simps |
|
4452 abacus.lex_triple_def abacus.lex_pair_def lex_square_def |
|
4453 split: if_splits) |
4492 split: if_splits) |
4454 done |
4493 done |
4455 qed |
|
4456 qed |
|
4457 next |
4494 next |
4458 show "?Q (?f 0)" |
4495 show "?Q (?f 0)" |
4459 apply(simp add: steps.simps wadjust_inv.simps wadjust_start.simps) |
4496 apply(simp add: steps.simps wadjust_inv.simps wadjust_start.simps, auto) |
4460 apply(rule_tac x = ln in exI,auto) |
|
4461 done |
4497 done |
4462 next |
4498 next |
4463 show "\<not> ?P (?f 0)" |
4499 show "\<not> ?P (?f 0)" |
4464 apply(simp add: steps.simps) |
4500 apply(simp add: steps.simps) |
4465 done |
4501 done |
4466 qed |
4502 qed |
4467 thus "?thesis" |
4503 thus"?thesis" |
4468 apply(auto) |
4504 apply(simp) |
4469 done |
4505 done |
4470 qed |
4506 qed |
4471 |
4507 |
4472 lemma [intro]: "t_correct t_wcode_adjust" |
4508 lemma [intro]: "tm_wf (t_wcode_adjust, 0)" |
4473 apply(auto simp: t_wcode_adjust_def t_correct.simps iseven_def) |
4509 apply(auto simp: t_wcode_adjust_def tm_wf.simps) |
4474 apply(rule_tac x = 11 in exI, simp) |
4510 done |
|
4511 |
|
4512 declare tm_comp.simps[simp del] |
|
4513 |
|
4514 lemma [simp]: "args \<noteq> [] \<Longrightarrow> bl_bin (<args::nat list>) > 0" |
|
4515 apply(case_tac args) |
|
4516 apply(auto simp: tape_of_nl_cons bl_bin.simps split: if_splits) |
4475 done |
4517 done |
4476 |
4518 |
4477 lemma wcode_lemma_pre': |
4519 lemma wcode_lemma_pre': |
4478 "args \<noteq> [] \<Longrightarrow> |
4520 "args \<noteq> [] \<Longrightarrow> |
4479 \<exists> stp rn. steps (Suc 0, [], <m # args>) |
4521 \<exists> stp rn. steps0 (Suc 0, [], <m # args>) |
4480 ((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) stp |
4522 ((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) stp |
4481 = (0, [Bk], Oc\<^bsup>Suc m\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
4523 = (0, [Bk], Oc\<up>(Suc m) @ Bk # Oc\<up>(Suc (bl_bin (<args>))) @ Bk\<up>(rn))" |
4482 proof - |
4524 proof - |
4483 let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <m # args>" |
4525 let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <m # args>" |
4484 let ?Q1 = "\<lambda>(l, r). l = Bk # Oc\<^bsup>Suc m\<^esup> \<and> |
4526 let ?Q1 = "\<lambda>(l, r). l = Bk # Oc\<up>(Suc m) \<and> |
4485 (\<exists>ln rn. r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
4527 (\<exists>ln rn. r = Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin (<args>)) @ Bk\<up>(rn))" |
4486 let ?P2 = ?Q1 |
4528 let ?P2 = ?Q1 |
4487 let ?Q2 = "\<lambda> (l, r). (wadjust_stop m (bl_bin (<args>) - 1) (l, r))" |
4529 let ?Q2 = "\<lambda> (l, r). (wadjust_stop m (bl_bin (<args>) - 1) (l, r))" |
4488 let ?P3 = "\<lambda> tp. False" |
4530 let ?P3 = "\<lambda> tp. False" |
4489 assume h: "args \<noteq> []" |
4531 assume h: "args \<noteq> []" |
4490 have "?P1 \<turnstile>-> \<lambda> tp. (\<exists> stp tp'. steps (Suc 0, tp) |
4532 hence a: "bl_bin (<args>) > 0" |
4491 ((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) stp = (0, tp') \<and> ?Q2 tp')" |
4533 using h by simp |
4492 proof(rule_tac turing_merge.t_merge_halt[of "t_wcode_prepare |+| t_wcode_main" |
4534 hence "{?P1} (t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust {?Q2}" |
4493 t_wcode_adjust ?P1 ?P2 ?P3 ?P3 ?Q1 ?Q2], |
4535 proof(rule_tac Hoare_plus_halt) |
4494 auto simp: turing_merge_def) |
4536 show "?Q1 \<mapsto> ?P2" |
4495 |
4537 by(simp add: assert_imp_def) |
4496 show "\<exists>stp. case steps (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) stp of |
4538 next |
4497 (st, tp') \<Rightarrow> st = 0 \<and> (case tp' of (l, r) \<Rightarrow> l = Bk # Oc\<^bsup>Suc m\<^esup> \<and> |
4539 show "tm_wf (t_wcode_prepare |+| t_wcode_main, 0)" |
4498 (\<exists>ln rn. r = Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>))" |
4540 apply(rule_tac tm_wf_comp, auto) |
4499 using h prepare_mainpart_lemma[of args m] |
|
4500 apply(auto) |
|
4501 apply(rule_tac x = stp in exI, simp) |
|
4502 apply(rule_tac x = ln in exI, auto) |
|
4503 done |
4541 done |
4504 next |
4542 next |
4505 fix ln rn |
4543 show "{?P1} t_wcode_prepare |+| t_wcode_main {?Q1}" |
4506 show "\<exists>stp. case steps (Suc 0, Bk # Oc\<^bsup>Suc m\<^esup>, Bk # Oc # Bk\<^bsup>ln\<^esup> @ Bk # Bk # |
4544 proof(rule_tac HoareI, auto) |
4507 Oc\<^bsup>bl_bin (<args>)\<^esup> @ Bk\<^bsup>rn\<^esup>) t_wcode_adjust stp of |
4545 show |
4508 (st, tp') \<Rightarrow> st = 0 \<and> wadjust_stop m (bl_bin (<args>) - Suc 0) tp'" |
4546 "\<exists>n. is_final (steps0 (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) n) \<and> |
4509 using wadjust_correctness[of m "bl_bin (<args>) - 1" "Suc ln" rn] |
4547 (\<lambda>(l, r). l = Bk # Oc # Oc \<up> m \<and> |
4510 apply(subgoal_tac "bl_bin (<args>) > 0", auto simp: wadjust_inv.simps) |
4548 (\<exists>ln rn. r = Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> bl_bin (<args>) @ Bk \<up> rn)) |
4511 apply(rule_tac x = n in exI, simp add: exp_ind) |
4549 holds_for steps0 (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) n" |
4512 using h |
4550 using h prepare_mainpart_lemma[of args m] |
4513 apply(case_tac args, simp_all, case_tac list, |
4551 apply(auto) |
4514 simp_all add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def |
4552 apply(rule_tac x = stp in exI, simp) |
4515 bl_bin.simps) |
4553 apply(rule_tac x = ln in exI, auto) |
4516 done |
4554 done |
|
4555 qed |
4517 next |
4556 next |
4518 show "?Q1 \<turnstile>-> ?P2" |
4557 show "{?P2} t_wcode_adjust {?Q2}" |
4519 by(simp add: t_imply_def) |
4558 proof(rule_tac HoareI, auto del: replicate_Suc) |
|
4559 fix ln rn |
|
4560 show "\<exists>n. is_final (steps0 (Suc 0, Bk # Oc # Oc \<up> m, |
|
4561 Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> bl_bin (<args>) @ Bk \<up> rn) t_wcode_adjust n) \<and> |
|
4562 wadjust_stop m (bl_bin (<args>) - Suc 0) holds_for steps0 |
|
4563 (Suc 0, Bk # Oc # Oc \<up> m, Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> bl_bin (<args>) @ Bk \<up> rn) t_wcode_adjust n" |
|
4564 using wadjust_correctness[of m "bl_bin (<args>) - 1" "Suc ln" rn] |
|
4565 apply(simp del: replicate_Suc add: replicate_Suc[THEN sym] exp_ind, auto) |
|
4566 apply(rule_tac x = n in exI) |
|
4567 using a |
|
4568 apply(case_tac "bl_bin (<args>)", simp, simp del: replicate_Suc add: exp_ind wadjust_inv.simps) |
|
4569 done |
|
4570 qed |
4520 qed |
4571 qed |
4521 thus "\<exists>stp rn. steps (Suc 0, [], <m # args>) ((t_wcode_prepare |+| t_wcode_main) |+| |
4572 thus "?thesis" |
4522 t_wcode_adjust) stp = (0, [Bk], Oc\<^bsup>Suc m\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
4573 apply(simp add: Hoare_def, auto) |
4523 apply(simp add: t_imply_def) |
4574 apply(case_tac "(steps0 (Suc 0, [], <(m::nat) # args>) |
4524 apply(erule_tac exE)+ |
4575 ((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) n)") |
4525 apply(subgoal_tac "bl_bin (<args>) > 0", auto simp: wadjust_stop.simps) |
4576 apply(rule_tac x = n in exI, auto simp: wadjust_stop.simps) |
4526 using h |
4577 using a |
4527 apply(case_tac args, simp_all, case_tac list, |
4578 apply(case_tac "bl_bin (<args>)", simp_all) |
4528 simp_all add: tape_of_nl_abv tape_of_nat_list.simps exp_ind_def |
|
4529 bl_bin.simps) |
|
4530 done |
4579 done |
4531 qed |
4580 qed |
4532 |
4581 |
4533 text {* |
4582 text {* |
4534 The initialization TM @{text "t_wcode"}. |
4583 The initialization TM @{text "t_wcode"}. |
4535 *} |
4584 *} |
4536 definition t_wcode :: "tprog" |
4585 definition t_wcode :: "instr list" |
4537 where |
4586 where |
4538 "t_wcode = (t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust" |
4587 "t_wcode = (t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust" |
4539 |
4588 |
4540 |
4589 |
4541 text {* |
4590 text {* |
4542 The correctness of @{text "t_wcode"}. |
4591 The correctness of @{text "t_wcode"}. |
4543 *} |
4592 *} |
|
4593 |
4544 lemma wcode_lemma_1: |
4594 lemma wcode_lemma_1: |
4545 "args \<noteq> [] \<Longrightarrow> |
4595 "args \<noteq> [] \<Longrightarrow> |
4546 \<exists> stp ln rn. steps (Suc 0, [], <m # args>) (t_wcode) stp = |
4596 \<exists> stp ln rn. steps0 (Suc 0, [], <m # args>) (t_wcode) stp = |
4547 (0, [Bk], Oc\<^bsup>Suc m\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
4597 (0, [Bk], Oc\<up>(Suc m) @ Bk # Oc\<up>(Suc (bl_bin (<args>))) @ Bk\<up>(rn))" |
4548 apply(simp add: wcode_lemma_pre' t_wcode_def) |
4598 apply(simp add: wcode_lemma_pre' t_wcode_def del: replicate_Suc) |
4549 done |
4599 done |
4550 |
4600 |
4551 lemma wcode_lemma: |
4601 lemma wcode_lemma: |
4552 "args \<noteq> [] \<Longrightarrow> |
4602 "args \<noteq> [] \<Longrightarrow> |
4553 \<exists> stp ln rn. steps (Suc 0, [], <m # args>) (t_wcode) stp = |
4603 \<exists> stp ln rn. steps0 (Suc 0, [], <m # args>) (t_wcode) stp = |
4554 (0, [Bk], <[m ,bl_bin (<args>)]> @ Bk\<^bsup>rn\<^esup>)" |
4604 (0, [Bk], <[m ,bl_bin (<args>)]> @ Bk\<up>(rn))" |
4555 using wcode_lemma_1[of args m] |
4605 using wcode_lemma_1[of args m] |
4556 apply(simp add: t_wcode_def tape_of_nl_abv tape_of_nat_list.simps) |
4606 apply(simp add: t_wcode_def tape_of_nl_abv tape_of_nat_list.simps) |
4557 done |
4607 done |
4558 |
4608 |
4559 section {* The universal TM *} |
4609 section {* The universal TM *} |
4625 declare tape_of_nl_abv_cons[simp del] |
4674 declare tape_of_nl_abv_cons[simp del] |
4626 |
4675 |
4627 lemma t_utm_halt_eq': |
4676 lemma t_utm_halt_eq': |
4628 "\<lbrakk>turing_basic.t_correct tp; |
4677 "\<lbrakk>turing_basic.t_correct tp; |
4629 0 < rs; |
4678 0 < rs; |
4630 steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>)\<rbrakk> |
4679 steps (Suc 0, Bk\<up>(l), <lm::nat list>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(n))\<rbrakk> |
4631 \<Longrightarrow> \<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp = |
4680 \<Longrightarrow> \<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp = |
4632 (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" |
4681 (0, Bk\<up>(m), Oc\<up>(rs) @ Bk\<up>(n))" |
4633 apply(drule_tac l = l in F_abc_halt_eq, simp, simp, simp) |
4682 apply(drule_tac l = l in F_abc_halt_eq, simp, simp, simp) |
4634 apply(erule_tac exE, erule_tac exE) |
4683 apply(erule_tac exE, erule_tac exE) |
4635 apply(rule_tac F_abc_utm_halt_eq, simp_all) |
4684 apply(rule_tac F_abc_utm_halt_eq, simp_all) |
4636 done |
4685 done |
4637 |
4686 *) |
4638 lemma [simp]: "tinres xs (xs @ Bk\<^bsup>i\<^esup>)" |
4687 (* |
|
4688 lemma [simp]: "tinres xs (xs @ Bk\<up>(i))" |
4639 apply(auto simp: tinres_def) |
4689 apply(auto simp: tinres_def) |
4640 done |
4690 done |
4641 |
4691 |
4642 lemma [elim]: "\<lbrakk>rs > 0; Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup> = c @ Bk\<^bsup>n\<^esup>\<rbrakk> |
4692 lemma [elim]: "\<lbrakk>rs > 0; Oc\<up>(rs) @ Bk\<up>(na) = c @ Bk\<up>(n)\<rbrakk> |
4643 \<Longrightarrow> \<exists>n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>" |
4693 \<Longrightarrow> \<exists>n. c = Oc\<up>(rs) @ Bk\<up>(n)" |
4644 apply(case_tac "na > n") |
4694 apply(case_tac "na > n") |
4645 apply(subgoal_tac "\<exists> d. na = d + n", auto simp: exp_add) |
4695 apply(subgoal_tac "\<exists> d. na = d + n", auto simp: exp_add) |
4646 apply(rule_tac x = "na - n" in exI, simp) |
4696 apply(rule_tac x = "na - n" in exI, simp) |
4647 apply(subgoal_tac "\<exists> d. n = d + na", auto simp: exp_add) |
4697 apply(subgoal_tac "\<exists> d. n = d + na", auto simp: exp_add) |
4648 apply(case_tac rs, simp_all add: exp_ind, case_tac d, |
4698 apply(case_tac rs, simp_all add: exp_ind, case_tac d, |
4649 simp_all add: exp_ind) |
4699 simp_all add: exp_ind) |
4650 apply(rule_tac x = "n - na" in exI, simp) |
4700 apply(rule_tac x = "n - na" in exI, simp) |
4651 done |
4701 done |
4652 |
4702 *) |
4653 |
4703 (* |
4654 lemma t_utm_halt_eq'': |
4704 lemma t_utm_halt_eq'': |
4655 "\<lbrakk>turing_basic.t_correct tp; |
4705 "\<lbrakk>turing_basic.t_correct tp; |
4656 0 < rs; |
4706 0 < rs; |
4657 steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>)\<rbrakk> |
4707 steps (Suc 0, Bk\<up>(l), <lm::nat list>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(n))\<rbrakk> |
4658 \<Longrightarrow> \<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stp = |
4708 \<Longrightarrow> \<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>(i)) t_utm stp = |
4659 (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" |
4709 (0, Bk\<up>(m), Oc\<up>(rs) @ Bk\<up>(n))" |
4660 apply(drule_tac t_utm_halt_eq', simp_all) |
4710 apply(drule_tac t_utm_halt_eq', simp_all) |
4661 apply(erule_tac exE)+ |
4711 apply(erule_tac exE)+ |
4662 proof - |
4712 proof - |
4663 fix stpa ma na |
4713 fix stpa ma na |
4664 assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stpa = (0, Bk\<^bsup>ma\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>)" |
4714 assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stpa = (0, Bk\<up>(ma), Oc\<up>(rs) @ Bk\<up>(na))" |
4665 and gr: "rs > 0" |
4715 and gr: "rs > 0" |
4666 thus "\<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" |
4716 thus "\<exists>stp m n. steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>(i)) t_utm stp = (0, Bk\<up>(m), Oc\<up>(rs) @ Bk\<up>(n))" |
4667 apply(rule_tac x = stpa in exI, rule_tac x = ma in exI, simp) |
4717 apply(rule_tac x = stpa in exI, rule_tac x = ma in exI, simp) |
4668 proof(case_tac "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa", simp) |
4718 proof(case_tac "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>(i)) t_utm stpa", simp) |
4669 fix a b c |
4719 fix a b c |
4670 assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stpa = (0, Bk\<^bsup>ma\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>)" |
4720 assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stpa = (0, Bk\<up>(ma), Oc\<up>(rs) @ Bk\<up>(na))" |
4671 "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (a, b, c)" |
4721 "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>(i)) t_utm stpa = (a, b, c)" |
4672 thus " a = 0 \<and> b = Bk\<^bsup>ma\<^esup> \<and> (\<exists>n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" |
4722 thus " a = 0 \<and> b = Bk\<up>(ma) \<and> (\<exists>n. c = Oc\<up>(rs) @ Bk\<up>(n))" |
4673 using tinres_steps2[of "<[code tp, bl2wc (<lm>)]>" "<[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>" |
4723 using tinres_steps2[of "<[code tp, bl2wc (<lm>)]>" "<[code tp, bl2wc (<lm>)]> @ Bk\<up>(i)" |
4674 "Suc 0" " [Bk, Bk]" t_utm stpa 0 "Bk\<^bsup>ma\<^esup>" "Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>" a b c] |
4724 "Suc 0" " [Bk, Bk]" t_utm stpa 0 "Bk\<up>(ma)" "Oc\<up>(rs) @ Bk\<up>(na)" a b c] |
4675 apply(simp) |
4725 apply(simp) |
4676 using gr |
4726 using gr |
4677 apply(simp only: tinres_def, auto) |
4727 apply(simp only: tinres_def, auto) |
4678 apply(rule_tac x = "na + n" in exI, simp add: exp_add) |
4728 apply(rule_tac x = "na + n" in exI, simp add: exp_add) |
4679 done |
4729 done |
4682 |
4732 |
4683 lemma [simp]: "tinres [Bk, Bk] [Bk]" |
4733 lemma [simp]: "tinres [Bk, Bk] [Bk]" |
4684 apply(auto simp: tinres_def) |
4734 apply(auto simp: tinres_def) |
4685 done |
4735 done |
4686 |
4736 |
4687 lemma [elim]: "Bk\<^bsup>ma\<^esup> = b @ Bk\<^bsup>n\<^esup> \<Longrightarrow> \<exists>m. b = Bk\<^bsup>m\<^esup>" |
4737 lemma [elim]: "Bk\<up>(ma) = b @ Bk\<up>(n) \<Longrightarrow> \<exists>m. b = Bk\<up>(m)" |
4688 apply(subgoal_tac "ma = length b + n") |
4738 apply(subgoal_tac "ma = length b + n") |
4689 apply(rule_tac x = "ma - n" in exI, simp add: exp_add) |
4739 apply(rule_tac x = "ma - n" in exI, simp add: exp_add) |
4690 apply(drule_tac length_equal) |
4740 apply(drule_tac length_equal) |
4691 apply(simp) |
4741 apply(simp) |
4692 done |
4742 done |
|
4743 *) |
|
4744 |
|
4745 |
|
4746 |
|
4747 lemma tinres_step1: |
|
4748 "\<lbrakk>tinres l l'; step (ss, l, r) (t, 0) = (sa, la, ra); |
|
4749 step (ss, l', r) (t, 0) = (sb, lb, rb)\<rbrakk> |
|
4750 \<Longrightarrow> tinres la lb \<and> ra = rb \<and> sa = sb" |
|
4751 apply(case_tac ss, case_tac [!]r, case_tac [!] "a::cell") |
|
4752 apply(auto simp: step.simps fetch.simps nth_of.simps |
|
4753 split: if_splits ) |
|
4754 apply(case_tac [!] "t ! (2 * nat)", |
|
4755 auto simp: tinres_def split: if_splits) |
|
4756 apply(case_tac [1-8] a, auto split: if_splits) |
|
4757 apply(case_tac [!] "t ! (2 * nat)", |
|
4758 auto simp: tinres_def split: if_splits) |
|
4759 apply(case_tac [1-4] a, auto split: if_splits) |
|
4760 apply(case_tac [!] "t ! Suc (2 * nat)", |
|
4761 auto simp: if_splits) |
|
4762 apply(case_tac [!] aa, auto split: if_splits) |
|
4763 done |
|
4764 |
|
4765 lemma tinres_steps1: |
|
4766 "\<lbrakk>tinres l l'; steps (ss, l, r) (t, 0) stp = (sa, la, ra); |
|
4767 steps (ss, l', r) (t, 0) stp = (sb, lb, rb)\<rbrakk> |
|
4768 \<Longrightarrow> tinres la lb \<and> ra = rb \<and> sa = sb" |
|
4769 apply(induct stp arbitrary: sa la ra sb lb rb, simp add: steps.simps) |
|
4770 apply(simp add: step_red) |
|
4771 apply(case_tac "(steps (ss, l, r) (t, 0) stp)") |
|
4772 apply(case_tac "(steps (ss, l', r) (t, 0) stp)") |
|
4773 proof - |
|
4774 fix stp sa la ra sb lb rb a b c aa ba ca |
|
4775 assume ind: "\<And>sa la ra sb lb rb. \<lbrakk>steps (ss, l, r) (t, 0) stp = (sa, (la::cell list), ra); |
|
4776 steps (ss, l', r) (t, 0) stp = (sb, lb, rb)\<rbrakk> \<Longrightarrow> tinres la lb \<and> ra = rb \<and> sa = sb" |
|
4777 and h: " tinres l l'" "step (steps (ss, l, r) (t, 0) stp) (t, 0) = (sa, la, ra)" |
|
4778 "step (steps (ss, l', r) (t, 0) stp) (t, 0) = (sb, lb, rb)" "steps (ss, l, r) (t, 0) stp = (a, b, c)" |
|
4779 "steps (ss, l', r) (t, 0) stp = (aa, ba, ca)" |
|
4780 have "tinres b ba \<and> c = ca \<and> a = aa" |
|
4781 apply(rule_tac ind, simp_all add: h) |
|
4782 done |
|
4783 thus "tinres la lb \<and> ra = rb \<and> sa = sb" |
|
4784 apply(rule_tac l = b and l' = ba and r = c and ss = a |
|
4785 and t = t in tinres_step1) |
|
4786 using h |
|
4787 apply(simp, simp, simp) |
|
4788 done |
|
4789 qed |
|
4790 |
|
4791 lemma [simp]: |
|
4792 "tinres (Bk \<up> m @ [Bk, Bk]) la \<Longrightarrow> \<exists>m. la = Bk \<up> m" |
|
4793 apply(auto simp: tinres_def) |
|
4794 apply(case_tac n, simp add: exp_ind) |
|
4795 apply(rule_tac x ="Suc (Suc m)" in exI, simp only: exp_ind, simp) |
|
4796 apply(simp add: exp_ind del: replicate_Suc) |
|
4797 apply(case_tac nat, simp add: exp_ind) |
|
4798 apply(rule_tac x = "Suc m" in exI, simp only: exp_ind) |
|
4799 apply(simp only: exp_ind, simp) |
|
4800 apply(subgoal_tac "m = length la + nata") |
|
4801 apply(rule_tac x = "m - nata" in exI, simp add: exp_add) |
|
4802 apply(drule_tac length_equal, simp) |
|
4803 apply(simp only: exp_ind[THEN sym] replicate_Suc[THEN sym] replicate_add[THEN sym]) |
|
4804 apply(rule_tac x = "m + Suc (Suc n)" in exI, simp) |
|
4805 done |
4693 |
4806 |
4694 lemma t_utm_halt_eq: |
4807 lemma t_utm_halt_eq: |
4695 "\<lbrakk>turing_basic.t_correct tp; |
4808 assumes tm_wf: "tm_wf (tp, 0)" |
4696 0 < rs; |
4809 and exec: "steps0 (Suc 0, Bk\<up>(l), <lm::nat list>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(n))" |
4697 steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>n\<^esup>)\<rbrakk> |
4810 and resutl: "0 < rs" |
4698 \<Longrightarrow> \<exists>stp m n. steps (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stp = |
4811 shows "\<exists>stp m n. steps0 (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>(i)) t_utm stp = |
4699 (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" |
4812 (0, Bk\<up>(m), Oc\<up>(rs) @ Bk\<up>(n))" |
4700 apply(drule_tac i = i in t_utm_halt_eq'', simp_all) |
|
4701 apply(erule_tac exE)+ |
|
4702 proof - |
4813 proof - |
4703 fix stpa ma na |
4814 obtain ap arity fp where a: "rec_ci rec_F = (ap, arity, fp)" |
4704 assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (0, Bk\<^bsup>ma\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>)" |
4815 by (metis prod_cases3) |
4705 and gr: "rs > 0" |
4816 moreover have b: "rec_calc_rel rec_F [code tp, (bl2wc (<lm>))] (rs - Suc 0)" |
4706 thus "\<exists>stp m n. steps (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" |
4817 using assms |
4707 apply(rule_tac x = stpa in exI) |
4818 apply(rule_tac F_correct, simp_all) |
4708 proof(case_tac "steps (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa", simp) |
4819 done |
4709 fix a b c |
4820 have "\<exists> stp m l. steps0 (Suc 0, Bk # Bk # [], <[code tp, bl2wc (<lm>)]> @ Bk\<up>i) |
4710 assume "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (0, Bk\<^bsup>ma\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>)" |
4821 (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp |
4711 "steps (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>) t_utm stpa = (a, b, c)" |
4822 = (0, Bk\<up>m @ Bk # Bk # [], Oc\<up>Suc (rs - 1) @ Bk\<up>l)" |
4712 thus "a = 0 \<and> (\<exists>m. b = Bk\<^bsup>m\<^esup>) \<and> (\<exists>n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" |
4823 proof(rule_tac recursive_compile_to_tm_correct) |
4713 using tinres_steps[of "[Bk, Bk]" "[Bk]" "Suc 0" "<[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>i\<^esup>" t_utm stpa 0 |
4824 show "rec_ci rec_F = (ap, arity, fp)" using a by simp |
4714 "Bk\<^bsup>ma\<^esup>" "Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>na\<^esup>" a b c] |
4825 next |
4715 apply(simp) |
4826 show "rec_calc_rel rec_F [code tp, bl2wc (<lm>)] (rs - 1)" |
4716 apply(auto simp: tinres_def) |
4827 using b by simp |
4717 apply(rule_tac x = "ma + n" in exI, simp add: exp_add) |
4828 next |
|
4829 show "length [code tp, bl2wc (<lm>)] = 2" by simp |
|
4830 next |
|
4831 show "layout_of (ap [+] dummy_abc 2) = layout_of (ap [+] dummy_abc 2)" |
|
4832 by simp |
|
4833 next |
|
4834 show "F_tprog = tm_of (ap [+] dummy_abc 2)" |
|
4835 using a |
|
4836 apply(simp add: F_tprog_def F_aprog_def numeral_2_eq_2) |
4718 done |
4837 done |
4719 qed |
4838 qed |
4720 qed |
4839 then obtain stp m l where |
4721 |
4840 "steps0 (Suc 0, Bk # Bk # [], <[code tp, bl2wc (<lm>)]> @ Bk\<up>i) |
4722 lemma [intro]: "t_correct t_wcode" |
4841 (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp |
4723 apply(simp add: t_wcode_def) |
4842 = (0, Bk\<up>m @ Bk # Bk # [], Oc\<up>Suc (rs - 1) @ Bk\<up>l)" by blast |
4724 apply(auto) |
4843 hence "\<exists> m. steps0 (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>i) |
4725 done |
4844 (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp |
4726 |
4845 = (0, Bk\<up>m, Oc\<up>Suc (rs - 1) @ Bk\<up>l)" |
4727 lemma [intro]: "t_correct t_utm" |
4846 proof - |
4728 apply(simp add: t_utm_def F_tprog_def) |
4847 assume g: "steps0 (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk \<up> i) |
4729 apply(rule_tac t_compiled_correct, auto) |
4848 (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp = |
4730 done |
4849 (0, Bk \<up> m @ [Bk, Bk], Oc \<up> Suc (rs - 1) @ Bk \<up> l)" |
4731 |
4850 moreover have "tinres [Bk, Bk] [Bk]" |
4732 lemma UTM_halt_lemma_pre: |
4851 apply(auto simp: tinres_def) |
4733 "\<lbrakk>turing_basic.t_correct tp; |
4852 done |
4734 0 < rs; |
4853 moreover obtain sa la ra where "steps0 (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>i) |
4735 args \<noteq> []; |
4854 (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp = (sa, la, ra)" |
4736 steps (Suc 0, Bk\<^bsup>i\<^esup>, <args::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>k\<^esup>)\<rbrakk> |
4855 apply(case_tac "steps0 (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>i) |
4737 \<Longrightarrow> \<exists>stp m n. steps (Suc 0, [], <code tp # args>) UTM_pre stp = |
4856 (F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp", auto) |
4738 (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" |
|
4739 proof - |
|
4740 let ?Q2 = "\<lambda> (l, r). (\<exists> ln rn. l = Bk\<^bsup>ln\<^esup> \<and> r = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>rn\<^esup>)" |
|
4741 term ?Q2 |
|
4742 let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <code tp # args>" |
|
4743 let ?Q1 = "\<lambda> (l, r). (l = [Bk] \<and> |
|
4744 (\<exists> rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>))" |
|
4745 let ?P2 = ?Q1 |
|
4746 let ?P3 = "\<lambda> (l, r). False" |
|
4747 assume h: "turing_basic.t_correct tp" "0 < rs" |
|
4748 "args \<noteq> []" "steps (Suc 0, Bk\<^bsup>i\<^esup>, <args::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>k\<^esup>)" |
|
4749 have "?P1 \<turnstile>-> \<lambda> tp. (\<exists> stp tp'. steps (Suc 0, tp) |
|
4750 (t_wcode |+| t_utm) stp = (0, tp') \<and> ?Q2 tp')" |
|
4751 proof(rule_tac turing_merge.t_merge_halt [of "t_wcode" "t_utm" |
|
4752 ?P1 ?P2 ?P3 ?P3 ?Q1 ?Q2], auto simp: turing_merge_def) |
|
4753 show "\<exists>stp. case steps (Suc 0, [], <code tp # args>) t_wcode stp of (st, tp') \<Rightarrow> |
|
4754 st = 0 \<and> (case tp' of (l, r) \<Rightarrow> l = [Bk] \<and> |
|
4755 (\<exists>rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>))" |
|
4756 using wcode_lemma_1[of args "code tp"] h |
|
4757 apply(simp, auto) |
|
4758 apply(rule_tac x = stpa in exI, auto) |
|
4759 done |
|
4760 next |
|
4761 fix rn |
|
4762 show "\<exists>stp. case steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @ |
|
4763 Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp of |
|
4764 (st, tp') \<Rightarrow> st = 0 \<and> (case tp' of (l, r) \<Rightarrow> |
|
4765 (\<exists>ln. l = Bk\<^bsup>ln\<^esup>) \<and> (\<exists>rn. r = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>rn\<^esup>))" |
|
4766 using t_utm_halt_eq[of tp rs i args stp m k rn] h |
|
4767 apply(auto) |
|
4768 apply(rule_tac x = stpa in exI, simp add: bin_wc_eq |
|
4769 tape_of_nat_list.simps tape_of_nl_abv) |
|
4770 apply(auto) |
|
4771 done |
4857 done |
4772 next |
4858 ultimately show "?thesis" |
4773 show "?Q1 \<turnstile>-> ?P2" |
4859 apply(drule_tac tinres_steps1, auto) |
4774 apply(simp add: t_imply_def) |
|
4775 done |
4860 done |
4776 qed |
4861 qed |
4777 thus "?thesis" |
4862 thus "?thesis" |
4778 apply(simp add: t_imply_def) |
4863 apply(auto) |
4779 apply(auto simp: UTM_pre_def) |
4864 apply(rule_tac x = stp in exI, simp add: t_utm_def) |
|
4865 using assms |
|
4866 apply(case_tac rs, simp_all add: numeral_2_eq_2) |
|
4867 done |
|
4868 qed |
|
4869 |
|
4870 lemma [intro]: "tm_wf (t_wcode, 0)" |
|
4871 apply(simp add: t_wcode_def) |
|
4872 apply(rule_tac tm_wf_comp) |
|
4873 apply(rule_tac tm_wf_comp, auto) |
|
4874 done |
|
4875 |
|
4876 lemma [intro]: "tm_wf (t_utm, 0)" |
|
4877 apply(simp only: t_utm_def F_tprog_def) |
|
4878 apply(rule_tac t_compiled_correct, auto) |
|
4879 done |
|
4880 |
|
4881 lemma UTM_halt_lemma_pre: |
|
4882 assumes wf_tm: "tm_wf (tp, 0)" |
|
4883 and result: "0 < rs" |
|
4884 and args: "args \<noteq> []" |
|
4885 and exec: "steps0 (Suc 0, Bk\<up>(i), <args::nat list>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(k))" |
|
4886 shows "\<exists>stp m n. steps0 (Suc 0, [], <code tp # args>) UTM_pre stp = |
|
4887 (0, Bk\<up>(m), Oc\<up>(rs) @ Bk\<up>(n))" |
|
4888 proof - |
|
4889 let ?Q2 = "\<lambda> (l, r). (\<exists> ln rn. l = Bk\<up>(ln) \<and> r = Oc\<up>(rs) @ Bk\<up>(rn))" |
|
4890 let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <code tp # args>" |
|
4891 let ?Q1 = "\<lambda> (l, r). (l = [Bk] \<and> |
|
4892 (\<exists> rn. r = Oc\<up>(Suc (code tp)) @ Bk # Oc\<up>(Suc (bl_bin (<args>))) @ Bk\<up>(rn)))" |
|
4893 let ?P2 = ?Q1 |
|
4894 let ?P3 = "\<lambda> (l, r). False" |
|
4895 have "{?P1} (t_wcode |+| t_utm) {?Q2}" |
|
4896 proof(rule_tac Hoare_plus_halt) |
|
4897 show "?Q1 \<mapsto> ?P2" |
|
4898 by(simp add: assert_imp_def) |
|
4899 next |
|
4900 show "tm_wf (t_wcode, 0)" by auto |
|
4901 next |
|
4902 show "{?P1} t_wcode {?Q1}" |
|
4903 apply(rule_tac HoareI, auto) |
|
4904 using wcode_lemma_1[of args "code tp"] args |
|
4905 apply(auto) |
|
4906 apply(rule_tac x = stp in exI, simp) |
|
4907 done |
|
4908 next |
|
4909 show "{?P2} t_utm {?Q2}" |
|
4910 proof(rule_tac HoareI, auto) |
|
4911 fix rn |
|
4912 show "\<exists>n. is_final (steps0 (Suc 0, [Bk], Oc # Oc \<up> code tp @ Bk # Oc # Oc \<up> bl_bin (<args>) @ Bk \<up> rn) t_utm n) \<and> |
|
4913 (\<lambda>(l, r). (\<exists>ln. l = Bk \<up> ln) \<and> |
|
4914 (\<exists>rn. r = Oc \<up> rs @ Bk \<up> rn)) holds_for steps0 (Suc 0, [Bk], |
|
4915 Oc # Oc \<up> code tp @ Bk # Oc # Oc \<up> bl_bin (<args>) @ Bk \<up> rn) t_utm n" |
|
4916 using t_utm_halt_eq[of tp i "args" stp m rs k rn] assms |
|
4917 apply(auto simp: bin_wc_eq) |
|
4918 apply(rule_tac x = stpa in exI, simp add: tape_of_nl_abv) |
|
4919 done |
|
4920 qed |
|
4921 qed |
|
4922 thus "?thesis" |
|
4923 apply(auto simp: Hoare_def UTM_pre_def) |
|
4924 apply(case_tac "steps0 (Suc 0, [], <code tp # args>) (t_wcode |+| t_utm) n") |
|
4925 apply(rule_tac x = n in exI, simp) |
4780 done |
4926 done |
4781 qed |
4927 qed |
4782 |
4928 |
4783 text {* |
4929 text {* |
4784 The correctness of @{text "UTM"}, the halt case. |
4930 The correctness of @{text "UTM"}, the halt case. |
4785 *} |
4931 *} |
4786 lemma UTM_halt_lemma: |
4932 lemma UTM_halt_lemma: |
4787 "\<lbrakk>turing_basic.t_correct tp; |
4933 assumes tm_wf: "tm_wf (tp, 0)" |
4788 0 < rs; |
4934 and result: "0 < rs" |
4789 args \<noteq> []; |
4935 and args: "args \<noteq> []" |
4790 steps (Suc 0, Bk\<^bsup>i\<^esup>, <args::nat list>) tp stp = (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup>@Bk\<^bsup>k\<^esup>)\<rbrakk> |
4936 and exec: "steps0 (Suc 0, Bk\<up>(i), <args::nat list>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(k))" |
4791 \<Longrightarrow> \<exists>stp m n. steps (Suc 0, [], <code tp # args>) UTM stp = |
4937 shows "\<exists>stp m n. steps0 (Suc 0, [], <code tp # args>) UTM stp = |
4792 (0, Bk\<^bsup>m\<^esup>, Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>)" |
4938 (0, Bk\<up>(m), Oc\<up>(rs) @ Bk\<up>(n))" |
4793 using UTM_halt_lemma_pre[of tp rs args i stp m k] |
4939 using UTM_halt_lemma_pre[of tp rs args i stp m k] assms |
4794 apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def) |
4940 apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def) |
4795 apply(case_tac "rec_ci rec_F", simp) |
4941 apply(case_tac "rec_ci rec_F", simp) |
4796 done |
4942 done |
4797 |
4943 |
4798 definition TSTD:: "t_conf \<Rightarrow> bool" |
4944 definition TSTD:: "config \<Rightarrow> bool" |
4799 where |
4945 where |
4800 "TSTD c = (let (st, l, r) = c in |
4946 "TSTD c = (let (st, l, r) = c in |
4801 st = 0 \<and> (\<exists> m. l = Bk\<^bsup>m\<^esup>) \<and> (\<exists> rs n. r = Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>))" |
4947 st = 0 \<and> (\<exists> m. l = Bk\<up>(m)) \<and> (\<exists> rs n. r = Oc\<up>(Suc rs) @ Bk\<up>(n)))" |
4802 |
|
4803 thm abacus_turing_eq_uhalt |
|
4804 |
4948 |
4805 lemma nstd_case1: "0 < a \<Longrightarrow> NSTD (trpl_code (a, b, c))" |
4949 lemma nstd_case1: "0 < a \<Longrightarrow> NSTD (trpl_code (a, b, c))" |
4806 apply(simp add: NSTD.simps trpl_code.simps) |
4950 apply(simp add: NSTD.simps trpl_code.simps) |
4807 done |
4951 done |
4808 |
4952 |
4809 lemma [simp]: "\<forall>m. b \<noteq> Bk\<^bsup>m\<^esup> \<Longrightarrow> 0 < bl2wc b" |
4953 lemma [simp]: "\<forall>m. b \<noteq> Bk\<up>(m) \<Longrightarrow> 0 < bl2wc b" |
4810 apply(rule classical, simp) |
4954 apply(rule classical, simp) |
4811 apply(induct b, erule_tac x = 0 in allE, simp) |
4955 apply(induct b, erule_tac x = 0 in allE, simp) |
4812 apply(simp add: bl2wc.simps, case_tac a, simp_all |
4956 apply(simp add: bl2wc.simps, case_tac a, simp_all |
4813 add: bl2nat.simps bl2nat_double) |
4957 add: bl2nat.simps bl2nat_double) |
4814 apply(case_tac "\<exists> m. b = Bk\<^bsup>m\<^esup>", erule exE) |
4958 apply(case_tac "\<exists> m. b = Bk\<up>(m)", erule exE) |
4815 apply(erule_tac x = "Suc m" in allE, simp add: exp_ind_def, simp) |
4959 apply(erule_tac x = "Suc m" in allE, simp add: , simp) |
4816 done |
4960 done |
4817 lemma nstd_case2: "\<forall>m. b \<noteq> Bk\<^bsup>m\<^esup> \<Longrightarrow> NSTD (trpl_code (a, b, c))" |
4961 |
|
4962 lemma nstd_case2: "\<forall>m. b \<noteq> Bk\<up>(m) \<Longrightarrow> NSTD (trpl_code (a, b, c))" |
4818 apply(simp add: NSTD.simps trpl_code.simps) |
4963 apply(simp add: NSTD.simps trpl_code.simps) |
4819 done |
4964 done |
4820 |
|
4821 thm lg.simps |
|
4822 thm lgR.simps |
|
4823 |
4965 |
4824 lemma [elim]: "Suc (2 * x) = 2 * y \<Longrightarrow> RR" |
4966 lemma [elim]: "Suc (2 * x) = 2 * y \<Longrightarrow> RR" |
4825 apply(induct x arbitrary: y, simp, simp) |
4967 apply(induct x arbitrary: y, simp, simp) |
4826 apply(case_tac y, simp, simp) |
4968 apply(case_tac y, simp, simp) |
4827 done |
4969 done |
4828 |
4970 |
4829 lemma bl2nat_zero_eq[simp]: "(bl2nat c 0 = 0) = (\<exists>n. c = Bk\<^bsup>n\<^esup>)" |
4971 declare replicate_Suc[simp del] |
|
4972 |
|
4973 lemma bl2nat_zero_eq[simp]: "(bl2nat c 0 = 0) = (\<exists>n. c = Bk\<up>(n))" |
4830 apply(auto) |
4974 apply(auto) |
4831 apply(induct c, simp add: bl2nat.simps) |
4975 apply(induct c, simp_all add: bl2nat.simps) |
4832 apply(rule_tac x = 0 in exI, simp) |
|
4833 apply(case_tac a, auto simp: bl2nat.simps bl2nat_double) |
4976 apply(case_tac a, auto simp: bl2nat.simps bl2nat_double) |
4834 done |
4977 done |
4835 |
4978 |
4836 lemma bl2wc_exp_ex: |
4979 lemma bl2wc_exp_ex: |
4837 "\<lbrakk>Suc (bl2wc c) = 2 ^ m\<rbrakk> \<Longrightarrow> \<exists> rs n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>" |
4980 "\<lbrakk>Suc (bl2wc c) = 2 ^ m\<rbrakk> \<Longrightarrow> \<exists> rs n. c = Oc\<up>(rs) @ Bk\<up>(n)" |
4838 apply(induct c arbitrary: m, simp add: bl2wc.simps bl2nat.simps) |
4981 apply(induct c arbitrary: m, simp add: bl2wc.simps bl2nat.simps) |
4839 apply(case_tac a, auto) |
4982 apply(case_tac a, auto) |
4840 apply(case_tac m, simp_all add: bl2wc.simps, auto) |
4983 apply(case_tac m, simp_all add: bl2wc.simps, auto) |
4841 apply(rule_tac x = 0 in exI, rule_tac x = "Suc n" in exI, |
4984 apply(rule_tac x = 0 in exI, rule_tac x = "Suc n" in exI, |
4842 simp add: exp_ind_def) |
4985 simp add: replicate_Suc) |
4843 apply(simp add: bl2wc.simps bl2nat.simps bl2nat_double) |
4986 apply(simp add: bl2wc.simps bl2nat.simps bl2nat_double) |
4844 apply(case_tac m, simp, simp) |
4987 apply(case_tac m, simp, simp) |
4845 proof - |
4988 proof - |
4846 fix c m nat |
4989 fix c m nat |
4847 assume ind: |
4990 assume ind: |
4848 "\<And>m. Suc (bl2nat c 0) = 2 ^ m \<Longrightarrow> \<exists>rs n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>" |
4991 "\<And>m. Suc (bl2nat c 0) = 2 ^ m \<Longrightarrow> \<exists>rs n. c = Oc\<up>(rs) @ Bk\<up>(n)" |
4849 and h: |
4992 and h: |
4850 "Suc (Suc (2 * bl2nat c 0)) = 2 * 2 ^ nat" |
4993 "Suc (Suc (2 * bl2nat c 0)) = 2 * 2 ^ nat" |
4851 have "\<exists>rs n. c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>" |
4994 have "\<exists>rs n. c = Oc\<up>(rs) @ Bk\<up>(n)" |
4852 apply(rule_tac m = nat in ind) |
4995 apply(rule_tac m = nat in ind) |
4853 using h |
4996 using h |
4854 apply(simp) |
4997 apply(simp) |
4855 done |
4998 done |
4856 from this obtain rs n where " c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>" by blast |
4999 from this obtain rs n where " c = Oc\<up>(rs) @ Bk\<up>(n)" by blast |
4857 thus "\<exists>rs n. Oc # c = Oc\<^bsup>rs\<^esup> @ Bk\<^bsup>n\<^esup>" |
5000 thus "\<exists>rs n. Oc # c = Oc\<up>(rs) @ Bk\<up>(n)" |
4858 apply(rule_tac x = "Suc rs" in exI, simp add: exp_ind_def) |
5001 apply(rule_tac x = "Suc rs" in exI, simp add: ) |
4859 apply(rule_tac x = n in exI, simp) |
5002 apply(rule_tac x = n in exI, simp add: replicate_Suc) |
4860 done |
5003 done |
4861 qed |
5004 qed |
4862 |
5005 |
4863 lemma [elim]: |
5006 lemma lg_bin: |
4864 "\<lbrakk>\<forall>rs n. c \<noteq> Oc\<^bsup>Suc rs\<^esup> @ Bk\<^bsup>n\<^esup>; |
5007 "\<lbrakk>\<forall>rs n. c \<noteq> Oc\<up>(Suc rs) @ Bk\<up>(n); |
4865 bl2wc c = 2 ^ lg (Suc (bl2wc c)) 2 - Suc 0\<rbrakk> \<Longrightarrow> bl2wc c = 0" |
5008 bl2wc c = 2 ^ lg (Suc (bl2wc c)) 2 - Suc 0\<rbrakk> \<Longrightarrow> bl2wc c = 0" |
4866 apply(subgoal_tac "\<exists> m. Suc (bl2wc c) = 2^m", erule_tac exE) |
5009 apply(subgoal_tac "\<exists> m. Suc (bl2wc c) = 2^m", erule_tac exE) |
4867 apply(drule_tac bl2wc_exp_ex, simp, erule_tac exE, erule_tac exE) |
5010 apply(drule_tac bl2wc_exp_ex, simp, erule_tac exE, erule_tac exE) |
4868 apply(case_tac rs, simp, simp, erule_tac x = nat in allE, |
5011 apply(case_tac rs, simp, simp, erule_tac x = nat in allE, |
4869 erule_tac x = n in allE, simp) |
5012 erule_tac x = n in allE, simp) |
5032 apply(simp) |
5173 apply(simp) |
5033 done |
5174 done |
5034 qed |
5175 qed |
5035 |
5176 |
5036 lemma tutm_uhalt': |
5177 lemma tutm_uhalt': |
5037 "\<lbrakk>turing_basic.t_correct tp; |
5178 assumes tm_wf: "tm_wf (tp,0)" |
5038 \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <lm>) tp stp))\<rbrakk> |
5179 and unhalt: "\<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <lm>) tp stp))" |
5039 \<Longrightarrow> \<forall> stp. \<not> isS0 (steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp)" |
5180 shows "\<forall> stp. \<not> is_final (steps0 (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp)" |
5040 using abacus_turing_eq_uhalt[of "layout_of (F_aprog)" |
5181 apply(simp add: t_utm_def) |
5041 "F_aprog" "F_tprog" "[code tp, bl2wc (<lm>)]" |
5182 proof(rule_tac compile_correct_unhalt) |
5042 "start_of (layout_of (F_aprog )) (length (F_aprog))" |
5183 show "layout_of F_aprog = layout_of F_aprog" by simp |
5043 "Suc (Suc 0)"] |
5184 next |
5044 apply(simp add: F_tprog_def) |
5185 show "F_tprog = tm_of F_aprog" |
5045 apply(subgoal_tac "\<forall>stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)]) |
5186 by(simp add: F_tprog_def) |
5046 (F_aprog) stp of (as, am) \<Rightarrow> as < length (F_aprog)", simp) |
5187 next |
5047 thm abacus_turing_eq_uhalt |
5188 show "crsp (layout_of F_aprog) (0, [code tp, bl2wc (<lm>)]) (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) []" |
5048 apply(simp add: t_utm_def F_tprog_def) |
5189 by(auto simp: crsp.simps start_of.simps) |
5049 apply(rule_tac uabc_uhalt, simp_all) |
5190 next |
5050 done |
5191 show "length F_tprog div 2 = length F_tprog div 2" by simp |
5051 |
5192 next |
|
5193 show "\<forall>stp. case abc_steps_l (0, [code tp, bl2wc (<lm>)]) F_aprog stp of (as, am) \<Rightarrow> as < length F_aprog" |
|
5194 using assms |
|
5195 apply(erule_tac uabc_uhalt, simp) |
|
5196 done |
|
5197 qed |
|
5198 |
|
5199 |
5052 lemma tinres_commute: "tinres r r' \<Longrightarrow> tinres r' r" |
5200 lemma tinres_commute: "tinres r r' \<Longrightarrow> tinres r' r" |
5053 apply(auto simp: tinres_def) |
5201 apply(auto simp: tinres_def) |
5054 done |
5202 done |
5055 |
5203 |
5056 lemma inres_tape: |
5204 lemma inres_tape: |
5057 "\<lbrakk>steps (st, l, r) tp stp = (a, b, c); steps (st, l', r') tp stp = (a', b', c'); |
5205 "\<lbrakk>steps0 (st, l, r) tp stp = (a, b, c); steps0 (st, l', r') tp stp = (a', b', c'); |
5058 tinres l l'; tinres r r'\<rbrakk> |
5206 tinres l l'; tinres r r'\<rbrakk> |
5059 \<Longrightarrow> a = a' \<and> tinres b b' \<and> tinres c c'" |
5207 \<Longrightarrow> a = a' \<and> tinres b b' \<and> tinres c c'" |
5060 proof(case_tac "steps (st, l', r) tp stp") |
5208 proof(case_tac "steps0 (st, l', r) tp stp") |
5061 fix aa ba ca |
5209 fix aa ba ca |
5062 assume h: "steps (st, l, r) tp stp = (a, b, c)" |
5210 assume h: "steps0 (st, l, r) tp stp = (a, b, c)" |
5063 "steps (st, l', r') tp stp = (a', b', c')" |
5211 "steps0 (st, l', r') tp stp = (a', b', c')" |
5064 "tinres l l'" "tinres r r'" |
5212 "tinres l l'" "tinres r r'" |
5065 "steps (st, l', r) tp stp = (aa, ba, ca)" |
5213 "steps0 (st, l', r) tp stp = (aa, ba, ca)" |
5066 have "tinres b ba \<and> c = ca \<and> a = aa" |
5214 have "tinres b ba \<and> c = ca \<and> a = aa" |
5067 using h |
5215 using h |
5068 apply(rule_tac tinres_steps, auto) |
5216 apply(rule_tac tinres_steps1, auto) |
5069 done |
5217 done |
5070 |
|
5071 thm tinres_steps2 |
|
5072 moreover have "b' = ba \<and> tinres c' ca \<and> a' = aa" |
5218 moreover have "b' = ba \<and> tinres c' ca \<and> a' = aa" |
5073 using h |
5219 using h |
5074 apply(rule_tac tinres_steps2, auto intro: tinres_commute) |
5220 apply(rule_tac tinres_steps2, auto intro: tinres_commute) |
5075 done |
5221 done |
5076 ultimately show "?thesis" |
5222 ultimately show "?thesis" |
5077 apply(auto intro: tinres_commute) |
5223 apply(auto intro: tinres_commute) |
5078 done |
5224 done |
5079 qed |
5225 qed |
5080 |
5226 |
5081 lemma tape_normalize: "\<forall> stp. \<not> isS0 (steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp) |
5227 lemma tape_normalize: "\<forall> stp. \<not> is_final(steps0 (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp) |
5082 \<Longrightarrow> \<forall> stp. \<not> isS0 (steps (Suc 0, Bk\<^bsup>m\<^esup>, <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>n\<^esup>) t_utm stp)" |
5228 \<Longrightarrow> \<forall> stp. \<not> is_final (steps0 (Suc 0, Bk\<up>(m), <[code tp, bl2wc (<lm>)]> @ Bk\<up>(n)) t_utm stp)" |
5083 apply(rule_tac allI, case_tac "(steps (Suc 0, Bk\<^bsup>m\<^esup>, |
5229 apply(rule_tac allI, case_tac "(steps0 (Suc 0, Bk\<up>(m), |
5084 <[code tp, bl2wc (<lm>)]> @ Bk\<^bsup>n\<^esup>) t_utm stp)", simp add: isS0_def) |
5230 <[code tp, bl2wc (<lm>)]> @ Bk\<up>(n)) t_utm stp)", simp) |
5085 apply(erule_tac x = stp in allE) |
5231 apply(erule_tac x = stp in allE) |
5086 apply(case_tac "steps (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp", simp) |
5232 apply(case_tac "steps0 (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp", simp) |
5087 apply(drule_tac inres_tape, auto) |
5233 apply(drule_tac inres_tape, auto) |
5088 apply(auto simp: tinres_def) |
5234 apply(auto simp: tinres_def) |
5089 apply(case_tac "m > Suc (Suc 0)") |
5235 apply(case_tac "m > Suc (Suc 0)") |
5090 apply(rule_tac x = "m - Suc (Suc 0)" in exI) |
5236 apply(rule_tac x = "m - Suc (Suc 0)" in exI) |
5091 apply(case_tac m, simp_all add: exp_ind_def, case_tac nat, simp_all add: exp_ind_def) |
5237 apply(case_tac m, simp_all add: , case_tac nat, simp_all add: replicate_Suc) |
5092 apply(rule_tac x = "2 - m" in exI, simp add: exp_ind_def[THEN sym] exp_add[THEN sym]) |
5238 apply(rule_tac x = "2 - m" in exI, simp add: exp_add[THEN sym]) |
5093 apply(simp only: numeral_2_eq_2, simp add: exp_ind_def) |
5239 apply(simp only: numeral_2_eq_2, simp add: replicate_Suc) |
5094 done |
5240 done |
5095 |
5241 |
5096 lemma tutm_uhalt: |
5242 lemma tutm_uhalt: |
5097 "\<lbrakk>turing_basic.t_correct tp; |
5243 "\<lbrakk>tm_wf (tp,0); |
5098 \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <args>) tp stp))\<rbrakk> |
5244 \<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <args>) tp stp))\<rbrakk> |
5099 \<Longrightarrow> \<forall> stp. \<not> isS0 (steps (Suc 0, Bk\<^bsup>m\<^esup>, <[code tp, bl2wc (<args>)]> @ Bk\<^bsup>n\<^esup>) t_utm stp)" |
5245 \<Longrightarrow> \<forall> stp. \<not> is_final (steps0 (Suc 0, Bk\<up>(m), <[code tp, bl2wc (<args>)]> @ Bk\<up>(n)) t_utm stp)" |
5100 apply(rule_tac tape_normalize) |
5246 apply(rule_tac tape_normalize) |
5101 apply(rule_tac tutm_uhalt', simp_all) |
5247 apply(rule_tac tutm_uhalt', simp_all) |
5102 done |
5248 done |
5103 |
5249 |
5104 lemma UTM_uhalt_lemma_pre: |
5250 lemma UTM_uhalt_lemma_pre: |
5105 "\<lbrakk>turing_basic.t_correct tp; |
5251 assumes tm_wf: "tm_wf (tp, 0)" |
5106 \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <args>) tp stp)); |
5252 and exec: "\<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <args>) tp stp))" |
5107 args \<noteq> []\<rbrakk> |
5253 and args: "args \<noteq> []" |
5108 \<Longrightarrow> \<forall> stp. \<not> isS0 (steps (Suc 0, [], <code tp # args>) UTM_pre stp)" |
5254 shows "\<forall> stp. \<not> is_final (steps0 (Suc 0, [], <code tp # args>) UTM_pre stp)" |
5109 proof - |
5255 proof - |
5110 let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <code tp # args>" |
5256 let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <code tp # args>" |
5111 let ?Q1 = "\<lambda> (l, r). (l = [Bk] \<and> |
5257 let ?Q1 = "\<lambda> (l, r). (l = [Bk] \<and> |
5112 (\<exists> rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>))" |
5258 (\<exists> rn. r = Oc\<up>(Suc (code tp)) @ Bk # Oc\<up>(Suc (bl_bin (<args>))) @ Bk\<up>(rn)))" |
5113 let ?P4 = ?Q1 |
5259 let ?P2 = ?Q1 |
5114 let ?P3 = "\<lambda> (l, r). False" |
5260 have "{?P1} (t_wcode |+| t_utm) \<up>" |
5115 assume h: "turing_basic.t_correct tp" "\<forall>stp. \<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <args>) tp stp)" |
5261 proof(rule_tac Hoare_plus_unhalt) |
5116 "args \<noteq> []" |
5262 show "?Q1 \<mapsto> ?P2" |
5117 have "?P1 \<turnstile>-> \<lambda> tp. \<not> (\<exists> stp. isS0 (steps (Suc 0, tp) (t_wcode |+| t_utm) stp))" |
5263 by(simp add: assert_imp_def) |
5118 proof(rule_tac turing_merge.t_merge_uhalt [of "t_wcode" "t_utm" |
|
5119 ?P1 ?P3 ?P3 ?P4 ?Q1 ?P3], auto simp: turing_merge_def) |
|
5120 show "\<exists>stp. case steps (Suc 0, [], <code tp # args>) t_wcode stp of (st, tp') \<Rightarrow> |
|
5121 st = 0 \<and> (case tp' of (l, r) \<Rightarrow> l = [Bk] \<and> |
|
5122 (\<exists>rn. r = Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>))" |
|
5123 using wcode_lemma_1[of args "code tp"] h |
|
5124 apply(simp, auto) |
|
5125 apply(rule_tac x = stp in exI, auto) |
|
5126 done |
|
5127 next |
5264 next |
5128 fix rn stp |
5265 show "tm_wf (t_wcode, 0)" by auto |
5129 show " isS0 (steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp) |
5266 next |
5130 \<Longrightarrow> False" |
5267 show "{?P1} t_wcode {?Q1}" |
5131 using tutm_uhalt[of tp l args "Suc 0" rn] h |
5268 apply(rule_tac HoareI, auto) |
5132 apply(simp) |
5269 using wcode_lemma_1[of args "code tp"] args |
5133 apply(erule_tac x = stp in allE) |
5270 apply(auto) |
5134 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps bin_wc_eq) |
5271 apply(rule_tac x = stp in exI, simp) |
5135 done |
5272 done |
5136 next |
5273 next |
5137 fix rn stp |
5274 show "{?P2} t_utm \<up>" |
5138 show "isS0 (steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp) \<Longrightarrow> |
5275 proof(rule_tac Hoare_unhalt_I, auto) |
5139 isS0 (steps (Suc 0, [Bk], Oc\<^bsup>Suc (code tp)\<^esup> @ Bk # Oc\<^bsup>Suc (bl_bin (<args>))\<^esup> @ Bk\<^bsup>rn\<^esup>) t_utm stp)" |
5276 fix n rn |
5140 by simp |
5277 assume h: "is_final (steps0 (Suc 0, [Bk], Oc \<up> Suc (code tp) @ Bk # Oc \<up> Suc (bl_bin (<args>)) @ Bk \<up> rn) t_utm n)" |
5141 next |
5278 have "\<forall> stp. \<not> is_final (steps0 (Suc 0, Bk\<up>(Suc 0), <[code tp, bl2wc (<args>)]> @ Bk\<up>(rn)) t_utm stp)" |
5142 show "?Q1 \<turnstile>-> ?P4" |
5279 using assms |
5143 apply(simp add: t_imply_def) |
5280 apply(rule_tac tutm_uhalt, simp_all) |
5144 done |
5281 done |
|
5282 thus "False" |
|
5283 using h |
|
5284 apply(erule_tac x = n in allE) |
|
5285 apply(simp add: tape_of_nl_abv bin_wc_eq) |
|
5286 done |
|
5287 qed |
5145 qed |
5288 qed |
5146 thus "?thesis" |
5289 thus "?thesis" |
5147 apply(simp add: t_imply_def UTM_pre_def) |
5290 apply(simp add: Hoare_unhalt_def UTM_pre_def) |
5148 done |
5291 done |
5149 qed |
5292 qed |
5150 |
5293 |
5151 text {* |
5294 text {* |
5152 The correctness of @{text "UTM"}, the unhalt case. |
5295 The correctness of @{text "UTM"}, the unhalt case. |
5153 *} |
5296 *} |
5154 |
5297 |
5155 lemma UTM_uhalt_lemma: |
5298 lemma UTM_uhalt_lemma: |
5156 "\<lbrakk>turing_basic.t_correct tp; |
5299 assumes tm_wf: "tm_wf (tp, 0)" |
5157 \<forall> stp. (\<not> TSTD (steps (Suc 0, Bk\<^bsup>l\<^esup>, <args>) tp stp)); |
5300 and unhalt: "\<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <args>) tp stp))" |
5158 args \<noteq> []\<rbrakk> |
5301 and args: "args \<noteq> []" |
5159 \<Longrightarrow> \<forall> stp. \<not> isS0 (steps (Suc 0, [], <code tp # args>) UTM stp)" |
5302 shows " \<forall> stp. \<not> is_final (steps0 (Suc 0, [], <code tp # args>) UTM stp)" |
5160 using UTM_uhalt_lemma_pre[of tp l args] |
5303 using UTM_uhalt_lemma_pre[of tp l args] assms |
5161 apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def) |
5304 apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def) |
5162 apply(case_tac "rec_ci rec_F", simp) |
5305 apply(case_tac "rec_ci rec_F", simp) |
5163 done |
5306 done |
5164 |
5307 |
5165 end |
5308 end |