changeset 63 | b620a2a0806a |
parent 62 | 031d2ae9c9b8 |
child 65 | 633b1fc8631b |
62:031d2ae9c9b8 | 63:b620a2a0806a |
---|---|
7 begin |
7 begin |
8 |
8 |
9 text {* @{text "the_preced"} is also the same as @{text "preced"}, the only |
9 text {* @{text "the_preced"} is also the same as @{text "preced"}, the only |
10 difference is the order of arguemts. *} |
10 difference is the order of arguemts. *} |
11 definition "the_preced s th = preced th s" |
11 definition "the_preced s th = preced th s" |
12 |
|
13 lemma inj_the_preced: |
|
14 "inj_on (the_preced s) (threads s)" |
|
15 by (metis inj_onI preced_unique the_preced_def) |
|
12 |
16 |
13 text {* @{term "the_thread"} extracts thread out of RAG node. *} |
17 text {* @{term "the_thread"} extracts thread out of RAG node. *} |
14 fun the_thread :: "node \<Rightarrow> thread" where |
18 fun the_thread :: "node \<Rightarrow> thread" where |
15 "the_thread (Th th) = th" |
19 "the_thread (Th th) = th" |
16 |
20 |
85 assumes "vt s'" |
89 assumes "vt s'" |
86 assumes "RAG s = RAG s' \<union> {(Th th, Cs cs)}" |
90 assumes "RAG s = RAG s' \<union> {(Th th, Cs cs)}" |
87 and "(Cs cs, Th th'') \<in> RAG s'" |
91 and "(Cs cs, Th th'') \<in> RAG s'" |
88 shows "tRAG s = tRAG s' \<union> {(Th th, Th th'')}" (is "?L = ?R") |
92 shows "tRAG s = tRAG s' \<union> {(Th th, Th th'')}" (is "?L = ?R") |
89 proof - |
93 proof - |
94 interpret vt_s': valid_trace "s'" using assms(1) |
|
95 by (unfold_locales, simp) |
|
90 interpret rtree: rtree "RAG s'" |
96 interpret rtree: rtree "RAG s'" |
91 proof |
97 proof |
92 show "single_valued (RAG s')" |
98 show "single_valued (RAG s')" |
93 apply (intro_locales) |
99 apply (intro_locales) |
94 by (unfold single_valued_def, |
100 by (unfold single_valued_def, |
95 auto intro:unique_RAG[OF assms(1)]) |
101 auto intro:vt_s'.unique_RAG) |
96 |
102 |
97 show "acyclic (RAG s')" |
103 show "acyclic (RAG s')" |
98 by (rule acyclic_RAG[OF assms(1)]) |
104 by (rule vt_s'.acyclic_RAG) |
99 qed |
105 qed |
100 { fix n1 n2 |
106 { fix n1 n2 |
101 assume "(n1, n2) \<in> ?L" |
107 assume "(n1, n2) \<in> ?L" |
102 from this[unfolded tRAG_alt_def] |
108 from this[unfolded tRAG_alt_def] |
103 obtain th1 th2 cs' where |
109 obtain th1 th2 cs' where |
146 by (unfold eq_n tRAG_alt_def, auto) |
152 by (unfold eq_n tRAG_alt_def, auto) |
147 qed |
153 qed |
148 } ultimately show ?thesis by auto |
154 } ultimately show ?thesis by auto |
149 qed |
155 qed |
150 |
156 |
157 context valid_trace |
|
158 begin |
|
159 |
|
160 lemmas RAG_tRAG_transfer = RAG_tRAG_transfer[OF vt] |
|
161 |
|
162 end |
|
163 |
|
151 lemma cp_alt_def: |
164 lemma cp_alt_def: |
152 "cp s th = |
165 "cp s th = |
153 Max ((the_preced s) ` {th'. Th th' \<in> (subtree (RAG s) (Th th))})" |
166 Max ((the_preced s) ` {th'. Th th' \<in> (subtree (RAG s) (Th th))})" |
154 proof - |
167 proof - |
155 have "Max (the_preced s ` ({th} \<union> dependants (wq s) th)) = |
168 have "Max (the_preced s ` ({th} \<union> dependants (wq s) th)) = |
215 have "(a, x) \<in> (RAG s)^*" by auto |
228 have "(a, x) \<in> (RAG s)^*" by auto |
216 hence "a \<in> subtree (RAG s) x" by (auto simp:subtree_def) |
229 hence "a \<in> subtree (RAG s) x" by (auto simp:subtree_def) |
217 } thus ?thesis by auto |
230 } thus ?thesis by auto |
218 qed |
231 qed |
219 |
232 |
233 lemma tRAG_trancl_eq: |
|
234 "{th'. (Th th', Th th) \<in> (tRAG s)^+} = |
|
235 {th'. (Th th', Th th) \<in> (RAG s)^+}" |
|
236 (is "?L = ?R") |
|
237 proof - |
|
238 { fix th' |
|
239 assume "th' \<in> ?L" |
|
240 hence "(Th th', Th th) \<in> (tRAG s)^+" by auto |
|
241 from tranclD[OF this] |
|
242 obtain z where h: "(Th th', z) \<in> tRAG s" "(z, Th th) \<in> (tRAG s)\<^sup>*" by auto |
|
243 from tRAG_subtree_RAG[of s] and this(2) |
|
244 have "(z, Th th) \<in> (RAG s)^*" by (meson subsetCE tRAG_star_RAG) |
|
245 moreover from h(1) have "(Th th', z) \<in> (RAG s)^+" using tRAG_alt_def by auto |
|
246 ultimately have "th' \<in> ?R" by auto |
|
247 } moreover |
|
248 { fix th' |
|
249 assume "th' \<in> ?R" |
|
250 hence "(Th th', Th th) \<in> (RAG s)^+" by (auto) |
|
251 from plus_rpath[OF this] |
|
252 obtain xs where rp: "rpath (RAG s) (Th th') xs (Th th)" "xs \<noteq> []" by auto |
|
253 hence "(Th th', Th th) \<in> (tRAG s)^+" |
|
254 proof(induct xs arbitrary:th' th rule:length_induct) |
|
255 case (1 xs th' th) |
|
256 then obtain x1 xs1 where Cons1: "xs = x1#xs1" by (cases xs, auto) |
|
257 show ?case |
|
258 proof(cases "xs1") |
|
259 case Nil |
|
260 from 1(2)[unfolded Cons1 Nil] |
|
261 have rp: "rpath (RAG s) (Th th') [x1] (Th th)" . |
|
262 hence "(Th th', x1) \<in> (RAG s)" by (cases, simp) |
|
263 then obtain cs where "x1 = Cs cs" |
|
264 by (unfold s_RAG_def, auto) |
|
265 from rpath_nnl_lastE[OF rp[unfolded this]] |
|
266 show ?thesis by auto |
|
267 next |
|
268 case (Cons x2 xs2) |
|
269 from 1(2)[unfolded Cons1[unfolded this]] |
|
270 have rp: "rpath (RAG s) (Th th') (x1 # x2 # xs2) (Th th)" . |
|
271 from rpath_edges_on[OF this] |
|
272 have eds: "edges_on (Th th' # x1 # x2 # xs2) \<subseteq> RAG s" . |
|
273 have "(Th th', x1) \<in> edges_on (Th th' # x1 # x2 # xs2)" |
|
274 by (simp add: edges_on_unfold) |
|
275 with eds have rg1: "(Th th', x1) \<in> RAG s" by auto |
|
276 then obtain cs1 where eq_x1: "x1 = Cs cs1" by (unfold s_RAG_def, auto) |
|
277 have "(x1, x2) \<in> edges_on (Th th' # x1 # x2 # xs2)" |
|
278 by (simp add: edges_on_unfold) |
|
279 from this eds |
|
280 have rg2: "(x1, x2) \<in> RAG s" by auto |
|
281 from this[unfolded eq_x1] |
|
282 obtain th1 where eq_x2: "x2 = Th th1" by (unfold s_RAG_def, auto) |
|
283 from rg1[unfolded eq_x1] rg2[unfolded eq_x1 eq_x2] |
|
284 have rt1: "(Th th', Th th1) \<in> tRAG s" by (unfold tRAG_alt_def, auto) |
|
285 from rp have "rpath (RAG s) x2 xs2 (Th th)" |
|
286 by (elim rpath_ConsE, simp) |
|
287 from this[unfolded eq_x2] have rp': "rpath (RAG s) (Th th1) xs2 (Th th)" . |
|
288 show ?thesis |
|
289 proof(cases "xs2 = []") |
|
290 case True |
|
291 from rpath_nilE[OF rp'[unfolded this]] |
|
292 have "th1 = th" by auto |
|
293 from rt1[unfolded this] show ?thesis by auto |
|
294 next |
|
295 case False |
|
296 from 1(1)[rule_format, OF _ rp' this, unfolded Cons1 Cons] |
|
297 have "(Th th1, Th th) \<in> (tRAG s)\<^sup>+" by simp |
|
298 with rt1 show ?thesis by auto |
|
299 qed |
|
300 qed |
|
301 qed |
|
302 hence "th' \<in> ?L" by auto |
|
303 } ultimately show ?thesis by blast |
|
304 qed |
|
305 |
|
306 lemma tRAG_trancl_eq_Th: |
|
307 "{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = |
|
308 {Th th' | th'. (Th th', Th th) \<in> (RAG s)^+}" |
|
309 using tRAG_trancl_eq by auto |
|
310 |
|
311 lemma dependants_alt_def: |
|
312 "dependants s th = {th'. (Th th', Th th) \<in> (tRAG s)^+}" |
|
313 by (metis eq_RAG s_dependants_def tRAG_trancl_eq) |
|
314 |
|
315 context valid_trace |
|
316 begin |
|
317 |
|
318 lemma count_eq_tRAG_plus: |
|
319 assumes "cntP s th = cntV s th" |
|
320 shows "{th'. (Th th', Th th) \<in> (tRAG s)^+} = {}" |
|
321 using assms count_eq_dependants dependants_alt_def eq_dependants by auto |
|
322 |
|
323 lemma count_eq_RAG_plus: |
|
324 assumes "cntP s th = cntV s th" |
|
325 shows "{th'. (Th th', Th th) \<in> (RAG s)^+} = {}" |
|
326 using assms count_eq_dependants cs_dependants_def eq_RAG by auto |
|
327 |
|
328 lemma count_eq_RAG_plus_Th: |
|
329 assumes "cntP s th = cntV s th" |
|
330 shows "{Th th' | th'. (Th th', Th th) \<in> (RAG s)^+} = {}" |
|
331 using count_eq_RAG_plus[OF assms] by auto |
|
332 |
|
333 lemma count_eq_tRAG_plus_Th: |
|
334 assumes "cntP s th = cntV s th" |
|
335 shows "{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = {}" |
|
336 using count_eq_tRAG_plus[OF assms] by auto |
|
337 |
|
338 end |
|
339 |
|
220 lemma tRAG_subtree_eq: |
340 lemma tRAG_subtree_eq: |
221 "(subtree (tRAG s) (Th th)) = {Th th' | th'. Th th' \<in> (subtree (RAG s) (Th th))}" |
341 "(subtree (tRAG s) (Th th)) = {Th th' | th'. Th th' \<in> (subtree (RAG s) (Th th))}" |
222 (is "?L = ?R") |
342 (is "?L = ?R") |
223 proof - |
343 proof - |
224 { fix n |
344 { fix n |
225 assume "n \<in> ?L" |
345 assume h: "n \<in> ?L" |
226 with subtree_nodeE[OF this] |
346 hence "n \<in> ?R" |
227 obtain th' where "n = Th th'" "Th th' \<in> subtree (tRAG s) (Th th)" by auto |
347 by (smt mem_Collect_eq subsetCE subtree_def subtree_nodeE tRAG_subtree_RAG) |
228 with tRAG_subtree_RAG[of s "Th th"] |
|
229 have "n \<in> ?R" by auto |
|
230 } moreover { |
348 } moreover { |
231 fix n |
349 fix n |
232 assume "n \<in> ?R" |
350 assume "n \<in> ?R" |
233 then obtain th' where h: "n = Th th'" "(Th th', Th th) \<in> (RAG s)^*" |
351 then obtain th' where h: "n = Th th'" "(Th th', Th th) \<in> (RAG s)^*" |
234 by (auto simp:subtree_def) |
352 by (auto simp:subtree_def) |
235 from star_rpath[OF this(2)] |
353 from rtranclD[OF this(2)] |
236 obtain xs where "rpath (RAG s) (Th th') xs (Th th)" by auto |
354 have "n \<in> ?L" |
237 hence "Th th' \<in> subtree (tRAG s) (Th th)" |
355 proof |
238 proof(induct xs arbitrary:th' th rule:length_induct) |
356 assume "Th th' \<noteq> Th th \<and> (Th th', Th th) \<in> (RAG s)\<^sup>+" |
239 case (1 xs th' th) |
357 with h have "n \<in> {Th th' | th'. (Th th', Th th) \<in> (RAG s)^+}" by auto |
240 show ?case |
358 thus ?thesis using subtree_def tRAG_trancl_eq by fastforce |
241 proof(cases xs) |
359 qed (insert h, auto simp:subtree_def) |
242 case Nil |
|
243 from rpath_nilE[OF 1(2)[unfolded this]] |
|
244 have "th' = th" by auto |
|
245 thus ?thesis by (auto simp:subtree_def) |
|
246 next |
|
247 case (Cons x1 xs1) note Cons1 = Cons |
|
248 show ?thesis |
|
249 proof(cases "xs1") |
|
250 case Nil |
|
251 from 1(2)[unfolded Cons[unfolded this]] |
|
252 have rp: "rpath (RAG s) (Th th') [x1] (Th th)" . |
|
253 hence "(Th th', x1) \<in> (RAG s)" by (cases, simp) |
|
254 then obtain cs where "x1 = Cs cs" |
|
255 by (unfold s_RAG_def, auto) |
|
256 from rpath_nnl_lastE[OF rp[unfolded this]] |
|
257 show ?thesis by auto |
|
258 next |
|
259 case (Cons x2 xs2) |
|
260 from 1(2)[unfolded Cons1[unfolded this]] |
|
261 have rp: "rpath (RAG s) (Th th') (x1 # x2 # xs2) (Th th)" . |
|
262 from rpath_edges_on[OF this] |
|
263 have eds: "edges_on (Th th' # x1 # x2 # xs2) \<subseteq> RAG s" . |
|
264 have "(Th th', x1) \<in> edges_on (Th th' # x1 # x2 # xs2)" |
|
265 by (simp add: edges_on_unfold) |
|
266 with eds have rg1: "(Th th', x1) \<in> RAG s" by auto |
|
267 then obtain cs1 where eq_x1: "x1 = Cs cs1" by (unfold s_RAG_def, auto) |
|
268 have "(x1, x2) \<in> edges_on (Th th' # x1 # x2 # xs2)" |
|
269 by (simp add: edges_on_unfold) |
|
270 from this eds |
|
271 have rg2: "(x1, x2) \<in> RAG s" by auto |
|
272 from this[unfolded eq_x1] |
|
273 obtain th1 where eq_x2: "x2 = Th th1" by (unfold s_RAG_def, auto) |
|
274 from rp have "rpath (RAG s) x2 xs2 (Th th)" |
|
275 by (elim rpath_ConsE, simp) |
|
276 from this[unfolded eq_x2] have rp': "rpath (RAG s) (Th th1) xs2 (Th th)" . |
|
277 from 1(1)[rule_format, OF _ this, unfolded Cons1 Cons] |
|
278 have "Th th1 \<in> subtree (tRAG s) (Th th)" by simp |
|
279 moreover have "(Th th', Th th1) \<in> (tRAG s)^*" |
|
280 proof - |
|
281 from rg1[unfolded eq_x1] rg2[unfolded eq_x1 eq_x2] |
|
282 show ?thesis by (unfold RAG_split tRAG_def wRAG_def hRAG_def, auto) |
|
283 qed |
|
284 ultimately show ?thesis by (auto simp:subtree_def) |
|
285 qed |
|
286 qed |
|
287 qed |
|
288 from this[folded h(1)] |
|
289 have "n \<in> ?L" . |
|
290 } ultimately show ?thesis by auto |
360 } ultimately show ?thesis by auto |
291 qed |
361 qed |
292 |
362 |
293 lemma threads_set_eq: |
363 lemma threads_set_eq: |
294 "the_thread ` (subtree (tRAG s) (Th th)) = |
364 "the_thread ` (subtree (tRAG s) (Th th)) = |
319 obtain th where eq_a: "a = Th th" by auto |
389 obtain th where eq_a: "a = Th th" by auto |
320 show "cp_gen s a = (cp s \<circ> the_thread) a" |
390 show "cp_gen s a = (cp s \<circ> the_thread) a" |
321 by (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp) |
391 by (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp) |
322 qed |
392 qed |
323 |
393 |
324 locale valid_trace = |
|
325 fixes s |
|
326 assumes vt : "vt s" |
|
327 |
394 |
328 context valid_trace |
395 context valid_trace |
329 begin |
396 begin |
397 |
|
398 lemma RAG_threads: |
|
399 assumes "(Th th) \<in> Field (RAG s)" |
|
400 shows "th \<in> threads s" |
|
401 using assms |
|
402 by (metis Field_def UnE dm_RAG_threads range_in vt) |
|
403 |
|
330 |
404 |
331 lemma readys_root: |
405 lemma readys_root: |
332 assumes "th \<in> readys s" |
406 assumes "th \<in> readys s" |
333 shows "root (RAG s) (Th th)" |
407 shows "root (RAG s) (Th th)" |
334 proof - |
408 proof - |
363 lemma not_in_thread_isolated: |
437 lemma not_in_thread_isolated: |
364 assumes "th \<notin> threads s" |
438 assumes "th \<notin> threads s" |
365 shows "(Th th) \<notin> Field (RAG s)" |
439 shows "(Th th) \<notin> Field (RAG s)" |
366 proof |
440 proof |
367 assume "(Th th) \<in> Field (RAG s)" |
441 assume "(Th th) \<in> Field (RAG s)" |
368 with dm_RAG_threads[OF vt] and range_in[OF vt] assms |
442 with dm_RAG_threads and range_in assms |
369 show False by (unfold Field_def, blast) |
443 show False by (unfold Field_def, blast) |
370 qed |
444 qed |
371 |
445 |
372 lemma wf_RAG: "wf (RAG s)" |
446 lemma wf_RAG: "wf (RAG s)" |
373 proof(rule finite_acyclic_wf) |
447 proof(rule finite_acyclic_wf) |
374 from finite_RAG[OF vt] show "finite (RAG s)" . |
448 from finite_RAG show "finite (RAG s)" . |
375 next |
449 next |
376 from acyclic_RAG[OF vt] show "acyclic (RAG s)" . |
450 from acyclic_RAG show "acyclic (RAG s)" . |
377 qed |
451 qed |
378 |
452 |
379 lemma sgv_wRAG: "single_valued (wRAG s)" |
453 lemma sgv_wRAG: "single_valued (wRAG s)" |
380 using waiting_unique[OF vt] |
454 using waiting_unique |
381 by (unfold single_valued_def wRAG_def, auto) |
455 by (unfold single_valued_def wRAG_def, auto) |
382 |
456 |
383 lemma sgv_hRAG: "single_valued (hRAG s)" |
457 lemma sgv_hRAG: "single_valued (hRAG s)" |
384 using holding_unique |
458 using holding_unique |
385 by (unfold single_valued_def hRAG_def, auto) |
459 by (unfold single_valued_def hRAG_def, auto) |
388 by (unfold tRAG_def, rule single_valued_relcomp, |
462 by (unfold tRAG_def, rule single_valued_relcomp, |
389 insert sgv_wRAG sgv_hRAG, auto) |
463 insert sgv_wRAG sgv_hRAG, auto) |
390 |
464 |
391 lemma acyclic_tRAG: "acyclic (tRAG s)" |
465 lemma acyclic_tRAG: "acyclic (tRAG s)" |
392 proof(unfold tRAG_def, rule acyclic_compose) |
466 proof(unfold tRAG_def, rule acyclic_compose) |
393 show "acyclic (RAG s)" using acyclic_RAG[OF vt] . |
467 show "acyclic (RAG s)" using acyclic_RAG . |
394 next |
468 next |
395 show "wRAG s \<subseteq> RAG s" unfolding RAG_split by auto |
469 show "wRAG s \<subseteq> RAG s" unfolding RAG_split by auto |
396 next |
470 next |
397 show "hRAG s \<subseteq> RAG s" unfolding RAG_split by auto |
471 show "hRAG s \<subseteq> RAG s" unfolding RAG_split by auto |
398 qed |
472 qed |
399 |
473 |
400 lemma sgv_RAG: "single_valued (RAG s)" |
474 lemma sgv_RAG: "single_valued (RAG s)" |
401 using unique_RAG[OF vt] by (auto simp:single_valued_def) |
475 using unique_RAG by (auto simp:single_valued_def) |
402 |
476 |
403 lemma rtree_RAG: "rtree (RAG s)" |
477 lemma rtree_RAG: "rtree (RAG s)" |
404 using sgv_RAG acyclic_RAG[OF vt] |
478 using sgv_RAG acyclic_RAG |
405 by (unfold rtree_def rtree_axioms_def sgv_def, auto) |
479 by (unfold rtree_def rtree_axioms_def sgv_def, auto) |
480 |
|
406 end |
481 end |
407 |
482 |
408 |
483 |
409 sublocale valid_trace < rtree_RAG: rtree "RAG s" |
484 sublocale valid_trace < rtree_RAG: rtree "RAG s" |
410 proof |
485 proof |
411 show "single_valued (RAG s)" |
486 show "single_valued (RAG s)" |
412 apply (intro_locales) |
487 apply (intro_locales) |
413 by (unfold single_valued_def, |
488 by (unfold single_valued_def, |
414 auto intro:unique_RAG[OF vt]) |
489 auto intro:unique_RAG) |
415 |
490 |
416 show "acyclic (RAG s)" |
491 show "acyclic (RAG s)" |
417 by (rule acyclic_RAG[OF vt]) |
492 by (rule acyclic_RAG) |
418 qed |
493 qed |
419 |
494 |
420 sublocale valid_trace < rtree_s: rtree "tRAG s" |
495 sublocale valid_trace < rtree_s: rtree "tRAG s" |
421 proof(unfold_locales) |
496 proof(unfold_locales) |
422 from sgv_tRAG show "single_valued (tRAG s)" . |
497 from sgv_tRAG show "single_valued (tRAG s)" . |
426 |
501 |
427 sublocale valid_trace < fsbtRAGs : fsubtree "RAG s" |
502 sublocale valid_trace < fsbtRAGs : fsubtree "RAG s" |
428 proof - |
503 proof - |
429 show "fsubtree (RAG s)" |
504 show "fsubtree (RAG s)" |
430 proof(intro_locales) |
505 proof(intro_locales) |
431 show "fbranch (RAG s)" using finite_fbranchI[OF finite_RAG[OF vt]] . |
506 show "fbranch (RAG s)" using finite_fbranchI[OF finite_RAG] . |
432 next |
507 next |
433 show "fsubtree_axioms (RAG s)" |
508 show "fsubtree_axioms (RAG s)" |
434 proof(unfold fsubtree_axioms_def) |
509 proof(unfold fsubtree_axioms_def) |
435 find_theorems wf RAG |
510 find_theorems wf RAG |
436 from wf_RAG show "wf (RAG s)" . |
511 from wf_RAG show "wf (RAG s)" . |
444 proof - |
519 proof - |
445 have "fbranch (tRAG s)" |
520 have "fbranch (tRAG s)" |
446 proof(unfold tRAG_def, rule fbranch_compose) |
521 proof(unfold tRAG_def, rule fbranch_compose) |
447 show "fbranch (wRAG s)" |
522 show "fbranch (wRAG s)" |
448 proof(rule finite_fbranchI) |
523 proof(rule finite_fbranchI) |
449 from finite_RAG[OF vt] show "finite (wRAG s)" |
524 from finite_RAG show "finite (wRAG s)" |
450 by (unfold RAG_split, auto) |
525 by (unfold RAG_split, auto) |
451 qed |
526 qed |
452 next |
527 next |
453 show "fbranch (hRAG s)" |
528 show "fbranch (hRAG s)" |
454 proof(rule finite_fbranchI) |
529 proof(rule finite_fbranchI) |
455 from finite_RAG[OF vt] |
530 from finite_RAG |
456 show "finite (hRAG s)" by (unfold RAG_split, auto) |
531 show "finite (hRAG s)" by (unfold RAG_split, auto) |
457 qed |
532 qed |
458 qed |
533 qed |
459 moreover have "wf (tRAG s)" |
534 moreover have "wf (tRAG s)" |
460 proof(rule wf_subset) |
535 proof(rule wf_subset) |
590 "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto |
665 "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto |
591 thus ?thesis |
666 thus ?thesis |
592 by (unfold cs_holding_def, auto) |
667 by (unfold cs_holding_def, auto) |
593 qed |
668 qed |
594 |
669 |
670 context valid_trace |
|
671 begin |
|
672 |
|
595 lemma next_th_waiting: |
673 lemma next_th_waiting: |
596 assumes vt: "vt s" |
674 assumes nxt: "next_th s th cs th'" |
597 and nxt: "next_th s th cs th'" |
|
598 shows "waiting (wq s) th' cs" |
675 shows "waiting (wq s) th' cs" |
599 proof - |
676 proof - |
600 from nxt[unfolded next_th_def] |
677 from nxt[unfolded next_th_def] |
601 obtain rest where h: "wq s cs = th # rest" |
678 obtain rest where h: "wq s cs = th # rest" |
602 "rest \<noteq> []" |
679 "rest \<noteq> []" |
603 "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto |
680 "th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto |
604 from wq_distinct[OF vt, of cs, unfolded h] |
681 from wq_distinct[of cs, unfolded h] |
605 have dst: "distinct (th # rest)" . |
682 have dst: "distinct (th # rest)" . |
606 have in_rest: "th' \<in> set rest" |
683 have in_rest: "th' \<in> set rest" |
607 proof(unfold h, rule someI2) |
684 proof(unfold h, rule someI2) |
608 show "distinct rest \<and> set rest = set rest" using dst by auto |
685 show "distinct rest \<and> set rest = set rest" using dst by auto |
609 next |
686 next |
616 by (unfold h(1), insert in_rest dst, auto) |
693 by (unfold h(1), insert in_rest dst, auto) |
617 ultimately show ?thesis by (auto simp:cs_waiting_def) |
694 ultimately show ?thesis by (auto simp:cs_waiting_def) |
618 qed |
695 qed |
619 |
696 |
620 lemma next_th_RAG: |
697 lemma next_th_RAG: |
621 assumes vt: "vt s" |
698 assumes nxt: "next_th (s::event list) th cs th'" |
622 and nxt: "next_th s th cs th'" |
|
623 shows "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s" |
699 shows "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s" |
624 using assms next_th_holding next_th_waiting |
700 using vt assms next_th_holding next_th_waiting |
625 by (unfold s_RAG_def, simp) |
701 by (unfold s_RAG_def, simp) |
702 |
|
703 end |
|
626 |
704 |
627 -- {* A useless definition *} |
705 -- {* A useless definition *} |
628 definition cps:: "state \<Rightarrow> (thread \<times> precedence) set" |
706 definition cps:: "state \<Rightarrow> (thread \<times> precedence) set" |
629 where "cps s = {(th, cp s th) | th . th \<in> threads s}" |
707 where "cps s = {(th, cp s th) | th . th \<in> threads s}" |
630 |
708 |
903 | |
981 | |
904 th7 ----| |
982 th7 ----| |
905 *) |
983 *) |
906 |
984 |
907 lemma sub_RAGs': "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s'" |
985 lemma sub_RAGs': "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s'" |
908 using next_th_RAG[OF vat_s'.vt nt] . |
986 using next_th_RAG[OF nt] . |
909 |
987 |
910 lemma ancestors_th': |
988 lemma ancestors_th': |
911 "ancestors (RAG s') (Th th') = {Th th, Cs cs}" |
989 "ancestors (RAG s') (Th th') = {Th th, Cs cs}" |
912 proof - |
990 proof - |
913 have "ancestors (RAG s') (Th th') = ancestors (RAG s') (Cs cs) \<union> {Cs cs}" |
991 have "ancestors (RAG s') (Th th') = ancestors (RAG s') (Cs cs) \<union> {Cs cs}" |
1169 thus ?thesis by (unfold RAG_split, auto) |
1247 thus ?thesis by (unfold RAG_split, auto) |
1170 qed |
1248 qed |
1171 |
1249 |
1172 lemma tRAG_s: |
1250 lemma tRAG_s: |
1173 "tRAG s = tRAG s' \<union> {(Th th, Th th')}" |
1251 "tRAG s = tRAG s' \<union> {(Th th, Th th')}" |
1174 using RAG_tRAG_transfer[OF step_back_vt[OF vt_s[unfolded s_def]] RAG_s cs_held] . |
1252 using RAG_tRAG_transfer[OF RAG_s cs_held] . |
1175 |
1253 |
1176 lemma cp_kept: |
1254 lemma cp_kept: |
1177 assumes "Th th'' \<notin> ancestors (tRAG s) (Th th)" |
1255 assumes "Th th'' \<notin> ancestors (tRAG s) (Th th)" |
1178 shows "cp s th'' = cp s' th''" |
1256 shows "cp s th'' = cp s' th''" |
1179 proof - |
1257 proof - |