62 - TM program (nat \<rightharpoonup> tm_inst) |
62 - TM program (nat \<rightharpoonup> tm_inst) |
63 - current state (nat) |
63 - current state (nat) |
64 - position of head (int) |
64 - position of head (int) |
65 - tape (int \<rightharpoonup> Block) |
65 - tape (int \<rightharpoonup> Block) |
66 *) |
66 *) |
67 type_synonym tconf = "nat \<times> (nat \<Rightarrow>f tm_inst option) \<times> nat \<times> int \<times> (int \<Rightarrow>f Block option)" |
67 type_synonym tconf = "nat \<times> (nat f\<rightharpoonup> tm_inst) \<times> nat \<times> int \<times> (int f\<rightharpoonup> Block)" |
68 |
68 |
69 (* updates the position/tape according to an action *) |
69 (* updates the position/tape according to an action *) |
70 fun |
70 fun |
71 next_tape :: "taction \<Rightarrow> (int \<times> (int \<Rightarrow>f Block option)) \<Rightarrow> (int \<times> (int \<Rightarrow>f Block option))" |
71 next_tape :: "taction \<Rightarrow> (int \<times> (int f\<rightharpoonup> Block)) \<Rightarrow> (int \<times> (int f\<rightharpoonup> Block))" |
72 where |
72 where |
73 "next_tape W0 (pos, m) = (pos, m(pos $:= Some Bk))" | |
73 "next_tape W0 (pos, m) = (pos, m(pos f\<mapsto> Bk))" | |
74 "next_tape W1 (pos, m) = (pos, m(pos $:= Some Oc))" | |
74 "next_tape W1 (pos, m) = (pos, m(pos f\<mapsto> Oc))" | |
75 "next_tape L (pos, m) = (pos - 1, m)" | |
75 "next_tape L (pos, m) = (pos - 1, m)" | |
76 "next_tape R (pos, m) = (pos + 1, m)" |
76 "next_tape R (pos, m) = (pos + 1, m)" |
77 |
77 |
78 fun tstep :: "tconf \<Rightarrow> tconf" |
78 fun tstep :: "tconf \<Rightarrow> tconf" |
79 where "tstep (faults, prog, pc, pos, m) = |
79 where "tstep (faults, prog, pc, pos, m) = |
80 (case (prog $ pc) of |
80 (case (prog f! pc) of |
81 Some ((action0, pc0), (action1, pc1)) \<Rightarrow> |
81 Some ((action0, pc0), (action1, pc1)) \<Rightarrow> |
82 case (m $ pos) of |
82 case (m f! pos) of |
83 Some Bk \<Rightarrow> (faults, prog, pc0, next_tape action0 (pos, m)) | |
83 Some Bk \<Rightarrow> (faults, prog, pc0, next_tape action0 (pos, m)) | |
84 Some Oc \<Rightarrow> (faults, prog, pc1, next_tape action1 (pos, m)) | |
84 Some Oc \<Rightarrow> (faults, prog, pc1, next_tape action1 (pos, m)) | |
85 None \<Rightarrow> (faults + 1, prog, pc, pos, m) |
85 None \<Rightarrow> (faults + 1, prog, pc, pos, m) |
86 | None \<Rightarrow> (faults + 1, prog, pc, pos, m))" |
86 | None \<Rightarrow> (faults + 1, prog, pc, pos, m))" |
87 |
87 |
|
88 |
|
89 fun tstep :: "tconf \<Rightarrow> tconf" |
|
90 where "tstep (faults, prog, pc, pos, m) = |
|
91 (case (prog f! pc) of |
|
92 Some ((action0, pc0), (action1, pc1)) \<Rightarrow> |
|
93 case (m f! pos) of |
|
94 Some Bk \<Rightarrow> (faults, prog, pc0, next_tape action0 (pos, m)) | |
|
95 Some Oc \<Rightarrow> (faults, prog, pc1, next_tape action1 (pos, m)) | |
|
96 None \<Rightarrow> (faults + 1, prog, pc, pos, m) |
|
97 | None \<Rightarrow> (faults + 1, prog, pc, pos, m))" |
|
98 |
88 lemma tstep_splits: |
99 lemma tstep_splits: |
89 "(P (tstep s)) = ((\<forall> faults prog pc pos m action0 pc0 action1 pc1. |
100 "(P (tstep s)) = ((\<forall> faults prog pc pos m action0 pc0 action1 pc1. |
90 s = (faults, prog, pc, pos, m) \<longrightarrow> |
101 s = (faults, prog, pc, pos, m) \<longrightarrow> |
91 prog $ pc = Some ((action0, pc0), (action1, pc1)) \<longrightarrow> |
102 prog f! pc = Some ((action0, pc0), (action1, pc1)) \<longrightarrow> |
92 m $ pos = Some Bk \<longrightarrow> P (faults, prog, pc0, next_tape action0 (pos, m))) \<and> |
103 m f! pos = Some Bk \<longrightarrow> P (faults, prog, pc0, next_tape action0 (pos, m))) \<and> |
93 (\<forall> faults prog pc pos m action0 pc0 action1 pc1. |
104 (\<forall> faults prog pc pos m action0 pc0 action1 pc1. |
94 s = (faults, prog, pc, pos, m) \<longrightarrow> |
105 s = (faults, prog, pc, pos, m) \<longrightarrow> |
95 prog $ pc = Some ((action0, pc0), (action1, pc1)) \<longrightarrow> |
106 prog f! pc = Some ((action0, pc0), (action1, pc1)) \<longrightarrow> |
96 m $ pos = Some Oc \<longrightarrow> P (faults, prog, pc1, next_tape action1 (pos, m))) \<and> |
107 m f! pos = Some Oc \<longrightarrow> P (faults, prog, pc1, next_tape action1 (pos, m))) \<and> |
97 (\<forall> faults prog pc pos m action0 pc0 action1 pc1. |
108 (\<forall> faults prog pc pos m action0 pc0 action1 pc1. |
98 s = (faults, prog, pc, pos, m) \<longrightarrow> |
109 s = (faults, prog, pc, pos, m) \<longrightarrow> |
99 prog $ pc = Some ((action0, pc0), (action1, pc1)) \<longrightarrow> |
110 prog f! pc = Some ((action0, pc0), (action1, pc1)) \<longrightarrow> |
100 m $ pos = None \<longrightarrow> P (faults + 1, prog, pc, pos, m)) \<and> |
111 m f! pos = None \<longrightarrow> P (faults + 1, prog, pc, pos, m)) \<and> |
101 (\<forall> faults prog pc pos m . |
112 (\<forall> faults prog pc pos m . |
102 s = (faults, prog, pc, pos, m) \<longrightarrow> |
113 s = (faults, prog, pc, pos, m) \<longrightarrow> |
103 prog $ pc = None \<longrightarrow> P (faults + 1, prog, pc, pos, m)) |
114 prog f! pc = None \<longrightarrow> P (faults + 1, prog, pc, pos, m)) |
104 )" |
115 )" |
105 by (cases s) (auto split: option.splits Block.splits) |
116 by (cases s) (auto split: option.splits Block.splits) |
106 |
117 |
107 datatype tresource = |
118 datatype tresource = |
108 TM int Block (* at the position int of the tape is a Bl or Oc *) |
119 TM int Block (* at the position int of the tape is a Bl or Oc *) |
109 | TC nat tm_inst (* at the address nat is a certain instruction *) |
120 | TC nat tm_inst (* at the address nat is a certain instruction *) |
110 | TAt nat (* TM is at state nat*) |
121 | TAt nat (* TM is at state nat*) |
111 | TPos int (* head of TM is at position int *) |
122 | TPos int (* head of TM is at position int *) |
112 | TFaults nat (* there are nat faults *) |
123 | TFaults nat (* there are nat faults *) |
113 |
124 |
114 definition "tprog_set prog = {TC i inst | i inst. prog $ i = Some inst}" |
125 definition "tprog_set prog = {TC i inst | i inst. prog f! i = Some inst}" |
115 definition "tpc_set pc = {TAt pc}" |
126 definition "tpc_set pc = {TAt pc}" |
116 definition "tmem_set m = {TM i n | i n. m $ i = Some n}" |
127 definition "tmem_set m = {TM i n | i n. m f! i = Some n}" |
117 definition "tpos_set pos = {TPos pos}" |
128 definition "tpos_set pos = {TPos pos}" |
118 definition "tfaults_set faults = {TFaults faults}" |
129 definition "tfaults_set faults = {TFaults faults}" |
119 |
130 |
120 lemmas tpn_set_def = tprog_set_def tpc_set_def tmem_set_def tfaults_set_def tpos_set_def |
131 lemmas tpn_set_def = tprog_set_def tpc_set_def tmem_set_def tfaults_set_def tpos_set_def |
121 |
132 |
403 by simp |
414 by simp |
404 thus ?thesis |
415 thus ?thesis |
405 by (unfold tpn_set_def, auto) |
416 by (unfold tpn_set_def, auto) |
406 qed |
417 qed |
407 |
418 |
408 lemma codeD: "(st i ** sg {TC i inst} ** r) (trset_of (ft, prog, i', pos, mem)) |
419 lemma codeD: "(st i \<and>* sg {TC i inst} ** r) (trset_of (ft, prog, i', pos, mem)) |
409 \<Longrightarrow> prog $ i = Some inst" |
420 \<Longrightarrow> prog f! i = Some inst" |
410 proof - |
421 proof - |
411 assume "(st i ** sg {TC i inst} ** r) (trset_of (ft, prog, i', pos, mem))" |
422 assume "(st i \<and>* sg {TC i inst} \<and>* r) (trset_of (ft, prog, i', pos, mem))" |
412 thus ?thesis |
423 thus ?thesis |
413 apply(unfold sep_conj_def set_ins_def sg_def trset_of.simps tpn_set_def) |
424 apply(unfold sep_conj_def set_ins_def sg_def trset_of.simps tpn_set_def) |
414 by auto |
425 by auto |
415 qed |
426 qed |
416 |
427 |
417 lemma memD: "((tm a v) ** r) (trset_of (ft, prog, i, pos, mem)) \<Longrightarrow> mem $ a = Some v" |
428 lemma memD: "((tm a v) ** r) (trset_of (ft, prog, i, pos, mem)) \<Longrightarrow> mem f! a = Some v" |
418 proof - |
429 proof - |
419 assume "((tm a v) ** r) (trset_of (ft, prog, i, pos, mem))" |
430 assume "((tm a v) ** r) (trset_of (ft, prog, i, pos, mem))" |
420 from stimes_sgD[OF this[unfolded trset_of.simps tpn_set_def tm_def]] |
431 from stimes_sgD[OF this[unfolded trset_of.simps tpn_set_def tm_def]] |
421 have "{TM a v} \<subseteq> {TC i inst |i inst. prog $ i = Some inst} \<union> {TAt i} \<union> |
432 have "{TM a v} \<subseteq> {TC i inst |i inst. prog f! i = Some inst} \<union> {TAt i} \<union> |
422 {TPos pos} \<union> {TM i n |i n. mem $ i = Some n} \<union> {TFaults ft}" by simp |
433 {TPos pos} \<union> {TM i n |i n. mem f! i = Some n} \<union> {TFaults ft}" by simp |
423 thus ?thesis by auto |
434 thus ?thesis by auto |
424 qed |
435 qed |
425 |
436 |
426 lemma t_hoare_seq: |
437 lemma t_hoare_seq: |
427 assumes a1: "\<And> i j. \<lbrace>st i \<and>* p\<rbrace> i:[c1]:j \<lbrace>st j \<and>* q\<rbrace>" |
438 assumes a1: "\<And> i j. \<lbrace>st i \<and>* p\<rbrace> i:[c1]:j \<lbrace>st j \<and>* q\<rbrace>" |
1210 ({TM a v} \<subseteq> tmem_set mem)" |
1221 ({TM a v} \<subseteq> tmem_set mem)" |
1211 by (unfold tpn_set_def, auto) |
1222 by (unfold tpn_set_def, auto) |
1212 |
1223 |
1213 lemma tmem_set_upd: |
1224 lemma tmem_set_upd: |
1214 "{TM a v} \<subseteq> tmem_set mem \<Longrightarrow> |
1225 "{TM a v} \<subseteq> tmem_set mem \<Longrightarrow> |
1215 tmem_set (mem(a $:=Some v')) = ((tmem_set mem) - {TM a v}) \<union> {TM a v'}" |
1226 tmem_set (mem(a f\<mapsto> Some v')) = ((tmem_set mem) - {TM a v}) \<union> {TM a v'}" |
1216 apply(unfold tpn_set_def) |
1227 apply(unfold tpn_set_def) |
1217 apply(auto) |
1228 apply(auto) |
1218 apply (metis finfun_upd_apply option.inject) |
1229 apply (metis the.simps the_lookup_fmap_upd the_lookup_fmap_upd_other) |
1219 apply (metis finfun_upd_apply_other) |
1230 apply (metis the_lookup_fmap_upd_other) |
1220 by (metis finfun_upd_apply_other option.inject) |
1231 by (metis option.inject the_lookup_fmap_upd_other) |
1221 |
|
1222 |
1232 |
1223 lemma tmem_set_disj: "{TM a v} \<subseteq> tmem_set mem \<Longrightarrow> |
1233 lemma tmem_set_disj: "{TM a v} \<subseteq> tmem_set mem \<Longrightarrow> |
1224 {TM a v'} \<inter> (tmem_set mem - {TM a v}) = {}" |
1234 {TM a v'} \<inter> (tmem_set mem - {TM a v}) = {}" |
1225 by (unfold tpn_set_def, auto) |
1235 by (unfold tpn_set_def, auto) |
1226 |
1236 |
1227 lemma smem_upd: "((tm a v) ** r) (trset_of (f, x, y, z, mem)) \<Longrightarrow> |
1237 lemma smem_upd: "((tm a v) ** r) (trset_of (f, x, y, z, mem)) \<Longrightarrow> |
1228 ((tm a v') ** r) (trset_of (f, x, y, z, mem(a $:= Some v')))" |
1238 ((tm a v') ** r) (trset_of (f, x, y, z, mem(a f\<mapsto> Some v')))" |
1229 proof - |
1239 proof - |
1230 have eq_s: "(tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> tmem_set mem \<union> tfaults_set f - {TM a v}) = |
1240 have eq_s: "(tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> tmem_set mem \<union> tfaults_set f - {TM a v}) = |
1231 (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f)" |
1241 (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f)" |
1232 by (unfold tpn_set_def, auto) |
1242 by (unfold tpn_set_def, auto) |
1233 assume g: "(tm a v \<and>* r) (trset_of (f, x, y, z, mem))" |
1243 assume g: "(tm a v \<and>* r) (trset_of (f, x, y, z, mem))" |
1236 hence h0: "r (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f)" |
1246 hence h0: "r (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f)" |
1237 by(sep_drule stimes_sgD, clarify, unfold eq_s, auto) |
1247 by(sep_drule stimes_sgD, clarify, unfold eq_s, auto) |
1238 from h TM_in_simp have "{TM a v} \<subseteq> tmem_set mem" |
1248 from h TM_in_simp have "{TM a v} \<subseteq> tmem_set mem" |
1239 by(sep_drule stimes_sgD, auto) |
1249 by(sep_drule stimes_sgD, auto) |
1240 from tmem_set_upd [OF this] tmem_set_disj[OF this] |
1250 from tmem_set_upd [OF this] tmem_set_disj[OF this] |
1241 have h2: "tmem_set (mem(a $:= Some v')) = {TM a v'} \<union> (tmem_set mem - {TM a v})" |
1251 have h2: "tmem_set (mem(a f\<mapsto> Some v')) = {TM a v'} \<union> (tmem_set mem - {TM a v})" |
1242 "{TM a v'} \<inter> (tmem_set mem - {TM a v}) = {}" by auto |
1252 "{TM a v'} \<inter> (tmem_set mem - {TM a v}) = {}" by auto |
1243 show ?thesis |
1253 show ?thesis |
1244 proof - |
1254 proof - |
1245 have "(tm a v' ** r) (tmem_set (mem(a $:= Some v')) \<union> tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> tfaults_set f)" |
1255 have "(tm a v' ** r) (tmem_set (mem(a f\<mapsto> Some v')) \<union> tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> tfaults_set f)" |
1246 proof(rule sep_conjI) |
1256 proof(rule sep_conjI) |
1247 show "tm a v' ({TM a v'})" by (unfold tm_def sg_def, simp) |
1257 show "tm a v' ({TM a v'})" by (unfold tm_def sg_def, simp) |
1248 next |
1258 next |
1249 from h0 show "r (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f)" . |
1259 from h0 show "r (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f)" . |
1250 next |
1260 next |
1251 show "{TM a v'} ## tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f" |
1261 show "{TM a v'} ## tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f" |
1252 proof - |
1262 proof - |
1253 from g have " mem $ a = Some v" |
1263 from g have " mem f! a = Some v" |
1254 by(sep_frule memD, simp) |
1264 by(sep_frule memD, simp) |
1255 thus "?thesis" |
1265 thus "?thesis" |
1256 by(unfold tpn_set_def set_ins_def, auto) |
1266 by(unfold tpn_set_def set_ins_def, auto) |
1257 qed |
1267 qed |
1258 next |
1268 next |
1259 show "tmem_set (mem(a $:= Some v')) \<union> tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> tfaults_set f = |
1269 show "tmem_set (mem(a f\<mapsto> Some v')) \<union> tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> tfaults_set f = |
1260 {TM a v'} + (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f)" |
1270 {TM a v'} + (tprog_set x \<union> tpc_set y \<union> tpos_set z \<union> (tmem_set mem - {TM a v}) \<union> tfaults_set f)" |
1261 by (unfold h2(1) set_ins_def eq_s, auto) |
1271 by (unfold h2(1) set_ins_def eq_s, auto) |
1262 qed |
1272 qed |
1263 thus ?thesis |
1273 thus ?thesis |
1264 apply (unfold trset_of.simps) |
1274 apply (unfold trset_of.simps) |
1279 \<lbrace>st (Suc i) \<and>* ps p \<and>* tm p Bk\<rbrace>" |
1289 \<lbrace>st (Suc i) \<and>* ps p \<and>* tm p Bk\<rbrace>" |
1280 proof(fold eq_j, unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) |
1290 proof(fold eq_j, unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) |
1281 fix ft prog cs i' mem r |
1291 fix ft prog cs i' mem r |
1282 assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v \<and>* sg {TC i ((W0, j), W0, j)}) |
1292 assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v \<and>* sg {TC i ((W0, j), W0, j)}) |
1283 (trset_of (ft, prog, cs, i', mem))" |
1293 (trset_of (ft, prog, cs, i', mem))" |
1284 from h have "prog $ i = Some ((W0, j), W0, j)" |
1294 from h have "prog f! i = Some ((W0, j), W0, j)" |
1285 apply(rule_tac r = "r \<and>* ps p \<and>* tm p v" in codeD) |
1295 apply(rule_tac r = "r \<and>* ps p \<and>* tm p v" in codeD) |
1286 by(simp add: sep_conj_ac) |
1296 by(simp add: sep_conj_ac) |
1287 from h and this have stp: |
1297 from h and this have stp: |
1288 "tm.run 1 (ft, prog, cs, i', mem) = (ft, prog, j, i', mem(i' $:= Some Bk))" (is "?x = ?y") |
1298 "tm.run 1 (ft, prog, cs, i', mem) = (ft, prog, j, i', mem(i' f\<mapsto> Some Bk))" (is "?x = ?y") |
1289 apply(sep_frule psD) |
1299 apply(sep_frule psD) |
1290 apply(sep_frule stD) |
1300 apply(sep_frule stD) |
1291 apply(sep_frule memD, simp) |
1301 apply(sep_frule memD, simp) |
1292 by(cases v, unfold tm.run_def, auto) |
1302 by(cases v, unfold tm.run_def, auto) |
1293 from h have "i' = p" |
1303 from h have "i' = p" |
1329 \<lbrace>ps p \<and>* st ?j \<and>* tm p Oc\<rbrace>" |
1339 \<lbrace>ps p \<and>* st ?j \<and>* tm p Oc\<rbrace>" |
1330 proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) |
1340 proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) |
1331 fix ft prog cs i' mem r |
1341 fix ft prog cs i' mem r |
1332 assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v \<and>* sg {TC i ((W1, ?j), W1, ?j)}) |
1342 assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v \<and>* sg {TC i ((W1, ?j), W1, ?j)}) |
1333 (trset_of (ft, prog, cs, i', mem))" |
1343 (trset_of (ft, prog, cs, i', mem))" |
1334 from h have "prog $ i = Some ((W1, ?j), W1, ?j)" |
1344 from h have "prog f! i = Some ((W1, ?j), W1, ?j)" |
1335 apply(rule_tac r = "r \<and>* ps p \<and>* tm p v" in codeD) |
1345 apply(rule_tac r = "r \<and>* ps p \<and>* tm p v" in codeD) |
1336 by(simp add: sep_conj_ac) |
1346 by(simp add: sep_conj_ac) |
1337 from h and this have stp: |
1347 from h and this have stp: |
1338 "tm.run 1 (ft, prog, cs, i', mem) = |
1348 "tm.run 1 (ft, prog, cs, i', mem) = |
1339 (ft, prog, ?j, i', mem(i' $:= Some Oc))" (is "?x = ?y") |
1349 (ft, prog, ?j, i', mem(i' f\<mapsto> Some Oc))" (is "?x = ?y") |
1340 apply(sep_frule psD) |
1350 apply(sep_frule psD) |
1341 apply(sep_frule stD) |
1351 apply(sep_frule stD) |
1342 apply(sep_frule memD, simp) |
1352 apply(sep_frule memD, simp) |
1343 by(cases v, unfold tm.run_def, auto) |
1353 by(cases v, unfold tm.run_def, auto) |
1344 from h have "i' = p" |
1354 from h have "i' = p" |
1381 \<lbrace>st ?j \<and>* tm p v2 \<and>* ps (p - 1)\<rbrace>" |
1391 \<lbrace>st ?j \<and>* tm p v2 \<and>* ps (p - 1)\<rbrace>" |
1382 proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) |
1392 proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) |
1383 fix ft prog cs i' mem r |
1393 fix ft prog cs i' mem r |
1384 assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v2 \<and>* sg {TC i ((L, ?j), L, ?j)}) |
1394 assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v2 \<and>* sg {TC i ((L, ?j), L, ?j)}) |
1385 (trset_of (ft, prog, cs, i', mem))" |
1395 (trset_of (ft, prog, cs, i', mem))" |
1386 from h have "prog $ i = Some ((L, ?j), L, ?j)" |
1396 from h have "prog f! i = Some ((L, ?j), L, ?j)" |
1387 apply(rule_tac r = "r \<and>* ps p \<and>* tm p v2" in codeD) |
1397 apply(rule_tac r = "r \<and>* ps p \<and>* tm p v2" in codeD) |
1388 by(simp add: sep_conj_ac) |
1398 by(simp add: sep_conj_ac) |
1389 from h and this have stp: |
1399 from h and this have stp: |
1390 "tm.run 1 (ft, prog, cs, i', mem) = (ft, prog, ?j, i' - 1, mem)" (is "?x = ?y") |
1400 "tm.run 1 (ft, prog, cs, i', mem) = (ft, prog, ?j, i' - 1, mem)" (is "?x = ?y") |
1391 apply(sep_frule psD) |
1401 apply(sep_frule psD) |
1439 \<lbrace>st ?j \<and>* tm p v1 \<and>* ps (p + 1)\<rbrace>" |
1449 \<lbrace>st ?j \<and>* tm p v1 \<and>* ps (p + 1)\<rbrace>" |
1440 proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) |
1450 proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) |
1441 fix ft prog cs i' mem r |
1451 fix ft prog cs i' mem r |
1442 assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v1 \<and>* sg {TC i ((R, ?j), R, ?j)}) |
1452 assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v1 \<and>* sg {TC i ((R, ?j), R, ?j)}) |
1443 (trset_of (ft, prog, cs, i', mem))" |
1453 (trset_of (ft, prog, cs, i', mem))" |
1444 from h have "prog $ i = Some ((R, ?j), R, ?j)" |
1454 from h have "prog f! i = Some ((R, ?j), R, ?j)" |
1445 apply(rule_tac r = "r \<and>* ps p \<and>* tm p v1" in codeD) |
1455 apply(rule_tac r = "r \<and>* ps p \<and>* tm p v1" in codeD) |
1446 by(simp add: sep_conj_ac) |
1456 by(simp add: sep_conj_ac) |
1447 from h and this have stp: |
1457 from h and this have stp: |
1448 "tm.run 1 (ft, prog, cs, i', mem) = (ft, prog, ?j, i'+ 1, mem)" (is "?x = ?y") |
1458 "tm.run 1 (ft, prog, cs, i', mem) = (ft, prog, ?j, i'+ 1, mem)" (is "?x = ?y") |
1449 apply(sep_frule psD) |
1459 apply(sep_frule psD) |
1495 show "\<lbrace>one p \<and>* ps p \<and>* st i\<rbrace> sg {TC i ((W0, ?j), W1, e)} \<lbrace>one p \<and>* ps p \<and>* st e\<rbrace>" |
1505 show "\<lbrace>one p \<and>* ps p \<and>* st i\<rbrace> sg {TC i ((W0, ?j), W1, e)} \<lbrace>one p \<and>* ps p \<and>* st e\<rbrace>" |
1496 proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) |
1506 proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) |
1497 fix ft prog cs pc mem r |
1507 fix ft prog cs pc mem r |
1498 assume h: "(r \<and>* one p \<and>* ps p \<and>* st i \<and>* sg {TC i ((W0, ?j), W1, e)}) |
1508 assume h: "(r \<and>* one p \<and>* ps p \<and>* st i \<and>* sg {TC i ((W0, ?j), W1, e)}) |
1499 (trset_of (ft, prog, cs, pc, mem))" |
1509 (trset_of (ft, prog, cs, pc, mem))" |
1500 from h have k1: "prog $ i = Some ((W0, ?j), W1, e)" |
1510 from h have k1: "prog f! i = Some ((W0, ?j), W1, e)" |
1501 apply(rule_tac r = "r \<and>* one p \<and>* ps p" in codeD) |
1511 apply(rule_tac r = "r \<and>* one p \<and>* ps p" in codeD) |
1502 by(simp add: sep_conj_ac) |
1512 by(simp add: sep_conj_ac) |
1503 from h have k2: "pc = p" |
1513 from h have k2: "pc = p" |
1504 by(sep_drule psD, simp) |
1514 by(sep_drule psD, simp) |
1505 from h and this have k3: "mem $ pc = Some Oc" |
1515 from h and this have k3: "mem f! pc = Some Oc" |
1506 apply(unfold one_def) |
1516 apply(unfold one_def) |
1507 by(sep_drule memD, simp) |
1517 by(sep_drule memD, simp) |
1508 from h k1 k2 k3 have stp: |
1518 from h k1 k2 k3 have stp: |
1509 "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, e, pc, mem)" (is "?x = ?y") |
1519 "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, e, pc, mem)" (is "?x = ?y") |
1510 apply(sep_drule stD) |
1520 apply(sep_drule stD) |
1511 apply(unfold tm.run_def) |
1521 apply(unfold tm.run_def) |
1512 apply(auto) |
1522 apply(auto) |
|
1523 apply(rule fmap_eqI) |
|
1524 apply(simp) |
|
1525 apply(subgoal_tac "p \<in> fdom mem") |
|
1526 apply(simp add: insert_absorb) |
|
1527 apply(simp add: fdomIff) |
|
1528 apply(rule_tac x="Some Oc" in exI) |
|
1529 apply(auto)[1] |
|
1530 apply(simp add: fun_eq_iff) |
1513 by (metis finfun_upd_triv) |
1531 by (metis finfun_upd_triv) |
1514 |
1532 |
1515 from h k2 |
1533 from h k2 |
1516 have "(r \<and>* one p \<and>* ps p \<and>* st e \<and>* sg {TC i ((W0, ?j), W1, e)}) (trset_of ?x)" |
1534 have "(r \<and>* one p \<and>* ps p \<and>* st e \<and>* sg {TC i ((W0, ?j), W1, e)}) (trset_of ?x)" |
1517 apply(unfold stp) |
1535 apply(unfold stp) |
1582 show "\<lbrace>ps p \<and>* st i \<and>* zero p\<rbrace> sg {TC i ((W0, ?j), W1, e)} \<lbrace>ps p \<and>* zero p \<and>* st ?j \<rbrace>" |
1600 show "\<lbrace>ps p \<and>* st i \<and>* zero p\<rbrace> sg {TC i ((W0, ?j), W1, e)} \<lbrace>ps p \<and>* zero p \<and>* st ?j \<rbrace>" |
1583 proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) |
1601 proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) |
1584 fix ft prog cs pc mem r |
1602 fix ft prog cs pc mem r |
1585 assume h: "(r \<and>* ps p \<and>* st i \<and>* zero p \<and>* sg {TC i ((W0, ?j), W1, e)}) |
1603 assume h: "(r \<and>* ps p \<and>* st i \<and>* zero p \<and>* sg {TC i ((W0, ?j), W1, e)}) |
1586 (trset_of (ft, prog, cs, pc, mem))" |
1604 (trset_of (ft, prog, cs, pc, mem))" |
1587 from h have k1: "prog $ i = Some ((W0, ?j), W1, e)" |
1605 from h have k1: "prog f! i = Some ((W0, ?j), W1, e)" |
1588 apply(rule_tac r = "r \<and>* zero p \<and>* ps p" in codeD) |
1606 apply(rule_tac r = "r \<and>* zero p \<and>* ps p" in codeD) |
1589 by(simp add: sep_conj_ac) |
1607 by(simp add: sep_conj_ac) |
1590 from h have k2: "pc = p" |
1608 from h have k2: "pc = p" |
1591 by(sep_drule psD, simp) |
1609 by(sep_drule psD, simp) |
1592 from h and this have k3: "mem $ pc = Some Bk" |
1610 from h and this have k3: "mem f! pc = Some Bk" |
1593 apply(unfold zero_def) |
1611 apply(unfold zero_def) |
1594 by(sep_drule memD, simp) |
1612 by(sep_drule memD, simp) |
1595 from h k1 k2 k3 have stp: |
1613 from h k1 k2 k3 have stp: |
1596 "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, ?j, pc, mem)" (is "?x = ?y") |
1614 "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, ?j, pc, mem)" (is "?x = ?y") |
1597 apply(sep_drule stD) |
1615 apply(sep_drule stD) |
1631 show "\<lbrace>ps p \<and>* st i \<and>* zero p\<rbrace> sg {TC i ((W0, e), W1, ?j)} \<lbrace>ps p \<and>* st e \<and>* zero p\<rbrace>" |
1649 show "\<lbrace>ps p \<and>* st i \<and>* zero p\<rbrace> sg {TC i ((W0, e), W1, ?j)} \<lbrace>ps p \<and>* st e \<and>* zero p\<rbrace>" |
1632 proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) |
1650 proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) |
1633 fix ft prog cs pc mem r |
1651 fix ft prog cs pc mem r |
1634 assume h: "(r \<and>* ps p \<and>* st i \<and>* zero p \<and>* sg {TC i ((W0, e), W1, ?j)}) |
1652 assume h: "(r \<and>* ps p \<and>* st i \<and>* zero p \<and>* sg {TC i ((W0, e), W1, ?j)}) |
1635 (trset_of (ft, prog, cs, pc, mem))" |
1653 (trset_of (ft, prog, cs, pc, mem))" |
1636 from h have k1: "prog $ i = Some ((W0, e), W1, ?j)" |
1654 from h have k1: "prog f! i = Some ((W0, e), W1, ?j)" |
1637 apply(rule_tac r = "r \<and>* zero p \<and>* ps p" in codeD) |
1655 apply(rule_tac r = "r \<and>* zero p \<and>* ps p" in codeD) |
1638 by(simp add: sep_conj_ac) |
1656 by(simp add: sep_conj_ac) |
1639 from h have k2: "pc = p" |
1657 from h have k2: "pc = p" |
1640 by(sep_drule psD, simp) |
1658 by(sep_drule psD, simp) |
1641 from h and this have k3: "mem $ pc = Some Bk" |
1659 from h and this have k3: "mem f! pc = Some Bk" |
1642 apply(unfold zero_def) |
1660 apply(unfold zero_def) |
1643 by(sep_drule memD, simp) |
1661 by(sep_drule memD, simp) |
1644 from h k1 k2 k3 have stp: |
1662 from h k1 k2 k3 have stp: |
1645 "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, e, pc, mem)" (is "?x = ?y") |
1663 "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, e, pc, mem)" (is "?x = ?y") |
1646 apply(sep_drule stD) |
1664 apply(sep_drule stD) |
1710 \<lbrace>ps p \<and>* st ?j \<and>* tm p Oc\<rbrace>" |
1728 \<lbrace>ps p \<and>* st ?j \<and>* tm p Oc\<rbrace>" |
1711 proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) |
1729 proof(unfold tassemble_to.simps tm.Hoare_gen_def sep_conj_ac, clarify) |
1712 fix ft prog cs pc mem r |
1730 fix ft prog cs pc mem r |
1713 assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p Oc \<and>* sg {TC i ((W0, e), W1, ?j)}) |
1731 assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p Oc \<and>* sg {TC i ((W0, e), W1, ?j)}) |
1714 (trset_of (ft, prog, cs, pc, mem))" |
1732 (trset_of (ft, prog, cs, pc, mem))" |
1715 from h have k1: "prog $ i = Some ((W0, e), W1, ?j)" |
1733 from h have k1: "prog f! i = Some ((W0, e), W1, ?j)" |
1716 apply(rule_tac r = "r \<and>* tm p Oc \<and>* ps p" in codeD) |
1734 apply(rule_tac r = "r \<and>* tm p Oc \<and>* ps p" in codeD) |
1717 by(simp add: sep_conj_ac) |
1735 by(simp add: sep_conj_ac) |
1718 from h have k2: "pc = p" |
1736 from h have k2: "pc = p" |
1719 by(sep_drule psD, simp) |
1737 by(sep_drule psD, simp) |
1720 from h and this have k3: "mem $ pc = Some Oc" |
1738 from h and this have k3: "mem f! pc = Some Oc" |
1721 by(sep_drule memD, simp) |
1739 by(sep_drule memD, simp) |
1722 from h k1 k2 k3 have stp: |
1740 from h k1 k2 k3 have stp: |
1723 "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, ?j, pc, mem)" (is "?x = ?y") |
1741 "tm.run 1 (ft, prog, cs, pc, mem) = (ft, prog, ?j, pc, mem)" (is "?x = ?y") |
1724 apply(sep_drule stD) |
1742 apply(sep_drule stD) |
1725 apply(unfold tm.run_def) |
1743 apply(unfold tm.run_def) |
1751 "\<lbrace>st i \<and>* ps p \<and>* tm p v\<rbrace> i:[jmp e]:j \<lbrace>st e \<and>* ps p \<and>* tm p v\<rbrace>" |
1769 "\<lbrace>st i \<and>* ps p \<and>* tm p v\<rbrace> i:[jmp e]:j \<lbrace>st e \<and>* ps p \<and>* tm p v\<rbrace>" |
1752 proof(unfold jmp_def tm.Hoare_gen_def tassemble_to.simps sep_conj_ac, clarify) |
1770 proof(unfold jmp_def tm.Hoare_gen_def tassemble_to.simps sep_conj_ac, clarify) |
1753 fix ft prog cs pos mem r |
1771 fix ft prog cs pos mem r |
1754 assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v \<and>* <(j = i + 1)> \<and>* sg {TC i ((W0, e), W1, e)}) |
1772 assume h: "(r \<and>* ps p \<and>* st i \<and>* tm p v \<and>* <(j = i + 1)> \<and>* sg {TC i ((W0, e), W1, e)}) |
1755 (trset_of (ft, prog, cs, pos, mem))" |
1773 (trset_of (ft, prog, cs, pos, mem))" |
1756 from h have k1: "prog $ i = Some ((W0, e), W1, e)" |
1774 from h have k1: "prog f! i = Some ((W0, e), W1, e)" |
1757 apply(rule_tac r = "r \<and>* <(j = i + 1)> \<and>* tm p v \<and>* ps p" in codeD) |
1775 apply(rule_tac r = "r \<and>* <(j = i + 1)> \<and>* tm p v \<and>* ps p" in codeD) |
1758 by(simp add: sep_conj_ac) |
1776 by(simp add: sep_conj_ac) |
1759 from h have k2: "p = pos" |
1777 from h have k2: "p = pos" |
1760 by(sep_drule psD, simp) |
1778 by(sep_drule psD, simp) |
1761 from h k2 have k3: "mem $ pos = Some v" |
1779 from h k2 have k3: "mem f! pos = Some v" |
1762 by(sep_drule memD, simp) |
1780 by(sep_drule memD, simp) |
1763 from h k1 k2 k3 have |
1781 from h k1 k2 k3 have |
1764 stp: "tm.run 1 (ft, prog, cs, pos, mem) = (ft, prog, e, pos, mem)" (is "?x = ?y") |
1782 stp: "tm.run 1 (ft, prog, cs, pos, mem) = (ft, prog, e, pos, mem)" (is "?x = ?y") |
1765 apply(sep_drule stD) |
1783 apply(sep_drule stD) |
1766 apply(unfold tm.run_def) |
1784 apply(unfold tm.run_def) |
1767 apply(cases "mem $ pos") |
1785 apply(cases "mem f! pos") |
1768 apply(simp) |
1786 apply(simp) |
1769 apply(cases v) |
1787 apply(cases v) |
1770 apply(auto) |
1788 apply(auto) |
1771 apply (metis finfun_upd_triv) |
1789 apply (metis finfun_upd_triv) |
1772 by (metis finfun_upd_triv) |
1790 by (metis finfun_upd_triv) |