553 done |
1040 done |
554 qed |
1041 qed |
555 |
1042 |
556 subsection {* Crsp of Inc*} |
1043 subsection {* Crsp of Inc*} |
557 |
1044 |
|
1045 fun at_begin_fst_bwtn :: "inc_inv_t" |
|
1046 where |
|
1047 "at_begin_fst_bwtn (as, lm) (s, l, r) ires = |
|
1048 (\<exists> lm1 tn rn. lm1 = (lm @ 0\<up>tn) \<and> length lm1 = s \<and> |
|
1049 (if lm1 = [] then l = Bk # Bk # ires |
|
1050 else l = [Bk]@<rev lm1>@Bk#Bk#ires) \<and> r = Bk\<up>rn)" |
|
1051 |
|
1052 |
|
1053 fun at_begin_fst_awtn :: "inc_inv_t" |
|
1054 where |
|
1055 "at_begin_fst_awtn (as, lm) (s, l, r) ires = |
|
1056 (\<exists> lm1 tn rn. lm1 = (lm @ 0\<up>tn) \<and> length lm1 = s \<and> |
|
1057 (if lm1 = [] then l = Bk # Bk # ires |
|
1058 else l = [Bk]@<rev lm1>@Bk#Bk#ires) \<and> r = [Oc]@Bk\<up>rn)" |
|
1059 |
|
1060 fun at_begin_norm :: "inc_inv_t" |
|
1061 where |
|
1062 "at_begin_norm (as, lm) (s, l, r) ires= |
|
1063 (\<exists> lm1 lm2 rn. lm = lm1 @ lm2 \<and> length lm1 = s \<and> |
|
1064 (if lm1 = [] then l = Bk # Bk # ires |
|
1065 else l = Bk # <rev lm1> @ Bk # Bk # ires ) \<and> r = <lm2>@Bk\<up>rn)" |
|
1066 |
|
1067 fun in_middle :: "inc_inv_t" |
|
1068 where |
|
1069 "in_middle (as, lm) (s, l, r) ires = |
|
1070 (\<exists> lm1 lm2 tn m ml mr rn. lm @ 0\<up>tn = lm1 @ [m] @ lm2 |
|
1071 \<and> length lm1 = s \<and> m + 1 = ml + mr \<and> |
|
1072 ml \<noteq> 0 \<and> tn = s + 1 - length lm \<and> |
|
1073 (if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires |
|
1074 else l = Oc\<up>ml@[Bk]@<rev lm1>@ |
|
1075 Bk # Bk # ires) \<and> (r = Oc\<up>mr @ [Bk] @ <lm2>@ Bk\<up>rn \<or> |
|
1076 (lm2 = [] \<and> r = Oc\<up>mr)) |
|
1077 )" |
|
1078 |
|
1079 fun inv_locate_a :: "inc_inv_t" |
|
1080 where "inv_locate_a (as, lm) (s, l, r) ires = |
|
1081 (at_begin_norm (as, lm) (s, l, r) ires \<or> |
|
1082 at_begin_fst_bwtn (as, lm) (s, l, r) ires \<or> |
|
1083 at_begin_fst_awtn (as, lm) (s, l, r) ires |
|
1084 )" |
|
1085 |
|
1086 fun inv_locate_b :: "inc_inv_t" |
|
1087 where "inv_locate_b (as, lm) (s, l, r) ires = |
|
1088 (in_middle (as, lm) (s, l, r)) ires " |
|
1089 |
|
1090 fun inv_after_write :: "inc_inv_t" |
|
1091 where "inv_after_write (as, lm) (s, l, r) ires = |
|
1092 (\<exists> rn m lm1 lm2. lm = lm1 @ m # lm2 \<and> |
|
1093 (if lm1 = [] then l = Oc\<up>m @ Bk # Bk # ires |
|
1094 else Oc # l = Oc\<up>Suc m@ Bk # <rev lm1> @ |
|
1095 Bk # Bk # ires) \<and> r = [Oc] @ <lm2> @ Bk\<up>rn)" |
|
1096 |
|
1097 fun inv_after_move :: "inc_inv_t" |
|
1098 where "inv_after_move (as, lm) (s, l, r) ires = |
|
1099 (\<exists> rn m lm1 lm2. lm = lm1 @ m # lm2 \<and> |
|
1100 (if lm1 = [] then l = Oc\<up>Suc m @ Bk # Bk # ires |
|
1101 else l = Oc\<up>Suc m@ Bk # <rev lm1> @ Bk # Bk # ires) \<and> |
|
1102 r = <lm2> @ Bk\<up>rn)" |
|
1103 |
|
1104 fun inv_after_clear :: "inc_inv_t" |
|
1105 where "inv_after_clear (as, lm) (s, l, r) ires = |
|
1106 (\<exists> rn m lm1 lm2 r'. lm = lm1 @ m # lm2 \<and> |
|
1107 (if lm1 = [] then l = Oc\<up>Suc m @ Bk # Bk # ires |
|
1108 else l = Oc\<up>Suc m @ Bk # <rev lm1> @ Bk # Bk # ires) \<and> |
|
1109 r = Bk # r' \<and> Oc # r' = <lm2> @ Bk\<up>rn)" |
|
1110 |
|
1111 fun inv_on_right_moving :: "inc_inv_t" |
|
1112 where "inv_on_right_moving (as, lm) (s, l, r) ires = |
|
1113 (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> |
|
1114 ml + mr = m \<and> |
|
1115 (if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires |
|
1116 else l = Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> |
|
1117 ((r = Oc\<up>mr @ [Bk] @ <lm2> @ Bk\<up>rn) \<or> |
|
1118 (r = Oc\<up>mr \<and> lm2 = [])))" |
|
1119 |
|
1120 fun inv_on_left_moving_norm :: "inc_inv_t" |
|
1121 where "inv_on_left_moving_norm (as, lm) (s, l, r) ires = |
|
1122 (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> |
|
1123 ml + mr = Suc m \<and> mr > 0 \<and> (if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires |
|
1124 else l = Oc\<up>ml @ Bk # <rev lm1> @ Bk # Bk # ires) |
|
1125 \<and> (r = Oc\<up>mr @ Bk # <lm2> @ Bk\<up>rn \<or> |
|
1126 (lm2 = [] \<and> r = Oc\<up>mr)))" |
|
1127 |
|
1128 fun inv_on_left_moving_in_middle_B:: "inc_inv_t" |
|
1129 where "inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires = |
|
1130 (\<exists> lm1 lm2 rn. lm = lm1 @ lm2 \<and> |
|
1131 (if lm1 = [] then l = Bk # ires |
|
1132 else l = <rev lm1> @ Bk # Bk # ires) \<and> |
|
1133 r = Bk # <lm2> @ Bk\<up>rn)" |
|
1134 |
|
1135 fun inv_on_left_moving :: "inc_inv_t" |
|
1136 where "inv_on_left_moving (as, lm) (s, l, r) ires = |
|
1137 (inv_on_left_moving_norm (as, lm) (s, l, r) ires \<or> |
|
1138 inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires)" |
|
1139 |
|
1140 |
|
1141 fun inv_check_left_moving_on_leftmost :: "inc_inv_t" |
|
1142 where "inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires = |
|
1143 (\<exists> rn. l = ires \<and> r = [Bk, Bk] @ <lm> @ Bk\<up>rn)" |
|
1144 |
|
1145 fun inv_check_left_moving_in_middle :: "inc_inv_t" |
|
1146 where "inv_check_left_moving_in_middle (as, lm) (s, l, r) ires = |
|
1147 (\<exists> lm1 lm2 r' rn. lm = lm1 @ lm2 \<and> |
|
1148 (Oc # l = <rev lm1> @ Bk # Bk # ires) \<and> r = Oc # Bk # r' \<and> |
|
1149 r' = <lm2> @ Bk\<up>rn)" |
|
1150 |
|
1151 fun inv_check_left_moving :: "inc_inv_t" |
|
1152 where "inv_check_left_moving (as, lm) (s, l, r) ires = |
|
1153 (inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires \<or> |
|
1154 inv_check_left_moving_in_middle (as, lm) (s, l, r) ires)" |
|
1155 |
|
1156 fun inv_after_left_moving :: "inc_inv_t" |
|
1157 where "inv_after_left_moving (as, lm) (s, l, r) ires= |
|
1158 (\<exists> rn. l = Bk # ires \<and> r = Bk # <lm> @ Bk\<up>rn)" |
|
1159 |
|
1160 fun inv_stop :: "inc_inv_t" |
|
1161 where "inv_stop (as, lm) (s, l, r) ires= |
|
1162 (\<exists> rn. l = Bk # Bk # ires \<and> r = <lm> @ Bk\<up>rn)" |
|
1163 |
|
1164 |
|
1165 lemma halt_lemma2': |
|
1166 "\<lbrakk>wf LE; \<forall> n. ((\<not> P (f n) \<and> Q (f n)) \<longrightarrow> |
|
1167 (Q (f (Suc n)) \<and> (f (Suc n), (f n)) \<in> LE)); Q (f 0)\<rbrakk> |
|
1168 \<Longrightarrow> \<exists> n. P (f n)" |
|
1169 apply(intro exCI, simp) |
|
1170 apply(subgoal_tac "\<forall> n. Q (f n)", simp) |
|
1171 apply(drule_tac f = f in wf_inv_image) |
|
1172 apply(simp add: inv_image_def) |
|
1173 apply(erule wf_induct, simp) |
|
1174 apply(erule_tac x = x in allE) |
|
1175 apply(erule_tac x = n in allE, erule_tac x = n in allE) |
|
1176 apply(erule_tac x = "Suc x" in allE, simp) |
|
1177 apply(rule_tac allI) |
|
1178 apply(induct_tac n, simp) |
|
1179 apply(erule_tac x = na in allE, simp) |
|
1180 done |
|
1181 |
|
1182 lemma halt_lemma2'': |
|
1183 "\<lbrakk>P (f n); \<not> P (f (0::nat))\<rbrakk> \<Longrightarrow> |
|
1184 \<exists> n. (P (f n) \<and> (\<forall> i < n. \<not> P (f i)))" |
|
1185 apply(induct n rule: nat_less_induct, auto) |
|
1186 done |
|
1187 |
|
1188 lemma halt_lemma2''': |
|
1189 "\<lbrakk>\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> LE; |
|
1190 Q (f 0); \<forall>i<na. \<not> P (f i)\<rbrakk> \<Longrightarrow> Q (f na)" |
|
1191 apply(induct na, simp, simp) |
|
1192 done |
|
1193 |
|
1194 lemma halt_lemma2: |
|
1195 "\<lbrakk>wf LE; |
|
1196 Q (f 0); \<not> P (f 0); |
|
1197 \<forall> n. ((\<not> P (f n) \<and> Q (f n)) \<longrightarrow> (Q (f (Suc n)) \<and> (f (Suc n), (f n)) \<in> LE))\<rbrakk> |
|
1198 \<Longrightarrow> \<exists> n. P (f n) \<and> Q (f n)" |
|
1199 apply(insert halt_lemma2' [of LE P f Q], simp, erule_tac exE) |
|
1200 apply(subgoal_tac "\<exists> n. (P (f n) \<and> (\<forall> i < n. \<not> P (f i)))") |
|
1201 apply(erule_tac exE)+ |
|
1202 apply(rule_tac x = na in exI, auto) |
|
1203 apply(rule halt_lemma2''', simp, simp, simp) |
|
1204 apply(erule_tac halt_lemma2'', simp) |
|
1205 done |
|
1206 |
|
1207 |
|
1208 fun findnth_inv :: "layout \<Rightarrow> nat \<Rightarrow> inc_inv_t" |
|
1209 where |
|
1210 "findnth_inv ly n (as, lm) (s, l, r) ires = |
|
1211 (if s = 0 then False |
|
1212 else if s \<le> Suc (2*n) then |
|
1213 if s mod 2 = 1 then inv_locate_a (as, lm) ((s - 1) div 2, l, r) ires |
|
1214 else inv_locate_b (as, lm) ((s - 1) div 2, l, r) ires |
|
1215 else False)" |
|
1216 |
|
1217 |
|
1218 fun findnth_state :: "config \<Rightarrow> nat \<Rightarrow> nat" |
|
1219 where |
|
1220 "findnth_state (s, l, r) n = (Suc (2*n) - s)" |
|
1221 |
|
1222 fun findnth_step :: "config \<Rightarrow> nat \<Rightarrow> nat" |
|
1223 where |
|
1224 "findnth_step (s, l, r) n = |
|
1225 (if s mod 2 = 1 then |
|
1226 (if (r \<noteq> [] \<and> hd r = Oc) then 0 |
|
1227 else 1) |
|
1228 else length r)" |
|
1229 |
|
1230 fun findnth_measure :: "config \<times> nat \<Rightarrow> nat \<times> nat" |
|
1231 where |
|
1232 "findnth_measure (c, n) = |
|
1233 (findnth_state c n, findnth_step c n)" |
|
1234 |
|
1235 definition lex_pair :: "((nat \<times> nat) \<times> nat \<times> nat) set" |
|
1236 where |
|
1237 "lex_pair \<equiv> less_than <*lex*> less_than" |
|
1238 |
|
1239 definition findnth_LE :: "((config \<times> nat) \<times> (config \<times> nat)) set" |
|
1240 where |
|
1241 "findnth_LE \<equiv> (inv_image lex_pair findnth_measure)" |
|
1242 |
|
1243 lemma wf_findnth_LE: "wf findnth_LE" |
|
1244 by(auto intro:wf_inv_image simp: findnth_LE_def lex_pair_def) |
|
1245 |
|
1246 declare findnth_inv.simps[simp del] |
|
1247 |
|
1248 lemma [simp]: |
|
1249 "\<lbrakk>x < Suc (Suc (2 * n)); Suc x mod 2 = Suc 0; \<not> x < 2 * n\<rbrakk> |
|
1250 \<Longrightarrow> x = 2*n" |
|
1251 by arith |
|
1252 |
|
1253 lemma [simp]: |
|
1254 "\<lbrakk>0 < a; a < Suc (2 * n); a mod 2 = Suc 0\<rbrakk> |
|
1255 \<Longrightarrow> fetch (findnth n) a Bk = (W1, a)" |
|
1256 apply(case_tac a, simp_all) |
|
1257 apply(induct n, auto simp: findnth.simps length_findnth nth_append) |
|
1258 apply arith |
|
1259 done |
|
1260 |
|
1261 lemma [simp]: |
|
1262 "\<lbrakk>0 < a; a < Suc (2 * n); a mod 2 = Suc 0\<rbrakk> |
|
1263 \<Longrightarrow> fetch (findnth n) a Oc = (R, Suc a)" |
|
1264 apply(case_tac a, simp_all) |
|
1265 apply(induct n, auto simp: findnth.simps length_findnth nth_append) |
|
1266 apply(subgoal_tac "nat = 2 * n", simp) |
|
1267 by arith |
|
1268 |
|
1269 lemma [simp]: |
|
1270 "\<lbrakk>0 < a; a < Suc (2*n); a mod 2 \<noteq> Suc 0\<rbrakk> |
|
1271 \<Longrightarrow> fetch (findnth n) a Oc = (R, a)" |
|
1272 apply(case_tac a, simp_all) |
|
1273 apply(induct n, auto simp: findnth.simps length_findnth nth_append) |
|
1274 apply(subgoal_tac "nat = Suc (2 * n)", simp) |
|
1275 apply arith |
|
1276 done |
|
1277 |
|
1278 lemma [simp]: |
|
1279 "\<lbrakk>0 < a; a < Suc (2*n); a mod 2 \<noteq> Suc 0\<rbrakk> |
|
1280 \<Longrightarrow> fetch (findnth n) a Bk = (R, Suc a)" |
|
1281 apply(case_tac a, simp_all) |
|
1282 apply(induct n, auto simp: findnth.simps length_findnth nth_append) |
|
1283 apply(subgoal_tac "nat = Suc (2 * n)", simp) |
|
1284 by arith |
|
1285 |
|
1286 declare at_begin_norm.simps[simp del] at_begin_fst_bwtn.simps[simp del] |
|
1287 at_begin_fst_awtn.simps[simp del] in_middle.simps[simp del] |
|
1288 abc_lm_s.simps[simp del] abc_lm_v.simps[simp del] |
|
1289 ci.simps[simp del] inv_after_move.simps[simp del] |
|
1290 inv_on_left_moving_norm.simps[simp del] |
|
1291 inv_on_left_moving_in_middle_B.simps[simp del] |
|
1292 inv_after_clear.simps[simp del] |
|
1293 inv_after_write.simps[simp del] inv_on_left_moving.simps[simp del] |
|
1294 inv_on_right_moving.simps[simp del] |
|
1295 inv_check_left_moving.simps[simp del] |
|
1296 inv_check_left_moving_in_middle.simps[simp del] |
|
1297 inv_check_left_moving_on_leftmost.simps[simp del] |
|
1298 inv_after_left_moving.simps[simp del] |
|
1299 inv_stop.simps[simp del] inv_locate_a.simps[simp del] |
|
1300 inv_locate_b.simps[simp del] |
|
1301 |
|
1302 lemma [intro]: "\<exists>rn. [Bk] = Bk \<up> rn" |
|
1303 apply(rule_tac x = "Suc 0" in exI, simp) |
|
1304 done |
|
1305 |
|
1306 lemma [intro]: "at_begin_norm (as, am) (q, aaa, []) ires |
|
1307 \<Longrightarrow> at_begin_norm (as, am) (q, aaa, [Bk]) ires" |
|
1308 apply(simp add: at_begin_norm.simps, erule_tac exE, erule_tac exE) |
|
1309 apply(rule_tac x = lm1 in exI, simp, auto) |
|
1310 done |
|
1311 |
|
1312 lemma [intro]: "at_begin_fst_bwtn (as, am) (q, aaa, []) ires |
|
1313 \<Longrightarrow> at_begin_fst_bwtn (as, am) (q, aaa, [Bk]) ires" |
|
1314 apply(simp only: at_begin_fst_bwtn.simps, erule_tac exE, erule_tac exE, erule_tac exE) |
|
1315 apply(rule_tac x = "am @ 0\<up>tn" in exI, auto) |
|
1316 done |
|
1317 |
|
1318 lemma [intro]: "at_begin_fst_awtn (as, am) (q, aaa, []) ires |
|
1319 \<Longrightarrow> at_begin_fst_awtn (as, am) (q, aaa, [Bk]) ires" |
|
1320 apply(auto simp: at_begin_fst_awtn.simps) |
|
1321 done |
|
1322 |
|
1323 lemma [intro]: "inv_locate_a (as, am) (q, aaa, []) ires |
|
1324 \<Longrightarrow> inv_locate_a (as, am) (q, aaa, [Bk]) ires" |
|
1325 apply(simp only: inv_locate_a.simps) |
|
1326 apply(erule disj_forward) |
|
1327 defer |
|
1328 apply(erule disj_forward, auto) |
|
1329 done |
|
1330 |
|
1331 lemma tape_of_nl_cons: "<m # lm> = (if lm = [] then Oc\<up>(Suc m) |
|
1332 else Oc\<up>(Suc m) @ Bk # <lm>)" |
|
1333 apply(case_tac lm, simp_all add: tape_of_nl_abv split: if_splits) |
|
1334 done |
|
1335 |
|
1336 |
|
1337 lemma locate_a_2_locate_a[simp]: "inv_locate_a (as, am) (q, aaa, Bk # xs) ires |
|
1338 \<Longrightarrow> inv_locate_a (as, am) (q, aaa, Oc # xs) ires" |
|
1339 apply(simp only: inv_locate_a.simps at_begin_norm.simps |
|
1340 at_begin_fst_bwtn.simps at_begin_fst_awtn.simps) |
|
1341 apply(erule_tac disjE, erule exE, erule exE, erule exE, |
|
1342 rule disjI2, rule disjI2) |
|
1343 defer |
|
1344 apply(erule_tac disjE, erule exE, erule exE, |
|
1345 erule exE, rule disjI2, rule disjI2) |
|
1346 prefer 2 |
|
1347 apply(simp) |
|
1348 proof- |
|
1349 fix lm1 tn rn |
|
1350 assume k: "lm1 = am @ 0\<up>tn \<and> length lm1 = q \<and> (if lm1 = [] then aaa = Bk # Bk # |
|
1351 ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> Bk # xs = Bk\<up>rn" |
|
1352 thus "\<exists>lm1 tn rn. lm1 = am @ 0 \<up> tn \<and> length lm1 = q \<and> |
|
1353 (if lm1 = [] then aaa = Bk # Bk # ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> Oc # xs = [Oc] @ Bk \<up> rn" |
|
1354 (is "\<exists>lm1 tn rn. ?P lm1 tn rn") |
|
1355 proof - |
|
1356 from k have "?P lm1 tn (rn - 1)" |
|
1357 apply(auto simp: Oc_def) |
|
1358 by(case_tac [!] "rn::nat", auto) |
|
1359 thus ?thesis by blast |
|
1360 qed |
|
1361 next |
|
1362 fix lm1 lm2 rn |
|
1363 assume h1: "am = lm1 @ lm2 \<and> length lm1 = q \<and> (if lm1 = [] |
|
1364 then aaa = Bk # Bk # ires else aaa = Bk # <rev lm1> @ Bk # Bk # ires) \<and> |
|
1365 Bk # xs = <lm2> @ Bk\<up>rn" |
|
1366 from h1 have h2: "lm2 = []" |
|
1367 apply(auto split: if_splits) |
|
1368 apply(case_tac [!] lm2, simp_all add: tape_of_nl_cons split: if_splits) |
|
1369 done |
|
1370 from h1 and h2 show "\<exists>lm1 tn rn. lm1 = am @ 0\<up>tn \<and> length lm1 = q \<and> |
|
1371 (if lm1 = [] then aaa = Bk # Bk # ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> |
|
1372 Oc # xs = [Oc] @ Bk\<up>rn" |
|
1373 (is "\<exists>lm1 tn rn. ?P lm1 tn rn") |
|
1374 proof - |
|
1375 from h1 and h2 have "?P lm1 0 (rn - 1)" |
|
1376 apply(auto simp: Oc_def |
|
1377 tape_of_nl_abv tape_of_nat_list.simps) |
|
1378 by(case_tac "rn::nat", simp, simp) |
|
1379 thus ?thesis by blast |
|
1380 qed |
|
1381 qed |
|
1382 |
|
1383 lemma [simp]: "inv_locate_a (as, am) (q, aaa, []) ires \<Longrightarrow> |
|
1384 inv_locate_a (as, am) (q, aaa, [Oc]) ires" |
|
1385 apply(insert locate_a_2_locate_a [of as am q aaa "[]"]) |
|
1386 apply(subgoal_tac "inv_locate_a (as, am) (q, aaa, [Bk]) ires", auto) |
|
1387 done |
|
1388 |
|
1389 (*inv: from locate_b to locate_b*) |
|
1390 lemma [simp]: "inv_locate_b (as, am) (q, aaa, Oc # xs) ires |
|
1391 \<Longrightarrow> inv_locate_b (as, am) (q, Oc # aaa, xs) ires" |
|
1392 apply(simp only: inv_locate_b.simps in_middle.simps) |
|
1393 apply(erule exE)+ |
|
1394 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
1395 rule_tac x = tn in exI, rule_tac x = m in exI) |
|
1396 apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - 1" in exI, |
|
1397 rule_tac x = rn in exI) |
|
1398 apply(case_tac mr, simp_all, auto) |
|
1399 done |
|
1400 |
|
1401 (* |
|
1402 lemma zero_and_nil[intro]: "(Bk # Bk\<^bsup>n\<^esup> = Oc\<^bsup>mr\<^esup> @ Bk # <lm::nat list> @ |
|
1403 Bk\<^bsup>rn \<^esup>) \<or> (lm2 = [] \<and> Bk # Bk\<^bsup>n\<^esup> = Oc\<^bsup>mr\<^esup>) |
|
1404 \<Longrightarrow> mr = 0 \<and> lm = []" |
|
1405 apply(rule context_conjI) |
|
1406 apply(case_tac mr, auto simp:exponent_def) |
|
1407 apply(insert BkCons_nil[of "replicate (n - 1) Bk" lm rn]) |
|
1408 apply(case_tac n, auto simp: exponent_def Bk_def tape_of_nl_nil_eq) |
|
1409 done |
|
1410 |
|
1411 lemma tape_of_nat_def: "<[m::nat]> = Oc # Oc\<^bsup>m\<^esup>" |
|
1412 apply(simp add: tape_of_nl_abv tape_of_nat_list.simps) |
|
1413 done |
|
1414 *) |
|
1415 lemma [simp]: "<[x::nat]> = Oc\<up>(Suc x)" |
|
1416 apply(simp add: tape_of_nat_abv tape_of_nl_abv) |
|
1417 done |
|
1418 |
|
1419 lemma [simp]: " <([]::nat list)> = []" |
|
1420 apply(simp add: tape_of_nl_abv) |
|
1421 done |
|
1422 |
|
1423 lemma [simp]: "\<lbrakk>inv_locate_b (as, am) (q, aaa, Bk # xs) ires; \<exists>n. xs = Bk\<up>n\<rbrakk> |
|
1424 \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires" |
|
1425 apply(simp add: inv_locate_b.simps inv_locate_a.simps) |
|
1426 apply(rule_tac disjI2, rule_tac disjI1) |
|
1427 apply(simp only: in_middle.simps at_begin_fst_bwtn.simps) |
|
1428 apply(erule_tac exE)+ |
|
1429 apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = tn in exI, simp split: if_splits) |
|
1430 apply(case_tac mr, simp_all) |
|
1431 apply(case_tac "length am", simp_all, case_tac tn, simp_all) |
|
1432 apply(case_tac lm2, simp_all add: tape_of_nl_cons split: if_splits) |
|
1433 apply(case_tac am, simp_all) |
|
1434 apply(case_tac n, simp_all) |
|
1435 apply(case_tac n, simp_all) |
|
1436 apply(case_tac mr, simp_all) |
|
1437 apply(case_tac lm2, simp_all add: tape_of_nl_cons split: if_splits, auto) |
|
1438 apply(case_tac [!] n, simp_all) |
|
1439 done |
|
1440 |
|
1441 lemma [simp]: "(Oc # r = Bk \<up> rn) = False" |
|
1442 apply(case_tac rn, simp_all) |
|
1443 done |
|
1444 |
|
1445 lemma [simp]: "(\<exists>rna. Bk \<up> rn = Bk # Bk \<up> rna) \<or> rn = 0" |
|
1446 apply(case_tac rn, auto) |
|
1447 done |
|
1448 |
|
1449 lemma [simp]: "(\<forall> x. a \<noteq> x) = False" |
|
1450 by auto |
|
1451 |
|
1452 lemma exp_ind: "a\<up>(Suc x) = a\<up>x @ [a]" |
|
1453 apply(induct x, auto) |
|
1454 done |
|
1455 |
|
1456 lemma [simp]: |
|
1457 "inv_locate_a (as, lm) (q, l, Oc # r) ires |
|
1458 \<Longrightarrow> inv_locate_b (as, lm) (q, Oc # l, r) ires" |
|
1459 apply(simp only: inv_locate_a.simps inv_locate_b.simps in_middle.simps |
|
1460 at_begin_norm.simps at_begin_fst_bwtn.simps |
|
1461 at_begin_fst_awtn.simps) |
|
1462 apply(erule disjE, erule exE, erule exE, erule exE) |
|
1463 apply(rule_tac x = lm1 in exI, rule_tac x = "tl lm2" in exI, simp) |
|
1464 apply(rule_tac x = 0 in exI, rule_tac x = "hd lm2" in exI) |
|
1465 apply(case_tac lm2, auto simp: tape_of_nl_cons ) |
|
1466 apply(rule_tac x = 1 in exI, rule_tac x = a in exI, auto) |
|
1467 apply(case_tac list, simp_all) |
|
1468 apply(case_tac rn, simp_all) |
|
1469 apply(rule_tac x = "lm @ replicate tn 0" in exI, |
|
1470 rule_tac x = "[]" in exI, |
|
1471 rule_tac x = "Suc tn" in exI, |
|
1472 rule_tac x = 0 in exI, auto) |
|
1473 apply(simp only: replicate_Suc[THEN sym] exp_ind) |
|
1474 apply(rule_tac x = "Suc 0" in exI, auto) |
|
1475 done |
|
1476 |
|
1477 lemma length_equal: "xs = ys \<Longrightarrow> length xs = length ys" |
|
1478 by auto |
|
1479 |
|
1480 lemma [simp]: "\<lbrakk>inv_locate_b (as, am) (q, aaa, Bk # xs) ires; |
|
1481 \<not> (\<exists>n. xs = Bk\<up>n)\<rbrakk> |
|
1482 \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires" |
|
1483 apply(simp add: inv_locate_b.simps inv_locate_a.simps) |
|
1484 apply(rule_tac disjI1) |
|
1485 apply(simp only: in_middle.simps at_begin_norm.simps) |
|
1486 apply(erule_tac exE)+ |
|
1487 apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = lm2 in exI, simp) |
|
1488 apply(subgoal_tac "tn = 0", simp , auto split: if_splits) |
|
1489 apply(case_tac [!] mr, simp_all, auto) |
|
1490 apply(simp add: tape_of_nl_cons) |
|
1491 apply(drule_tac length_equal, simp) |
|
1492 apply(case_tac "length am", simp_all, erule_tac x = rn in allE, simp) |
|
1493 apply(drule_tac length_equal, simp) |
|
1494 apply(case_tac "(Suc (length lm1) - length am)", simp_all) |
|
1495 apply(case_tac lm2, simp, simp) |
|
1496 done |
|
1497 |
|
1498 lemma locate_b_2_a[intro]: |
|
1499 "inv_locate_b (as, am) (q, aaa, Bk # xs) ires |
|
1500 \<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires" |
|
1501 apply(case_tac "\<exists> n. xs = Bk\<up>n", simp, simp) |
|
1502 done |
|
1503 |
|
1504 |
|
1505 lemma [simp]: "inv_locate_b (as, am) (q, l, []) ires |
|
1506 \<Longrightarrow> inv_locate_b (as, am) (q, l, [Bk]) ires" |
|
1507 apply(simp only: inv_locate_b.simps in_middle.simps) |
|
1508 apply(erule exE)+ |
|
1509 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
1510 rule_tac x = tn in exI, rule_tac x = m in exI, |
|
1511 rule_tac x = ml in exI, rule_tac x = mr in exI) |
|
1512 apply(auto) |
|
1513 done |
|
1514 |
|
1515 (*inv: from locate_b to after_write*) |
|
1516 |
|
1517 lemma [simp]: "(a mod 2 \<noteq> Suc 0) = (a mod 2 = 0) " |
|
1518 by arith |
|
1519 |
|
1520 lemma [simp]: "(a mod 2 \<noteq> 0) = (a mod 2 = Suc 0) " |
|
1521 by arith |
|
1522 |
|
1523 lemma mod_ex1: "(a mod 2 = Suc 0) = (\<exists> q. a = Suc (2 * q))" |
|
1524 by arith |
|
1525 |
|
1526 lemma mod_ex2: "(a mod (2::nat) = 0) = (\<exists> q. a = 2 * q)" |
|
1527 by arith |
|
1528 |
|
1529 lemma [simp]: "(2*q - Suc 0) div 2 = (q - 1)" |
|
1530 by arith |
|
1531 |
|
1532 lemma [simp]: "(Suc (2*q)) div 2 = q" |
|
1533 by arith |
|
1534 |
|
1535 lemma mod_2: "x mod 2 = 0 \<or> x mod 2 = Suc 0" |
|
1536 by arith |
|
1537 |
|
1538 lemma [simp]: "x mod 2 = 0 \<Longrightarrow> Suc x mod 2 = Suc 0" |
|
1539 by arith |
|
1540 |
|
1541 lemma [simp]: "x mod 2 = Suc 0 \<Longrightarrow> Suc x mod 2 = 0" |
|
1542 by arith |
|
1543 |
|
1544 lemma [simp]: "inv_locate_b (as, am) (q, l, []) ires |
|
1545 \<Longrightarrow> inv_locate_b (as, am) (q, l, [Bk]) ires" |
|
1546 apply(simp only: inv_locate_b.simps in_middle.simps) |
|
1547 apply(erule exE)+ |
|
1548 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
1549 rule_tac x = tn in exI, rule_tac x = m in exI, |
|
1550 rule_tac x = ml in exI, rule_tac x = mr in exI) |
|
1551 apply(auto) |
|
1552 done |
|
1553 |
|
1554 lemma locate_b_2_locate_a[simp]: |
|
1555 "\<lbrakk>q > 0; inv_locate_b (as, am) (q - Suc 0, aaa, Bk # xs) ires\<rbrakk> |
|
1556 \<Longrightarrow> inv_locate_a (as, am) (q, Bk # aaa, xs) ires" |
|
1557 apply(insert locate_b_2_a [of as am "q - 1" aaa xs ires], simp) |
|
1558 done |
|
1559 |
|
1560 |
|
1561 lemma [simp]: "inv_locate_b (as, am) (q, l, []) ires |
|
1562 \<Longrightarrow> inv_locate_b (as, am) (q, l, [Bk]) ires" |
|
1563 apply(simp only: inv_locate_b.simps in_middle.simps) |
|
1564 apply(erule exE)+ |
|
1565 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
1566 rule_tac x = tn in exI, rule_tac x = m in exI, |
|
1567 rule_tac x = ml in exI, rule_tac x = mr in exI) |
|
1568 apply(auto) |
|
1569 done |
|
1570 |
|
1571 (*inv: from locate_b to after_write*) |
|
1572 |
|
1573 lemma [simp]: |
|
1574 "crsp (layout_of ap) (as, lm) (s, l, r) ires |
|
1575 \<Longrightarrow> findnth_inv (layout_of ap) n (as, lm) (Suc 0, l, r) ires" |
|
1576 apply(auto simp: crsp.simps findnth_inv.simps inv_locate_a.simps |
|
1577 at_begin_norm.simps at_begin_fst_awtn.simps at_begin_fst_bwtn.simps) |
|
1578 done |
|
1579 |
|
1580 lemma findnth_correct_pre: |
|
1581 assumes layout: "ly = layout_of ap" |
|
1582 and crsp: "crsp ly (as, lm) (s, l, r) ires" |
|
1583 and not0: "n > 0" |
|
1584 and f: "f = (\<lambda> stp. (steps (Suc 0, l, r) (findnth n, 0) stp, n))" |
|
1585 and P: "P = (\<lambda> ((s, l, r), n). s = Suc (2 * n))" |
|
1586 and Q: "Q = (\<lambda> ((s, l, r), n). findnth_inv ly n (as, lm) (s, l, r) ires)" |
|
1587 shows "\<exists> stp. P (f stp) \<and> Q (f stp)" |
|
1588 thm halt_lemma2 |
|
1589 proof(rule_tac LE = findnth_LE in halt_lemma2) |
|
1590 show "wf findnth_LE" by(intro wf_findnth_LE) |
|
1591 next |
|
1592 show "Q (f 0)" |
|
1593 using crsp layout |
|
1594 apply(simp add: f P Q steps.simps) |
|
1595 done |
|
1596 next |
|
1597 show "\<not> P (f 0)" |
|
1598 using not0 |
|
1599 apply(simp add: f P steps.simps) |
|
1600 done |
|
1601 next |
|
1602 show "\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) |
|
1603 \<in> findnth_LE" |
|
1604 proof(rule_tac allI, rule_tac impI, simp add: f, |
|
1605 case_tac "steps (Suc 0, l, r) (findnth n, 0) na", simp add: P) |
|
1606 fix na a b c |
|
1607 assume "a \<noteq> Suc (2 * n) \<and> Q ((a, b, c), n)" |
|
1608 thus "Q (step (a, b, c) (findnth n, 0), n) \<and> |
|
1609 ((step (a, b, c) (findnth n, 0), n), (a, b, c), n) \<in> findnth_LE" |
|
1610 apply(case_tac c, case_tac [2] aa) |
|
1611 apply(simp_all add: step.simps findnth_LE_def Q findnth_inv.simps mod_2 lex_pair_def split: if_splits) |
|
1612 apply(auto simp: mod_ex1 mod_ex2) |
|
1613 done |
|
1614 qed |
|
1615 qed |
|
1616 |
|
1617 lemma [intro]: "inv_locate_a (as, lm) (0, Bk # Bk # ires, <lm> @ Bk \<up> x) ires" |
|
1618 apply(auto simp: crsp.simps inv_locate_a.simps at_begin_norm.simps) |
|
1619 done |
|
1620 lemma [simp]: "crsp ly (as, lm) (s, l, r) ires \<Longrightarrow> inv_locate_a (as, lm) (0, l, r) ires" |
|
1621 apply(auto simp: crsp.simps inv_locate_a.simps at_begin_norm.simps) |
|
1622 done |
|
1623 |
|
1624 lemma findnth_correct: |
|
1625 assumes layout: "ly = layout_of ap" |
|
1626 and crsp: "crsp ly (as, lm) (s, l, r) ires" |
|
1627 shows "\<exists> stp l' r'. steps (Suc 0, l, r) (findnth n, 0) stp = (Suc (2 * n), l', r') |
|
1628 \<and> inv_locate_a (as, lm) (n, l', r') ires" |
|
1629 using crsp |
|
1630 apply(case_tac "n = 0") |
|
1631 apply(rule_tac x = 0 in exI, auto simp: steps.simps) |
|
1632 using assms |
|
1633 apply(drule_tac findnth_correct_pre, auto) |
|
1634 apply(rule_tac x = stp in exI, simp add: findnth_inv.simps) |
|
1635 done |
|
1636 |
|
1637 |
|
1638 fun inc_inv :: "nat \<Rightarrow> inc_inv_t" |
|
1639 where |
|
1640 "inc_inv n (as, lm) (s, l, r) ires = |
|
1641 (let lm' = abc_lm_s lm n (Suc (abc_lm_v lm n)) in |
|
1642 if s = 0 then False |
|
1643 else if s = 1 then |
|
1644 inv_locate_a (as, lm) (n, l, r) ires |
|
1645 else if s = 2 then |
|
1646 inv_locate_b (as, lm) (n, l, r) ires |
|
1647 else if s = 3 then |
|
1648 inv_after_write (as, lm') (s, l, r) ires |
|
1649 else if s = Suc 3 then |
|
1650 inv_after_move (as, lm') (s, l, r) ires |
|
1651 else if s = Suc 4 then |
|
1652 inv_after_clear (as, lm') (s, l, r) ires |
|
1653 else if s = Suc (Suc 4) then |
|
1654 inv_on_right_moving (as, lm') (s, l, r) ires |
|
1655 else if s = Suc (Suc 5) then |
|
1656 inv_on_left_moving (as, lm') (s, l, r) ires |
|
1657 else if s = Suc (Suc (Suc 5)) then |
|
1658 inv_check_left_moving (as, lm') (s, l, r) ires |
|
1659 else if s = Suc (Suc (Suc (Suc 5))) then |
|
1660 inv_after_left_moving (as, lm') (s, l, r) ires |
|
1661 else if s = Suc (Suc (Suc (Suc (Suc 5)))) then |
|
1662 inv_stop (as, lm') (s, l, r) ires |
|
1663 else False)" |
|
1664 |
|
1665 |
|
1666 fun abc_inc_stage1 :: "config \<Rightarrow> nat" |
|
1667 where |
|
1668 "abc_inc_stage1 (s, l, r) = |
|
1669 (if s = 0 then 0 |
|
1670 else if s \<le> 2 then 5 |
|
1671 else if s \<le> 6 then 4 |
|
1672 else if s \<le> 8 then 3 |
|
1673 else if s = 9 then 2 |
|
1674 else 1)" |
|
1675 |
|
1676 fun abc_inc_stage2 :: "config \<Rightarrow> nat" |
|
1677 where |
|
1678 "abc_inc_stage2 (s, l, r) = |
|
1679 (if s = 1 then 2 |
|
1680 else if s = 2 then 1 |
|
1681 else if s = 3 then length r |
|
1682 else if s = 4 then length r |
|
1683 else if s = 5 then length r |
|
1684 else if s = 6 then |
|
1685 if r \<noteq> [] then length r |
|
1686 else 1 |
|
1687 else if s = 7 then length l |
|
1688 else if s = 8 then length l |
|
1689 else 0)" |
|
1690 |
|
1691 fun abc_inc_stage3 :: "config \<Rightarrow> nat" |
|
1692 where |
|
1693 "abc_inc_stage3 (s, l, r) = ( |
|
1694 if s = 4 then 4 |
|
1695 else if s = 5 then 3 |
|
1696 else if s = 6 then |
|
1697 if r \<noteq> [] \<and> hd r = Oc then 2 |
|
1698 else 1 |
|
1699 else if s = 3 then 0 |
|
1700 else if s = 2 then length r |
|
1701 else if s = 1 then |
|
1702 if (r \<noteq> [] \<and> hd r = Oc) then 0 |
|
1703 else 1 |
|
1704 else 10 - s)" |
|
1705 |
|
1706 |
|
1707 definition inc_measure :: "config \<Rightarrow> nat \<times> nat \<times> nat" |
|
1708 where |
|
1709 "inc_measure c = |
|
1710 (abc_inc_stage1 c, abc_inc_stage2 c, abc_inc_stage3 c)" |
|
1711 |
|
1712 definition lex_triple :: |
|
1713 "((nat \<times> (nat \<times> nat)) \<times> (nat \<times> (nat \<times> nat))) set" |
|
1714 where "lex_triple \<equiv> less_than <*lex*> lex_pair" |
|
1715 |
|
1716 definition inc_LE :: "(config \<times> config) set" |
|
1717 where |
|
1718 "inc_LE \<equiv> (inv_image lex_triple inc_measure)" |
|
1719 |
|
1720 declare inc_inv.simps[simp del] |
|
1721 |
|
1722 lemma wf_inc_le[intro]: "wf inc_LE" |
|
1723 by(auto intro:wf_inv_image simp: inc_LE_def lex_triple_def lex_pair_def) |
|
1724 |
|
1725 lemma numeral_5_eq_5: "5 = Suc (Suc (Suc (Suc (Suc 0))))" |
|
1726 by arith |
|
1727 |
|
1728 lemma numeral_6_eq_6: "6 = Suc (Suc (Suc (Suc (Suc 1))))" |
|
1729 by arith |
|
1730 |
|
1731 lemma numeral_7_eq_7: "7 = Suc (Suc (Suc (Suc (Suc 2))))" |
|
1732 by arith |
|
1733 |
|
1734 lemma numeral_8_eq_8: "8 = Suc (Suc (Suc (Suc (Suc 3))))" |
|
1735 by arith |
|
1736 |
|
1737 lemma numeral_9_eq_9: "9 = Suc (Suc (Suc (Suc (Suc (Suc 3)))))" |
|
1738 by arith |
|
1739 |
|
1740 lemma numeral_10_eq_10: "10 = Suc (Suc (Suc (Suc (Suc (Suc (Suc 3))))))" |
|
1741 by arith |
|
1742 |
|
1743 lemma inv_locate_b_2_after_write[simp]: |
|
1744 "inv_locate_b (as, am) (n, aaa, Bk # xs) ires |
|
1745 \<Longrightarrow> inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) |
|
1746 (s, aaa, Oc # xs) ires" |
|
1747 apply(auto simp: in_middle.simps inv_after_write.simps |
|
1748 abc_lm_v.simps abc_lm_s.simps inv_locate_b.simps) |
|
1749 apply(case_tac [!] mr, auto split: if_splits) |
|
1750 apply(rule_tac x = rn in exI, rule_tac x = "Suc m" in exI, |
|
1751 rule_tac x = "lm1" in exI, simp) |
|
1752 apply(rule_tac x = "lm2" in exI, simp) |
|
1753 apply(simp only: Suc_diff_le exp_ind) |
|
1754 apply(subgoal_tac "lm2 = []", simp) |
|
1755 apply(drule_tac length_equal, simp) |
|
1756 done |
|
1757 |
|
1758 lemma [simp]: "inv_locate_b (as, am) (n, aaa, []) ires \<Longrightarrow> |
|
1759 inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) |
|
1760 (s, aaa, [Oc]) ires" |
|
1761 apply(insert inv_locate_b_2_after_write [of as am n aaa "[]"]) |
|
1762 by(simp) |
|
1763 |
|
1764 |
|
1765 |
|
1766 (*inv: from after_write to after_move*) |
|
1767 lemma [simp]: "inv_after_write (as, lm) (x, l, Oc # r) ires |
|
1768 \<Longrightarrow> inv_after_move (as, lm) (y, Oc # l, r) ires" |
|
1769 apply(auto simp:inv_after_move.simps inv_after_write.simps split: if_splits) |
|
1770 done |
|
1771 |
|
1772 lemma [simp]: "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n) |
|
1773 )) (x, aaa, Bk # xs) ires = False" |
|
1774 apply(simp add: inv_after_write.simps ) |
|
1775 done |
|
1776 |
|
1777 lemma [simp]: |
|
1778 "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) |
|
1779 (x, aaa, []) ires = False" |
|
1780 apply(simp add: inv_after_write.simps ) |
|
1781 done |
|
1782 |
|
1783 (*inv: from after_move to after_clear*) |
|
1784 lemma [simp]: "inv_after_move (as, lm) (s, l, Oc # r) ires |
|
1785 \<Longrightarrow> inv_after_clear (as, lm) (s', l, Bk # r) ires" |
|
1786 apply(auto simp: inv_after_move.simps inv_after_clear.simps split: if_splits) |
|
1787 done |
|
1788 |
|
1789 (*inv: from after_move to on_leftmoving*) |
|
1790 lemma [intro]: "Bk \<up> rn = Bk # Bk \<up> (rn - Suc 0) \<or> rn = 0" |
|
1791 apply(case_tac rn, auto) |
|
1792 done |
|
1793 |
|
1794 lemma inv_after_move_2_inv_on_left_moving[simp]: |
|
1795 "inv_after_move (as, lm) (s, l, Bk # r) ires |
|
1796 \<Longrightarrow> (l = [] \<longrightarrow> |
|
1797 inv_on_left_moving (as, lm) (s', [], Bk # Bk # r) ires) \<and> |
|
1798 (l \<noteq> [] \<longrightarrow> |
|
1799 inv_on_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)" |
|
1800 apply(simp only: inv_after_move.simps inv_on_left_moving.simps) |
|
1801 apply(subgoal_tac "l \<noteq> []", rule conjI, simp, rule impI, |
|
1802 rule disjI1, simp only: inv_on_left_moving_norm.simps) |
|
1803 apply(erule exE)+ |
|
1804 apply(subgoal_tac "lm2 = []") |
|
1805 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
1806 rule_tac x = m in exI, rule_tac x = m in exI, |
|
1807 rule_tac x = 1 in exI, |
|
1808 rule_tac x = "rn - 1" in exI, auto) |
|
1809 apply(auto split: if_splits) |
|
1810 apply(case_tac [1-2] rn, simp_all) |
|
1811 apply(case_tac [!] lm2, simp_all add: tape_of_nl_cons split: if_splits) |
|
1812 done |
|
1813 |
|
1814 |
|
1815 lemma inv_after_move_2_inv_on_left_moving_B[simp]: |
|
1816 "inv_after_move (as, lm) (s, l, []) ires |
|
1817 \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s', [], [Bk]) ires) \<and> |
|
1818 (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s', tl l, [hd l]) ires)" |
|
1819 apply(simp only: inv_after_move.simps inv_on_left_moving.simps) |
|
1820 apply(subgoal_tac "l \<noteq> []", rule conjI, simp, rule impI, rule disjI1, |
|
1821 simp only: inv_on_left_moving_norm.simps) |
|
1822 apply(erule exE)+ |
|
1823 apply(subgoal_tac "lm2 = []") |
|
1824 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
1825 rule_tac x = m in exI, rule_tac x = m in exI, |
|
1826 rule_tac x = 1 in exI, rule_tac x = "rn - 1" in exI, simp, case_tac rn) |
|
1827 apply(auto split: if_splits) |
|
1828 apply(case_tac [!] lm2, auto simp: tape_of_nl_cons split: if_splits) |
|
1829 done |
|
1830 |
|
1831 (*inv: from after_clear to on_right_moving*) |
|
1832 lemma [simp]: "Oc # r = replicate rn Bk = False" |
|
1833 apply(case_tac rn, simp, simp) |
|
1834 done |
|
1835 |
|
1836 lemma inv_after_clear_2_inv_on_right_moving[simp]: |
|
1837 "inv_after_clear (as, lm) (x, l, Bk # r) ires |
|
1838 \<Longrightarrow> inv_on_right_moving (as, lm) (y, Bk # l, r) ires" |
|
1839 apply(auto simp: inv_after_clear.simps inv_on_right_moving.simps ) |
|
1840 apply(subgoal_tac "lm2 \<noteq> []") |
|
1841 apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI, |
|
1842 rule_tac x = "hd lm2" in exI, simp) |
|
1843 apply(rule_tac x = 0 in exI, rule_tac x = "hd lm2" in exI) |
|
1844 apply(simp, rule conjI) |
|
1845 apply(case_tac [!] "lm2::nat list", auto) |
|
1846 apply(case_tac rn, auto split: if_splits simp: tape_of_nl_cons) |
|
1847 apply(case_tac [!] rn, simp_all) |
|
1848 done |
|
1849 |
|
1850 lemma [simp]: "inv_after_clear (as, lm) (x, l, []) ires\<Longrightarrow> |
|
1851 inv_after_clear (as, lm) (y, l, [Bk]) ires" |
|
1852 by(auto simp: inv_after_clear.simps) |
|
1853 |
|
1854 lemma [simp]: "inv_after_clear (as, lm) (x, l, []) ires |
|
1855 \<Longrightarrow> inv_on_right_moving (as, lm) (y, Bk # l, []) ires" |
|
1856 by(insert |
|
1857 inv_after_clear_2_inv_on_right_moving[of as lm n l "[]"], simp) |
|
1858 |
|
1859 (*inv: from on_right_moving to on_right_movign*) |
|
1860 lemma [simp]: "inv_on_right_moving (as, lm) (x, l, Oc # r) ires |
|
1861 \<Longrightarrow> inv_on_right_moving (as, lm) (y, Oc # l, r) ires" |
|
1862 apply(auto simp: inv_on_right_moving.simps) |
|
1863 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
1864 rule_tac x = "ml + mr" in exI, simp) |
|
1865 apply(rule_tac x = "Suc ml" in exI, |
|
1866 rule_tac x = "mr - 1" in exI, simp) |
|
1867 apply(case_tac mr, auto) |
|
1868 apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, |
|
1869 rule_tac x = "ml + mr" in exI, simp) |
|
1870 apply(rule_tac x = "Suc ml" in exI, |
|
1871 rule_tac x = "mr - 1" in exI, simp) |
|
1872 apply(case_tac mr, auto split: if_splits) |
|
1873 done |
|
1874 |
|
1875 lemma inv_on_right_moving_2_inv_on_right_moving[simp]: |
|
1876 "inv_on_right_moving (as, lm) (x, l, Bk # r) ires |
|
1877 \<Longrightarrow> inv_after_write (as, lm) (y, l, Oc # r) ires" |
|
1878 apply(auto simp: inv_on_right_moving.simps inv_after_write.simps ) |
|
1879 apply(case_tac mr, auto simp: split: if_splits) |
|
1880 apply(case_tac [!] mr, simp_all) |
|
1881 done |
|
1882 |
|
1883 lemma [simp]: "inv_on_right_moving (as, lm) (x, l, []) ires\<Longrightarrow> |
|
1884 inv_on_right_moving (as, lm) (y, l, [Bk]) ires" |
|
1885 apply(auto simp: inv_on_right_moving.simps) |
|
1886 apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, simp) |
|
1887 done |
|
1888 |
|
1889 (*inv: from on_right_moving to after_write*) |
|
1890 lemma [simp]: "inv_on_right_moving (as, lm) (x, l, []) ires |
|
1891 \<Longrightarrow> inv_after_write (as, lm) (y, l, [Oc]) ires" |
|
1892 apply(rule_tac inv_on_right_moving_2_inv_on_right_moving, simp) |
|
1893 done |
|
1894 |
|
1895 (*inv: from on_left_moving to on_left_moving*) |
|
1896 lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) |
|
1897 (s, l, Oc # r) ires = False" |
|
1898 apply(auto simp: inv_on_left_moving_in_middle_B.simps ) |
|
1899 done |
|
1900 |
|
1901 lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, l, Bk # r) ires |
|
1902 = False" |
|
1903 apply(auto simp: inv_on_left_moving_norm.simps) |
|
1904 apply(case_tac [!] mr, auto simp: ) |
|
1905 done |
|
1906 |
|
1907 lemma [simp]: |
|
1908 "\<lbrakk>inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires; |
|
1909 hd l = Bk; l \<noteq> []\<rbrakk> \<Longrightarrow> |
|
1910 inv_on_left_moving_in_middle_B (as, lm) (s, tl l, Bk # Oc # r) ires" |
|
1911 apply(case_tac l, simp, simp) |
|
1912 apply(simp only: inv_on_left_moving_norm.simps |
|
1913 inv_on_left_moving_in_middle_B.simps) |
|
1914 apply(erule_tac exE)+ |
|
1915 apply(rule_tac x = lm1 in exI, rule_tac x = "m # lm2" in exI, auto) |
|
1916 apply(case_tac [!] ml, auto) |
|
1917 apply(auto simp: tape_of_nl_cons split: if_splits) |
|
1918 apply(rule_tac [!] x = "Suc rn" in exI, simp_all) |
|
1919 done |
|
1920 |
|
1921 lemma [simp]: "\<lbrakk>inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires; |
|
1922 hd l = Oc; l \<noteq> []\<rbrakk> |
|
1923 \<Longrightarrow> inv_on_left_moving_norm (as, lm) |
|
1924 (s, tl l, Oc # Oc # r) ires" |
|
1925 apply(simp only: inv_on_left_moving_norm.simps) |
|
1926 apply(erule exE)+ |
|
1927 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
1928 rule_tac x = m in exI, rule_tac x = "ml - 1" in exI, |
|
1929 rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI, simp) |
|
1930 apply(case_tac ml, auto simp: split: if_splits) |
|
1931 done |
|
1932 |
|
1933 lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, [], Oc # r) ires |
|
1934 \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # Oc # r) ires" |
|
1935 apply(auto simp: inv_on_left_moving_norm.simps |
|
1936 inv_on_left_moving_in_middle_B.simps split: if_splits) |
|
1937 done |
|
1938 |
|
1939 lemma [simp]:"inv_on_left_moving (as, lm) (s, l, Oc # r) ires |
|
1940 \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s, [], Bk # Oc # r) ires) |
|
1941 \<and> (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s, tl l, hd l # Oc # r) ires)" |
|
1942 apply(simp add: inv_on_left_moving.simps) |
|
1943 apply(case_tac "l \<noteq> []", rule conjI, simp, simp) |
|
1944 apply(case_tac "hd l", simp, simp, simp) |
|
1945 done |
|
1946 |
|
1947 (*inv: from on_left_moving to check_left_moving*) |
|
1948 lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) |
|
1949 (s, Bk # list, Bk # r) ires |
|
1950 \<Longrightarrow> inv_check_left_moving_on_leftmost (as, lm) |
|
1951 (s', list, Bk # Bk # r) ires" |
|
1952 apply(auto simp: inv_on_left_moving_in_middle_B.simps |
|
1953 inv_check_left_moving_on_leftmost.simps split: if_splits) |
|
1954 apply(case_tac [!] "rev lm1", simp_all) |
|
1955 apply(case_tac [!] lista, simp_all add: tape_of_nl_abv tape_of_nat_list.simps) |
|
1956 done |
|
1957 |
|
1958 lemma [simp]: |
|
1959 "inv_check_left_moving_in_middle (as, lm) (s, l, Bk # r) ires= False" |
|
1960 by(auto simp: inv_check_left_moving_in_middle.simps ) |
|
1961 |
|
1962 lemma [simp]: |
|
1963 "inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # r) ires\<Longrightarrow> |
|
1964 inv_check_left_moving_on_leftmost (as, lm) (s', [], Bk # Bk # r) ires" |
|
1965 apply(auto simp: inv_on_left_moving_in_middle_B.simps |
|
1966 inv_check_left_moving_on_leftmost.simps split: if_splits) |
|
1967 done |
|
1968 |
|
1969 lemma [simp]: "inv_check_left_moving_on_leftmost (as, lm) |
|
1970 (s, list, Oc # r) ires= False" |
|
1971 by(auto simp: inv_check_left_moving_on_leftmost.simps split: if_splits) |
|
1972 |
|
1973 lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm) |
|
1974 (s, Oc # list, Bk # r) ires |
|
1975 \<Longrightarrow> inv_check_left_moving_in_middle (as, lm) (s', list, Oc # Bk # r) ires" |
|
1976 apply(auto simp: inv_on_left_moving_in_middle_B.simps |
|
1977 inv_check_left_moving_in_middle.simps split: if_splits) |
|
1978 done |
|
1979 |
|
1980 lemma inv_on_left_moving_2_check_left_moving[simp]: |
|
1981 "inv_on_left_moving (as, lm) (s, l, Bk # r) ires |
|
1982 \<Longrightarrow> (l = [] \<longrightarrow> inv_check_left_moving (as, lm) (s', [], Bk # Bk # r) ires) |
|
1983 \<and> (l \<noteq> [] \<longrightarrow> |
|
1984 inv_check_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)" |
|
1985 apply(simp add: inv_on_left_moving.simps inv_check_left_moving.simps) |
|
1986 apply(case_tac l, simp, simp) |
|
1987 apply(case_tac a, simp, simp) |
|
1988 done |
|
1989 |
|
1990 lemma [simp]: "inv_on_left_moving_norm (as, lm) (s, l, []) ires = False" |
|
1991 apply(auto simp: inv_on_left_moving_norm.simps) |
|
1992 done |
|
1993 |
|
1994 lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires\<Longrightarrow> |
|
1995 inv_on_left_moving (as, lm) (6 + 2 * n, l, [Bk]) ires" |
|
1996 apply(simp add: inv_on_left_moving.simps) |
|
1997 apply(auto simp: inv_on_left_moving_in_middle_B.simps) |
|
1998 done |
|
1999 |
|
2000 lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires = False" |
|
2001 apply(simp add: inv_on_left_moving.simps) |
|
2002 apply(simp add: inv_on_left_moving_in_middle_B.simps) |
|
2003 done |
|
2004 |
|
2005 lemma [simp]: "inv_on_left_moving (as, lm) (s, l, []) ires |
|
2006 \<Longrightarrow> (l = [] \<longrightarrow> inv_check_left_moving (as, lm) (s', [], [Bk]) ires) \<and> |
|
2007 (l \<noteq> [] \<longrightarrow> inv_check_left_moving (as, lm) (s', tl l, [hd l]) ires)" |
|
2008 by simp |
|
2009 |
|
2010 lemma [intro]: "\<exists>rna. Bk # Bk \<up> rn = Bk \<up> rna" |
|
2011 apply(rule_tac x = "Suc rn" in exI, simp) |
|
2012 done |
|
2013 |
|
2014 lemma |
|
2015 inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[simp]: |
|
2016 "inv_check_left_moving_in_middle (as, lm) (s, Bk # list, Oc # r) ires |
|
2017 \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s', list, Bk # Oc # r) ires" |
|
2018 apply(simp only: inv_check_left_moving_in_middle.simps |
|
2019 inv_on_left_moving_in_middle_B.simps) |
|
2020 apply(erule_tac exE)+ |
|
2021 apply(rule_tac x = "rev (tl (rev lm1))" in exI, |
|
2022 rule_tac x = "[hd (rev lm1)] @ lm2" in exI, auto) |
|
2023 apply(case_tac [!] "rev lm1",simp_all add: tape_of_nl_abv tape_of_nat_list.simps) |
|
2024 apply(case_tac [!] a, simp_all) |
|
2025 apply(case_tac [1] lm2, simp_all add: tape_of_nat_list.simps, auto) |
|
2026 apply(case_tac [3] lm2, simp_all add: tape_of_nat_list.simps, auto) |
|
2027 apply(case_tac [!] lista, simp_all add: tape_of_nat_list.simps) |
|
2028 done |
|
2029 |
|
2030 lemma [simp]: |
|
2031 "inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires\<Longrightarrow> |
|
2032 inv_check_left_moving_in_middle (as, lm) (s', [Bk], Oc # r) ires" |
|
2033 apply(auto simp: inv_check_left_moving_in_middle.simps ) |
|
2034 done |
|
2035 |
|
2036 lemma [simp]: |
|
2037 "inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires |
|
2038 \<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s', [], Bk # Oc # r) ires" |
|
2039 apply(insert |
|
2040 inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[of |
|
2041 as lm n "[]" r], simp) |
|
2042 done |
|
2043 |
|
2044 lemma [simp]: "inv_check_left_moving_in_middle (as, lm) |
|
2045 (s, Oc # list, Oc # r) ires |
|
2046 \<Longrightarrow> inv_on_left_moving_norm (as, lm) (s', list, Oc # Oc # r) ires" |
|
2047 apply(auto simp: inv_check_left_moving_in_middle.simps |
|
2048 inv_on_left_moving_norm.simps) |
|
2049 apply(rule_tac x = "rev (tl (rev lm1))" in exI, |
|
2050 rule_tac x = lm2 in exI, rule_tac x = "hd (rev lm1)" in exI) |
|
2051 apply(rule_tac conjI) |
|
2052 apply(case_tac "rev lm1", simp, simp) |
|
2053 apply(rule_tac x = "hd (rev lm1) - 1" in exI, auto) |
|
2054 apply(rule_tac [!] x = "Suc (Suc 0)" in exI, simp) |
|
2055 apply(case_tac [!] "rev lm1", simp_all) |
|
2056 apply(case_tac [!] a, simp_all add: tape_of_nl_cons split: if_splits) |
|
2057 done |
|
2058 |
|
2059 lemma [simp]: "inv_check_left_moving (as, lm) (s, l, Oc # r) ires |
|
2060 \<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s', [], Bk # Oc # r) ires) \<and> |
|
2061 (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s', tl l, hd l # Oc # r) ires)" |
|
2062 apply(case_tac l, |
|
2063 auto simp: inv_check_left_moving.simps inv_on_left_moving.simps) |
|
2064 apply(case_tac a, simp, simp) |
|
2065 done |
|
2066 |
|
2067 (*inv: check_left_moving to after_left_moving*) |
|
2068 lemma [simp]: "inv_check_left_moving (as, lm) (s, l, Bk # r) ires |
|
2069 \<Longrightarrow> inv_after_left_moving (as, lm) (s', Bk # l, r) ires" |
|
2070 apply(auto simp: inv_check_left_moving.simps |
|
2071 inv_check_left_moving_on_leftmost.simps inv_after_left_moving.simps) |
|
2072 done |
|
2073 |
|
2074 |
|
2075 lemma [simp]:"inv_check_left_moving (as, lm) (s, l, []) ires |
|
2076 \<Longrightarrow> inv_after_left_moving (as, lm) (s', Bk # l, []) ires" |
|
2077 by(simp add: inv_check_left_moving.simps |
|
2078 inv_check_left_moving_in_middle.simps |
|
2079 inv_check_left_moving_on_leftmost.simps) |
|
2080 |
|
2081 (*inv: after_left_moving to inv_stop*) |
|
2082 lemma [simp]: "inv_after_left_moving (as, lm) (s, l, Bk # r) ires |
|
2083 \<Longrightarrow> inv_stop (as, lm) (s', Bk # l, r) ires" |
|
2084 apply(auto simp: inv_after_left_moving.simps inv_stop.simps) |
|
2085 done |
|
2086 |
|
2087 lemma [simp]: "inv_after_left_moving (as, lm) (s, l, []) ires |
|
2088 \<Longrightarrow> inv_stop (as, lm) (s', Bk # l, []) ires" |
|
2089 by(auto simp: inv_after_left_moving.simps) |
|
2090 |
|
2091 (*inv: stop to stop*) |
|
2092 lemma [simp]: "inv_stop (as, lm) (x, l, r) ires \<Longrightarrow> |
|
2093 inv_stop (as, lm) (y, l, r) ires" |
|
2094 apply(simp add: inv_stop.simps) |
|
2095 done |
|
2096 |
|
2097 lemma [simp]: "inv_after_clear (as, lm) (s, aaa, Oc # xs) ires= False" |
|
2098 apply(auto simp: inv_after_clear.simps ) |
|
2099 done |
|
2100 |
|
2101 lemma [simp]: |
|
2102 "inv_after_left_moving (as, lm) (s, aaa, Oc # xs) ires = False" |
|
2103 by(auto simp: inv_after_left_moving.simps ) |
|
2104 |
|
2105 lemma [simp]: "inv_after_clear (as, abc_lm_s lm n (Suc (abc_lm_v lm n))) (s, b, []) ires = False" |
|
2106 apply(auto simp: inv_after_clear.simps) |
|
2107 done |
|
2108 |
|
2109 lemma [simp]: "inv_on_left_moving (as, abc_lm_s lm n (Suc (abc_lm_v lm n))) |
|
2110 (s, b, Oc # list) ires \<Longrightarrow> b \<noteq> []" |
|
2111 apply(auto simp: inv_on_left_moving.simps inv_on_left_moving_norm.simps split: if_splits) |
|
2112 done |
|
2113 |
|
2114 lemma [simp]: "inv_check_left_moving (as, abc_lm_s lm n (Suc (abc_lm_v lm n))) (s, b, Oc # list) ires \<Longrightarrow> b \<noteq> []" |
|
2115 apply(auto simp: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps split: if_splits) |
|
2116 done |
|
2117 |
|
2118 lemma tinc_correct_pre: |
|
2119 assumes layout: "ly = layout_of ap" |
|
2120 and inv_start: "inv_locate_a (as, lm) (n, l, r) ires" |
|
2121 and lm': "lm' = abc_lm_s lm n (Suc (abc_lm_v lm n))" |
|
2122 and f: "f = steps (Suc 0, l, r) (tinc_b, 0)" |
|
2123 and P: "P = (\<lambda> (s, l, r). s = 10)" |
|
2124 and Q: "Q = (\<lambda> (s, l, r). inc_inv n (as, lm) (s, l, r) ires)" |
|
2125 shows "\<exists> stp. P (f stp) \<and> Q (f stp)" |
|
2126 proof(rule_tac LE = inc_LE in halt_lemma2) |
|
2127 show "wf inc_LE" by(auto) |
|
2128 next |
|
2129 show "Q (f 0)" |
|
2130 using inv_start |
|
2131 apply(simp add: f P Q steps.simps inc_inv.simps) |
|
2132 done |
|
2133 next |
|
2134 show "\<not> P (f 0)" |
|
2135 apply(simp add: f P steps.simps) |
|
2136 done |
|
2137 next |
|
2138 show "\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) |
|
2139 \<in> inc_LE" |
|
2140 proof(rule_tac allI, rule_tac impI, simp add: f, |
|
2141 case_tac "steps (Suc 0, l, r) (tinc_b, 0) n", simp add: P) |
|
2142 fix n a b c |
|
2143 assume "a \<noteq> 10 \<and> Q (a, b, c)" |
|
2144 thus "Q (step (a, b, c) (tinc_b, 0)) \<and> (step (a, b, c) (tinc_b, 0), a, b, c) \<in> inc_LE" |
|
2145 apply(simp add:Q) |
|
2146 apply(simp add: inc_inv.simps) |
|
2147 apply(case_tac c, case_tac [2] aa) |
|
2148 apply(auto simp: Let_def step.simps tinc_b_def numeral_2_eq_2 numeral_3_eq_3 split: if_splits) |
|
2149 apply(simp_all add: inc_inv.simps inc_LE_def lex_triple_def lex_pair_def inc_measure_def numeral_5_eq_5 |
|
2150 numeral_6_eq_6 numeral_7_eq_7 numeral_8_eq_8 numeral_9_eq_9) |
|
2151 done |
|
2152 qed |
|
2153 qed |
|
2154 |
|
2155 |
|
2156 lemma tinc_correct: |
|
2157 assumes layout: "ly = layout_of ap" |
|
2158 and inv_start: "inv_locate_a (as, lm) (n, l, r) ires" |
|
2159 and lm': "lm' = abc_lm_s lm n (Suc (abc_lm_v lm n))" |
|
2160 shows "\<exists> stp l' r'. steps (Suc 0, l, r) (tinc_b, 0) stp = (10, l', r') |
|
2161 \<and> inv_stop (as, lm') (10, l', r') ires" |
|
2162 using assms |
|
2163 apply(drule_tac tinc_correct_pre, auto) |
|
2164 apply(rule_tac x = stp in exI, simp) |
|
2165 apply(simp add: inc_inv.simps) |
|
2166 done |
|
2167 |
|
2168 declare inv_locate_a.simps[simp del] abc_lm_s.simps[simp del] |
|
2169 abc_lm_v.simps[simp del] |
|
2170 |
|
2171 lemma [simp]: "(4::nat) * n mod 2 = 0" |
|
2172 apply(arith) |
|
2173 done |
|
2174 |
|
2175 lemma crsp_step_inc_pre: |
|
2176 assumes layout: "ly = layout_of ap" |
|
2177 and crsp: "crsp ly (as, lm) (s, l, r) ires" |
|
2178 and aexec: "abc_step_l (as, lm) (Some (Inc n)) = (asa, lma)" |
|
2179 shows "\<exists> stp k. steps (Suc 0, l, r) (findnth n @ shift tinc_b (2 * n), 0) stp |
|
2180 = (2*n + 10, Bk # Bk # ires, <lma> @ Bk\<up>k) \<and> stp > 0" |
|
2181 proof - |
|
2182 thm tm_append_steps |
|
2183 have "\<exists> stp l' r'. steps (Suc 0, l, r) (findnth n, 0) stp = (Suc (2 * n), l', r') |
|
2184 \<and> inv_locate_a (as, lm) (n, l', r') ires" |
|
2185 using assms |
|
2186 apply(rule_tac findnth_correct, simp_all add: crsp layout) |
|
2187 done |
|
2188 from this obtain stp l' r' where a: |
|
2189 "steps (Suc 0, l, r) (findnth n, 0) stp = (Suc (2 * n), l', r') |
|
2190 \<and> inv_locate_a (as, lm) (n, l', r') ires" by blast |
|
2191 moreover have |
|
2192 "\<exists> stp la ra. steps (Suc 0, l', r') (tinc_b, 0) stp = (10, la, ra) |
|
2193 \<and> inv_stop (as, lma) (10, la, ra) ires" |
|
2194 using assms a |
|
2195 proof(rule_tac lm' = lma and n = n and lm = lm and ly = ly and ap = ap in tinc_correct, |
|
2196 simp, simp) |
|
2197 show "lma = abc_lm_s lm n (Suc (abc_lm_v lm n))" |
|
2198 using aexec |
|
2199 apply(simp add: abc_step_l.simps) |
|
2200 done |
|
2201 qed |
|
2202 from this obtain stpa la ra where b: |
|
2203 "steps (Suc 0, l', r') (tinc_b, 0) stpa = (10, la, ra) |
|
2204 \<and> inv_stop (as, lma) (10, la, ra) ires" by blast |
|
2205 from a b show "\<exists>stp k. steps (Suc 0, l, r) (findnth n @ shift tinc_b (2 * n), 0) stp |
|
2206 = (2 * n + 10, Bk # Bk # ires, <lma> @ Bk \<up> k) \<and> stp > 0" |
|
2207 apply(rule_tac x = "stp + stpa" in exI) |
|
2208 using tm_append_steps[of "Suc 0" l r "findnth n" stp l' r' tinc_b stpa 10 la ra "length (findnth n) div 2"] |
|
2209 apply(simp add: length_findnth inv_stop.simps) |
|
2210 apply(case_tac stpa, simp_all add: steps.simps) |
|
2211 done |
|
2212 qed |
|
2213 |
558 lemma crsp_step_inc: |
2214 lemma crsp_step_inc: |
559 assumes layout: "ly = layout_of ap" |
2215 assumes layout: "ly = layout_of ap" |
560 and crsp: "crsp ly (as, lm) (s, l, r) ires" |
2216 and crsp: "crsp ly (as, lm) (s, l, r) ires" |
561 and fetch: "abc_fetch as ap = Some (Inc n)" |
2217 and fetch: "abc_fetch as ap = Some (Inc n)" |
562 shows "\<exists>stp. crsp ly (abc_step_l (as, lm) (Some (Inc n))) |
2218 shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Inc n))) |
563 (steps (s, l, r) (ci ly (start_of ly as) (Inc n), start_of ly as - Suc 0) stp) ires" |
2219 (steps (s, l, r) (ci ly (start_of ly as) (Inc n), start_of ly as - Suc 0) stp) ires" |
564 sorry |
2220 proof(case_tac "(abc_step_l (as, lm) (Some (Inc n)))") |
|
2221 fix a b |
|
2222 assume aexec: "abc_step_l (as, lm) (Some (Inc n)) = (a, b)" |
|
2223 then have "\<exists> stp k. steps (Suc 0, l, r) (findnth n @ shift tinc_b (2 * n), 0) stp |
|
2224 = (2*n + 10, Bk # Bk # ires, <b> @ Bk\<up>k) \<and> stp > 0" |
|
2225 using assms |
|
2226 apply(rule_tac crsp_step_inc_pre, simp_all) |
|
2227 done |
|
2228 thus "?thesis" |
|
2229 using assms aexec |
|
2230 apply(erule_tac exE) |
|
2231 apply(erule_tac exE) |
|
2232 apply(erule_tac conjE) |
|
2233 apply(rule_tac x = stp in exI, simp add: ci.simps tm_shift_eq_steps) |
|
2234 apply(drule_tac off = "(start_of (layout_of ap) as - Suc 0)" in tm_shift_eq_steps) |
|
2235 apply(auto simp: crsp.simps abc_step_l.simps fetch start_of_Suc1) |
|
2236 done |
|
2237 qed |
565 |
2238 |
566 subsection{* Crsp of Dec n e*} |
2239 subsection{* Crsp of Dec n e*} |
|
2240 declare sete.simps[simp del] |
|
2241 |
|
2242 type_synonym dec_inv_t = "(nat * nat list) \<Rightarrow> config \<Rightarrow> cell list \<Rightarrow> bool" |
|
2243 |
|
2244 fun dec_first_on_right_moving :: "nat \<Rightarrow> dec_inv_t" |
|
2245 where |
|
2246 "dec_first_on_right_moving n (as, lm) (s, l, r) ires = |
|
2247 (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> |
|
2248 ml + mr = Suc m \<and> length lm1 = n \<and> ml > 0 \<and> m > 0 \<and> |
|
2249 (if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires |
|
2250 else l = Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> |
|
2251 ((r = Oc\<up>mr @ [Bk] @ <lm2> @ Bk\<up>rn) \<or> (r = Oc\<up>mr \<and> lm2 = [])))" |
|
2252 |
|
2253 fun dec_on_right_moving :: "dec_inv_t" |
|
2254 where |
|
2255 "dec_on_right_moving (as, lm) (s, l, r) ires = |
|
2256 (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> |
|
2257 ml + mr = Suc (Suc m) \<and> |
|
2258 (if lm1 = [] then l = Oc\<up>ml@ Bk # Bk # ires |
|
2259 else l = Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> |
|
2260 ((r = Oc\<up>mr @ [Bk] @ <lm2> @ Bk\<up>rn) \<or> (r = Oc\<up>mr \<and> lm2 = [])))" |
|
2261 |
|
2262 fun dec_after_clear :: "dec_inv_t" |
|
2263 where |
|
2264 "dec_after_clear (as, lm) (s, l, r) ires = |
|
2265 (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> |
|
2266 ml + mr = Suc m \<and> ml = Suc m \<and> r \<noteq> [] \<and> r \<noteq> [] \<and> |
|
2267 (if lm1 = [] then l = Oc\<up>ml@ Bk # Bk # ires |
|
2268 else l = Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> |
|
2269 (tl r = Bk # <lm2> @ Bk\<up>rn \<or> tl r = [] \<and> lm2 = []))" |
|
2270 |
|
2271 fun dec_after_write :: "dec_inv_t" |
|
2272 where |
|
2273 "dec_after_write (as, lm) (s, l, r) ires = |
|
2274 (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> |
|
2275 ml + mr = Suc m \<and> ml = Suc m \<and> lm2 \<noteq> [] \<and> |
|
2276 (if lm1 = [] then l = Bk # Oc\<up>ml @ Bk # Bk # ires |
|
2277 else l = Bk # Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> |
|
2278 tl r = <lm2> @ Bk\<up>rn)" |
|
2279 |
|
2280 fun dec_right_move :: "dec_inv_t" |
|
2281 where |
|
2282 "dec_right_move (as, lm) (s, l, r) ires = |
|
2283 (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 |
|
2284 \<and> ml = Suc m \<and> mr = (0::nat) \<and> |
|
2285 (if lm1 = [] then l = Bk # Oc\<up>ml @ Bk # Bk # ires |
|
2286 else l = Bk # Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) |
|
2287 \<and> (r = Bk # <lm2> @ Bk\<up>rn \<or> r = [] \<and> lm2 = []))" |
|
2288 |
|
2289 fun dec_check_right_move :: "dec_inv_t" |
|
2290 where |
|
2291 "dec_check_right_move (as, lm) (s, l, r) ires = |
|
2292 (\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and> |
|
2293 ml = Suc m \<and> mr = (0::nat) \<and> |
|
2294 (if lm1 = [] then l = Bk # Bk # Oc\<up>ml @ Bk # Bk # ires |
|
2295 else l = Bk # Bk # Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> |
|
2296 r = <lm2> @ Bk\<up>rn)" |
|
2297 |
|
2298 fun dec_left_move :: "dec_inv_t" |
|
2299 where |
|
2300 "dec_left_move (as, lm) (s, l, r) ires = |
|
2301 (\<exists> lm1 m rn. (lm::nat list) = lm1 @ [m::nat] \<and> |
|
2302 rn > 0 \<and> |
|
2303 (if lm1 = [] then l = Bk # Oc\<up>Suc m @ Bk # Bk # ires |
|
2304 else l = Bk # Oc\<up>Suc m @ Bk # <rev lm1> @ Bk # Bk # ires) \<and> r = Bk\<up>rn)" |
|
2305 |
|
2306 declare |
|
2307 dec_on_right_moving.simps[simp del] dec_after_clear.simps[simp del] |
|
2308 dec_after_write.simps[simp del] dec_left_move.simps[simp del] |
|
2309 dec_check_right_move.simps[simp del] dec_right_move.simps[simp del] |
|
2310 dec_first_on_right_moving.simps[simp del] |
|
2311 |
|
2312 fun inv_locate_n_b :: "inc_inv_t" |
|
2313 where |
|
2314 "inv_locate_n_b (as, lm) (s, l, r) ires= |
|
2315 (\<exists> lm1 lm2 tn m ml mr rn. lm @ 0\<up>tn = lm1 @ [m] @ lm2 \<and> |
|
2316 length lm1 = s \<and> m + 1 = ml + mr \<and> |
|
2317 ml = 1 \<and> tn = s + 1 - length lm \<and> |
|
2318 (if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires |
|
2319 else l = Oc\<up>ml @ Bk # <rev lm1> @ Bk # Bk # ires) \<and> |
|
2320 (r = Oc\<up>mr @ [Bk] @ <lm2>@ Bk\<up>rn \<or> (lm2 = [] \<and> r = Oc\<up>mr)) |
|
2321 )" |
|
2322 (* |
|
2323 fun dec_inv_1 :: "layout \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> dec_inv_t" |
|
2324 where |
|
2325 "dec_inv_1 ly n e (as, am) (s, l, r) ires = |
|
2326 (let ss = start_of ly as in |
|
2327 let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in |
|
2328 let am'' = abc_lm_s am n (abc_lm_v am n) in |
|
2329 if s = start_of ly e then inv_stop (as, am'') (s, l, r) ires |
|
2330 else if s = ss then False |
|
2331 else if s = ss + 2 * n then |
|
2332 inv_locate_a (as, am) (n, l, r) ires |
|
2333 \<or> inv_locate_a (as, am'') (n, l, r) ires |
|
2334 else if s = ss + 2 * n + 1 then |
|
2335 inv_locate_b (as, am) (n, l, r) ires |
|
2336 else if s = ss + 2 * n + 13 then |
|
2337 inv_on_left_moving (as, am'') (s, l, r) ires |
|
2338 else if s = ss + 2 * n + 14 then |
|
2339 inv_check_left_moving (as, am'') (s, l, r) ires |
|
2340 else if s = ss + 2 * n + 15 then |
|
2341 inv_after_left_moving (as, am'') (s, l, r) ires |
|
2342 else False)" |
|
2343 *) |
|
2344 |
|
2345 |
|
2346 fun dec_inv_1 :: "layout \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> dec_inv_t" |
|
2347 where |
|
2348 "dec_inv_1 ly n e (as, am) (s, l, r) ires = |
|
2349 (let ss = start_of ly as in |
|
2350 let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in |
|
2351 let am'' = abc_lm_s am n (abc_lm_v am n) in |
|
2352 if s = start_of ly e then inv_stop (as, am'') (s, l, r) ires |
|
2353 else if s = ss then False |
|
2354 else if s = ss + 2 * n + 1 then |
|
2355 inv_locate_b (as, am) (n, l, r) ires |
|
2356 else if s = ss + 2 * n + 13 then |
|
2357 inv_on_left_moving (as, am'') (s, l, r) ires |
|
2358 else if s = ss + 2 * n + 14 then |
|
2359 inv_check_left_moving (as, am'') (s, l, r) ires |
|
2360 else if s = ss + 2 * n + 15 then |
|
2361 inv_after_left_moving (as, am'') (s, l, r) ires |
|
2362 else False)" |
|
2363 |
|
2364 declare fetch.simps[simp del] |
|
2365 lemma [simp]: |
|
2366 "fetch (ci ly (start_of ly as) (Dec n e)) (Suc (2 * n)) Bk = (W1, start_of ly as + 2 *n)" |
|
2367 apply(auto simp: fetch.simps length_ci_dec) |
|
2368 apply(auto simp: ci.simps nth_append length_findnth sete.simps shift.simps tdec_b_def) |
|
2369 using startof_not0[of ly as] by simp |
|
2370 |
|
2371 lemma [simp]: |
|
2372 "fetch (ci ly (start_of ly as) (Dec n e)) (Suc (2 * n)) Oc = (R, Suc (start_of ly as) + 2 *n)" |
|
2373 apply(auto simp: fetch.simps length_ci_dec) |
|
2374 apply(auto simp: ci.simps nth_append length_findnth sete.simps shift.simps tdec_b_def) |
|
2375 done |
|
2376 |
|
2377 lemma [simp]: |
|
2378 "\<lbrakk>r = [] \<or> hd r = Bk; inv_locate_a (as, lm) (n, l, r) ires\<rbrakk> |
|
2379 \<Longrightarrow> \<exists>stp la ra. |
|
2380 steps (start_of ly as + 2 * n, l, r) (ci ly (start_of ly as) (Dec n e), |
|
2381 start_of ly as - Suc 0) stp = (Suc (start_of ly as + 2 * n), la, ra) \<and> |
|
2382 inv_locate_b (as, lm) (n, la, ra) ires" |
|
2383 apply(rule_tac x = "Suc (Suc 0)" in exI) |
|
2384 apply(auto simp: steps.simps step.simps length_ci_dec) |
|
2385 apply(case_tac r, simp_all) |
|
2386 done |
|
2387 |
|
2388 lemma [simp]: |
|
2389 "\<lbrakk>inv_locate_a (as, lm) (n, l, r) ires; r \<noteq> [] \<and> hd r \<noteq> Bk\<rbrakk> |
|
2390 \<Longrightarrow> \<exists>stp la ra. |
|
2391 steps (start_of ly as + 2 * n, l, r) (ci ly (start_of ly as) (Dec n e), |
|
2392 start_of ly as - Suc 0) stp = (Suc (start_of ly as + 2 * n), la, ra) \<and> |
|
2393 inv_locate_b (as, lm) (n, la, ra) ires" |
|
2394 apply(rule_tac x = "(Suc 0)" in exI, case_tac "hd r", simp_all) |
|
2395 apply(auto simp: steps.simps step.simps length_ci_dec) |
|
2396 apply(case_tac r, simp_all) |
|
2397 done |
|
2398 |
|
2399 fun abc_dec_1_stage1:: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
2400 where |
|
2401 "abc_dec_1_stage1 (s, l, r) ss n = |
|
2402 (if s > ss \<and> s \<le> ss + 2*n + 1 then 4 |
|
2403 else if s = ss + 2 * n + 13 \<or> s = ss + 2*n + 14 then 3 |
|
2404 else if s = ss + 2*n + 15 then 2 |
|
2405 else 0)" |
|
2406 |
|
2407 fun abc_dec_1_stage2:: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
2408 where |
|
2409 "abc_dec_1_stage2 (s, l, r) ss n = |
|
2410 (if s \<le> ss + 2 * n + 1 then (ss + 2 * n + 16 - s) |
|
2411 else if s = ss + 2*n + 13 then length l |
|
2412 else if s = ss + 2*n + 14 then length l |
|
2413 else 0)" |
|
2414 |
|
2415 fun abc_dec_1_stage3 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
2416 where |
|
2417 "abc_dec_1_stage3 (s, l, r) ss n = |
|
2418 (if s \<le> ss + 2*n + 1 then |
|
2419 if (s - ss) mod 2 = 0 then |
|
2420 if r \<noteq> [] \<and> hd r = Oc then 0 else 1 |
|
2421 else length r |
|
2422 else if s = ss + 2 * n + 13 then |
|
2423 if r \<noteq> [] \<and> hd r = Oc then 2 |
|
2424 else 1 |
|
2425 else if s = ss + 2 * n + 14 then |
|
2426 if r \<noteq> [] \<and> hd r = Oc then 3 else 0 |
|
2427 else 0)" |
|
2428 |
|
2429 fun abc_dec_1_measure :: "(config \<times> nat \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)" |
|
2430 where |
|
2431 "abc_dec_1_measure (c, ss, n) = (abc_dec_1_stage1 c ss n, |
|
2432 abc_dec_1_stage2 c ss n, abc_dec_1_stage3 c ss n)" |
|
2433 |
|
2434 definition abc_dec_1_LE :: |
|
2435 "((config \<times> nat \<times> |
|
2436 nat) \<times> (config \<times> nat \<times> nat)) set" |
|
2437 where "abc_dec_1_LE \<equiv> (inv_image lex_triple abc_dec_1_measure)" |
|
2438 |
|
2439 lemma wf_dec_le: "wf abc_dec_1_LE" |
|
2440 by(auto intro:wf_inv_image simp:abc_dec_1_LE_def lex_triple_def lex_pair_def) |
|
2441 |
|
2442 lemma startof_Suc2: |
|
2443 "abc_fetch as ap = Some (Dec n e) \<Longrightarrow> |
|
2444 start_of (layout_of ap) (Suc as) = |
|
2445 start_of (layout_of ap) as + 2 * n + 16" |
|
2446 apply(auto simp: start_of.simps layout_of.simps |
|
2447 length_of.simps abc_fetch.simps |
|
2448 take_Suc_conv_app_nth split: if_splits) |
|
2449 done |
|
2450 |
|
2451 lemma start_of_less_2: |
|
2452 "start_of ly e \<le> start_of ly (Suc e)" |
|
2453 thm take_Suc |
|
2454 apply(case_tac "e < length ly") |
|
2455 apply(auto simp: start_of.simps take_Suc take_Suc_conv_app_nth) |
|
2456 done |
|
2457 |
|
2458 lemma start_of_less_1: "start_of ly e \<le> start_of ly (e + d)" |
|
2459 proof(induct d) |
|
2460 case 0 thus "?case" by simp |
|
2461 next |
|
2462 case (Suc d) |
|
2463 have "start_of ly e \<le> start_of ly (e + d)" by fact |
|
2464 moreover have "start_of ly (e + d) \<le> start_of ly (Suc (e + d))" |
|
2465 by(rule_tac start_of_less_2) |
|
2466 ultimately show"?case" |
|
2467 by(simp) |
|
2468 qed |
|
2469 |
|
2470 lemma start_of_less: |
|
2471 assumes "e < as" |
|
2472 shows "start_of ly e \<le> start_of ly as" |
|
2473 proof - |
|
2474 obtain d where " as = e + d" |
|
2475 using assms by (metis less_imp_add_positive) |
|
2476 thus "?thesis" |
|
2477 by(simp add: start_of_less_1) |
|
2478 qed |
|
2479 |
|
2480 lemma start_of_ge: |
|
2481 assumes fetch: "abc_fetch as ap = Some (Dec n e)" |
|
2482 and layout: "ly = layout_of ap" |
|
2483 and great: "e > as" |
|
2484 shows "start_of ly e \<ge> start_of ly as + 2*n + 16" |
|
2485 proof(cases "e = Suc as") |
|
2486 case True |
|
2487 have "e = Suc as" by fact |
|
2488 moreover hence "start_of ly (Suc as) = start_of ly as + 2*n + 16" |
|
2489 using layout fetch |
|
2490 by(simp add: startof_Suc2) |
|
2491 ultimately show "?thesis" by (simp) |
|
2492 next |
|
2493 case False |
|
2494 have "e \<noteq> Suc as" by fact |
|
2495 then have "e > Suc as" using great by arith |
|
2496 then have "start_of ly (Suc as) \<le> start_of ly e" |
|
2497 by(simp add: start_of_less) |
|
2498 moreover have "start_of ly (Suc as) = start_of ly as + 2*n + 16" |
|
2499 using layout fetch |
|
2500 by(simp add: startof_Suc2) |
|
2501 ultimately show "?thesis" |
|
2502 by arith |
|
2503 qed |
|
2504 |
|
2505 lemma [elim]: "\<lbrakk>abc_fetch as ap = Some (Dec n e); as < e; |
|
2506 Suc (start_of (layout_of ap) as + 2 * n) = start_of (layout_of ap) e\<rbrakk> \<Longrightarrow> RR" |
|
2507 apply(drule_tac start_of_ge, simp_all) |
|
2508 apply(auto) |
|
2509 done |
|
2510 |
|
2511 lemma [elim]: "\<lbrakk>abc_fetch as ap = Some (Dec n e); as > e; |
|
2512 Suc (start_of (layout_of ap) as + 2 * n) = start_of (layout_of ap) e\<rbrakk> \<Longrightarrow> RR" |
|
2513 apply(drule_tac ly = "layout_of ap" in start_of_less[of]) |
|
2514 apply(arith) |
|
2515 done |
|
2516 |
|
2517 lemma [elim]: "\<lbrakk>abc_fetch as ap = Some (Dec n e); |
|
2518 Suc (start_of (layout_of ap) as + 2 * n) = start_of (layout_of ap) e\<rbrakk> \<Longrightarrow> RR" |
|
2519 apply(subgoal_tac "as = e \<or> as < e \<or> as > e", auto) |
|
2520 done |
|
2521 |
|
2522 lemma [simp]:"fetch (ci (ly) (start_of ly as) (Dec n e)) (Suc (2 * n)) Oc |
|
2523 = (R, start_of ly as + 2*n + 1)" |
|
2524 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2525 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2526 done |
|
2527 |
|
2528 lemma [simp]: "(start_of ly as = 0) = False" |
|
2529 apply(simp add: start_of.simps) |
|
2530 done |
|
2531 |
|
2532 lemma [simp]: "fetch (ci (ly) |
|
2533 (start_of ly as) (Dec n e)) (Suc (2 * n)) Bk |
|
2534 = (W1, start_of ly as + 2*n)" |
|
2535 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2536 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2537 done |
|
2538 |
|
2539 lemma [simp]: |
|
2540 "fetch (ci (ly) |
|
2541 (start_of ly as) (Dec n e)) (Suc (Suc (2 * n))) Oc |
|
2542 = (R, start_of ly as + 2*n + 2)" |
|
2543 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2544 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2545 done |
|
2546 |
|
2547 |
|
2548 lemma [simp]: "fetch (ci (ly) |
|
2549 (start_of ly as) (Dec n e)) (Suc (Suc (2 * n))) Bk |
|
2550 = (L, start_of ly as + 2*n + 13)" |
|
2551 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2552 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2553 done |
|
2554 |
|
2555 |
|
2556 lemma [simp]: "fetch (ci (ly) |
|
2557 (start_of ly as) (Dec n e)) (Suc (Suc (Suc (2 * n)))) Oc |
|
2558 = (R, start_of ly as + 2*n + 2)" |
|
2559 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2560 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2561 done |
|
2562 |
|
2563 |
|
2564 lemma [simp]: "fetch (ci (ly) (start_of ly as) (Dec n e)) |
|
2565 (Suc (Suc (Suc (2 * n)))) Bk |
|
2566 = (L, start_of ly as + 2*n + 3)" |
|
2567 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2568 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2569 done |
|
2570 |
|
2571 lemma [simp]: |
|
2572 "fetch (ci (ly) |
|
2573 (start_of ly as) (Dec n e)) (2 * n + 4) Oc |
|
2574 = (W0, start_of ly as + 2*n + 3)" |
|
2575 apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps) |
|
2576 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2577 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2578 done |
|
2579 |
|
2580 lemma [simp]: "fetch (ci (ly) |
|
2581 (start_of ly as) (Dec n e)) (2 * n + 4) Bk |
|
2582 = (R, start_of ly as + 2*n + 4)" |
|
2583 apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps) |
|
2584 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2585 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2586 done |
|
2587 |
|
2588 lemma [simp]:"fetch (ci (ly) |
|
2589 (start_of ly as) (Dec n e)) (2 * n + 5) Bk |
|
2590 = (R, start_of ly as + 2*n + 5)" |
|
2591 apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps) |
|
2592 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2593 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2594 done |
|
2595 |
|
2596 |
|
2597 lemma [simp]: |
|
2598 "fetch (ci (ly) |
|
2599 (start_of ly as) (Dec n e)) (2 * n + 6) Bk |
|
2600 = (L, start_of ly as + 2*n + 6)" |
|
2601 apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps) |
|
2602 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2603 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2604 done |
|
2605 |
|
2606 lemma [simp]: |
|
2607 "fetch (ci (ly) (start_of ly as) |
|
2608 (Dec n e)) (2 * n + 6) Oc |
|
2609 = (L, start_of ly as + 2*n + 7)" |
|
2610 apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps) |
|
2611 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2612 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2613 done |
|
2614 |
|
2615 lemma [simp]:"fetch (ci (ly) |
|
2616 (start_of ly as) (Dec n e)) (2 * n + 7) Bk |
|
2617 = (L, start_of ly as + 2*n + 10)" |
|
2618 apply(subgoal_tac "2*n + 7 = Suc (2*n + 6)", simp only: fetch.simps) |
|
2619 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2620 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2621 done |
|
2622 |
|
2623 lemma [simp]: |
|
2624 "fetch (ci (ly) |
|
2625 (start_of ly as) (Dec n e)) (2 * n + 8) Bk |
|
2626 = (W1, start_of ly as + 2*n + 7)" |
|
2627 apply(subgoal_tac "2*n + 8 = Suc (2*n + 7)", simp only: fetch.simps) |
|
2628 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2629 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2630 done |
|
2631 |
|
2632 |
|
2633 lemma [simp]: |
|
2634 "fetch (ci (ly) |
|
2635 (start_of ly as) (Dec n e)) (2 * n + 8) Oc |
|
2636 = (R, start_of ly as + 2*n + 8)" |
|
2637 apply(subgoal_tac "2*n + 8 = Suc (2*n + 7)", simp only: fetch.simps) |
|
2638 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2639 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2640 done |
|
2641 |
|
2642 lemma [simp]: |
|
2643 "fetch (ci (ly) |
|
2644 (start_of ly as) (Dec n e)) (2 * n + 9) Bk |
|
2645 = (L, start_of ly as + 2*n + 9)" |
|
2646 apply(subgoal_tac "2*n + 9 = Suc (2*n + 8)", simp only: fetch.simps) |
|
2647 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2648 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2649 done |
|
2650 |
|
2651 lemma [simp]: |
|
2652 "fetch (ci (ly) |
|
2653 (start_of ly as) (Dec n e)) (2 * n + 9) Oc |
|
2654 = (R, start_of ly as + 2*n + 8)" |
|
2655 apply(subgoal_tac "2*n + 9 = Suc (2*n + 8)", simp only: fetch.simps) |
|
2656 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2657 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2658 done |
|
2659 |
|
2660 |
|
2661 lemma [simp]: |
|
2662 "fetch (ci (ly) |
|
2663 (start_of ly as) (Dec n e)) (2 * n + 10) Bk |
|
2664 = (R, start_of ly as + 2*n + 4)" |
|
2665 apply(subgoal_tac "2*n + 10 = Suc (2*n + 9)", simp only: fetch.simps) |
|
2666 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2667 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2668 done |
|
2669 |
|
2670 lemma [simp]: "fetch (ci (ly) |
|
2671 (start_of ly as) (Dec n e)) (2 * n + 10) Oc |
|
2672 = (W0, start_of ly as + 2*n + 9)" |
|
2673 apply(subgoal_tac "2*n + 10 = Suc (2*n + 9)", simp only: fetch.simps) |
|
2674 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2675 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2676 done |
|
2677 |
|
2678 |
|
2679 lemma [simp]: |
|
2680 "fetch (ci (ly) |
|
2681 (start_of ly as) (Dec n e)) (2 * n + 11) Oc |
|
2682 = (L, start_of ly as + 2*n + 10)" |
|
2683 apply(subgoal_tac "2*n + 11 = Suc (2*n + 10)", simp only: fetch.simps) |
|
2684 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2685 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2686 done |
|
2687 |
|
2688 |
|
2689 lemma [simp]: |
|
2690 "fetch (ci (ly) |
|
2691 (start_of ly as) (Dec n e)) (2 * n + 11) Bk |
|
2692 = (L, start_of ly as + 2*n + 11)" |
|
2693 apply(subgoal_tac "2*n + 11 = Suc (2*n + 10)", simp only: fetch.simps) |
|
2694 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2695 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2696 done |
|
2697 |
|
2698 lemma [simp]: |
|
2699 "fetch (ci (ly) |
|
2700 (start_of ly as) (Dec n e)) (2 * n + 12) Oc |
|
2701 = (L, start_of ly as + 2*n + 10)" |
|
2702 apply(subgoal_tac "2*n + 12 = Suc (2*n + 11)", simp only: fetch.simps) |
|
2703 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2704 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2705 done |
|
2706 |
|
2707 |
|
2708 lemma [simp]: |
|
2709 "fetch (ci (ly) |
|
2710 (start_of ly as) (Dec n e)) (2 * n + 12) Bk |
|
2711 = (R, start_of ly as + 2*n + 12)" |
|
2712 apply(subgoal_tac "2*n + 12 = Suc (2*n + 11)", simp only: fetch.simps) |
|
2713 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2714 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2715 done |
|
2716 |
|
2717 lemma [simp]: |
|
2718 "fetch (ci (ly) |
|
2719 (start_of ly as) (Dec n e)) (2 * n + 13) Bk |
|
2720 = (R, start_of ly as + 2*n + 16)" |
|
2721 apply(subgoal_tac "2*n + 13 = Suc (2*n + 12)", simp only: fetch.simps) |
|
2722 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2723 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2724 done |
|
2725 |
|
2726 |
|
2727 lemma [simp]: |
|
2728 "fetch (ci (ly) |
|
2729 (start_of ly as) (Dec n e)) (14 + 2 * n) Oc |
|
2730 = (L, start_of ly as + 2*n + 13)" |
|
2731 apply(subgoal_tac "14 + 2*n = Suc (2*n + 13)", simp only: fetch.simps) |
|
2732 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2733 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2734 done |
|
2735 |
|
2736 lemma [simp]: |
|
2737 "fetch (ci (ly) |
|
2738 (start_of ly as) (Dec n e)) (14 + 2 * n) Bk |
|
2739 = (L, start_of ly as + 2*n + 14)" |
|
2740 apply(subgoal_tac "14 + 2*n = Suc (2*n + 13)", simp only: fetch.simps) |
|
2741 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2742 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2743 done |
|
2744 |
|
2745 lemma [simp]: |
|
2746 "fetch (ci (ly) |
|
2747 (start_of ly as) (Dec n e)) (15 + 2 * n) Oc |
|
2748 = (L, start_of ly as + 2*n + 13)" |
|
2749 apply(subgoal_tac "15 + 2*n = Suc (2*n + 14)", simp only: fetch.simps) |
|
2750 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2751 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2752 done |
|
2753 |
|
2754 lemma [simp]: |
|
2755 "fetch (ci (ly) |
|
2756 (start_of ly as) (Dec n e)) (15 + 2 * n) Bk |
|
2757 = (R, start_of ly as + 2*n + 15)" |
|
2758 apply(subgoal_tac "15 + 2*n = Suc (2*n + 14)", simp only: fetch.simps) |
|
2759 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2760 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2761 done |
|
2762 |
|
2763 lemma [simp]: |
|
2764 "abc_fetch as aprog = Some (Dec n e) \<Longrightarrow> |
|
2765 fetch (ci (ly) (start_of (ly) as) |
|
2766 (Dec n e)) (16 + 2 * n) Bk |
|
2767 = (R, start_of (ly) e)" |
|
2768 apply(subgoal_tac "16 + 2*n = Suc (2*n + 15)", simp only: fetch.simps) |
|
2769 apply(auto simp: ci.simps findnth.simps fetch.simps |
|
2770 nth_of.simps shift.simps nth_append tdec_b_def length_findnth sete.simps) |
|
2771 done |
|
2772 |
|
2773 declare dec_inv_1.simps[simp del] |
|
2774 |
|
2775 |
|
2776 lemma [simp]: |
|
2777 "\<lbrakk>abc_fetch as aprog = Some (Dec n e); ly = layout_of aprog\<rbrakk> |
|
2778 \<Longrightarrow> (start_of ly e \<noteq> Suc (start_of ly as + 2 * n) \<and> |
|
2779 start_of ly e \<noteq> Suc (Suc (start_of ly as + 2 * n)) \<and> |
|
2780 start_of ly e \<noteq> start_of ly as + 2 * n + 3 \<and> |
|
2781 start_of ly e \<noteq> start_of ly as + 2 * n + 4 \<and> |
|
2782 start_of ly e \<noteq> start_of ly as + 2 * n + 5 \<and> |
|
2783 start_of ly e \<noteq> start_of ly as + 2 * n + 6 \<and> |
|
2784 start_of ly e \<noteq> start_of ly as + 2 * n + 7 \<and> |
|
2785 start_of ly e \<noteq> start_of ly as + 2 * n + 8 \<and> |
|
2786 start_of ly e \<noteq> start_of ly as + 2 * n + 9 \<and> |
|
2787 start_of ly e \<noteq> start_of ly as + 2 * n + 10 \<and> |
|
2788 start_of ly e \<noteq> start_of ly as + 2 * n + 11 \<and> |
|
2789 start_of ly e \<noteq> start_of ly as + 2 * n + 12 \<and> |
|
2790 start_of ly e \<noteq> start_of ly as + 2 * n + 13 \<and> |
|
2791 start_of ly e \<noteq> start_of ly as + 2 * n + 14 \<and> |
|
2792 start_of ly e \<noteq> start_of ly as + 2 * n + 15)" |
|
2793 using start_of_ge[of as aprog n e ly] start_of_less[of e as ly] |
|
2794 apply(case_tac "e < as", simp) |
|
2795 apply(case_tac "e = as", simp, simp) |
|
2796 done |
|
2797 |
|
2798 lemma [simp]: "\<lbrakk>abc_fetch as aprog = Some (Dec n e); ly = layout_of aprog\<rbrakk> |
|
2799 \<Longrightarrow> (Suc (start_of ly as + 2 * n) \<noteq> start_of ly e \<and> |
|
2800 Suc (Suc (start_of ly as + 2 * n)) \<noteq> start_of ly e \<and> |
|
2801 start_of ly as + 2 * n + 3 \<noteq> start_of ly e \<and> |
|
2802 start_of ly as + 2 * n + 4 \<noteq> start_of ly e \<and> |
|
2803 start_of ly as + 2 * n + 5 \<noteq>start_of ly e \<and> |
|
2804 start_of ly as + 2 * n + 6 \<noteq> start_of ly e \<and> |
|
2805 start_of ly as + 2 * n + 7 \<noteq> start_of ly e \<and> |
|
2806 start_of ly as + 2 * n + 8 \<noteq> start_of ly e \<and> |
|
2807 start_of ly as + 2 * n + 9 \<noteq> start_of ly e \<and> |
|
2808 start_of ly as + 2 * n + 10 \<noteq> start_of ly e \<and> |
|
2809 start_of ly as + 2 * n + 11 \<noteq> start_of ly e \<and> |
|
2810 start_of ly as + 2 * n + 12 \<noteq> start_of ly e \<and> |
|
2811 start_of ly as + 2 * n + 13 \<noteq> start_of ly e \<and> |
|
2812 start_of ly as + 2 * n + 14 \<noteq> start_of ly e \<and> |
|
2813 start_of ly as + 2 * n + 15 \<noteq> start_of ly e)" |
|
2814 using start_of_ge[of as aprog n e ly] start_of_less[of e as ly] |
|
2815 apply(case_tac "e < as", simp, simp) |
|
2816 apply(case_tac "e = as", simp, simp) |
|
2817 done |
|
2818 |
|
2819 lemma [simp]: "inv_locate_b (as, lm) (n, [], []) ires = False" |
|
2820 apply(auto simp: inv_locate_b.simps in_middle.simps split: if_splits) |
|
2821 done |
|
2822 |
|
2823 lemma [simp]: "inv_locate_b (as, lm) (n, [], Bk # list) ires = False" |
|
2824 apply(auto simp: inv_locate_b.simps in_middle.simps split: if_splits) |
|
2825 done |
|
2826 |
|
2827 (* |
|
2828 |
|
2829 lemma inv_locate_b_2_on_left_moving_b[simp]: |
|
2830 "inv_locate_b (as, am) (n, l, []) ires |
|
2831 \<Longrightarrow> inv_on_left_moving (as, |
|
2832 abc_lm_s am n (abc_lm_v am n)) (s, [], [Bk]) ires" |
|
2833 apply(auto simp: inv_locate_b.simps inv_on_left_moving.simps inv_on_left_moving_in_middle_B.simps |
|
2834 in_middle.simps split: if_splits) |
|
2835 apply(drule_tac length_equal, simp) |
|
2836 |
|
2837 apply(insert inv_locate_b_2_on_left_moving[of as am n l "[]" ires s]) |
|
2838 apply(simp only: inv_on_left_moving.simps, simp) |
|
2839 apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B |
|
2840 (as, abc_lm_s am n (abc_lm_v am n)) (s, tl l, [hd l]) ires", simp) |
|
2841 *) |
|
2842 |
|
2843 (* |
|
2844 lemma [simp]: |
|
2845 "inv_locate_b (as, am) (n, l, []) ires; l \<noteq> []\<rbrakk> |
|
2846 \<Longrightarrow> inv_on_left_moving (as, abc_lm_s am n |
|
2847 (abc_lm_v am n)) (s, tl l, [hd l]) ires" |
|
2848 apply(auto simp: inv_locate_b.simps inv_on_left_moving.simps inv_on_left_moving_in_middle_B.simps |
|
2849 in_middle.simps split: if_splits) |
|
2850 apply(drule_tac length_equal, simp) |
|
2851 |
|
2852 apply(insert inv_locate_b_2_on_left_moving[of as am n l "[]" ires s]) |
|
2853 apply(simp only: inv_on_left_moving.simps, simp) |
|
2854 apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B |
|
2855 (as, abc_lm_s am n (abc_lm_v am n)) (s, tl l, [hd l]) ires", simp) |
|
2856 |
|
2857 apply(insert inv_locate_b_2_on_left_moving[of as am n l "[]" ires s]) |
|
2858 apply(simp only: inv_on_left_moving.simps, simp) |
|
2859 apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B |
|
2860 (as, abc_lm_s am n (abc_lm_v am n)) (s, tl l, [hd l]) ires", simp) |
|
2861 apply(simp only: inv_on_left_moving_norm.simps) |
|
2862 apply(erule_tac exE)+ |
|
2863 apply(erule_tac conjE)+ |
|
2864 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
2865 rule_tac x = m in exI, rule_tac x = ml in exI, |
|
2866 rule_tac x = mr in exI, simp) |
|
2867 apply(case_tac mr, simp, simp, case_tac nat, auto intro: nil_2_nil) |
|
2868 done |
|
2869 *) |
|
2870 |
|
2871 lemma [simp]: |
|
2872 "\<lbrakk>dec_first_on_right_moving n (as, am) (s, aaa, Oc # xs) ires\<rbrakk> |
|
2873 \<Longrightarrow> dec_first_on_right_moving n (as, am) (s', Oc # aaa, xs) ires" |
|
2874 apply(simp only: dec_first_on_right_moving.simps) |
|
2875 apply(erule exE)+ |
|
2876 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
2877 rule_tac x = m in exI, simp) |
|
2878 apply(rule_tac x = "Suc ml" in exI, |
|
2879 rule_tac x = "mr - 1" in exI, auto) |
|
2880 apply(case_tac [!] mr, auto) |
|
2881 done |
|
2882 |
|
2883 lemma [simp]: |
|
2884 "dec_first_on_right_moving n (as, am) (s, l, Bk # xs) ires \<Longrightarrow> l \<noteq> []" |
|
2885 apply(auto simp: dec_first_on_right_moving.simps split: if_splits) |
|
2886 done |
|
2887 |
|
2888 lemma [elim]: |
|
2889 "\<lbrakk>\<not> length lm1 < length am; |
|
2890 am @ replicate (length lm1 - length am) 0 @ [0::nat] = |
|
2891 lm1 @ m # lm2; |
|
2892 0 < m\<rbrakk> |
|
2893 \<Longrightarrow> RR" |
|
2894 apply(subgoal_tac "lm2 = []", simp) |
|
2895 apply(drule_tac length_equal, simp) |
|
2896 done |
|
2897 |
|
2898 lemma [simp]: |
|
2899 "\<lbrakk>dec_first_on_right_moving n (as, |
|
2900 abc_lm_s am n (abc_lm_v am n)) (s, l, Bk # xs) ires\<rbrakk> |
|
2901 \<Longrightarrow> dec_after_clear (as, abc_lm_s am n |
|
2902 (abc_lm_v am n - Suc 0)) (s', tl l, hd l # Bk # xs) ires" |
|
2903 apply(simp only: dec_first_on_right_moving.simps |
|
2904 dec_after_clear.simps abc_lm_s.simps abc_lm_v.simps) |
|
2905 apply(erule_tac exE)+ |
|
2906 apply(case_tac "n < length am") |
|
2907 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
2908 rule_tac x = "m - 1" in exI, auto simp: ) |
|
2909 apply(case_tac [!] mr, auto) |
|
2910 done |
|
2911 |
|
2912 lemma [simp]: |
|
2913 "\<lbrakk>dec_first_on_right_moving n (as, |
|
2914 abc_lm_s am n (abc_lm_v am n)) (s, l, []) ires\<rbrakk> |
|
2915 \<Longrightarrow> (l = [] \<longrightarrow> dec_after_clear (as, |
|
2916 abc_lm_s am n (abc_lm_v am n - Suc 0)) (s', [], [Bk]) ires) \<and> |
|
2917 (l \<noteq> [] \<longrightarrow> dec_after_clear (as, abc_lm_s am n |
|
2918 (abc_lm_v am n - Suc 0)) (s', tl l, [hd l]) ires)" |
|
2919 apply(subgoal_tac "l \<noteq> []", |
|
2920 simp only: dec_first_on_right_moving.simps |
|
2921 dec_after_clear.simps abc_lm_s.simps abc_lm_v.simps) |
|
2922 apply(erule_tac exE)+ |
|
2923 apply(case_tac "n < length am", simp) |
|
2924 apply(rule_tac x = lm1 in exI, rule_tac x = "m - 1" in exI, auto) |
|
2925 apply(case_tac [1-2] m, auto) |
|
2926 apply(auto simp: dec_first_on_right_moving.simps split: if_splits) |
|
2927 done |
|
2928 |
|
2929 lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, Oc # r) ires\<rbrakk> |
|
2930 \<Longrightarrow> dec_after_clear (as, am) (s', l, Bk # r) ires" |
|
2931 apply(auto simp: dec_after_clear.simps) |
|
2932 done |
|
2933 |
|
2934 lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, Bk # r) ires\<rbrakk> |
|
2935 \<Longrightarrow> dec_right_move (as, am) (s', Bk # l, r) ires" |
|
2936 apply(auto simp: dec_after_clear.simps dec_right_move.simps split: if_splits) |
|
2937 done |
|
2938 |
|
2939 lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, []) ires\<rbrakk> |
|
2940 \<Longrightarrow> dec_right_move (as, am) (s', Bk # l, []) ires" |
|
2941 apply(auto simp: dec_after_clear.simps dec_right_move.simps ) |
|
2942 done |
|
2943 |
|
2944 lemma [simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, []) ires\<rbrakk> |
|
2945 \<Longrightarrow> dec_right_move (as, am) (s', Bk # l, [Bk]) ires" |
|
2946 apply(auto simp: dec_after_clear.simps dec_right_move.simps split: if_splits) |
|
2947 done |
|
2948 |
|
2949 lemma [simp]:"dec_right_move (as, am) (s, l, Oc # r) ires = False" |
|
2950 apply(auto simp: dec_right_move.simps) |
|
2951 done |
|
2952 |
|
2953 lemma dec_right_move_2_check_right_move[simp]: |
|
2954 "\<lbrakk>dec_right_move (as, am) (s, l, Bk # r) ires\<rbrakk> |
|
2955 \<Longrightarrow> dec_check_right_move (as, am) (s', Bk # l, r) ires" |
|
2956 apply(auto simp: dec_right_move.simps dec_check_right_move.simps split: if_splits) |
|
2957 done |
|
2958 |
|
2959 lemma [simp]: "(<lm::nat list> = []) = (lm = [])" |
|
2960 apply(case_tac lm, simp_all add: tape_of_nl_cons) |
|
2961 done |
|
2962 |
|
2963 lemma [simp]: |
|
2964 "dec_right_move (as, am) (s, l, []) ires= |
|
2965 dec_right_move (as, am) (s, l, [Bk]) ires" |
|
2966 apply(simp add: dec_right_move.simps) |
|
2967 done |
|
2968 |
|
2969 lemma [simp]: "\<lbrakk>dec_right_move (as, am) (s, l, []) ires\<rbrakk> |
|
2970 \<Longrightarrow> dec_check_right_move (as, am) (s, Bk # l, []) ires" |
|
2971 apply(insert dec_right_move_2_check_right_move[of as am s l "[]" s'], |
|
2972 simp) |
|
2973 done |
|
2974 |
|
2975 lemma [simp]: "dec_check_right_move (as, am) (s, l, r) ires\<Longrightarrow> l \<noteq> []" |
|
2976 apply(auto simp: dec_check_right_move.simps split: if_splits) |
|
2977 done |
|
2978 |
|
2979 lemma [simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, Oc # r) ires\<rbrakk> |
|
2980 \<Longrightarrow> dec_after_write (as, am) (s', tl l, hd l # Oc # r) ires" |
|
2981 apply(auto simp: dec_check_right_move.simps dec_after_write.simps) |
|
2982 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
2983 rule_tac x = m in exI, auto) |
|
2984 done |
|
2985 |
|
2986 |
|
2987 |
|
2988 lemma [simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, Bk # r) ires\<rbrakk> |
|
2989 \<Longrightarrow> dec_left_move (as, am) (s', tl l, hd l # Bk # r) ires" |
|
2990 apply(auto simp: dec_check_right_move.simps |
|
2991 dec_left_move.simps inv_after_move.simps) |
|
2992 apply(rule_tac x = lm1 in exI, rule_tac x = m in exI, auto split: if_splits) |
|
2993 apply(case_tac [!] lm2, simp_all add: tape_of_nl_cons split: if_splits) |
|
2994 apply(rule_tac [!] x = "(Suc rn)" in exI, simp_all) |
|
2995 done |
|
2996 |
|
2997 lemma [simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, []) ires\<rbrakk> |
|
2998 \<Longrightarrow> dec_left_move (as, am) (s', tl l, [hd l]) ires" |
|
2999 apply(auto simp: dec_check_right_move.simps |
|
3000 dec_left_move.simps inv_after_move.simps) |
|
3001 apply(rule_tac x = lm1 in exI, rule_tac x = m in exI, auto) |
|
3002 done |
|
3003 |
|
3004 lemma [simp]: "dec_left_move (as, am) (s, aaa, Oc # xs) ires = False" |
|
3005 apply(auto simp: dec_left_move.simps inv_after_move.simps) |
|
3006 done |
|
3007 |
|
3008 lemma [simp]: "dec_left_move (as, am) (s, l, r) ires |
|
3009 \<Longrightarrow> l \<noteq> []" |
|
3010 apply(auto simp: dec_left_move.simps split: if_splits) |
|
3011 done |
|
3012 |
|
3013 lemma [simp]: "inv_on_left_moving_in_middle_B (as, [m]) |
|
3014 (s', Oc # Oc\<up>m @ Bk # Bk # ires, Bk # Bk\<up>rn) ires" |
|
3015 apply(simp add: inv_on_left_moving_in_middle_B.simps) |
|
3016 apply(rule_tac x = "[m]" in exI, auto) |
|
3017 done |
|
3018 |
|
3019 lemma [simp]: "inv_on_left_moving_in_middle_B (as, [m]) |
|
3020 (s', Oc # Oc\<up>m @ Bk # Bk # ires, [Bk]) ires" |
|
3021 apply(simp add: inv_on_left_moving_in_middle_B.simps) |
|
3022 done |
|
3023 |
|
3024 |
|
3025 lemma [simp]: "lm1 \<noteq> [] \<Longrightarrow> |
|
3026 inv_on_left_moving_in_middle_B (as, lm1 @ [m]) (s', |
|
3027 Oc # Oc\<up>m @ Bk # <rev lm1> @ Bk # Bk # ires, Bk # Bk\<up>rn) ires" |
|
3028 apply(simp only: inv_on_left_moving_in_middle_B.simps) |
|
3029 apply(rule_tac x = "lm1 @ [m ]" in exI, rule_tac x = "[]" in exI, simp) |
|
3030 apply(simp add: tape_of_nl_cons split: if_splits) |
|
3031 done |
|
3032 |
|
3033 lemma [simp]: "lm1 \<noteq> [] \<Longrightarrow> |
|
3034 inv_on_left_moving_in_middle_B (as, lm1 @ [m]) (s', |
|
3035 Oc # Oc\<up> m @ Bk # <rev lm1> @ Bk # Bk # ires, [Bk]) ires" |
|
3036 apply(simp only: inv_on_left_moving_in_middle_B.simps) |
|
3037 apply(rule_tac x = "lm1 @ [m ]" in exI, rule_tac x = "[]" in exI, simp) |
|
3038 apply(simp add: tape_of_nl_cons split: if_splits) |
|
3039 done |
|
3040 |
|
3041 lemma [simp]: "dec_left_move (as, am) (s, l, Bk # r) ires |
|
3042 \<Longrightarrow> inv_on_left_moving (as, am) (s', tl l, hd l # Bk # r) ires" |
|
3043 apply(auto simp: dec_left_move.simps inv_on_left_moving.simps split: if_splits) |
|
3044 done |
|
3045 |
|
3046 (* |
|
3047 lemma [simp]: "inv_on_left_moving_in_middle_B (as, lm1 @ [m]) |
|
3048 (s', Oc # Oc\<^bsup>m\<^esup> @ Bk # <rev lm1> @ Bk\<^bsup>ln\<^esup>, [Bk]) ires" |
|
3049 apply(auto simp: inv_on_left_moving_in_middle_B.simps) |
|
3050 apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "[]" in exI, auto) |
|
3051 done |
|
3052 *) |
|
3053 |
|
3054 lemma [simp]: "dec_left_move (as, am) (s, l, []) ires |
|
3055 \<Longrightarrow> inv_on_left_moving (as, am) (s', tl l, [hd l]) ires" |
|
3056 apply(auto simp: dec_left_move.simps inv_on_left_moving.simps split: if_splits) |
|
3057 done |
|
3058 |
|
3059 lemma [simp]: "dec_after_write (as, am) (s, l, Oc # r) ires |
|
3060 \<Longrightarrow> dec_on_right_moving (as, am) (s', Oc # l, r) ires" |
|
3061 apply(auto simp: dec_after_write.simps dec_on_right_moving.simps) |
|
3062 apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI, |
|
3063 rule_tac x = "hd lm2" in exI, simp) |
|
3064 apply(rule_tac x = "Suc 0" in exI,rule_tac x = "Suc (hd lm2)" in exI) |
|
3065 apply(case_tac lm2, auto split: if_splits simp: tape_of_nl_cons) |
|
3066 done |
|
3067 |
|
3068 lemma [simp]: "dec_after_write (as, am) (s, l, Bk # r) ires |
|
3069 \<Longrightarrow> dec_after_write (as, am) (s', l, Oc # r) ires" |
|
3070 apply(auto simp: dec_after_write.simps) |
|
3071 done |
|
3072 |
|
3073 lemma [simp]: "dec_after_write (as, am) (s, aaa, []) ires |
|
3074 \<Longrightarrow> dec_after_write (as, am) (s', aaa, [Oc]) ires" |
|
3075 apply(auto simp: dec_after_write.simps) |
|
3076 done |
|
3077 |
|
3078 lemma [simp]: "dec_on_right_moving (as, am) (s, l, Oc # r) ires |
|
3079 \<Longrightarrow> dec_on_right_moving (as, am) (s', Oc # l, r) ires" |
|
3080 apply(simp only: dec_on_right_moving.simps) |
|
3081 apply(erule_tac exE)+ |
|
3082 apply(erule conjE)+ |
|
3083 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
3084 rule_tac x = "m" in exI, rule_tac x = "Suc ml" in exI, |
|
3085 rule_tac x = "mr - 1" in exI, simp) |
|
3086 apply(case_tac mr, auto) |
|
3087 done |
|
3088 |
|
3089 lemma [simp]: "dec_on_right_moving (as, am) (s, l, r) ires\<Longrightarrow> l \<noteq> []" |
|
3090 apply(auto simp: dec_on_right_moving.simps split: if_splits) |
|
3091 done |
|
3092 |
|
3093 lemma [simp]: "dec_on_right_moving (as, am) (s, l, Bk # r) ires |
|
3094 \<Longrightarrow> dec_after_clear (as, am) (s', tl l, hd l # Bk # r) ires" |
|
3095 apply(auto simp: dec_on_right_moving.simps dec_after_clear.simps) |
|
3096 apply(case_tac [!] mr, auto split: if_splits) |
|
3097 done |
|
3098 |
|
3099 lemma [simp]: "dec_on_right_moving (as, am) (s, l, []) ires |
|
3100 \<Longrightarrow> dec_after_clear (as, am) (s', tl l, [hd l]) ires" |
|
3101 apply(auto simp: dec_on_right_moving.simps dec_after_clear.simps) |
|
3102 apply(simp_all split: if_splits) |
|
3103 apply(rule_tac x = lm1 in exI, simp) |
|
3104 done |
|
3105 |
|
3106 lemma [simp]: |
|
3107 "inv_stop (as, abc_lm_s am n (abc_lm_v am n)) (s, l, r) ires \<Longrightarrow> l \<noteq> []" |
|
3108 apply(auto simp: inv_stop.simps) |
|
3109 done |
|
3110 |
|
3111 lemma dec_false_1[simp]: |
|
3112 "\<lbrakk>abc_lm_v am n = 0; inv_locate_b (as, am) (n, aaa, Oc # xs) ires\<rbrakk> |
|
3113 \<Longrightarrow> False" |
|
3114 apply(auto simp: inv_locate_b.simps in_middle.simps) |
|
3115 apply(case_tac "length lm1 \<ge> length am", auto) |
|
3116 apply(subgoal_tac "lm2 = []", simp, subgoal_tac "m = 0", simp) |
|
3117 apply(case_tac mr, auto simp: ) |
|
3118 apply(subgoal_tac "Suc (length lm1) - length am = |
|
3119 Suc (length lm1 - length am)", |
|
3120 simp add: exp_ind del: replicate.simps, simp) |
|
3121 apply(drule_tac xs = "am @ replicate (Suc (length lm1) - length am) 0" |
|
3122 and ys = "lm1 @ m # lm2" in length_equal, simp) |
|
3123 apply(case_tac mr, auto simp: abc_lm_v.simps) |
|
3124 apply(case_tac "mr = 0", simp_all split: if_splits) |
|
3125 apply(subgoal_tac "Suc (length lm1) - length am = |
|
3126 Suc (length lm1 - length am)", |
|
3127 simp add: exp_ind del: replicate.simps, simp) |
|
3128 done |
|
3129 |
|
3130 lemma [simp]: |
|
3131 "\<lbrakk>inv_locate_b (as, am) (n, aaa, Bk # xs) ires; |
|
3132 abc_lm_v am n = 0\<rbrakk> |
|
3133 \<Longrightarrow> inv_on_left_moving (as, abc_lm_s am n 0) |
|
3134 (s, tl aaa, hd aaa # Bk # xs) ires" |
|
3135 apply(simp add: inv_on_left_moving.simps) |
|
3136 apply(simp only: inv_locate_b.simps in_middle.simps) |
|
3137 apply(erule_tac exE)+ |
|
3138 apply(simp add: inv_on_left_moving.simps) |
|
3139 apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B |
|
3140 (as, abc_lm_s am n 0) (s, tl aaa, hd aaa # Bk # xs) ires", simp) |
|
3141 apply(simp only: inv_on_left_moving_norm.simps) |
|
3142 apply(erule_tac conjE)+ |
|
3143 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
3144 rule_tac x = m in exI, rule_tac x = m in exI, |
|
3145 rule_tac x = "Suc 0" in exI, simp add: abc_lm_s.simps) |
|
3146 apply(case_tac mr, simp_all, auto simp: abc_lm_v.simps) |
|
3147 apply(simp only: exp_ind[THEN sym] replicate_Suc Nat.Suc_diff_le) |
|
3148 apply(auto simp: inv_on_left_moving_in_middle_B.simps split: if_splits) |
|
3149 done |
|
3150 |
|
3151 |
|
3152 lemma [simp]: |
|
3153 "\<lbrakk>abc_lm_v am n = 0; inv_locate_b (as, am) (n, aaa, []) ires\<rbrakk> |
|
3154 \<Longrightarrow> inv_on_left_moving (as, abc_lm_s am n 0) (s, tl aaa, [hd aaa]) ires" |
|
3155 apply(simp add: inv_on_left_moving.simps) |
|
3156 apply(simp only: inv_locate_b.simps in_middle.simps) |
|
3157 apply(erule_tac exE)+ |
|
3158 apply(simp add: inv_on_left_moving.simps) |
|
3159 apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B |
|
3160 (as, abc_lm_s am n 0) (s, tl aaa, [hd aaa]) ires", simp) |
|
3161 apply(simp only: inv_on_left_moving_norm.simps) |
|
3162 apply(erule_tac conjE)+ |
|
3163 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
3164 rule_tac x = m in exI, rule_tac x = m in exI, |
|
3165 rule_tac x = "Suc 0" in exI, simp add: abc_lm_s.simps) |
|
3166 apply(case_tac mr, simp_all, auto simp: abc_lm_v.simps) |
|
3167 apply(simp_all only: exp_ind Nat.Suc_diff_le del: replicate_Suc, simp_all) |
|
3168 apply(auto simp: inv_on_left_moving_in_middle_B.simps split: if_splits) |
|
3169 apply(case_tac [!] m, simp_all) |
|
3170 done |
|
3171 |
|
3172 lemma [simp]: "\<lbrakk>am ! n = (0::nat); n < length am\<rbrakk> \<Longrightarrow> am[n := 0] = am" |
|
3173 apply(simp add: list_update_same_conv) |
|
3174 done |
|
3175 |
|
3176 lemma [intro]: "\<lbrakk>abc_lm_v (a # list) 0 = 0\<rbrakk> \<Longrightarrow> a = 0" |
|
3177 apply(simp add: abc_lm_v.simps split: if_splits) |
|
3178 done |
|
3179 |
|
3180 lemma [simp]: |
|
3181 "inv_stop (as, abc_lm_s am n 0) |
|
3182 (start_of (layout_of aprog) e, aaa, Oc # xs) ires |
|
3183 \<Longrightarrow> inv_locate_a (as, abc_lm_s am n 0) (0, aaa, Oc # xs) ires" |
|
3184 apply(simp add: inv_locate_a.simps) |
|
3185 apply(rule disjI1) |
|
3186 apply(auto simp: inv_stop.simps at_begin_norm.simps) |
|
3187 done |
|
3188 |
|
3189 lemma [simp]: |
|
3190 "\<lbrakk>inv_stop (as, abc_lm_s am n 0) |
|
3191 (start_of (layout_of aprog) e, aaa, Oc # xs) ires\<rbrakk> |
|
3192 \<Longrightarrow> inv_locate_b (as, am) (0, Oc # aaa, xs) ires \<or> |
|
3193 inv_locate_b (as, abc_lm_s am n 0) (0, Oc # aaa, xs) ires" |
|
3194 apply(simp) |
|
3195 done |
|
3196 |
|
3197 lemma dec_false2: |
|
3198 "inv_stop (as, abc_lm_s am n 0) |
|
3199 (start_of (layout_of aprog) e, aaa, Bk # xs) ires = False" |
|
3200 apply(auto simp: inv_stop.simps abc_lm_s.simps) |
|
3201 apply(case_tac [!] am, auto) |
|
3202 apply(case_tac [!] n, auto simp: tape_of_nl_cons split: if_splits) |
|
3203 done |
|
3204 |
|
3205 lemma dec_false3: |
|
3206 "inv_stop (as, abc_lm_s am n 0) |
|
3207 (start_of (layout_of aprog) e, aaa, []) ires = False" |
|
3208 apply(auto simp: inv_stop.simps abc_lm_s.simps) |
|
3209 done |
|
3210 |
|
3211 lemma [simp]: |
|
3212 "fetch (ci (layout_of aprog) |
|
3213 (start_of (layout_of aprog) as) (Dec n e)) 0 b = (Nop, 0)" |
|
3214 by(simp add: fetch.simps) |
|
3215 |
|
3216 declare dec_inv_1.simps[simp del] |
|
3217 |
|
3218 declare inv_locate_n_b.simps [simp del] |
|
3219 |
|
3220 lemma [simp]: |
|
3221 "\<lbrakk>0 < abc_lm_v am n; 0 < n; |
|
3222 at_begin_fst_bwtn (as, am) (n, aaa, Oc # xs) ires\<rbrakk> |
|
3223 \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires" |
|
3224 apply(simp add: at_begin_fst_bwtn.simps inv_locate_n_b.simps ) |
|
3225 done |
|
3226 |
|
3227 lemma Suc_minus:"length am + tn = n |
|
3228 \<Longrightarrow> Suc tn = Suc n - length am " |
|
3229 apply(arith) |
|
3230 done |
|
3231 |
|
3232 lemma [simp]: |
|
3233 "\<lbrakk>0 < abc_lm_v am n; 0 < n; |
|
3234 at_begin_fst_awtn (as, am) (n, aaa, Oc # xs) ires\<rbrakk> |
|
3235 \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires" |
|
3236 apply(simp only: at_begin_fst_awtn.simps inv_locate_n_b.simps ) |
|
3237 apply(erule exE)+ |
|
3238 apply(erule conjE)+ |
|
3239 apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, |
|
3240 rule_tac x = "Suc tn" in exI, rule_tac x = 0 in exI) |
|
3241 apply(simp add: exp_ind del: replicate.simps) |
|
3242 apply(rule conjI)+ |
|
3243 apply(auto) |
|
3244 done |
|
3245 |
|
3246 lemma [simp]: |
|
3247 "\<lbrakk>inv_locate_n_b (as, am) (n, aaa, Oc # xs) ires\<rbrakk> |
|
3248 \<Longrightarrow> dec_first_on_right_moving n (as, abc_lm_s am n (abc_lm_v am n)) |
|
3249 (s, Oc # aaa, xs) ires" |
|
3250 apply(auto simp: inv_locate_n_b.simps dec_first_on_right_moving.simps |
|
3251 abc_lm_s.simps abc_lm_v.simps) |
|
3252 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
3253 rule_tac x = m in exI, simp) |
|
3254 apply(rule_tac x = "Suc (Suc 0)" in exI, |
|
3255 rule_tac x = "m - 1" in exI, simp) |
|
3256 apply(case_tac m, auto) |
|
3257 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
3258 rule_tac x = m in exI, |
|
3259 simp add: Suc_diff_le exp_ind del: replicate.simps) |
|
3260 apply(rule_tac x = "Suc (Suc 0)" in exI, |
|
3261 rule_tac x = "m - 1" in exI, simp) |
|
3262 apply(case_tac m, auto) |
|
3263 apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, |
|
3264 rule_tac x = m in exI, simp) |
|
3265 apply(rule_tac x = "Suc (Suc 0)" in exI, |
|
3266 rule_tac x = "m - 1" in exI, simp) |
|
3267 apply(case_tac m, auto) |
|
3268 apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, |
|
3269 rule_tac x = m in exI, |
|
3270 simp add: Suc_diff_le exp_ind del: replicate.simps, simp) |
|
3271 done |
|
3272 |
|
3273 lemma [simp]: "inv_on_left_moving (as, am) (s, [], r) ires |
|
3274 = False" |
|
3275 apply(simp add: inv_on_left_moving.simps inv_on_left_moving_norm.simps |
|
3276 inv_on_left_moving_in_middle_B.simps) |
|
3277 done |
|
3278 |
|
3279 lemma [simp]: |
|
3280 "inv_check_left_moving (as, abc_lm_s am n 0) |
|
3281 (start_of (layout_of aprog) as + 2 * n + 14, [], Oc # xs) ires |
|
3282 = False" |
|
3283 apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps) |
|
3284 done |
|
3285 |
|
3286 lemma [simp]: "inv_check_left_moving (as, abc_lm_s lm n (abc_lm_v lm n)) (s, [], Oc # list) ires = False" |
|
3287 apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps) |
|
3288 done |
|
3289 |
|
3290 lemma [elim]: "\<lbrakk>abc_fetch as ap = Some (Dec n e); |
|
3291 start_of (layout_of ap) as < start_of (layout_of ap) e; |
|
3292 start_of (layout_of ap) e \<le> Suc (start_of (layout_of ap) as + 2 * n)\<rbrakk> |
|
3293 \<Longrightarrow> RR" |
|
3294 using start_of_less[of e as "layout_of ap"] start_of_ge[of as ap n e "layout_of ap"] |
|
3295 apply(case_tac "as < e", simp) |
|
3296 apply(case_tac "as = e", simp, simp) |
|
3297 done |
|
3298 |
|
3299 lemma crsp_step_dec_b_e_pre': |
|
3300 assumes layout: "ly = layout_of ap" |
|
3301 and inv_start: "inv_locate_b (as, lm) (n, la, ra) ires" |
|
3302 and fetch: "abc_fetch as ap = Some (Dec n e)" |
|
3303 and dec_0: "abc_lm_v lm n = 0" |
|
3304 and f: "f = (\<lambda> stp. (steps (Suc (start_of ly as) + 2 * n, la, ra) (ci ly (start_of ly as) (Dec n e), |
|
3305 start_of ly as - Suc 0) stp, start_of ly as, n))" |
|
3306 and P: "P = (\<lambda> ((s, l, r), ss, x). s = start_of ly e)" |
|
3307 and Q: "Q = (\<lambda> ((s, l, r), ss, x). dec_inv_1 ly x e (as, lm) (s, l, r) ires)" |
|
3308 shows "\<exists> stp. P (f stp) \<and> Q (f stp)" |
|
3309 proof(rule_tac LE = abc_dec_1_LE in halt_lemma2) |
|
3310 show "wf abc_dec_1_LE" by(intro wf_dec_le) |
|
3311 next |
|
3312 show "Q (f 0)" |
|
3313 using layout fetch |
|
3314 apply(simp add: f steps.simps Q dec_inv_1.simps) |
|
3315 apply(subgoal_tac "e > as \<or> e = as \<or> e < as") |
|
3316 apply(auto simp: Let_def start_of_ge start_of_less inv_start) |
|
3317 done |
|
3318 next |
|
3319 show "\<not> P (f 0)" |
|
3320 using layout fetch |
|
3321 apply(simp add: f steps.simps P) |
|
3322 done |
|
3323 next |
|
3324 show "\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> abc_dec_1_LE" |
|
3325 using fetch |
|
3326 proof(rule_tac allI, rule_tac impI) |
|
3327 fix na |
|
3328 assume "\<not> P (f na) \<and> Q (f na)" |
|
3329 thus "Q (f (Suc na)) \<and> (f (Suc na), f na) \<in> abc_dec_1_LE" |
|
3330 apply(simp add: f) |
|
3331 apply(case_tac "steps (Suc (start_of ly as + 2 * n), la, ra) |
|
3332 (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) na", simp) |
|
3333 proof - |
|
3334 fix a b c |
|
3335 assume "\<not> P ((a, b, c), start_of ly as, n) \<and> Q ((a, b, c), start_of ly as, n)" |
|
3336 thus "Q (step (a, b, c) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0), start_of ly as, n) \<and> |
|
3337 ((step (a, b, c) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0), start_of ly as, n), |
|
3338 (a, b, c), start_of ly as, n) \<in> abc_dec_1_LE" |
|
3339 apply(simp add: Q) |
|
3340 apply(case_tac c, case_tac [2] aa) |
|
3341 apply(simp_all add: dec_inv_1.simps Let_def split: if_splits) |
|
3342 using fetch layout dec_0 |
|
3343 apply(auto simp: step.simps P dec_inv_1.simps Let_def abc_dec_1_LE_def lex_triple_def lex_pair_def) |
|
3344 using dec_0 |
|
3345 apply(drule_tac dec_false_1, simp_all) |
|
3346 done |
|
3347 qed |
|
3348 qed |
|
3349 qed |
|
3350 |
|
3351 lemma crsp_step_dec_b_e_pre: |
|
3352 assumes "ly = layout_of ap" |
|
3353 and inv_start: "inv_locate_b (as, lm) (n, la, ra) ires" |
|
3354 and dec_0: "abc_lm_v lm n = 0" |
|
3355 and fetch: "abc_fetch as ap = Some (Dec n e)" |
|
3356 shows "\<exists>stp lb rb. |
|
3357 steps (Suc (start_of ly as) + 2 * n, la, ra) (ci ly (start_of ly as) (Dec n e), |
|
3358 start_of ly as - Suc 0) stp = (start_of ly e, lb, rb) \<and> |
|
3359 dec_inv_1 ly n e (as, lm) (start_of ly e, lb, rb) ires" |
|
3360 using assms |
|
3361 apply(drule_tac crsp_step_dec_b_e_pre', auto) |
|
3362 apply(rule_tac x = stp in exI, simp) |
|
3363 done |
|
3364 |
|
3365 lemma [simp]: |
|
3366 "\<lbrakk>abc_lm_v lm n = 0; |
|
3367 inv_stop (as, abc_lm_s lm n (abc_lm_v lm n)) (start_of ly e, lb, rb) ires\<rbrakk> |
|
3368 \<Longrightarrow> crsp ly (abc_step_l (as, lm) (Some (Dec n e))) (start_of ly e, lb, rb) ires" |
|
3369 apply(auto simp: crsp.simps abc_step_l.simps inv_stop.simps) |
|
3370 done |
|
3371 |
|
3372 lemma crsp_step_dec_b_e: |
|
3373 assumes layout: "ly = layout_of ap" |
|
3374 and inv_start: "inv_locate_a (as, lm) (n, l, r) ires" |
|
3375 and dec_0: "abc_lm_v lm n = 0" |
|
3376 and fetch: "abc_fetch as ap = Some (Dec n e)" |
|
3377 shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e))) |
|
3378 (steps (start_of ly as + 2 * n, l, r) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires" |
|
3379 proof - |
|
3380 let ?P = "ci ly (start_of ly as) (Dec n e)" |
|
3381 let ?off = "start_of ly as - Suc 0" |
|
3382 have "\<exists> stp la ra. steps (start_of ly as + 2 * n, l, r) (?P, ?off) stp = (Suc (start_of ly as) + 2*n, la, ra) |
|
3383 \<and> inv_locate_b (as, lm) (n, la, ra) ires" |
|
3384 using inv_start |
|
3385 apply(case_tac "r = [] \<or> hd r = Bk", simp_all) |
|
3386 done |
|
3387 from this obtain stpa la ra where a: |
|
3388 "steps (start_of ly as + 2 * n, l, r) (?P, ?off) stpa = (Suc (start_of ly as) + 2*n, la, ra) |
|
3389 \<and> inv_locate_b (as, lm) (n, la, ra) ires" by blast |
|
3390 term dec_inv_1 |
|
3391 have "\<exists> stp lb rb. steps (Suc (start_of ly as) + 2 * n, la, ra) (?P, ?off) stp = (start_of ly e, lb, rb) |
|
3392 \<and> dec_inv_1 ly n e (as, lm) (start_of ly e, lb, rb) ires" |
|
3393 using assms a |
|
3394 apply(rule_tac crsp_step_dec_b_e_pre, auto) |
|
3395 done |
|
3396 from this obtain stpb lb rb where b: |
|
3397 "steps (Suc (start_of ly as) + 2 * n, la, ra) (?P, ?off) stpb = (start_of ly e, lb, rb) |
|
3398 \<and> dec_inv_1 ly n e (as, lm) (start_of ly e, lb, rb) ires" by blast |
|
3399 from a b show "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e))) |
|
3400 (steps (start_of ly as + 2 * n, l, r) (?P, ?off) stp) ires" |
|
3401 apply(rule_tac x = "stpa + stpb" in exI) |
|
3402 apply(simp add: steps_add) |
|
3403 using dec_0 |
|
3404 apply(simp add: dec_inv_1.simps) |
|
3405 apply(case_tac stpa, simp_all add: steps.simps) |
|
3406 done |
|
3407 qed |
|
3408 |
|
3409 fun dec_inv_2 :: "layout \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> dec_inv_t" |
|
3410 where |
|
3411 "dec_inv_2 ly n e (as, am) (s, l, r) ires = |
|
3412 (let ss = start_of ly as in |
|
3413 let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in |
|
3414 let am'' = abc_lm_s am n (abc_lm_v am n) in |
|
3415 if s = 0 then False |
|
3416 else if s = ss + 2 * n then |
|
3417 inv_locate_a (as, am) (n, l, r) ires |
|
3418 else if s = ss + 2 * n + 1 then |
|
3419 inv_locate_n_b (as, am) (n, l, r) ires |
|
3420 else if s = ss + 2 * n + 2 then |
|
3421 dec_first_on_right_moving n (as, am'') (s, l, r) ires |
|
3422 else if s = ss + 2 * n + 3 then |
|
3423 dec_after_clear (as, am') (s, l, r) ires |
|
3424 else if s = ss + 2 * n + 4 then |
|
3425 dec_right_move (as, am') (s, l, r) ires |
|
3426 else if s = ss + 2 * n + 5 then |
|
3427 dec_check_right_move (as, am') (s, l, r) ires |
|
3428 else if s = ss + 2 * n + 6 then |
|
3429 dec_left_move (as, am') (s, l, r) ires |
|
3430 else if s = ss + 2 * n + 7 then |
|
3431 dec_after_write (as, am') (s, l, r) ires |
|
3432 else if s = ss + 2 * n + 8 then |
|
3433 dec_on_right_moving (as, am') (s, l, r) ires |
|
3434 else if s = ss + 2 * n + 9 then |
|
3435 dec_after_clear (as, am') (s, l, r) ires |
|
3436 else if s = ss + 2 * n + 10 then |
|
3437 inv_on_left_moving (as, am') (s, l, r) ires |
|
3438 else if s = ss + 2 * n + 11 then |
|
3439 inv_check_left_moving (as, am') (s, l, r) ires |
|
3440 else if s = ss + 2 * n + 12 then |
|
3441 inv_after_left_moving (as, am') (s, l, r) ires |
|
3442 else if s = ss + 2 * n + 16 then |
|
3443 inv_stop (as, am') (s, l, r) ires |
|
3444 else False)" |
|
3445 |
|
3446 declare dec_inv_2.simps[simp del] |
|
3447 fun abc_dec_2_stage1 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
3448 where |
|
3449 "abc_dec_2_stage1 (s, l, r) ss n = |
|
3450 (if s \<le> ss + 2*n + 1 then 7 |
|
3451 else if s = ss + 2*n + 2 then 6 |
|
3452 else if s = ss + 2*n + 3 then 5 |
|
3453 else if s \<ge> ss + 2*n + 4 \<and> s \<le> ss + 2*n + 9 then 4 |
|
3454 else if s = ss + 2*n + 6 then 3 |
|
3455 else if s = ss + 2*n + 10 \<or> s = ss + 2*n + 11 then 2 |
|
3456 else if s = ss + 2*n + 12 then 1 |
|
3457 else 0)" |
|
3458 |
|
3459 fun abc_dec_2_stage2 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
3460 where |
|
3461 "abc_dec_2_stage2 (s, l, r) ss n = |
|
3462 (if s \<le> ss + 2 * n + 1 then (ss + 2 * n + 16 - s) |
|
3463 else if s = ss + 2*n + 10 then length l |
|
3464 else if s = ss + 2*n + 11 then length l |
|
3465 else if s = ss + 2*n + 4 then length r - 1 |
|
3466 else if s = ss + 2*n + 5 then length r |
|
3467 else if s = ss + 2*n + 7 then length r - 1 |
|
3468 else if s = ss + 2*n + 8 then |
|
3469 length r + length (takeWhile (\<lambda> a. a = Oc) l) - 1 |
|
3470 else if s = ss + 2*n + 9 then |
|
3471 length r + length (takeWhile (\<lambda> a. a = Oc) l) - 1 |
|
3472 else 0)" |
|
3473 |
|
3474 fun abc_dec_2_stage3 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
3475 where |
|
3476 "abc_dec_2_stage3 (s, l, r) ss n = |
|
3477 (if s \<le> ss + 2*n + 1 then |
|
3478 if (s - ss) mod 2 = 0 then if r \<noteq> [] \<and> |
|
3479 hd r = Oc then 0 else 1 |
|
3480 else length r |
|
3481 else if s = ss + 2 * n + 10 then |
|
3482 if r \<noteq> [] \<and> hd r = Oc then 2 |
|
3483 else 1 |
|
3484 else if s = ss + 2 * n + 11 then |
|
3485 if r \<noteq> [] \<and> hd r = Oc then 3 |
|
3486 else 0 |
|
3487 else (ss + 2 * n + 16 - s))" |
|
3488 |
|
3489 fun abc_dec_2_stage4 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" |
|
3490 where |
|
3491 "abc_dec_2_stage4 (s, l, r) ss n = |
|
3492 (if s = ss + 2*n + 2 then length r |
|
3493 else if s = ss + 2*n + 8 then length r |
|
3494 else if s = ss + 2*n + 3 then |
|
3495 if r \<noteq> [] \<and> hd r = Oc then 1 |
|
3496 else 0 |
|
3497 else if s = ss + 2*n + 7 then |
|
3498 if r \<noteq> [] \<and> hd r = Oc then 0 |
|
3499 else 1 |
|
3500 else if s = ss + 2*n + 9 then |
|
3501 if r \<noteq> [] \<and> hd r = Oc then 1 |
|
3502 else 0 |
|
3503 else 0)" |
|
3504 |
|
3505 fun abc_dec_2_measure :: "(config \<times> nat \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat \<times> nat)" |
|
3506 where |
|
3507 "abc_dec_2_measure (c, ss, n) = |
|
3508 (abc_dec_2_stage1 c ss n, |
|
3509 abc_dec_2_stage2 c ss n, abc_dec_2_stage3 c ss n, abc_dec_2_stage4 c ss n)" |
|
3510 |
|
3511 definition lex_square:: |
|
3512 "((nat \<times> nat \<times> nat \<times> nat) \<times> (nat \<times> nat \<times> nat \<times> nat)) set" |
|
3513 where "lex_square \<equiv> less_than <*lex*> lex_triple" |
|
3514 |
|
3515 definition abc_dec_2_LE :: |
|
3516 "((config \<times> nat \<times> |
|
3517 nat) \<times> (config \<times> nat \<times> nat)) set" |
|
3518 where "abc_dec_2_LE \<equiv> (inv_image lex_square abc_dec_2_measure)" |
|
3519 |
|
3520 lemma wf_dec2_le: "wf abc_dec_2_LE" |
|
3521 by(auto intro:wf_inv_image simp:abc_dec_2_LE_def lex_square_def lex_triple_def lex_pair_def) |
|
3522 |
|
3523 lemma fix_add: "fetch ap ((x::nat) + 2*n) b = fetch ap (2*n + x) b" |
|
3524 by (metis Suc_1 mult_2 nat_add_commute) |
|
3525 |
|
3526 lemma [elim]: |
|
3527 "\<lbrakk>0 < abc_lm_v am n; inv_locate_n_b (as, am) (n, aaa, Bk # xs) ires\<rbrakk> |
|
3528 \<Longrightarrow> RR" |
|
3529 apply(auto simp: inv_locate_n_b.simps abc_lm_v.simps split: if_splits) |
|
3530 apply(case_tac [!] m, auto) |
|
3531 done |
|
3532 |
|
3533 lemma [elim]: |
|
3534 "\<lbrakk>0 < abc_lm_v am n; inv_locate_n_b (as, am) |
|
3535 (n, aaa, []) ires\<rbrakk> \<Longrightarrow> RR" |
|
3536 apply(auto simp: inv_locate_n_b.simps abc_lm_v.simps split: if_splits) |
|
3537 done |
|
3538 |
|
3539 lemma [simp]: "dec_after_write (as, am) (s, aa, r) ires |
|
3540 \<Longrightarrow> takeWhile (\<lambda>a. a = Oc) aa = []" |
|
3541 apply(simp only : dec_after_write.simps) |
|
3542 apply(erule exE)+ |
|
3543 apply(erule_tac conjE)+ |
|
3544 apply(case_tac aa, simp) |
|
3545 apply(case_tac a, simp only: takeWhile.simps , simp_all split: if_splits) |
|
3546 done |
|
3547 |
|
3548 lemma [simp]: |
|
3549 "\<lbrakk>dec_on_right_moving (as, lm) (s, aa, []) ires; |
|
3550 length (takeWhile (\<lambda>a. a = Oc) (tl aa)) |
|
3551 \<noteq> length (takeWhile (\<lambda>a. a = Oc) aa) - Suc 0\<rbrakk> |
|
3552 \<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (tl aa)) < |
|
3553 length (takeWhile (\<lambda>a. a = Oc) aa) - Suc 0" |
|
3554 apply(simp only: dec_on_right_moving.simps) |
|
3555 apply(erule_tac exE)+ |
|
3556 apply(erule_tac conjE)+ |
|
3557 apply(case_tac mr, auto split: if_splits) |
|
3558 done |
|
3559 |
|
3560 lemma [simp]: |
|
3561 "dec_after_clear (as, abc_lm_s am n (abc_lm_v am n - Suc 0)) |
|
3562 (start_of (layout_of aprog) as + 2 * n + 9, aa, Bk # xs) ires |
|
3563 \<Longrightarrow> length xs - Suc 0 < length xs + |
|
3564 length (takeWhile (\<lambda>a. a = Oc) aa)" |
|
3565 apply(simp only: dec_after_clear.simps) |
|
3566 apply(erule_tac exE)+ |
|
3567 apply(erule conjE)+ |
|
3568 apply(simp split: if_splits ) |
|
3569 done |
|
3570 |
|
3571 lemma [simp]: |
|
3572 "\<lbrakk>dec_after_clear (as, abc_lm_s am n (abc_lm_v am n - Suc 0)) |
|
3573 (start_of (layout_of aprog) as + 2 * n + 9, aa, []) ires\<rbrakk> |
|
3574 \<Longrightarrow> Suc 0 < length (takeWhile (\<lambda>a. a = Oc) aa)" |
|
3575 apply(simp add: dec_after_clear.simps split: if_splits) |
|
3576 done |
|
3577 |
|
3578 lemma [elim]: |
|
3579 "inv_check_left_moving (as, lm) |
|
3580 (s, [], Oc # xs) ires |
|
3581 \<Longrightarrow> RR" |
|
3582 apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps) |
|
3583 done |
|
3584 |
|
3585 lemma [simp]: |
|
3586 "\<lbrakk>0 < abc_lm_v am n; |
|
3587 at_begin_norm (as, am) (n, aaa, Oc # xs) ires\<rbrakk> |
|
3588 \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires" |
|
3589 apply(simp only: at_begin_norm.simps inv_locate_n_b.simps) |
|
3590 apply(erule_tac exE)+ |
|
3591 apply(rule_tac x = lm1 in exI, simp) |
|
3592 apply(case_tac "length lm2", simp) |
|
3593 apply(case_tac "lm2", simp, simp) |
|
3594 apply(case_tac "lm2", auto simp: tape_of_nl_cons split: if_splits) |
|
3595 done |
|
3596 |
|
3597 lemma [simp]: |
|
3598 "\<lbrakk>0 < abc_lm_v am n; |
|
3599 at_begin_fst_awtn (as, am) (n, aaa, Oc # xs) ires\<rbrakk> |
|
3600 \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires" |
|
3601 apply(simp only: at_begin_fst_awtn.simps inv_locate_n_b.simps ) |
|
3602 apply(erule exE)+ |
|
3603 apply(erule conjE)+ |
|
3604 apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI, |
|
3605 rule_tac x = "Suc tn" in exI, rule_tac x = 0 in exI) |
|
3606 apply(simp add: exp_ind del: replicate.simps) |
|
3607 apply(rule conjI)+ |
|
3608 apply(auto) |
|
3609 done |
|
3610 |
|
3611 lemma [simp]: |
|
3612 "\<lbrakk>0 < abc_lm_v am n; inv_locate_a (as, am) (n, aaa, Oc # xs) ires\<rbrakk> |
|
3613 \<Longrightarrow> inv_locate_n_b (as, am) (n, Oc#aaa, xs) ires" |
|
3614 apply(auto simp: inv_locate_a.simps at_begin_fst_bwtn.simps) |
|
3615 done |
|
3616 |
|
3617 lemma [simp]: |
|
3618 "\<lbrakk>dec_on_right_moving (as, am) (s, aa, Bk # xs) ires; |
|
3619 Suc (length (takeWhile (\<lambda>a. a = Oc) (tl aa))) |
|
3620 \<noteq> length (takeWhile (\<lambda>a. a = Oc) aa)\<rbrakk> |
|
3621 \<Longrightarrow> Suc (length (takeWhile (\<lambda>a. a = Oc) (tl aa))) |
|
3622 < length (takeWhile (\<lambda>a. a = Oc) aa)" |
|
3623 apply(simp only: dec_on_right_moving.simps) |
|
3624 apply(erule exE)+ |
|
3625 apply(erule conjE)+ |
|
3626 apply(case_tac ml, auto split: if_splits ) |
|
3627 done |
|
3628 |
|
3629 lemma crsp_step_dec_b_suc_pre: |
|
3630 assumes layout: "ly = layout_of ap" |
|
3631 and crsp: "crsp ly (as, lm) (s, l, r) ires" |
|
3632 and inv_start: "inv_locate_a (as, lm) (n, la, ra) ires" |
|
3633 and fetch: "abc_fetch as ap = Some (Dec n e)" |
|
3634 and dec_suc: "0 < abc_lm_v lm n" |
|
3635 and f: "f = (\<lambda> stp. (steps (start_of ly as + 2 * n, la, ra) (ci ly (start_of ly as) (Dec n e), |
|
3636 start_of ly as - Suc 0) stp, start_of ly as, n))" |
|
3637 and P: "P = (\<lambda> ((s, l, r), ss, x). s = start_of ly as + 2*n + 16)" |
|
3638 and Q: "Q = (\<lambda> ((s, l, r), ss, x). dec_inv_2 ly x e (as, lm) (s, l, r) ires)" |
|
3639 shows "\<exists> stp. P (f stp) \<and> Q(f stp)" |
|
3640 proof(rule_tac LE = abc_dec_2_LE in halt_lemma2) |
|
3641 show "wf abc_dec_2_LE" by(intro wf_dec2_le) |
|
3642 next |
|
3643 show "Q (f 0)" |
|
3644 using layout fetch inv_start |
|
3645 apply(simp add: f steps.simps Q) |
|
3646 apply(simp only: dec_inv_2.simps) |
|
3647 apply(auto simp: Let_def start_of_ge start_of_less inv_start dec_inv_2.simps) |
|
3648 done |
|
3649 next |
|
3650 show "\<not> P (f 0)" |
|
3651 using layout fetch |
|
3652 apply(simp add: f steps.simps P) |
|
3653 done |
|
3654 next |
|
3655 show "\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> abc_dec_2_LE" |
|
3656 using fetch |
|
3657 proof(rule_tac allI, rule_tac impI) |
|
3658 fix na |
|
3659 assume "\<not> P (f na) \<and> Q (f na)" |
|
3660 thus "Q (f (Suc na)) \<and> (f (Suc na), f na) \<in> abc_dec_2_LE" |
|
3661 apply(simp add: f) |
|
3662 apply(case_tac "steps ((start_of ly as + 2 * n), la, ra) |
|
3663 (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) na", simp) |
|
3664 proof - |
|
3665 fix a b c |
|
3666 assume "\<not> P ((a, b, c), start_of ly as, n) \<and> Q ((a, b, c), start_of ly as, n)" |
|
3667 thus "Q (step (a, b, c) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0), start_of ly as, n) \<and> |
|
3668 ((step (a, b, c) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0), start_of ly as, n), |
|
3669 (a, b, c), start_of ly as, n) \<in> abc_dec_2_LE" |
|
3670 apply(simp add: Q) |
|
3671 apply(erule_tac conjE) |
|
3672 apply(case_tac c, case_tac [2] aa) |
|
3673 apply(simp_all add: dec_inv_2.simps Let_def) |
|
3674 apply(simp_all split: if_splits) |
|
3675 using fetch layout dec_suc |
|
3676 apply(auto simp: step.simps P dec_inv_2.simps Let_def abc_dec_2_LE_def lex_triple_def lex_pair_def lex_square_def |
|
3677 fix_add numeral_3_eq_3) |
|
3678 done |
|
3679 qed |
|
3680 qed |
|
3681 qed |
|
3682 |
|
3683 lemma [simp]: |
|
3684 "\<lbrakk>inv_stop (as, abc_lm_s lm n (abc_lm_v lm n - Suc 0)) |
|
3685 (start_of (layout_of ap) as + 2 * n + 16, a, b) ires; |
|
3686 abc_lm_v lm n > 0; |
|
3687 abc_fetch as ap = Some (Dec n e)\<rbrakk> |
|
3688 \<Longrightarrow> crsp (layout_of ap) (abc_step_l (as, lm) (Some (Dec n e))) |
|
3689 (start_of (layout_of ap) as + 2 * n + 16, a, b) ires" |
|
3690 apply(auto simp: inv_stop.simps crsp.simps abc_step_l.simps startof_Suc2) |
|
3691 apply(drule_tac startof_Suc2, simp) |
|
3692 done |
|
3693 |
|
3694 lemma crsp_step_dec_b_suc: |
|
3695 assumes layout: "ly = layout_of ap" |
|
3696 and crsp: "crsp ly (as, lm) (s, l, r) ires" |
|
3697 and inv_start: "inv_locate_a (as, lm) (n, la, ra) ires" |
|
3698 and fetch: "abc_fetch as ap = Some (Dec n e)" |
|
3699 and dec_suc: "0 < abc_lm_v lm n" |
|
3700 shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e))) |
|
3701 (steps (start_of ly as + 2 * n, la, ra) (ci (layout_of ap) |
|
3702 (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires" |
|
3703 using assms |
|
3704 apply(drule_tac crsp_step_dec_b_suc_pre, auto) |
|
3705 apply(rule_tac x = stp in exI, simp) |
|
3706 apply(simp add: dec_inv_2.simps) |
|
3707 apply(case_tac stp, simp_all add: steps.simps) |
|
3708 done |
|
3709 |
|
3710 lemma crsp_step_dec_b: |
|
3711 assumes layout: "ly = layout_of ap" |
|
3712 and crsp: "crsp ly (as, lm) (s, l, r) ires" |
|
3713 and inv_start: "inv_locate_a (as, lm) (n, la, ra) ires" |
|
3714 and fetch: "abc_fetch as ap = Some (Dec n e)" |
|
3715 shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e))) |
|
3716 (steps (start_of ly as + 2 * n, la, ra) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires" |
|
3717 using assms |
|
3718 apply(case_tac "abc_lm_v lm n = 0") |
|
3719 apply(rule_tac crsp_step_dec_b_e, simp_all) |
|
3720 apply(rule_tac crsp_step_dec_b_suc, simp_all) |
|
3721 done |
|
3722 |
567 lemma crsp_step_dec: |
3723 lemma crsp_step_dec: |
568 assumes layout: "ly = layout_of ap" |
3724 assumes layout: "ly = layout_of ap" |
569 and crsp: "crsp ly (as, lm) (s, l, r) ires" |
3725 and crsp: "crsp ly (as, lm) (s, l, r) ires" |
570 shows "\<exists>stp. crsp ly (abc_step_l (as, lm) (Some (Dec n e))) |
3726 and fetch: "abc_fetch as ap = Some (Dec n e)" |
|
3727 shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e))) |
571 (steps (s, l, r) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires" |
3728 (steps (s, l, r) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires" |
572 sorry |
3729 proof(simp add: ci.simps) |
573 |
3730 let ?off = "start_of ly as - Suc 0" |
|
3731 let ?A = "findnth n" |
|
3732 let ?B = "sete (shift (shift tdec_b (2 * n)) ?off) (start_of ly e)" |
|
3733 have "\<exists> stp la ra. steps (s, l, r) (shift ?A ?off @ ?B, ?off) stp = (start_of ly as + 2*n, la, ra) |
|
3734 \<and> inv_locate_a (as, lm) (n, la, ra) ires" |
|
3735 proof - |
|
3736 have "\<exists>stp l' r'. steps (Suc 0, l, r) (?A, 0) stp = (Suc (2 * n), l', r') \<and> |
|
3737 inv_locate_a (as, lm) (n, l', r') ires" |
|
3738 using assms |
|
3739 apply(rule_tac findnth_correct, simp_all) |
|
3740 done |
|
3741 then obtain stp l' r' where a: |
|
3742 "steps (Suc 0, l, r) (?A, 0) stp = (Suc (2 * n), l', r') \<and> |
|
3743 inv_locate_a (as, lm) (n, l', r') ires" by blast |
|
3744 then have "steps (Suc 0 + ?off, l, r) (shift ?A ?off, ?off) stp = (Suc (2 * n) + ?off, l', r')" |
|
3745 apply(rule_tac tm_shift_eq_steps, simp_all) |
|
3746 done |
|
3747 moreover have "s = start_of ly as" |
|
3748 using crsp |
|
3749 apply(auto simp: crsp.simps) |
|
3750 done |
|
3751 ultimately show "\<exists> stp la ra. steps (s, l, r) (shift ?A ?off @ ?B, ?off) stp = (start_of ly as + 2*n, la, ra) |
|
3752 \<and> inv_locate_a (as, lm) (n, la, ra) ires" |
|
3753 using a |
|
3754 apply(drule_tac B = ?B in tm_append_first_steps_eq, auto) |
|
3755 apply(rule_tac x = stp in exI, simp) |
|
3756 done |
|
3757 qed |
|
3758 from this obtain stpa la ra where a: |
|
3759 "steps (s, l, r) (shift ?A ?off @ ?B, ?off) stpa = (start_of ly as + 2*n, la, ra) |
|
3760 \<and> inv_locate_a (as, lm) (n, la, ra) ires" by blast |
|
3761 have "\<exists>stp. crsp ly (abc_step_l (as, lm) (Some (Dec n e))) |
|
3762 (steps (start_of ly as + 2*n, la, ra) (shift ?A ?off @ ?B, ?off) stp) ires \<and> stp > 0" |
|
3763 using assms a |
|
3764 apply(drule_tac crsp_step_dec_b, auto) |
|
3765 apply(rule_tac x = stp in exI, simp add: ci.simps) |
|
3766 done |
|
3767 then obtain stpb where b: |
|
3768 "crsp ly (abc_step_l (as, lm) (Some (Dec n e))) |
|
3769 (steps (start_of ly as + 2*n, la, ra) (shift ?A ?off @ ?B, ?off) stpb) ires \<and> stpb > 0" .. |
|
3770 from a b show "\<exists> stp>0. crsp ly (abc_step_l (as, lm) (Some (Dec n e))) |
|
3771 (steps (s, l, r) (shift ?A ?off @ ?B, ?off) stp) ires" |
|
3772 apply(rule_tac x = "stpa + stpb" in exI) |
|
3773 apply(simp add: steps_add) |
|
3774 done |
|
3775 qed |
|
3776 |
574 subsection{*Crsp of Goto*} |
3777 subsection{*Crsp of Goto*} |
|
3778 |
575 lemma crsp_step_goto: |
3779 lemma crsp_step_goto: |
576 assumes layout: "ly = layout_of ap" |
3780 assumes layout: "ly = layout_of ap" |
577 and crsp: "crsp ly (as, lm) (s, l, r) ires" |
3781 and crsp: "crsp ly (as, lm) (s, l, r) ires" |
578 shows "\<exists>stp. crsp ly (abc_step_l (as, lm) (Some (Goto n))) |
3782 shows "\<exists>stp>0. crsp ly (abc_step_l (as, lm) (Some (Goto n))) |
579 (steps (s, l, r) (ci ly (start_of ly as) (Goto n), |
3783 (steps (s, l, r) (ci ly (start_of ly as) (Goto n), |
580 start_of ly as - Suc 0) stp) ires" |
3784 start_of ly as - Suc 0) stp) ires" |
581 sorry |
3785 using crsp |
582 |
3786 apply(rule_tac x = "Suc 0" in exI) |
|
3787 apply(case_tac r, case_tac [2] a) |
|
3788 apply(simp_all add: ci.simps steps.simps step.simps crsp.simps fetch.simps |
|
3789 crsp.simps abc_step_l.simps) |
|
3790 done |
583 |
3791 |
584 lemma crsp_step_in: |
3792 lemma crsp_step_in: |
585 assumes layout: "ly = layout_of ap" |
3793 assumes layout: "ly = layout_of ap" |
586 and compile: "tp = tm_of ap" |
3794 and compile: "tp = tm_of ap" |
587 and crsp: "crsp ly (as, lm) (s, l, r) ires" |
3795 and crsp: "crsp ly (as, lm) (s, l, r) ires" |
588 and fetch: "abc_fetch as ap = Some ins" |
3796 and fetch: "abc_fetch as ap = Some ins" |
589 shows "\<exists> stp. crsp ly (abc_step_l (as, lm) (Some ins)) |
3797 shows "\<exists> stp>0. crsp ly (abc_step_l (as, lm) (Some ins)) |
590 (steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp) ires" |
3798 (steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp) ires" |
591 using assms |
3799 using assms |
592 apply(case_tac ins, simp_all) |
3800 apply(case_tac ins, simp_all) |
593 apply(rule crsp_step_inc, simp_all) |
3801 apply(rule crsp_step_inc, simp_all) |
594 apply(rule crsp_step_dec, simp_all) |
3802 apply(rule crsp_step_dec, simp_all) |
684 fun mopup :: "nat \<Rightarrow> instr list" |
4014 fun mopup :: "nat \<Rightarrow> instr list" |
685 where |
4015 where |
686 "mopup n = mopup_a n @ shift mopup_b (2*n)" |
4016 "mopup n = mopup_a n @ shift mopup_b (2*n)" |
687 (****) |
4017 (****) |
688 |
4018 |
|
4019 type_synonym mopup_type = "config \<Rightarrow> nat list \<Rightarrow> nat \<Rightarrow> cell list \<Rightarrow> bool" |
|
4020 |
|
4021 fun mopup_stop :: "mopup_type" |
|
4022 where |
|
4023 "mopup_stop (s, l, r) lm n ires= |
|
4024 (\<exists> ln rn. l = Bk\<up>ln @ Bk # Bk # ires \<and> r = <abc_lm_v lm n> @ Bk\<up>rn)" |
|
4025 |
|
4026 fun mopup_bef_erase_a :: "mopup_type" |
|
4027 where |
|
4028 "mopup_bef_erase_a (s, l, r) lm n ires= |
|
4029 (\<exists> ln m rn. l = Bk\<up>ln @ Bk # Bk # ires \<and> |
|
4030 r = Oc\<up>m@ Bk # <(drop ((s + 1) div 2) lm)> @ Bk\<up>rn)" |
|
4031 |
|
4032 fun mopup_bef_erase_b :: "mopup_type" |
|
4033 where |
|
4034 "mopup_bef_erase_b (s, l, r) lm n ires = |
|
4035 (\<exists> ln m rn. l = Bk\<up>ln @ Bk # Bk # ires \<and> r = Bk # Oc\<up>m @ Bk # |
|
4036 <(drop (s div 2) lm)> @ Bk\<up>rn)" |
|
4037 |
|
4038 fun mopup_jump_over1 :: "mopup_type" |
|
4039 where |
|
4040 "mopup_jump_over1 (s, l, r) lm n ires = |
|
4041 (\<exists> ln m1 m2 rn. m1 + m2 = Suc (abc_lm_v lm n) \<and> |
|
4042 l = Oc\<up>m1 @ Bk\<up>ln @ Bk # Bk # ires \<and> |
|
4043 (r = Oc\<up>m2 @ Bk # <(drop (Suc n) lm)> @ Bk\<up>rn \<or> |
|
4044 (r = Oc\<up>m2 \<and> (drop (Suc n) lm) = [])))" |
|
4045 |
|
4046 fun mopup_aft_erase_a :: "mopup_type" |
|
4047 where |
|
4048 "mopup_aft_erase_a (s, l, r) lm n ires = |
|
4049 (\<exists> lnl lnr rn (ml::nat list) m. |
|
4050 m = Suc (abc_lm_v lm n) \<and> l = Bk\<up>lnr @ Oc\<up>m @ Bk\<up>lnl @ Bk # Bk # ires \<and> |
|
4051 (r = <ml> @ Bk\<up>rn))" |
|
4052 |
|
4053 fun mopup_aft_erase_b :: "mopup_type" |
|
4054 where |
|
4055 "mopup_aft_erase_b (s, l, r) lm n ires= |
|
4056 (\<exists> lnl lnr rn (ml::nat list) m. |
|
4057 m = Suc (abc_lm_v lm n) \<and> |
|
4058 l = Bk\<up>lnr @ Oc\<up>m @ Bk\<up>lnl @ Bk # Bk # ires \<and> |
|
4059 (r = Bk # <ml> @ Bk\<up>rn \<or> |
|
4060 r = Bk # Bk # <ml> @ Bk\<up>rn))" |
|
4061 |
|
4062 fun mopup_aft_erase_c :: "mopup_type" |
|
4063 where |
|
4064 "mopup_aft_erase_c (s, l, r) lm n ires = |
|
4065 (\<exists> lnl lnr rn (ml::nat list) m. |
|
4066 m = Suc (abc_lm_v lm n) \<and> |
|
4067 l = Bk\<up>lnr @ Oc\<up>m @ Bk\<up>lnl @ Bk # Bk # ires \<and> |
|
4068 (r = <ml> @ Bk\<up>rn \<or> r = Bk # <ml> @ Bk\<up>rn))" |
|
4069 |
|
4070 fun mopup_left_moving :: "mopup_type" |
|
4071 where |
|
4072 "mopup_left_moving (s, l, r) lm n ires = |
|
4073 (\<exists> lnl lnr rn m. |
|
4074 m = Suc (abc_lm_v lm n) \<and> |
|
4075 ((l = Bk\<up>lnr @ Oc\<up>m @ Bk\<up>lnl @ Bk # Bk # ires \<and> r = Bk\<up>rn) \<or> |
|
4076 (l = Oc\<up>(m - 1) @ Bk\<up>lnl @ Bk # Bk # ires \<and> r = Oc # Bk\<up>rn)))" |
|
4077 |
|
4078 fun mopup_jump_over2 :: "mopup_type" |
|
4079 where |
|
4080 "mopup_jump_over2 (s, l, r) lm n ires = |
|
4081 (\<exists> ln rn m1 m2. |
|
4082 m1 + m2 = Suc (abc_lm_v lm n) |
|
4083 \<and> r \<noteq> [] |
|
4084 \<and> (hd r = Oc \<longrightarrow> (l = Oc\<up>m1 @ Bk\<up>ln @ Bk # Bk # ires \<and> r = Oc\<up>m2 @ Bk\<up>rn)) |
|
4085 \<and> (hd r = Bk \<longrightarrow> (l = Bk\<up>ln @ Bk # ires \<and> r = Bk # Oc\<up>(m1+m2)@ Bk\<up>rn)))" |
|
4086 |
|
4087 |
|
4088 fun mopup_inv :: "mopup_type" |
|
4089 where |
|
4090 "mopup_inv (s, l, r) lm n ires = |
|
4091 (if s = 0 then mopup_stop (s, l, r) lm n ires |
|
4092 else if s \<le> 2*n then |
|
4093 if s mod 2 = 1 then mopup_bef_erase_a (s, l, r) lm n ires |
|
4094 else mopup_bef_erase_b (s, l, r) lm n ires |
|
4095 else if s = 2*n + 1 then |
|
4096 mopup_jump_over1 (s, l, r) lm n ires |
|
4097 else if s = 2*n + 2 then mopup_aft_erase_a (s, l, r) lm n ires |
|
4098 else if s = 2*n + 3 then mopup_aft_erase_b (s, l, r) lm n ires |
|
4099 else if s = 2*n + 4 then mopup_aft_erase_c (s, l, r) lm n ires |
|
4100 else if s = 2*n + 5 then mopup_left_moving (s, l, r) lm n ires |
|
4101 else if s = 2*n + 6 then mopup_jump_over2 (s, l, r) lm n ires |
|
4102 else False)" |
|
4103 |
|
4104 lemma mopup_fetch_0[simp]: |
|
4105 "(fetch (mopup_a n @ shift mopup_b (2 * n)) 0 b) = (Nop, 0)" |
|
4106 by(simp add: fetch.simps) |
|
4107 |
|
4108 lemma mop_bef_length[simp]: "length (mopup_a n) = 4 * n" |
|
4109 apply(induct n, simp_all add: mopup_a.simps) |
|
4110 done |
|
4111 |
|
4112 lemma mopup_a_nth: |
|
4113 "\<lbrakk>q < n; x < 4\<rbrakk> \<Longrightarrow> mopup_a n ! (4 * q + x) = |
|
4114 mopup_a (Suc q) ! ((4 * q) + x)" |
|
4115 apply(induct n, simp) |
|
4116 apply(case_tac "q < n", simp add: mopup_a.simps, auto) |
|
4117 apply(simp add: nth_append) |
|
4118 apply(subgoal_tac "q = n", simp) |
|
4119 apply(arith) |
|
4120 done |
|
4121 |
|
4122 lemma fetch_bef_erase_a_o[simp]: |
|
4123 "\<lbrakk>0 < s; s \<le> 2 * n; s mod 2 = Suc 0\<rbrakk> |
|
4124 \<Longrightarrow> (fetch (mopup_a n @ shift mopup_b (2 * n)) s Oc) = (W0, s + 1)" |
|
4125 apply(subgoal_tac "\<exists> q. s = 2*q + 1", auto) |
|
4126 apply(subgoal_tac "length (mopup_a n) = 4*n") |
|
4127 apply(auto simp: fetch.simps nth_of.simps nth_append) |
|
4128 apply(subgoal_tac "mopup_a n ! (4 * q + 1) = |
|
4129 mopup_a (Suc q) ! ((4 * q) + 1)", |
|
4130 simp add: mopup_a.simps nth_append) |
|
4131 apply(rule mopup_a_nth, auto) |
|
4132 apply arith |
|
4133 done |
|
4134 |
|
4135 lemma fetch_bef_erase_a_b[simp]: |
|
4136 "\<lbrakk>0 < s; s \<le> 2 * n; s mod 2 = Suc 0\<rbrakk> |
|
4137 \<Longrightarrow> (fetch (mopup_a n @ shift mopup_b (2 * n)) s Bk) = (R, s + 2)" |
|
4138 apply(subgoal_tac "\<exists> q. s = 2*q + 1", auto) |
|
4139 apply(subgoal_tac "length (mopup_a n) = 4*n") |
|
4140 apply(auto simp: fetch.simps nth_of.simps nth_append) |
|
4141 apply(subgoal_tac "mopup_a n ! (4 * q + 0) = |
|
4142 mopup_a (Suc q) ! ((4 * q + 0))", |
|
4143 simp add: mopup_a.simps nth_append) |
|
4144 apply(rule mopup_a_nth, auto) |
|
4145 apply arith |
|
4146 done |
|
4147 |
|
4148 lemma fetch_bef_erase_b_b: |
|
4149 "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = 0\<rbrakk> \<Longrightarrow> |
|
4150 (fetch (mopup_a n @ shift mopup_b (2 * n)) s Bk) = (R, s - 1)" |
|
4151 apply(subgoal_tac "\<exists> q. s = 2 * q", auto) |
|
4152 apply(case_tac qa, simp, simp) |
|
4153 apply(auto simp: fetch.simps nth_of.simps nth_append) |
|
4154 apply(subgoal_tac "mopup_a n ! (4 * nat + 2) = |
|
4155 mopup_a (Suc nat) ! ((4 * nat) + 2)", |
|
4156 simp add: mopup_a.simps nth_append) |
|
4157 apply(rule mopup_a_nth, auto) |
|
4158 done |
|
4159 |
|
4160 lemma fetch_jump_over1_o: |
|
4161 "fetch (mopup_a n @ shift mopup_b (2 * n)) (Suc (2 * n)) Oc |
|
4162 = (R, Suc (2 * n))" |
|
4163 apply(subgoal_tac "length (mopup_a n) = 4 * n") |
|
4164 apply(auto simp: fetch.simps nth_of.simps mopup_b_def nth_append |
|
4165 shift.simps) |
|
4166 done |
|
4167 |
|
4168 lemma fetch_jump_over1_b: |
|
4169 "fetch (mopup_a n @ shift mopup_b (2 * n)) (Suc (2 * n)) Bk |
|
4170 = (R, Suc (Suc (2 * n)))" |
|
4171 apply(subgoal_tac "length (mopup_a n) = 4 * n") |
|
4172 apply(auto simp: fetch.simps nth_of.simps mopup_b_def |
|
4173 nth_append shift.simps) |
|
4174 done |
|
4175 |
|
4176 lemma fetch_aft_erase_a_o: |
|
4177 "fetch (mopup_a n @ shift mopup_b (2 * n)) (Suc (Suc (2 * n))) Oc |
|
4178 = (W0, Suc (2 * n + 2))" |
|
4179 apply(subgoal_tac "length (mopup_a n) = 4 * n") |
|
4180 apply(auto simp: fetch.simps nth_of.simps mopup_b_def |
|
4181 nth_append shift.simps) |
|
4182 done |
|
4183 |
|
4184 lemma fetch_aft_erase_a_b: |
|
4185 "fetch (mopup_a n @ shift mopup_b (2 * n)) (Suc (Suc (2 * n))) Bk |
|
4186 = (L, Suc (2 * n + 4))" |
|
4187 apply(subgoal_tac "length (mopup_a n) = 4 * n") |
|
4188 apply(auto simp: fetch.simps nth_of.simps mopup_b_def |
|
4189 nth_append shift.simps) |
|
4190 done |
|
4191 |
|
4192 lemma fetch_aft_erase_b_b: |
|
4193 "fetch (mopup_a n @ shift mopup_b (2 * n)) (2*n + 3) Bk |
|
4194 = (R, Suc (2 * n + 3))" |
|
4195 apply(subgoal_tac "length (mopup_a n) = 4 * n") |
|
4196 apply(subgoal_tac "2*n + 3 = Suc (2*n + 2)", simp only: fetch.simps) |
|
4197 apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps) |
|
4198 done |
|
4199 |
|
4200 lemma fetch_aft_erase_c_o: |
|
4201 "fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 4) Oc |
|
4202 = (W0, Suc (2 * n + 2))" |
|
4203 apply(subgoal_tac "length (mopup_a n) = 4 * n") |
|
4204 apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps) |
|
4205 apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps) |
|
4206 done |
|
4207 |
|
4208 lemma fetch_aft_erase_c_b: |
|
4209 "fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 4) Bk |
|
4210 = (R, Suc (2 * n + 1))" |
|
4211 apply(subgoal_tac "length (mopup_a n) = 4 * n") |
|
4212 apply(subgoal_tac "2*n + 4 = Suc (2*n + 3)", simp only: fetch.simps) |
|
4213 apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps) |
|
4214 done |
|
4215 |
|
4216 lemma fetch_left_moving_o: |
|
4217 "(fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 5) Oc) |
|
4218 = (L, 2*n + 6)" |
|
4219 apply(subgoal_tac "length (mopup_a n) = 4 * n") |
|
4220 apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps) |
|
4221 apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps) |
|
4222 done |
|
4223 |
|
4224 lemma fetch_left_moving_b: |
|
4225 "(fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 5) Bk) |
|
4226 = (L, 2*n + 5)" |
|
4227 apply(subgoal_tac "length (mopup_a n) = 4 * n") |
|
4228 apply(subgoal_tac "2*n + 5 = Suc (2*n + 4)", simp only: fetch.simps) |
|
4229 apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps) |
|
4230 done |
|
4231 |
|
4232 lemma fetch_jump_over2_b: |
|
4233 "(fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 6) Bk) |
|
4234 = (R, 0)" |
|
4235 apply(subgoal_tac "length (mopup_a n) = 4 * n") |
|
4236 apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps) |
|
4237 apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps) |
|
4238 done |
|
4239 |
|
4240 lemma fetch_jump_over2_o: |
|
4241 "(fetch (mopup_a n @ shift mopup_b (2 * n)) (2 * n + 6) Oc) |
|
4242 = (L, 2*n + 6)" |
|
4243 apply(subgoal_tac "length (mopup_a n) = 4 * n") |
|
4244 apply(subgoal_tac "2*n + 6 = Suc (2*n + 5)", simp only: fetch.simps) |
|
4245 apply(auto simp: nth_of.simps mopup_b_def nth_append shift.simps) |
|
4246 done |
|
4247 |
|
4248 lemmas mopupfetchs = |
|
4249 fetch_bef_erase_a_o fetch_bef_erase_a_b fetch_bef_erase_b_b |
|
4250 fetch_jump_over1_o fetch_jump_over1_b fetch_aft_erase_a_o |
|
4251 fetch_aft_erase_a_b fetch_aft_erase_b_b fetch_aft_erase_c_o |
|
4252 fetch_aft_erase_c_b fetch_left_moving_o fetch_left_moving_b |
|
4253 fetch_jump_over2_b fetch_jump_over2_o |
|
4254 |
|
4255 declare |
|
4256 mopup_jump_over2.simps[simp del] mopup_left_moving.simps[simp del] |
|
4257 mopup_aft_erase_c.simps[simp del] mopup_aft_erase_b.simps[simp del] |
|
4258 mopup_aft_erase_a.simps[simp del] mopup_jump_over1.simps[simp del] |
|
4259 mopup_bef_erase_a.simps[simp del] mopup_bef_erase_b.simps[simp del] |
|
4260 mopup_stop.simps[simp del] |
|
4261 |
|
4262 lemma [simp]: |
|
4263 "\<lbrakk>mopup_bef_erase_a (s, l, Oc # xs) lm n ires\<rbrakk> \<Longrightarrow> |
|
4264 mopup_bef_erase_b (Suc s, l, Bk # xs) lm n ires" |
|
4265 apply(auto simp: mopup_bef_erase_a.simps mopup_bef_erase_b.simps ) |
|
4266 apply(rule_tac x = "m - 1" in exI, rule_tac x = rn in exI) |
|
4267 apply(case_tac m, simp, simp) |
|
4268 done |
|
4269 |
|
4270 lemma mopup_false1: |
|
4271 "\<lbrakk>0 < s; s \<le> 2 * n; s mod 2 = Suc 0; \<not> Suc s \<le> 2 * n\<rbrakk> |
|
4272 \<Longrightarrow> RR" |
|
4273 apply(arith) |
|
4274 done |
|
4275 |
|
4276 lemma [simp]: |
|
4277 "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = Suc 0; |
|
4278 mopup_bef_erase_a (s, l, Oc # xs) lm n ires; r = Oc # xs\<rbrakk> |
|
4279 \<Longrightarrow> (Suc s \<le> 2 * n \<longrightarrow> mopup_bef_erase_b (Suc s, l, Bk # xs) lm n ires) \<and> |
|
4280 (\<not> Suc s \<le> 2 * n \<longrightarrow> mopup_jump_over1 (Suc s, l, Bk # xs) lm n ires) " |
|
4281 apply(auto elim: mopup_false1) |
|
4282 done |
|
4283 |
|
4284 lemma drop_tape_of_cons: |
|
4285 "\<lbrakk>Suc q < length lm; x = lm ! q\<rbrakk> \<Longrightarrow> <drop q lm> = Oc # Oc \<up> x @ Bk # <drop (Suc q) lm>" |
|
4286 by (metis Suc_lessD append_Cons list.simps(2) nth_drop' replicate_Suc tape_of_nl_cons) |
|
4287 |
|
4288 lemma erase2jumpover1: |
|
4289 "\<lbrakk>q < length list; |
|
4290 \<forall>rn. <drop q list> \<noteq> Oc # Oc \<up> abc_lm_v (a # list) (Suc q) @ Bk # <drop (Suc q) list> @ Bk \<up> rn\<rbrakk> |
|
4291 \<Longrightarrow> <drop q list> = Oc # Oc \<up> abc_lm_v (a # list) (Suc q)" |
|
4292 apply(erule_tac x = 0 in allE, simp) |
|
4293 apply(case_tac "Suc q < length list") |
|
4294 apply(erule_tac notE) |
|
4295 apply(rule_tac drop_tape_of_cons, simp_all add: abc_lm_v.simps) |
|
4296 apply(subgoal_tac "length list = Suc q", auto) |
|
4297 apply(subgoal_tac "drop q list = [list ! q]") |
|
4298 apply(simp add: tape_of_nl_abv) |
|
4299 by (metis append_Nil2 append_eq_conv_conj lessI nth_drop') |
|
4300 |
|
4301 lemma erase2jumpover2: |
|
4302 "\<lbrakk>q < length list; \<forall>rn. <drop q list> @ Bk # Bk \<up> n \<noteq> |
|
4303 Oc # Oc \<up> abc_lm_v (a # list) (Suc q) @ Bk # <drop (Suc q) list> @ Bk \<up> rn\<rbrakk> |
|
4304 \<Longrightarrow> RR" |
|
4305 apply(case_tac "Suc q < length list") |
|
4306 apply(erule_tac x = "Suc n" in allE, simp) |
|
4307 apply(erule_tac notE) |
|
4308 apply(rule_tac drop_tape_of_cons, simp_all add: abc_lm_v.simps) |
|
4309 apply(subgoal_tac "length list = Suc q", auto) |
|
4310 apply(erule_tac x = "n" in allE, simp) |
|
4311 by (metis append_Nil2 append_eq_conv_conj lessI nth_drop' |
|
4312 replicate_Suc tape_of_nat_list.simps(2) tape_of_nl_abv) |
|
4313 |
|
4314 lemma mopup_bef_erase_a_2_jump_over[simp]: |
|
4315 "\<lbrakk>n < length lm; 0 < s; s mod 2 = Suc 0; s \<le> 2 * n; |
|
4316 mopup_bef_erase_a (s, l, Bk # xs) lm n ires; \<not> (Suc (Suc s) \<le> 2 * n)\<rbrakk> |
|
4317 \<Longrightarrow> mopup_jump_over1 (s', Bk # l, xs) lm n ires" |
|
4318 apply(auto simp: mopup_bef_erase_a.simps mopup_jump_over1.simps) |
|
4319 apply(case_tac m, auto simp: mod_ex1) |
|
4320 apply(subgoal_tac "n = Suc q", auto) |
|
4321 apply(rule_tac x = "Suc ln" in exI, rule_tac x = 0 in exI, auto) |
|
4322 apply(case_tac [!] lm, simp_all) |
|
4323 apply(case_tac [!] rn, auto elim: erase2jumpover1 erase2jumpover2) |
|
4324 apply(erule_tac x = 0 in allE, simp) |
|
4325 apply(rule_tac classical, simp) |
|
4326 apply(erule_tac notE) |
|
4327 apply(rule_tac drop_tape_of_cons, simp_all add: abc_lm_v.simps) |
|
4328 done |
|
4329 |
|
4330 lemma Suc_Suc_div: "\<lbrakk>0 < s; s mod 2 = Suc 0; Suc (Suc s) \<le> 2 * n\<rbrakk> |
|
4331 \<Longrightarrow> (Suc (Suc (s div 2))) \<le> n" |
|
4332 apply(arith) |
|
4333 done |
|
4334 |
|
4335 lemma mopup_bef_erase_a_2_a[simp]: |
|
4336 "\<lbrakk>n < length lm; 0 < s; s mod 2 = Suc 0; |
|
4337 mopup_bef_erase_a (s, l, Bk # xs) lm n ires; |
|
4338 Suc (Suc s) \<le> 2 * n\<rbrakk> \<Longrightarrow> |
|
4339 mopup_bef_erase_a (Suc (Suc s), Bk # l, xs) lm n ires" |
|
4340 apply(auto simp: mopup_bef_erase_a.simps) |
|
4341 apply(subgoal_tac "drop (Suc (Suc (s div 2))) lm \<noteq> []") |
|
4342 apply(case_tac m, simp_all) |
|
4343 apply(rule_tac x = "Suc (abc_lm_v lm (Suc (s div 2)))" in exI, |
|
4344 rule_tac x = rn in exI, auto simp: mod_ex1) |
|
4345 apply(rule_tac drop_tape_of_cons) |
|
4346 apply arith |
|
4347 apply(simp add: abc_lm_v.simps) |
|
4348 done |
|
4349 |
|
4350 lemma mopup_false2: |
|
4351 "\<lbrakk>0 < s; s \<le> 2 * n; |
|
4352 s mod 2 = Suc 0; Suc s \<noteq> 2 * n; |
|
4353 \<not> Suc (Suc s) \<le> 2 * n\<rbrakk> \<Longrightarrow> RR" |
|
4354 apply(arith) |
|
4355 done |
|
4356 |
|
4357 lemma [simp]: "mopup_bef_erase_a (s, l, []) lm n ires \<Longrightarrow> |
|
4358 mopup_bef_erase_a (s, l, [Bk]) lm n ires" |
|
4359 apply(auto simp: mopup_bef_erase_a.simps) |
|
4360 done |
|
4361 |
|
4362 lemma [simp]: |
|
4363 "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; s mod 2 = Suc 0; \<not> Suc (Suc s) \<le> 2 *n; |
|
4364 mopup_bef_erase_a (s, l, []) lm n ires\<rbrakk> |
|
4365 \<Longrightarrow> mopup_jump_over1 (s', Bk # l, []) lm n ires" |
|
4366 by auto |
|
4367 |
|
4368 lemma "mopup_bef_erase_b (s, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []" |
|
4369 apply(auto simp: mopup_bef_erase_b.simps) |
|
4370 done |
|
4371 |
|
4372 lemma [simp]: "mopup_bef_erase_b (s, l, Oc # xs) lm n ires = False" |
|
4373 apply(auto simp: mopup_bef_erase_b.simps ) |
|
4374 done |
|
4375 |
|
4376 lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow> |
|
4377 (s - Suc 0) mod 2 = Suc 0" |
|
4378 apply(arith) |
|
4379 done |
|
4380 |
|
4381 lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow> |
|
4382 s - Suc 0 \<le> 2 * n" |
|
4383 apply(simp) |
|
4384 done |
|
4385 |
|
4386 lemma [simp]: "\<lbrakk>0 < s; s \<le> 2 *n; s mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow> \<not> s \<le> Suc 0" |
|
4387 apply(arith) |
|
4388 done |
|
4389 |
|
4390 lemma [simp]: "\<lbrakk>n < length lm; 0 < s; s \<le> 2 * n; |
|
4391 s mod 2 \<noteq> Suc 0; |
|
4392 mopup_bef_erase_b (s, l, Bk # xs) lm n ires; r = Bk # xs\<rbrakk> |
|
4393 \<Longrightarrow> mopup_bef_erase_a (s - Suc 0, Bk # l, xs) lm n ires" |
|
4394 apply(auto simp: mopup_bef_erase_b.simps mopup_bef_erase_a.simps) |
|
4395 done |
|
4396 |
|
4397 lemma [simp]: "\<lbrakk>mopup_bef_erase_b (s, l, []) lm n ires\<rbrakk> \<Longrightarrow> |
|
4398 mopup_bef_erase_a (s - Suc 0, Bk # l, []) lm n ires" |
|
4399 apply(auto simp: mopup_bef_erase_b.simps mopup_bef_erase_a.simps) |
|
4400 done |
|
4401 |
|
4402 lemma [simp]: |
|
4403 "\<lbrakk>n < length lm; |
|
4404 mopup_jump_over1 (Suc (2 * n), l, Oc # xs) lm n ires; |
|
4405 r = Oc # xs\<rbrakk> |
|
4406 \<Longrightarrow> mopup_jump_over1 (Suc (2 * n), Oc # l, xs) lm n ires" |
|
4407 apply(auto simp: mopup_jump_over1.simps) |
|
4408 apply(rule_tac x = ln in exI, rule_tac x = "Suc m1" in exI, |
|
4409 rule_tac x = "m2 - 1" in exI, simp) |
|
4410 apply(case_tac "m2", simp, simp) |
|
4411 apply(rule_tac x = ln in exI, rule_tac x = "Suc m1" in exI, |
|
4412 rule_tac x = "m2 - 1" in exI) |
|
4413 apply(case_tac m2, simp, simp) |
|
4414 done |
|
4415 |
|
4416 lemma mopup_jump_over1_2_aft_erase_a[simp]: |
|
4417 "\<lbrakk>n < length lm; mopup_jump_over1 (Suc (2 * n), l, Bk # xs) lm n ires\<rbrakk> |
|
4418 \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, xs) lm n ires" |
|
4419 apply(simp only: mopup_jump_over1.simps mopup_aft_erase_a.simps) |
|
4420 apply(erule_tac exE)+ |
|
4421 apply(rule_tac x = ln in exI, rule_tac x = "Suc 0" in exI) |
|
4422 apply(case_tac m2, simp) |
|
4423 apply(rule_tac x = rn in exI, rule_tac x = "drop (Suc n) lm" in exI, |
|
4424 simp) |
|
4425 apply(simp) |
|
4426 done |
|
4427 |
|
4428 lemma [simp]: |
|
4429 "\<lbrakk>n < length lm; mopup_jump_over1 (Suc (2 * n), l, []) lm n ires\<rbrakk> \<Longrightarrow> |
|
4430 mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, []) lm n ires" |
|
4431 apply(rule mopup_jump_over1_2_aft_erase_a, simp) |
|
4432 apply(auto simp: mopup_jump_over1.simps) |
|
4433 apply(rule_tac x = ln in exI, rule_tac x = "Suc (abc_lm_v lm n)" in exI, |
|
4434 rule_tac x = 0 in exI, simp add: ) |
|
4435 done |
|
4436 |
|
4437 |
|
4438 lemma [simp]: |
|
4439 "\<lbrakk>n < length lm; |
|
4440 mopup_aft_erase_a (Suc (Suc (2 * n)), l, Oc # xs) lm n ires\<rbrakk> |
|
4441 \<Longrightarrow> mopup_aft_erase_b (Suc (Suc (Suc (2 * n))), l, Bk # xs) lm n ires" |
|
4442 apply(auto simp: mopup_aft_erase_a.simps mopup_aft_erase_b.simps ) |
|
4443 apply(case_tac ml) |
|
4444 apply(simp_all add: tape_of_nl_cons split: if_splits) |
|
4445 apply(case_tac a, simp_all) |
|
4446 apply(rule_tac x = rn in exI, rule_tac x = "[]" in exI, simp) |
|
4447 apply(rule_tac x = rn in exI, rule_tac x = "[nat]" in exI, simp) |
|
4448 apply(case_tac a, simp_all) |
|
4449 apply(rule_tac x = rn in exI, rule_tac x = "list" in exI, simp) |
|
4450 apply(rule_tac x = rn in exI) |
|
4451 apply(rule_tac x = "nat # list" in exI, simp add: tape_of_nl_cons) |
|
4452 done |
|
4453 |
|
4454 lemma [simp]: |
|
4455 "mopup_aft_erase_a (Suc (Suc (2 * n)), l, Bk # xs) lm n ires \<Longrightarrow> l \<noteq> []" |
|
4456 apply(auto simp: mopup_aft_erase_a.simps) |
|
4457 done |
|
4458 |
|
4459 lemma [simp]: |
|
4460 "\<lbrakk>n < length lm; |
|
4461 mopup_aft_erase_a (Suc (Suc (2 * n)), l, Bk # xs) lm n ires\<rbrakk> |
|
4462 \<Longrightarrow> mopup_left_moving (5 + 2 * n, tl l, hd l # Bk # xs) lm n ires" |
|
4463 apply(simp only: mopup_aft_erase_a.simps mopup_left_moving.simps) |
|
4464 apply(erule exE)+ |
|
4465 apply(case_tac lnr, simp) |
|
4466 apply(case_tac ml, simp, simp add: tape_of_nl_cons split: if_splits) |
|
4467 apply(auto) |
|
4468 apply(case_tac ml, simp_all add: tape_of_nl_cons split: if_splits) |
|
4469 apply(rule_tac x = "Suc rn" in exI, simp) |
|
4470 done |
|
4471 |
|
4472 lemma [simp]: |
|
4473 "mopup_aft_erase_a (Suc (Suc (2 * n)), l, []) lm n ires \<Longrightarrow> l \<noteq> []" |
|
4474 apply(simp only: mopup_aft_erase_a.simps) |
|
4475 apply(erule exE)+ |
|
4476 apply(auto) |
|
4477 done |
|
4478 |
|
4479 lemma [simp]: |
|
4480 "\<lbrakk>n < length lm; mopup_aft_erase_a (Suc (Suc (2 * n)), l, []) lm n ires\<rbrakk> |
|
4481 \<Longrightarrow> mopup_left_moving (5 + 2 * n, tl l, [hd l]) lm n ires" |
|
4482 apply(simp only: mopup_aft_erase_a.simps mopup_left_moving.simps) |
|
4483 apply(erule exE)+ |
|
4484 apply(subgoal_tac "ml = [] \<and> rn = 0", erule conjE, erule conjE, simp) |
|
4485 apply(case_tac lnr, simp) |
|
4486 apply(rule_tac x = lnl in exI, simp) |
|
4487 apply(rule_tac x = 1 in exI, simp) |
|
4488 apply(case_tac ml, simp, simp) |
|
4489 done |
|
4490 |
|
4491 |
|
4492 lemma [simp]: "mopup_aft_erase_b (2 * n + 3, l, Oc # xs) lm n ires = False" |
|
4493 apply(auto simp: mopup_aft_erase_b.simps ) |
|
4494 done |
|
4495 |
|
4496 lemma tape_of_ex1[intro]: |
|
4497 "\<exists>rna ml. Oc \<up> a @ Bk \<up> rn = <ml::nat list> @ Bk \<up> rna \<or> Oc \<up> a @ Bk \<up> rn = Bk # <ml> @ Bk \<up> rna" |
|
4498 apply(case_tac a, simp_all) |
|
4499 apply(rule_tac x = rn in exI, rule_tac x = "[]" in exI, simp) |
|
4500 apply(rule_tac x = rn in exI, rule_tac x = "[nat]" in exI, simp) |
|
4501 done |
|
4502 |
|
4503 lemma [intro]: "\<exists>rna ml. Oc \<up> a @ Bk # <list::nat list> @ Bk \<up> rn = |
|
4504 <ml> @ Bk \<up> rna \<or> Oc \<up> a @ Bk # <list> @ Bk \<up> rn = Bk # <ml::nat list> @ Bk \<up> rna" |
|
4505 apply(case_tac "list = []", simp add: replicate_Suc[THEN sym] del: replicate_Suc) |
|
4506 apply(rule_tac rn = "Suc rn" in tape_of_ex1) |
|
4507 apply(case_tac a, simp) |
|
4508 apply(rule_tac x = rn in exI, rule_tac x = list in exI, simp) |
|
4509 apply(rule_tac x = rn in exI, rule_tac x = "nat # list" in exI) |
|
4510 apply(simp add: tape_of_nl_cons) |
|
4511 done |
|
4512 |
|
4513 lemma [simp]: |
|
4514 "\<lbrakk>n < length lm; |
|
4515 mopup_aft_erase_c (2 * n + 4, l, Oc # xs) lm n ires\<rbrakk> |
|
4516 \<Longrightarrow> mopup_aft_erase_b (Suc (Suc (Suc (2 * n))), l, Bk # xs) lm n ires" |
|
4517 apply(auto simp: mopup_aft_erase_c.simps mopup_aft_erase_b.simps ) |
|
4518 apply(case_tac ml, simp_all add: tape_of_nl_cons split: if_splits, auto) |
|
4519 done |
|
4520 |
|
4521 lemma mopup_aft_erase_c_aft_erase_a[simp]: |
|
4522 "\<lbrakk>n < length lm; mopup_aft_erase_c (2 * n + 4, l, Bk # xs) lm n ires\<rbrakk> |
|
4523 \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, xs) lm n ires" |
|
4524 apply(simp only: mopup_aft_erase_c.simps mopup_aft_erase_a.simps ) |
|
4525 apply(erule_tac exE)+ |
|
4526 apply(erule conjE, erule conjE, erule disjE) |
|
4527 apply(subgoal_tac "ml = []", simp, case_tac rn, |
|
4528 simp, simp, rule conjI) |
|
4529 apply(rule_tac x = lnl in exI, rule_tac x = "Suc lnr" in exI, simp) |
|
4530 apply(rule_tac x = nat in exI, rule_tac x = "[]" in exI, simp) |
|
4531 apply(case_tac ml, simp, simp add: tape_of_nl_cons split: if_splits) |
|
4532 apply(rule_tac x = lnl in exI, rule_tac x = "Suc lnr" in exI, simp) |
|
4533 apply(rule_tac x = rn in exI, rule_tac x = "ml" in exI, simp) |
|
4534 done |
|
4535 |
|
4536 lemma [simp]: |
|
4537 "\<lbrakk>n < length lm; mopup_aft_erase_c (2 * n + 4, l, []) lm n ires\<rbrakk> |
|
4538 \<Longrightarrow> mopup_aft_erase_a (Suc (Suc (2 * n)), Bk # l, []) lm n ires" |
|
4539 apply(rule mopup_aft_erase_c_aft_erase_a, simp) |
|
4540 apply(simp only: mopup_aft_erase_c.simps) |
|
4541 apply(erule exE)+ |
|
4542 apply(rule_tac x = lnl in exI, rule_tac x = lnr in exI, simp add: ) |
|
4543 apply(rule_tac x = 0 in exI, rule_tac x = "[]" in exI, simp) |
|
4544 done |
|
4545 |
|
4546 lemma mopup_aft_erase_b_2_aft_erase_c[simp]: |
|
4547 "\<lbrakk>n < length lm; mopup_aft_erase_b (2 * n + 3, l, Bk # xs) lm n ires\<rbrakk> |
|
4548 \<Longrightarrow> mopup_aft_erase_c (4 + 2 * n, Bk # l, xs) lm n ires" |
|
4549 apply(auto simp: mopup_aft_erase_b.simps mopup_aft_erase_c.simps) |
|
4550 apply(rule_tac x = "lnl" in exI, rule_tac x = "Suc lnr" in exI, simp) |
|
4551 apply(rule_tac x = "lnl" in exI, rule_tac x = "Suc lnr" in exI, simp) |
|
4552 done |
|
4553 |
|
4554 lemma [simp]: |
|
4555 "\<lbrakk>n < length lm; mopup_aft_erase_b (2 * n + 3, l, []) lm n ires\<rbrakk> |
|
4556 \<Longrightarrow> mopup_aft_erase_c (4 + 2 * n, Bk # l, []) lm n ires" |
|
4557 apply(rule_tac mopup_aft_erase_b_2_aft_erase_c, simp) |
|
4558 apply(simp add: mopup_aft_erase_b.simps) |
|
4559 done |
|
4560 |
|
4561 lemma [simp]: |
|
4562 "mopup_left_moving (2 * n + 5, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []" |
|
4563 apply(auto simp: mopup_left_moving.simps) |
|
4564 done |
|
4565 |
|
4566 lemma [simp]: |
|
4567 "\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, Oc # xs) lm n ires\<rbrakk> |
|
4568 \<Longrightarrow> mopup_jump_over2 (2 * n + 6, tl l, hd l # Oc # xs) lm n ires" |
|
4569 apply(simp only: mopup_left_moving.simps mopup_jump_over2.simps) |
|
4570 apply(erule_tac exE)+ |
|
4571 apply(erule conjE, erule disjE, erule conjE) |
|
4572 apply(case_tac rn, simp, simp add: ) |
|
4573 apply(case_tac "hd l", simp add: ) |
|
4574 apply(case_tac "abc_lm_v lm n", simp) |
|
4575 apply(rule_tac x = "lnl" in exI, rule_tac x = rn in exI, |
|
4576 rule_tac x = "Suc 0" in exI, rule_tac x = 0 in exI) |
|
4577 apply(case_tac lnl, simp, simp, simp add: exp_ind[THEN sym], simp) |
|
4578 apply(case_tac "abc_lm_v lm n", simp) |
|
4579 apply(case_tac lnl, simp, simp) |
|
4580 apply(rule_tac x = lnl in exI, rule_tac x = rn in exI) |
|
4581 apply(rule_tac x = nat in exI, rule_tac x = "Suc (Suc 0)" in exI, simp) |
|
4582 done |
|
4583 |
|
4584 lemma [simp]: "mopup_left_moving (2 * n + 5, l, xs) lm n ires \<Longrightarrow> l \<noteq> []" |
|
4585 apply(auto simp: mopup_left_moving.simps) |
|
4586 done |
|
4587 |
|
4588 lemma [simp]: |
|
4589 "\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, Bk # xs) lm n ires\<rbrakk> |
|
4590 \<Longrightarrow> mopup_left_moving (2 * n + 5, tl l, hd l # Bk # xs) lm n ires" |
|
4591 apply(simp only: mopup_left_moving.simps) |
|
4592 apply(erule exE)+ |
|
4593 apply(case_tac lnr, simp) |
|
4594 apply(rule_tac x = lnl in exI, rule_tac x = nat in exI, simp) |
|
4595 apply(rule_tac x = "Suc rn" in exI, simp) |
|
4596 done |
|
4597 |
|
4598 lemma [simp]: |
|
4599 "\<lbrakk>n < length lm; mopup_left_moving (2 * n + 5, l, []) lm n ires\<rbrakk> |
|
4600 \<Longrightarrow> mopup_left_moving (2 * n + 5, tl l, [hd l]) lm n ires" |
|
4601 apply(simp only: mopup_left_moving.simps) |
|
4602 apply(erule exE)+ |
|
4603 apply(case_tac lnr, auto) |
|
4604 done |
|
4605 |
|
4606 |
|
4607 lemma [simp]: |
|
4608 "mopup_jump_over2 (2 * n + 6, l, Oc # xs) lm n ires \<Longrightarrow> l \<noteq> []" |
|
4609 apply(auto simp: mopup_jump_over2.simps ) |
|
4610 done |
|
4611 |
|
4612 lemma [simp]: |
|
4613 "\<lbrakk>n < length lm; mopup_jump_over2 (2 * n + 6, l, Oc # xs) lm n ires\<rbrakk> |
|
4614 \<Longrightarrow> mopup_jump_over2 (2 * n + 6, tl l, hd l # Oc # xs) lm n ires" |
|
4615 apply(simp only: mopup_jump_over2.simps) |
|
4616 apply(erule_tac exE)+ |
|
4617 apply(simp add: , erule conjE, erule_tac conjE) |
|
4618 apply(case_tac m1, simp) |
|
4619 apply(rule_tac x = ln in exI, rule_tac x = rn in exI, |
|
4620 rule_tac x = 0 in exI, simp) |
|
4621 apply(case_tac ln, simp, simp, simp only: exp_ind[THEN sym], simp) |
|
4622 apply(rule_tac x = ln in exI, rule_tac x = rn in exI, |
|
4623 rule_tac x = nat in exI, rule_tac x = "Suc m2" in exI, simp) |
|
4624 done |
|
4625 |
|
4626 lemma [simp]: |
|
4627 "\<lbrakk>n < length lm; mopup_jump_over2 (2 * n + 6, l, Bk # xs) lm n ires\<rbrakk> |
|
4628 \<Longrightarrow> mopup_stop (0, Bk # l, xs) lm n ires" |
|
4629 apply(auto simp: mopup_jump_over2.simps mopup_stop.simps) |
|
4630 apply(simp_all add: tape_of_nat_abv exp_ind[THEN sym]) |
|
4631 done |
|
4632 |
|
4633 lemma [simp]: "mopup_jump_over2 (2 * n + 6, l, []) lm n ires = False" |
|
4634 apply(simp only: mopup_jump_over2.simps, simp) |
|
4635 done |
|
4636 |
|
4637 lemma mopup_inv_step: |
|
4638 "\<lbrakk>n < length lm; mopup_inv (s, l, r) lm n ires\<rbrakk> |
|
4639 \<Longrightarrow> mopup_inv (step (s, l, r) (mopup_a n @ shift mopup_b (2 * n), 0)) lm n ires" |
|
4640 apply(case_tac r, case_tac [2] a) |
|
4641 apply(auto split:if_splits simp add:step.simps) |
|
4642 apply(simp_all add: mopupfetchs) |
|
4643 done |
|
4644 |
|
4645 declare mopup_inv.simps[simp del] |
|
4646 lemma mopup_inv_steps: |
|
4647 "\<lbrakk>n < length lm; mopup_inv (s, l, r) lm n ires\<rbrakk> \<Longrightarrow> |
|
4648 mopup_inv (steps (s, l, r) (mopup_a n @ shift mopup_b (2 * n), 0) stp) lm n ires" |
|
4649 apply(induct_tac stp, simp add: steps.simps) |
|
4650 apply(simp add: step_red) |
|
4651 apply(case_tac "steps (s, l, r) |
|
4652 (mopup_a n @ shift mopup_b (2 * n), 0) na", simp) |
|
4653 apply(rule_tac mopup_inv_step, simp, simp) |
|
4654 done |
|
4655 |
|
4656 fun abc_mopup_stage1 :: "config \<Rightarrow> nat \<Rightarrow> nat" |
|
4657 where |
|
4658 "abc_mopup_stage1 (s, l, r) n = |
|
4659 (if s > 0 \<and> s \<le> 2*n then 6 |
|
4660 else if s = 2*n + 1 then 4 |
|
4661 else if s \<ge> 2*n + 2 \<and> s \<le> 2*n + 4 then 3 |
|
4662 else if s = 2*n + 5 then 2 |
|
4663 else if s = 2*n + 6 then 1 |
|
4664 else 0)" |
|
4665 |
|
4666 fun abc_mopup_stage2 :: "config \<Rightarrow> nat \<Rightarrow> nat" |
|
4667 where |
|
4668 "abc_mopup_stage2 (s, l, r) n = |
|
4669 (if s > 0 \<and> s \<le> 2*n then length r |
|
4670 else if s = 2*n + 1 then length r |
|
4671 else if s = 2*n + 5 then length l |
|
4672 else if s = 2*n + 6 then length l |
|
4673 else if s \<ge> 2*n + 2 \<and> s \<le> 2*n + 4 then length r |
|
4674 else 0)" |
|
4675 |
|
4676 fun abc_mopup_stage3 :: "config \<Rightarrow> nat \<Rightarrow> nat" |
|
4677 where |
|
4678 "abc_mopup_stage3 (s, l, r) n = |
|
4679 (if s > 0 \<and> s \<le> 2*n then |
|
4680 if hd r = Bk then 0 |
|
4681 else 1 |
|
4682 else if s = 2*n + 2 then 1 |
|
4683 else if s = 2*n + 3 then 0 |
|
4684 else if s = 2*n + 4 then 2 |
|
4685 else 0)" |
|
4686 |
|
4687 fun abc_mopup_measure :: "(config \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)" |
|
4688 where |
|
4689 "abc_mopup_measure (c, n) = |
|
4690 (abc_mopup_stage1 c n, abc_mopup_stage2 c n, |
|
4691 abc_mopup_stage3 c n)" |
|
4692 |
|
4693 definition abc_mopup_LE :: |
|
4694 "(((nat \<times> cell list \<times> cell list) \<times> nat) \<times> |
|
4695 ((nat \<times> cell list \<times> cell list) \<times> nat)) set" |
|
4696 where |
|
4697 "abc_mopup_LE \<equiv> (inv_image lex_triple abc_mopup_measure)" |
|
4698 |
|
4699 lemma wf_abc_mopup_le[intro]: "wf abc_mopup_LE" |
|
4700 by(auto intro:wf_inv_image simp:abc_mopup_LE_def lex_triple_def lex_pair_def) |
|
4701 |
|
4702 lemma [simp]: "mopup_bef_erase_a (a, aa, []) lm n ires = False" |
|
4703 apply(auto simp: mopup_bef_erase_a.simps) |
|
4704 done |
|
4705 |
|
4706 lemma [simp]: "mopup_bef_erase_b (a, aa, []) lm n ires = False" |
|
4707 apply(auto simp: mopup_bef_erase_b.simps) |
|
4708 done |
|
4709 |
|
4710 lemma [simp]: "mopup_aft_erase_b (2 * n + 3, aa, []) lm n ires = False" |
|
4711 apply(auto simp: mopup_aft_erase_b.simps) |
|
4712 done |
|
4713 |
|
4714 declare mopup_inv.simps[simp del] |
|
4715 term mopup_inv |
|
4716 |
|
4717 lemma [simp]: |
|
4718 "\<lbrakk>0 < q; q \<le> n\<rbrakk> \<Longrightarrow> |
|
4719 (fetch (mopup_a n @ shift mopup_b (2 * n)) (2*q) Bk) = (R, 2*q - 1)" |
|
4720 apply(case_tac q, simp, simp) |
|
4721 apply(auto simp: fetch.simps nth_of.simps nth_append) |
|
4722 apply(subgoal_tac "mopup_a n ! (4 * nat + 2) = |
|
4723 mopup_a (Suc nat) ! ((4 * nat) + 2)", |
|
4724 simp add: mopup_a.simps nth_append) |
|
4725 apply(rule mopup_a_nth, auto) |
|
4726 done |
|
4727 |
|
4728 lemma mopup_halt: |
|
4729 assumes |
|
4730 less: "n < length lm" |
|
4731 and inv: "mopup_inv (Suc 0, l, r) lm n ires" |
|
4732 and f: "f = (\<lambda> stp. (steps (Suc 0, l, r) (mopup_a n @ shift mopup_b (2 * n), 0) stp, n))" |
|
4733 and P: "P = (\<lambda> (c, n). is_final c)" |
|
4734 shows "\<exists> stp. P (f stp)" |
|
4735 proof(rule_tac LE = abc_mopup_LE in halt_lemma) |
|
4736 show "wf abc_mopup_LE" by(auto) |
|
4737 next |
|
4738 show "\<forall>n. \<not> P (f n) \<longrightarrow> (f (Suc n), f n) \<in> abc_mopup_LE" |
|
4739 proof(rule_tac allI, rule_tac impI) |
|
4740 fix na |
|
4741 assume h: "\<not> P (f na)" |
|
4742 show "(f (Suc na), f na) \<in> abc_mopup_LE" |
|
4743 proof(simp add: f) |
|
4744 obtain a b c where g:"steps (Suc 0, l, r) (mopup_a n @ shift mopup_b (2 * n), 0) na = (a, b, c)" |
|
4745 apply(case_tac "steps (Suc 0, l, r) (mopup_a n @ shift mopup_b (2 * n), 0) na", auto) |
|
4746 done |
|
4747 then have "mopup_inv (a, b, c) lm n ires" |
|
4748 thm mopup_inv_steps |
|
4749 using inv less mopup_inv_steps[of n lm "Suc 0" l r ires na] |
|
4750 apply(simp) |
|
4751 done |
|
4752 moreover have "a > 0" |
|
4753 using h g |
|
4754 apply(simp add: f P) |
|
4755 done |
|
4756 ultimately have "((step (a, b, c) (mopup_a n @ shift mopup_b (2 * n), 0), n), (a, b, c), n) \<in> abc_mopup_LE" |
|
4757 apply(case_tac c, case_tac [2] aa) |
|
4758 apply(auto split:if_splits simp add:step.simps mopup_inv.simps) |
|
4759 apply(simp_all add: mopupfetchs abc_mopup_LE_def lex_triple_def lex_pair_def ) |
|
4760 done |
|
4761 thus "((step (steps (Suc 0, l, r) (mopup_a n @ shift mopup_b (2 * n), 0) na) |
|
4762 (mopup_a n @ shift mopup_b (2 * n), 0), n), |
|
4763 steps (Suc 0, l, r) (mopup_a n @ shift mopup_b (2 * n), 0) na, n) |
|
4764 \<in> abc_mopup_LE" |
|
4765 using g by simp |
|
4766 qed |
|
4767 qed |
|
4768 qed |
|
4769 |
|
4770 lemma mopup_inv_start: |
|
4771 "n < length am \<Longrightarrow> mopup_inv (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) am n ires" |
|
4772 apply(auto simp: mopup_inv.simps mopup_bef_erase_a.simps mopup_jump_over1.simps) |
|
4773 apply(case_tac [!] am, auto split: if_splits simp: tape_of_nl_cons) |
|
4774 apply(rule_tac x = "Suc a" in exI, rule_tac x = k in exI, simp) |
|
4775 apply(case_tac [!] n, simp_all add: abc_lm_v.simps) |
|
4776 apply(case_tac k, simp, simp_all) |
|
4777 done |
|
4778 |
|
4779 lemma mopup_correct: |
|
4780 assumes less: "n < length (am::nat list)" |
|
4781 and rs: "abc_lm_v am n = rs" |
|
4782 shows "\<exists> stp i j. (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp) |
|
4783 = (0, Bk\<up>i @ Bk # Bk # ires, Oc # Oc\<up> rs @ Bk\<up>j)" |
|
4784 using less |
|
4785 proof - |
|
4786 have a: "mopup_inv (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) am n ires" |
|
4787 using less |
|
4788 apply(simp add: mopup_inv_start) |
|
4789 done |
|
4790 then have "\<exists> stp. is_final (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp)" |
|
4791 using less mopup_halt[of n am "Bk # Bk # ires" "<am> @ Bk \<up> k" ires |
|
4792 "(\<lambda>stp. (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp, n))" |
|
4793 "(\<lambda>(c, n). is_final c)"] |
|
4794 apply(simp) |
|
4795 done |
|
4796 from this obtain stp where b: |
|
4797 "is_final (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp)" .. |
|
4798 from a b have |
|
4799 "mopup_inv (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp) |
|
4800 am n ires" |
|
4801 apply(rule_tac mopup_inv_steps, simp_all add: less) |
|
4802 done |
|
4803 from b and this show "?thesis" |
|
4804 apply(rule_tac x = stp in exI, simp) |
|
4805 apply(case_tac "steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) |
|
4806 (mopup_a n @ shift mopup_b (2 * n), 0) stp") |
|
4807 apply(simp add: mopup_inv.simps mopup_stop.simps rs) |
|
4808 using rs |
|
4809 apply(simp add: tape_of_nat_abv) |
|
4810 done |
|
4811 qed |
|
4812 |
689 (*we can use Hoare_plus here*) |
4813 (*we can use Hoare_plus here*) |
|
4814 |
|
4815 lemma wf_mopup[intro]: "tm_wf (mopup n, 0)" |
|
4816 apply(induct n, simp add: mopup.simps shift.simps mopup_b_def tm_wf.simps) |
|
4817 apply(auto simp: mopup.simps shift.simps mopup_b_def tm_wf.simps) |
|
4818 done |
|
4819 |
|
4820 lemma length_tp: |
|
4821 "\<lbrakk>ly = layout_of ap; tp = tm_of ap\<rbrakk> \<Longrightarrow> |
|
4822 start_of ly (length ap) = Suc (length tp div 2)" |
|
4823 apply(frule_tac length_tp', simp_all) |
|
4824 apply(simp add: start_of.simps) |
|
4825 done |
|
4826 |
690 lemma compile_correct_halt: |
4827 lemma compile_correct_halt: |
691 assumes layout: "ly = layout_of ap" |
4828 assumes layout: "ly = layout_of ap" |
|
4829 and compile: "tp = tm_of ap" |
692 and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires" |
4830 and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires" |
693 and abc_halt: "abc_steps_l (0, lm) ap stp = (length ap, am)" |
4831 and abc_halt: "abc_steps_l (0, lm) ap stp = (length ap, am)" |
694 and rs_loc: "n < length am" |
4832 and rs_loc: "n < length am" |
695 and rs: "rs = abc_lm_v am n" |
4833 and rs: "abc_lm_v am n = rs" |
696 shows "\<exists> stp i j. steps (Suc 0, l, r) (tp_norm ap |+| mopup n, 0) stp = (0, Bk\<up>i @ Bk # Bk # ires, Oc\<up>Suc rs @ Bk\<up>j)" |
4834 and off: "off = length tp div 2" |
697 sorry |
4835 shows "\<exists> stp i j. steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stp = (0, Bk\<up>i @ Bk # Bk # ires, Oc\<up>Suc rs @ Bk\<up>j)" |
698 |
4836 proof - |
|
4837 have "\<exists> stp k. steps (Suc 0, l, r) (tp, 0) stp = (Suc off, Bk # Bk # ires, <am> @ Bk\<up>k)" |
|
4838 using assms tp_correct'[of ly ap tp lm l r ires stp am] |
|
4839 by(simp add: length_tp) |
|
4840 then obtain stp k where "steps (Suc 0, l, r) (tp, 0) stp = (Suc off, Bk # Bk # ires, <am> @ Bk\<up>k)" |
|
4841 by blast |
|
4842 then have a: "steps (Suc 0, l, r) (tp@shift (mopup n) off , 0) stp = (Suc off, Bk # Bk # ires, <am> @ Bk\<up>k)" |
|
4843 using assms |
|
4844 by(auto intro: tm_append_first_steps_eq) |
|
4845 have "\<exists> stp i j. (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp) |
|
4846 = (0, Bk\<up>i @ Bk # Bk # ires, Oc # Oc\<up> rs @ Bk\<up>j)" |
|
4847 using assms |
|
4848 by(auto intro: mopup_correct) |
|
4849 then obtain stpb i j where |
|
4850 "steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stpb |
|
4851 = (0, Bk\<up>i @ Bk # Bk # ires, Oc # Oc\<up> rs @ Bk\<up>j)" by blast |
|
4852 then have b: "steps (Suc 0 + off, Bk # Bk # ires, <am> @ Bk \<up> k) (tp @ shift (mopup n) off, 0) stpb |
|
4853 = (0, Bk\<up>i @ Bk # Bk # ires, Oc # Oc\<up> rs @ Bk\<up>j)" |
|
4854 using assms wf_mopup |
|
4855 thm tm_append_second_halt_eq |
|
4856 apply(drule_tac tm_append_second_halt_eq, auto) |
|
4857 done |
|
4858 from a b show "?thesis" |
|
4859 by(rule_tac x = "stp + stpb" in exI, simp add: steps_add) |
|
4860 qed |
|
4861 |
|
4862 declare mopup.simps[simp del] |
|
4863 lemma abc_step_red2: |
|
4864 "abc_steps_l (s, lm) p (Suc n) = (let (as', am') = abc_steps_l (s, lm) p n in |
|
4865 abc_step_l (as', am') (abc_fetch as' p))" |
|
4866 apply(case_tac "abc_steps_l (s, lm) p n", simp) |
|
4867 apply(drule_tac abc_step_red, simp) |
|
4868 done |
|
4869 |
|
4870 lemma crsp_steps2: |
|
4871 assumes |
|
4872 layout: "ly = layout_of ap" |
|
4873 and compile: "tp = tm_of ap" |
|
4874 and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires" |
|
4875 and nothalt: "as < length ap" |
|
4876 and aexec: "abc_steps_l (0, lm) ap stp = (as, am)" |
|
4877 shows "\<exists>stpa\<ge>stp. crsp ly (as, am) (steps (Suc 0, l, r) (tp, 0) stpa) ires" |
|
4878 using nothalt aexec |
|
4879 proof(induct stp arbitrary: as am) |
|
4880 case 0 |
|
4881 thus "?case" |
|
4882 using crsp |
|
4883 by(rule_tac x = 0 in exI, auto simp: abc_steps_l.simps steps.simps crsp) |
|
4884 next |
|
4885 case (Suc stp as am) |
|
4886 have ind: |
|
4887 "\<And> as am. \<lbrakk>as < length ap; abc_steps_l (0, lm) ap stp = (as, am)\<rbrakk> |
|
4888 \<Longrightarrow> \<exists>stpa\<ge>stp. crsp ly (as, am) (steps (Suc 0, l, r) (tp, 0) stpa) ires" by fact |
|
4889 have a: "as < length ap" by fact |
|
4890 have b: "abc_steps_l (0, lm) ap (Suc stp) = (as, am)" by fact |
|
4891 obtain as' am' where c: "abc_steps_l (0, lm) ap stp = (as', am')" |
|
4892 by(case_tac "abc_steps_l (0, lm) ap stp", auto) |
|
4893 then have d: "as' < length ap" |
|
4894 using a b |
|
4895 by(simp add: abc_step_red2, case_tac "as' < length ap", simp, |
|
4896 simp add: abc_fetch.simps abc_steps_l.simps abc_step_l.simps) |
|
4897 have "\<exists>stpa\<ge>stp. crsp ly (as', am') (steps (Suc 0, l, r) (tp, 0) stpa) ires" |
|
4898 using d c ind by simp |
|
4899 from this obtain stpa where e: |
|
4900 "stpa \<ge> stp \<and> crsp ly (as', am') (steps (Suc 0, l, r) (tp, 0) stpa) ires" |
|
4901 by blast |
|
4902 obtain s' l' r' where f: "steps (Suc 0, l, r) (tp, 0) stpa = (s', l', r')" |
|
4903 by(case_tac "steps (Suc 0, l, r) (tp, 0) stpa", auto) |
|
4904 obtain ins where g: "abc_fetch as' ap = Some ins" using d |
|
4905 by(case_tac "abc_fetch as' ap",auto simp: abc_fetch.simps) |
|
4906 then have "\<exists>stp> (0::nat). crsp ly (abc_step_l (as', am') (Some ins)) |
|
4907 (steps (s', l', r') (tp, 0) stp) ires " |
|
4908 using layout compile e f |
|
4909 by(rule_tac crsp_step, simp_all) |
|
4910 then obtain stpb where "stpb > 0 \<and> crsp ly (abc_step_l (as', am') (Some ins)) |
|
4911 (steps (s', l', r') (tp, 0) stpb) ires" .. |
|
4912 from this show "?case" using b e g f c |
|
4913 by(rule_tac x = "stpa + stpb" in exI, simp add: steps_add abc_step_red2) |
|
4914 qed |
|
4915 |
699 lemma compile_correct_unhalt: |
4916 lemma compile_correct_unhalt: |
700 assumes layout: "ly = layout_of ap" |
4917 assumes layout: "ly = layout_of ap" |
|
4918 and compile: "tp = tm_of ap" |
701 and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires" |
4919 and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires" |
|
4920 and off: "off = length tp div 2" |
702 and abc_unhalt: "\<forall> stp. (\<lambda> (as, am). as < length ap) (abc_steps_l (0, lm) ap stp)" |
4921 and abc_unhalt: "\<forall> stp. (\<lambda> (as, am). as < length ap) (abc_steps_l (0, lm) ap stp)" |
703 shows "\<forall> stp.\<not> is_final (steps (Suc 0, l, r) (tp_norm ap |+| mopup n, 0) stp)" |
4922 shows "\<forall> stp.\<not> is_final (steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stp)" |
704 sorry |
4923 using assms |
|
4924 proof(rule_tac allI, rule_tac notI) |
|
4925 fix stp |
|
4926 assume h: "is_final (steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stp)" |
|
4927 obtain as am where a: "abc_steps_l (0, lm) ap stp = (as, am)" |
|
4928 by(case_tac "abc_steps_l (0, lm) ap stp", auto) |
|
4929 then have b: "as < length ap" |
|
4930 using abc_unhalt |
|
4931 by(erule_tac x = stp in allE, simp) |
|
4932 have "\<exists> stpa\<ge>stp. crsp ly (as, am) (steps (Suc 0, l, r) (tp, 0) stpa) ires " |
|
4933 using assms b a |
|
4934 apply(rule_tac crsp_steps2, simp_all) |
|
4935 done |
|
4936 then obtain stpa where |
|
4937 "stpa\<ge>stp \<and> crsp ly (as, am) (steps (Suc 0, l, r) (tp, 0) stpa) ires" .. |
|
4938 then obtain s' l' r' where b: "(steps (Suc 0, l, r) (tp, 0) stpa) = (s', l', r') \<and> |
|
4939 stpa\<ge>stp \<and> crsp ly (as, am) (s', l', r') ires" |
|
4940 by(case_tac "steps (Suc 0, l, r) (tp, 0) stpa", auto) |
|
4941 hence c: |
|
4942 "(steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stpa) = (s', l', r')" |
|
4943 by(rule_tac tm_append_first_steps_eq, simp_all add: crsp.simps) |
|
4944 from b have d: "s' > 0 \<and> stpa \<ge> stp" |
|
4945 by(simp add: crsp.simps) |
|
4946 then obtain diff where e: "stpa = stp + diff" by (metis le_iff_add) |
|
4947 obtain s'' l'' r'' where f: |
|
4948 "steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stp = (s'', l'', r'') \<and> is_final (s'', l'', r'')" |
|
4949 using h |
|
4950 by(case_tac "steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stp", auto) |
|
4951 |
|
4952 then have "is_final (steps (s'', l'', r'') (tp @ shift (mopup n) off, 0) diff)" |
|
4953 by(auto intro: is_final_steps) |
|
4954 then have "is_final (steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stpa)" |
|
4955 using e |
|
4956 by(simp add: steps_add f) |
|
4957 from this and c d show "False" by simp |
|
4958 qed |
705 |
4959 |
706 end |
4960 end |
707 |
4961 |