180 {P1} A {Q1} {P2} B {Q2} Q1 \<mapsto> P2 |
203 {P1} A {Q1} {P2} B {Q2} Q1 \<mapsto> P2 |
181 ----------------------------------- |
204 ----------------------------------- |
182 {P1} A |+| B {Q2} |
205 {P1} A |+| B {Q2} |
183 *} |
206 *} |
184 |
207 |
|
208 lemma before_final: |
|
209 assumes "steps (1, tp) A n = (0, tp')" |
|
210 obtains n' where "\<not> is_final (steps (1, tp) A n')" and "steps (1, tp) A (Suc n') = (0, tp')" |
|
211 using assms |
|
212 apply(induct n) |
|
213 apply(auto) |
|
214 by (metis is_final.simps step_red steps.simps steps_0 surj_pair) |
|
215 |
|
216 lemma t_merge_fetch_pre: |
|
217 assumes "fetch A s b = (ac, ns)" "s \<le> length A div 2" "tm_wf A" "s \<noteq> 0" |
|
218 shows "fetch (adjustA |+| B) s b = (ac, if ns = 0 then Suc (length A div 2) else ns)" |
|
219 using assms |
|
220 unfolding tm_comp_def |
|
221 apply(induct A) |
|
222 apply(auto) |
|
223 apply(subgoal_tac "2 * (s - Suc 0) < length A \<and> Suc (2 * (s - Suc 0)) < length A") |
|
224 apply(auto simp: tm_wf_def iseven_def tm_comp_def split: if_splits cell.splits) |
|
225 oops |
|
226 |
|
227 lemma t_merge_pre_eq_step: |
|
228 "\<lbrakk>step (a, b, c) A = cf; tm_wf A; \<not> is_final cf\<rbrakk> |
|
229 \<Longrightarrow> step (a, b, c) (A |+| B) = cf" |
|
230 apply(subgoal_tac "a \<le> length A div 2 \<and> a \<noteq> 0") |
|
231 apply(simp) |
|
232 apply(case_tac "fetch A a (read c)", simp) |
|
233 apply(auto) |
|
234 oops |
|
235 |
|
236 lemma t_merge_pre_eq: |
|
237 "\<lbrakk>steps (Suc 0, tp) A stp = cf; \<not> is_final cf; tm_wf A\<rbrakk> |
|
238 \<Longrightarrow> steps (Suc 0, tp) (A |+| B) stp = cf" |
|
239 apply(induct stp arbitrary: cf) |
|
240 apply(auto)[1] |
|
241 apply(auto) |
|
242 oops |
|
243 |
|
244 lemma t_merge_pre_halt_same: |
|
245 assumes a_ht: "steps (1, tp) A n = (0, tp')" |
|
246 and a_wf: "t_correct A" |
|
247 obtains n' where "steps (1, tp) (A |+| B) n' = (Suc (length A div 2), tp')" |
|
248 proof - |
|
249 assume a: "\<And>n. steps (1, tp) (A |+| B) n = (Suc (length A div 2), tp') \<Longrightarrow> thesis" |
|
250 |
|
251 obtain stp' where "\<not> is_final (steps (1, tp) A stp')" and "steps (1, tp) A (Suc stp') = (0, tp')" |
|
252 using a_ht before_final by blast |
|
253 then have "steps (1, tp) (A |+| B) (Suc stp') = (Suc (length A div 2), tp')" |
|
254 sorry (*using a_wf t_merge_pre_halt_same' by blast*) |
|
255 with a show thesis by blast |
|
256 qed |
|
257 |
|
258 |
|
259 |
|
260 lemma steps_comp: |
|
261 assumes a1: "steps (1, l, r) A n1 = (s1, l1, r1)" |
|
262 and a2: "steps (1, l1, r1) B n2 = (s2, l2, r2)" |
|
263 shows "steps (1, l, r) (A |+| B) (n1 + n2) = (s2, l2, r2)" |
|
264 using assms |
|
265 apply(induct n2) |
|
266 apply(simp) |
|
267 apply(simp add: tm_comp_def) |
|
268 oops |
185 |
269 |
186 lemma Hoare_plus: |
270 lemma Hoare_plus: |
187 assumes aimpb: "Q1 \<mapsto> P2" |
271 assumes aimpb: "Q1 \<mapsto> P2" |
188 and A_wf : "tm_wf A" |
272 and A_wf : "tm_wf A" |
189 and B_wf : "tm_wf B" |
273 and B_wf : "tm_wf B" |
190 and A_halt : "{P1} A {Q1}" |
274 and A_halt : "{P1} A {Q1}" |
191 and B_halt : "{P2} B {Q2}" |
275 and B_halt : "{P2} B {Q2}" |
192 shows "{P1} A |+| B {Q2}" |
276 shows "{P1} A |+| B {Q2}" |
193 proof(rule HoareI) |
277 proof(rule HoareI) |
194 fix a b |
278 fix l r |
195 assume h: "P1 (a, b)" |
279 assume h: "P1 (l, r)" |
196 then have "\<exists>n. is_final (steps (1, a, b) A n) \<and> Q1 holds_for (steps (1, a, b) A n)" |
280 then obtain n1 where a: "is_final (steps (1, l, r) A n1)" and b: "Q1 holds_for (steps (1, l, r) A n1)" |
197 using A_halt unfolding Hoare_def by simp |
281 using A_halt unfolding Hoare_def by auto |
198 then obtain n1 where "is_final (steps (1, a, b) A n1) \<and> Q1 holds_for (steps (1, a, b) A stp1)" .. |
282 from b aimpb have "P2 holds_for (steps (1, l, r) A n1)" |
199 then show "\<exists>n. is_final (steps (1, a, b) (A |+| B) n) \<and> Q2 holds_for (steps (1, a, b) (A |+| B) n)" |
283 by (rule holds_for_imp) |
200 proof(case_tac "steps (Suc 0, a, b) A stp1", simp, erule_tac conjE) |
284 then obtain l' r' where "P2 (l', r')" |
201 fix aa ba c |
285 apply(auto) |
202 assume g1: "Q1 (ba, c)" |
286 apply(case_tac "steps (Suc 0, l, r) A n1") |
203 and g2: "steps (Suc 0, a, b) A stp1 = (0, ba, c)" |
287 apply(simp) |
204 hence "P2 (ba, c)" |
288 done |
205 using aimpb apply(simp add: assert_imp_def) |
289 then obtain n2 where a: "is_final (steps (1, l', r') B n2)" and b: "Q2 holds_for (steps (1, l', r') B n2)" |
206 done |
290 using B_halt unfolding Hoare_def by auto |
207 hence "\<exists> stp. is_final (steps (Suc 0, ba, c) B stp) \<and> Q2 holds_for (steps (Suc 0, ba, c) B stp)" |
291 show "\<exists>n. is_final (steps (1, l, r) (A |+| B) n) \<and> Q2 holds_for (steps (1, l, r) (A |+| B) n)" |
208 using B_halt unfolding Hoare_def by simp |
292 apply(rule_tac x="n1 + n2" in exI) |
209 from this obtain stp2 where |
293 apply(rule conjI) |
210 "is_final (steps (Suc 0, ba, c) B stp2) \<and> Q2 holds_for (steps (Suc 0, ba, c) B stp2)" .. |
294 apply(simp) |
211 thus "?thesis" |
295 apply(simp only: steps_add[symmetric]) |
212 proof(case_tac "steps (Suc 0, ba, c) B stp2", simp, erule_tac conjE) |
296 sorry |
213 fix aa bb ca |
|
214 assume g3: " Q2 (bb, ca)" "steps (Suc 0, ba, c) B stp2 = (0, bb, ca)" |
|
215 have "\<exists> stp. steps (Suc 0, a, b) (A |+| B) stp = (Suc (length A div 2), ba , c)" |
|
216 using g2 A_wf B_wf |
|
217 sorry |
|
218 moreover have "\<exists> stp. steps (Suc (length A div 2), ba, c) (A |+| B) stp = (0, bb, ca)" |
|
219 using g3 A_wf B_wf |
|
220 sorry |
|
221 ultimately show "\<exists>n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \<Rightarrow> s = 0 \<and> Q2 tp'" |
|
222 apply(erule_tac exE, erule_tac exE) |
|
223 apply(rule_tac x = "stp + stpa" in exI, simp) |
|
224 using g3 by simp |
|
225 qed |
|
226 qed |
|
227 qed |
297 qed |
228 |
298 |
229 lemma Hoare_plus2: |
299 |
230 assumes A_wf : "tm_wf A" |
|
231 and B_wf : "tm_wf B" |
|
232 and A_halt : "{P1} A {Q1}" |
|
233 and B_halt : "{P2} B {Q2}" |
|
234 and aimpb: "Q1 \<mapsto> P2" |
|
235 shows "{P1} A |+| B {Q2}" |
|
236 unfolding hoare_def |
|
237 proof(auto split: ) |
|
238 fix a b |
|
239 assume h: "P1 (a, b)" |
|
240 hence "\<exists>n. let (s, tp') = steps (Suc 0, a, b) A n in s = 0 \<and> Q1 tp'" |
|
241 using A_halt unfolding hoare_def by simp |
|
242 from this obtain stp1 where "let (s, tp') = steps (Suc 0, a, b) A stp1 in s = 0 \<and> Q1 tp'" .. |
|
243 then show "\<exists>n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \<Rightarrow> s = 0 \<and> Q2 tp'" |
|
244 proof(case_tac "steps (Suc 0, a, b) A stp1", simp, erule_tac conjE) |
|
245 fix aa ba c |
|
246 assume g1: "Q1 (ba, c)" |
|
247 and g2: "steps (Suc 0, a, b) A stp1 = (0, ba, c)" |
|
248 hence "P2 (ba, c)" |
|
249 using aimpb apply(simp add: assert_imp_def) |
|
250 done |
|
251 hence "\<exists> stp. let (s, tp') = steps (Suc 0, ba, c) B stp in s = 0 \<and> Q2 tp'" |
|
252 using B_halt unfolding hoare_def by simp |
|
253 from this obtain stp2 where "let (s, tp') = steps (Suc 0, ba, c) B stp2 in s = 0 \<and> Q2 tp'" .. |
|
254 thus "?thesis" |
|
255 proof(case_tac "steps (Suc 0, ba, c) B stp2", simp, erule_tac conjE) |
|
256 fix aa bb ca |
|
257 assume g3: " Q2 (bb, ca)" "steps (Suc 0, ba, c) B stp2 = (0, bb, ca)" |
|
258 have "\<exists> stp. steps (Suc 0, a, b) (A |+| B) stp = (Suc (length A div 2), ba , c)" |
|
259 using g2 A_wf B_wf |
|
260 sorry |
|
261 moreover have "\<exists> stp. steps (Suc (length A div 2), ba, c) (A |+| B) stp = (0, bb, ca)" |
|
262 using g3 A_wf B_wf |
|
263 sorry |
|
264 ultimately show "\<exists>n. case steps (Suc 0, a, b) (A |+| B) n of (s, tp') \<Rightarrow> s = 0 \<and> Q2 tp'" |
|
265 apply(erule_tac exE, erule_tac exE) |
|
266 apply(rule_tac x = "stp + stpa" in exI, simp) |
|
267 using g3 by simp |
|
268 qed |
|
269 qed |
|
270 qed |
|
271 |
300 |
272 |
301 |
273 |
302 |
274 locale turing_merge = |
303 locale turing_merge = |
275 fixes A :: "tprog" and B :: "tprog" and P1 :: "assert" |
304 fixes A :: "tprog" and B :: "tprog" and P1 :: "assert" |