author | zhangx |
Wed, 03 Feb 2016 21:51:57 +0800 | |
changeset 104 | 43482ab31341 |
parent 103 | d5e9653fbf19 |
parent 97 | c7ba70dc49bd |
child 106 | 5454387e42ce |
permissions | -rw-r--r-- |
93
524bd3caa6b6
The overwriten original .thy files are working now. The ones in last revision aren't.
zhangx
parents:
92
diff
changeset
|
1 |
theory PIPBasics |
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
63
diff
changeset
|
2 |
imports PIPDefs |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3 |
begin |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4 |
|
99 | 5 |
section {* Generic aulxiliary lemmas *} |
6 |
||
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7 |
lemma f_image_eq: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
8 |
assumes h: "\<And> a. a \<in> A \<Longrightarrow> f a = g a" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
9 |
shows "f ` A = g ` A" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
10 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
11 |
show "f ` A \<subseteq> g ` A" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
12 |
by(rule image_subsetI, auto intro:h) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
13 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
14 |
show "g ` A \<subseteq> f ` A" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
15 |
by (rule image_subsetI, auto intro:h[symmetric]) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
16 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
17 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
18 |
lemma Max_fg_mono: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
19 |
assumes "finite A" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
20 |
and "\<forall> a \<in> A. f a \<le> g a" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
21 |
shows "Max (f ` A) \<le> Max (g ` A)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
22 |
proof(cases "A = {}") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
23 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
24 |
thus ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
25 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
26 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
27 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
28 |
proof(rule Max.boundedI) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
29 |
from assms show "finite (f ` A)" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
30 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
31 |
from False show "f ` A \<noteq> {}" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
32 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
33 |
fix fa |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
34 |
assume "fa \<in> f ` A" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
35 |
then obtain a where h_fa: "a \<in> A" "fa = f a" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
36 |
show "fa \<le> Max (g ` A)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
37 |
proof(rule Max_ge_iff[THEN iffD2]) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
38 |
from assms show "finite (g ` A)" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
39 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
40 |
from False show "g ` A \<noteq> {}" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
41 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
42 |
from h_fa have "g a \<in> g ` A" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
43 |
moreover have "fa \<le> g a" using h_fa assms(2) by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
44 |
ultimately show "\<exists>a\<in>g ` A. fa \<le> a" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
45 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
46 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
47 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
48 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
49 |
lemma Max_f_mono: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
50 |
assumes seq: "A \<subseteq> B" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
51 |
and np: "A \<noteq> {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
52 |
and fnt: "finite B" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
53 |
shows "Max (f ` A) \<le> Max (f ` B)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
54 |
proof(rule Max_mono) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
55 |
from seq show "f ` A \<subseteq> f ` B" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
56 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
57 |
from np show "f ` A \<noteq> {}" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
58 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
59 |
from fnt and seq show "finite (f ` B)" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
60 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
61 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
62 |
lemma Max_UNION: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
63 |
assumes "finite A" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
64 |
and "A \<noteq> {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
65 |
and "\<forall> M \<in> f ` A. finite M" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
66 |
and "\<forall> M \<in> f ` A. M \<noteq> {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
67 |
shows "Max (\<Union>x\<in> A. f x) = Max (Max ` f ` A)" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
68 |
using assms[simp] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
69 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
70 |
have "?L = Max (\<Union>(f ` A))" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
71 |
by (fold Union_image_eq, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
72 |
also have "... = ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
73 |
by (subst Max_Union, simp+) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
74 |
finally show ?thesis . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
75 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
76 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
77 |
lemma max_Max_eq: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
78 |
assumes "finite A" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
79 |
and "A \<noteq> {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
80 |
and "x = y" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
81 |
shows "max x (Max A) = Max ({y} \<union> A)" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
82 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
83 |
have "?R = Max (insert y A)" by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
84 |
also from assms have "... = ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
85 |
by (subst Max.insert, simp+) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
86 |
finally show ?thesis by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
87 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
88 |
|
100 | 89 |
lemma rel_eqI: |
90 |
assumes "\<And> x y. (x,y) \<in> A \<Longrightarrow> (x,y) \<in> B" |
|
91 |
and "\<And> x y. (x,y) \<in> B \<Longrightarrow> (x, y) \<in> A" |
|
92 |
shows "A = B" |
|
93 |
using assms by auto |
|
94 |
||
99 | 95 |
section {* Lemmas do not depend on trace validity *} |
96 |
||
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
97 |
lemma birth_time_lt: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
98 |
assumes "s \<noteq> []" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
99 |
shows "last_set th s < length s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
100 |
using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
101 |
proof(induct s) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
102 |
case (Cons a s) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
103 |
show ?case |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
104 |
proof(cases "s \<noteq> []") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
105 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
106 |
thus ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
107 |
by (cases a, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
108 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
109 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
110 |
show ?thesis using Cons(1)[OF True] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
111 |
by (cases a, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
112 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
113 |
qed simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
114 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
115 |
lemma th_in_ne: "th \<in> threads s \<Longrightarrow> s \<noteq> []" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
116 |
by (induct s, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
117 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
118 |
lemma preced_tm_lt: "th \<in> threads s \<Longrightarrow> preced th s = Prc x y \<Longrightarrow> y < length s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
119 |
by (drule_tac th_in_ne, unfold preced_def, auto intro: birth_time_lt) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
120 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
121 |
lemma eq_RAG: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
122 |
"RAG (wq s) = RAG s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
123 |
by (unfold cs_RAG_def s_RAG_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
124 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
125 |
lemma waiting_holding: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
126 |
assumes "waiting (s::state) th cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
127 |
obtains th' where "holding s th' cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
128 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
129 |
from assms[unfolded s_waiting_def, folded wq_def] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
130 |
obtain th' where "th' \<in> set (wq s cs)" "th' = hd (wq s cs)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
131 |
by (metis empty_iff hd_in_set list.set(1)) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
132 |
hence "holding s th' cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
133 |
by (unfold s_holding_def, fold wq_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
134 |
from that[OF this] show ?thesis . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
135 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
136 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
137 |
lemma cp_eq_cpreced: "cp s th = cpreced (wq s) s th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
138 |
unfolding cp_def wq_def |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
139 |
apply(induct s rule: schs.induct) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
140 |
apply(simp add: Let_def cpreced_initial) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
141 |
apply(simp add: Let_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
142 |
apply(simp add: Let_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
143 |
apply(simp add: Let_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
144 |
apply(subst (2) schs.simps) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
145 |
apply(simp add: Let_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
146 |
apply(subst (2) schs.simps) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
147 |
apply(simp add: Let_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
148 |
done |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
149 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
150 |
lemma cp_alt_def: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
151 |
"cp s th = |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
152 |
Max ((the_preced s) ` {th'. Th th' \<in> (subtree (RAG s) (Th th))})" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
153 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
154 |
have "Max (the_preced s ` ({th} \<union> dependants (wq s) th)) = |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
155 |
Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
156 |
(is "Max (_ ` ?L) = Max (_ ` ?R)") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
157 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
158 |
have "?L = ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
159 |
by (auto dest:rtranclD simp:cs_dependants_def cs_RAG_def s_RAG_def subtree_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
160 |
thus ?thesis by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
161 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
162 |
thus ?thesis by (unfold cp_eq_cpreced cpreced_def, fold the_preced_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
163 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
164 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
165 |
lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
166 |
by (unfold s_RAG_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
167 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
168 |
lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
169 |
by (unfold s_waiting_def cs_waiting_def wq_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
170 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
171 |
lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
172 |
by (unfold s_holding_def wq_def cs_holding_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
173 |
|
101 | 174 |
lemma children_RAG_alt_def: |
175 |
"children (RAG (s::state)) (Th th) = Cs ` {cs. holding s th cs}" |
|
176 |
by (unfold s_RAG_def, auto simp:children_def holding_eq) |
|
177 |
||
178 |
lemma holdents_alt_def: |
|
179 |
"holdents s th = the_cs ` (children (RAG (s::state)) (Th th))" |
|
180 |
by (unfold children_RAG_alt_def holdents_def, simp add: image_image) |
|
181 |
||
182 |
lemma cntCS_alt_def: |
|
183 |
"cntCS s th = card (children (RAG s) (Th th))" |
|
184 |
apply (unfold children_RAG_alt_def cntCS_def holdents_def) |
|
185 |
by (rule card_image[symmetric], auto simp:inj_on_def) |
|
186 |
||
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
187 |
lemma runing_ready: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
188 |
shows "runing s \<subseteq> readys s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
189 |
unfolding runing_def readys_def |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
190 |
by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
191 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
192 |
lemma readys_threads: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
193 |
shows "readys s \<subseteq> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
194 |
unfolding readys_def |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
195 |
by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
196 |
|
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
197 |
lemma wq_v_neq [simp]: |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
198 |
"cs \<noteq> cs' \<Longrightarrow> wq (V thread cs#s) cs' = wq s cs'" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
199 |
by (auto simp:wq_def Let_def cp_def split:list.splits) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
200 |
|
68 | 201 |
lemma runing_head: |
202 |
assumes "th \<in> runing s" |
|
203 |
and "th \<in> set (wq_fun (schs s) cs)" |
|
204 |
shows "th = hd (wq_fun (schs s) cs)" |
|
205 |
using assms |
|
206 |
by (simp add:runing_def readys_def s_waiting_def wq_def) |
|
207 |
||
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
208 |
lemma runing_wqE: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
209 |
assumes "th \<in> runing s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
210 |
and "th \<in> set (wq s cs)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
211 |
obtains rest where "wq s cs = th#rest" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
212 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
213 |
from assms(2) obtain th' rest where eq_wq: "wq s cs = th'#rest" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
214 |
by (meson list.set_cases) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
215 |
have "th' = th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
216 |
proof(rule ccontr) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
217 |
assume "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
218 |
hence "th \<noteq> hd (wq s cs)" using eq_wq by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
219 |
with assms(2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
220 |
have "waiting s th cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
221 |
by (unfold s_waiting_def, fold wq_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
222 |
with assms show False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
223 |
by (unfold runing_def readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
224 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
225 |
with eq_wq that show ?thesis by metis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
226 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
227 |
|
100 | 228 |
lemma isP_E: |
229 |
assumes "isP e" |
|
230 |
obtains cs where "e = P (actor e) cs" |
|
231 |
using assms by (cases e, auto) |
|
232 |
||
233 |
lemma isV_E: |
|
234 |
assumes "isV e" |
|
235 |
obtains cs where "e = V (actor e) cs" |
|
236 |
using assms by (cases e, auto) |
|
237 |
||
238 |
||
239 |
text {* |
|
240 |
Every thread can only be blocked on one critical resource, |
|
241 |
symmetrically, every critical resource can only be held by one thread. |
|
242 |
This fact is much more easier according to our definition. |
|
243 |
*} |
|
244 |
lemma held_unique: |
|
245 |
assumes "holding (s::event list) th1 cs" |
|
246 |
and "holding s th2 cs" |
|
247 |
shows "th1 = th2" |
|
248 |
by (insert assms, unfold s_holding_def, auto) |
|
249 |
||
250 |
lemma last_set_lt: "th \<in> threads s \<Longrightarrow> last_set th s < length s" |
|
251 |
apply (induct s, auto) |
|
252 |
by (case_tac a, auto split:if_splits) |
|
253 |
||
254 |
lemma last_set_unique: |
|
255 |
"\<lbrakk>last_set th1 s = last_set th2 s; th1 \<in> threads s; th2 \<in> threads s\<rbrakk> |
|
256 |
\<Longrightarrow> th1 = th2" |
|
257 |
apply (induct s, auto) |
|
258 |
by (case_tac a, auto split:if_splits dest:last_set_lt) |
|
259 |
||
260 |
lemma preced_unique : |
|
261 |
assumes pcd_eq: "preced th1 s = preced th2 s" |
|
262 |
and th_in1: "th1 \<in> threads s" |
|
263 |
and th_in2: " th2 \<in> threads s" |
|
264 |
shows "th1 = th2" |
|
265 |
proof - |
|
266 |
from pcd_eq have "last_set th1 s = last_set th2 s" by (simp add:preced_def) |
|
267 |
from last_set_unique [OF this th_in1 th_in2] |
|
268 |
show ?thesis . |
|
269 |
qed |
|
270 |
||
271 |
lemma preced_linorder: |
|
272 |
assumes neq_12: "th1 \<noteq> th2" |
|
273 |
and th_in1: "th1 \<in> threads s" |
|
274 |
and th_in2: " th2 \<in> threads s" |
|
275 |
shows "preced th1 s < preced th2 s \<or> preced th1 s > preced th2 s" |
|
276 |
proof - |
|
277 |
from preced_unique [OF _ th_in1 th_in2] and neq_12 |
|
278 |
have "preced th1 s \<noteq> preced th2 s" by auto |
|
279 |
thus ?thesis by auto |
|
280 |
qed |
|
281 |
||
282 |
lemma in_RAG_E: |
|
283 |
assumes "(n1, n2) \<in> RAG (s::state)" |
|
284 |
obtains (waiting) th cs where "n1 = Th th" "n2 = Cs cs" "waiting s th cs" |
|
285 |
| (holding) th cs where "n1 = Cs cs" "n2 = Th th" "holding s th cs" |
|
286 |
using assms[unfolded s_RAG_def, folded waiting_eq holding_eq] |
|
287 |
by auto |
|
288 |
||
101 | 289 |
lemma count_rec1 [simp]: |
290 |
assumes "Q e" |
|
291 |
shows "count Q (e#es) = Suc (count Q es)" |
|
292 |
using assms |
|
293 |
by (unfold count_def, auto) |
|
294 |
||
295 |
lemma count_rec2 [simp]: |
|
296 |
assumes "\<not>Q e" |
|
297 |
shows "count Q (e#es) = (count Q es)" |
|
298 |
using assms |
|
299 |
by (unfold count_def, auto) |
|
300 |
||
301 |
lemma count_rec3 [simp]: |
|
302 |
shows "count Q [] = 0" |
|
303 |
by (unfold count_def, auto) |
|
304 |
||
305 |
lemma cntP_simp1[simp]: |
|
306 |
"cntP (P th cs'#s) th = cntP s th + 1" |
|
307 |
by (unfold cntP_def, simp) |
|
308 |
||
309 |
lemma cntP_simp2[simp]: |
|
310 |
assumes "th' \<noteq> th" |
|
311 |
shows "cntP (P th cs'#s) th' = cntP s th'" |
|
312 |
using assms |
|
313 |
by (unfold cntP_def, simp) |
|
314 |
||
315 |
lemma cntP_simp3[simp]: |
|
316 |
assumes "\<not> isP e" |
|
317 |
shows "cntP (e#s) th' = cntP s th'" |
|
318 |
using assms |
|
319 |
by (unfold cntP_def, cases e, simp+) |
|
320 |
||
321 |
lemma cntV_simp1[simp]: |
|
322 |
"cntV (V th cs'#s) th = cntV s th + 1" |
|
323 |
by (unfold cntV_def, simp) |
|
324 |
||
325 |
lemma cntV_simp2[simp]: |
|
326 |
assumes "th' \<noteq> th" |
|
327 |
shows "cntV (V th cs'#s) th' = cntV s th'" |
|
328 |
using assms |
|
329 |
by (unfold cntV_def, simp) |
|
330 |
||
331 |
lemma cntV_simp3[simp]: |
|
332 |
assumes "\<not> isV e" |
|
333 |
shows "cntV (e#s) th' = cntV s th'" |
|
334 |
using assms |
|
335 |
by (unfold cntV_def, cases e, simp+) |
|
336 |
||
337 |
lemma cntP_diff_inv: |
|
338 |
assumes "cntP (e#s) th \<noteq> cntP s th" |
|
339 |
shows "isP e \<and> actor e = th" |
|
340 |
proof(cases e) |
|
341 |
case (P th' pty) |
|
342 |
show ?thesis |
|
343 |
by (cases "(\<lambda>e. \<exists>cs. e = P th cs) (P th' pty)", |
|
344 |
insert assms P, auto simp:cntP_def) |
|
345 |
qed (insert assms, auto simp:cntP_def) |
|
346 |
||
347 |
lemma cntV_diff_inv: |
|
348 |
assumes "cntV (e#s) th \<noteq> cntV s th" |
|
349 |
shows "isV e \<and> actor e = th" |
|
350 |
proof(cases e) |
|
351 |
case (V th' pty) |
|
352 |
show ?thesis |
|
353 |
by (cases "(\<lambda>e. \<exists>cs. e = V th cs) (V th' pty)", |
|
354 |
insert assms V, auto simp:cntV_def) |
|
355 |
qed (insert assms, auto simp:cntV_def) |
|
356 |
||
357 |
lemma eq_dependants: "dependants (wq s) = dependants s" |
|
358 |
by (simp add: s_dependants_abv wq_def) |
|
359 |
||
360 |
lemma inj_the_preced: |
|
361 |
"inj_on (the_preced s) (threads s)" |
|
362 |
by (metis inj_onI preced_unique the_preced_def) |
|
363 |
||
103
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
364 |
lemma holding_next_thI: |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
365 |
assumes "holding s th cs" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
366 |
and "length (wq s cs) > 1" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
367 |
obtains th' where "next_th s th cs th'" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
368 |
proof - |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
369 |
from assms(1)[folded holding_eq, unfolded cs_holding_def] |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
370 |
have " th \<in> set (wq s cs) \<and> th = hd (wq s cs)" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
371 |
by (unfold s_holding_def, fold wq_def, auto) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
372 |
then obtain rest where h1: "wq s cs = th#rest" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
373 |
by (cases "wq s cs", auto) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
374 |
with assms(2) have h2: "rest \<noteq> []" by auto |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
375 |
let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
376 |
have "next_th s th cs ?th'" using h1(1) h2 |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
377 |
by (unfold next_th_def, auto) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
378 |
from that[OF this] show ?thesis . |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
379 |
qed |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
380 |
|
99 | 381 |
(* ccc *) |
382 |
||
383 |
section {* Locales used to investigate the execution of PIP *} |
|
384 |
||
385 |
text {* |
|
386 |
The following locale @{text valid_trace} is used to constrain the |
|
387 |
trace to be valid. All properties hold for valid traces are |
|
388 |
derived under this locale. |
|
389 |
*} |
|
104 | 390 |
======= |
391 |
>>>>>>> other |
|
63 | 392 |
locale valid_trace = |
393 |
fixes s |
|
394 |
assumes vt : "vt s" |
|
395 |
||
99 | 396 |
text {* |
397 |
The following locale @{text valid_trace_e} describes |
|
398 |
the valid extension of a valid trace. The event @{text "e"} |
|
399 |
represents an event in the system, which corresponds |
|
400 |
to a one step operation of the PIP protocol. |
|
401 |
It is required that @{text "e"} is an event eligible to happen |
|
402 |
under state @{text "s"}, which is already required to be valid |
|
403 |
by the parent locale @{text "valid_trace"}. |
|
404 |
||
405 |
This locale is used to investigate one step execution of PIP, |
|
406 |
properties concerning the effects of @{text "e"}'s execution, |
|
407 |
for example, how the values of observation functions are changed, |
|
408 |
or how desirable properties are kept invariant, are derived |
|
409 |
under this locale. The state before execution is @{text "s"}, while |
|
410 |
the state after execution is @{text "e#s"}. Therefore, the lemmas |
|
411 |
derived usually relate observations on @{text "e#s"} to those |
|
412 |
on @{text "s"}. |
|
413 |
*} |
|
414 |
||
63 | 415 |
locale valid_trace_e = valid_trace + |
416 |
fixes e |
|
417 |
assumes vt_e: "vt (e#s)" |
|
418 |
begin |
|
419 |
||
99 | 420 |
text {* |
421 |
The following lemma shows that @{text "e"} must be a |
|
422 |
eligible event (or a valid step) to be taken under |
|
423 |
the state represented by @{text "s"}. |
|
424 |
*} |
|
63 | 425 |
lemma pip_e: "PIP s e" |
426 |
using vt_e by (cases, simp) |
|
427 |
||
428 |
end |
|
429 |
||
104 | 430 |
<<<<<<< local |
99 | 431 |
text {* |
432 |
Because @{term "e#s"} is also a valid trace, properties |
|
433 |
derived for valid trace @{term s} also hold on @{term "e#s"}. |
|
434 |
*} |
|
435 |
sublocale valid_trace_e < vat_es!: valid_trace "e#s" |
|
436 |
using vt_e |
|
437 |
by (unfold_locales, simp) |
|
438 |
||
439 |
text {* |
|
440 |
For each specific event (or operation), there is a sublocale |
|
441 |
further constraining that the event @{text e} to be that |
|
442 |
particular event. |
|
443 |
||
444 |
For example, the following |
|
445 |
locale @{text "valid_trace_create"} is the sublocale for |
|
446 |
event @{term "Create"}: |
|
447 |
*} |
|
448 |
locale valid_trace_create = valid_trace_e + |
|
449 |
fixes th prio |
|
450 |
assumes is_create: "e = Create th prio" |
|
451 |
||
452 |
locale valid_trace_exit = valid_trace_e + |
|
453 |
fixes th |
|
454 |
assumes is_exit: "e = Exit th" |
|
455 |
||
456 |
locale valid_trace_p = valid_trace_e + |
|
457 |
fixes th cs |
|
458 |
assumes is_p: "e = P th cs" |
|
459 |
||
460 |
text {* |
|
461 |
locale @{text "valid_trace_p"} is divided further into two |
|
462 |
sublocales, namely, @{text "valid_trace_p_h"} |
|
463 |
and @{text "valid_trace_p_w"}. |
|
464 |
*} |
|
465 |
||
466 |
text {* |
|
467 |
The following two sublocales @{text "valid_trace_p_h"} |
|
468 |
and @{text "valid_trace_p_w"} represent two complementary |
|
469 |
cases under @{text "valid_trace_p"}, where |
|
470 |
@{text "valid_trace_p_h"} further constraints that |
|
471 |
@{text "wq s cs = []"}, which means the waiting queue of |
|
472 |
the requested resource @{text "cs"} is empty, in which |
|
473 |
case, the requesting thread @{text "th"} |
|
474 |
will take hold of @{text "cs"}. |
|
475 |
||
476 |
Opposite to @{text "valid_trace_p_h"}, |
|
477 |
@{text "valid_trace_p_w"} constraints that |
|
478 |
@{text "wq s cs \<noteq> []"}, which means the waiting queue of |
|
479 |
the requested resource @{text "cs"} is nonempty, in which |
|
480 |
case, the requesting thread @{text "th"} will be blocked |
|
481 |
on @{text "cs"}: |
|
482 |
||
483 |
Peculiar properties will be derived under respective |
|
484 |
locales. |
|
485 |
*} |
|
486 |
||
487 |
locale valid_trace_p_h = valid_trace_p + |
|
488 |
assumes we: "wq s cs = []" |
|
489 |
||
490 |
locale valid_trace_p_w = valid_trace_p + |
|
491 |
assumes wne: "wq s cs \<noteq> []" |
|
492 |
begin |
|
493 |
||
494 |
text {* |
|
495 |
The following @{text "holder"} designates |
|
496 |
the holder of @{text "cs"} before the @{text "P"}-operation. |
|
497 |
*} |
|
498 |
definition "holder = hd (wq s cs)" |
|
499 |
||
500 |
text {* |
|
501 |
The following @{text "waiters"} designates |
|
502 |
the list of threads waiting for @{text "cs"} |
|
503 |
before the @{text "P"}-operation. |
|
504 |
*} |
|
505 |
definition "waiters = tl (wq s cs)" |
|
506 |
end |
|
507 |
||
508 |
text {* |
|
509 |
@{text "valid_trace_v"} is set for the @{term V}-operation. |
|
510 |
*} |
|
511 |
locale valid_trace_v = valid_trace_e + |
|
512 |
fixes th cs |
|
513 |
assumes is_v: "e = V th cs" |
|
514 |
begin |
|
515 |
-- {* The following @{text "rest"} is the tail of |
|
516 |
waiting queue of the resource @{text "cs"} |
|
517 |
to be released by this @{text "V"}-operation. |
|
518 |
*} |
|
519 |
definition "rest = tl (wq s cs)" |
|
520 |
||
521 |
text {* |
|
522 |
The following @{text "wq'"} is the waiting |
|
523 |
queue of @{term "cs"} |
|
524 |
after the @{text "V"}-operation, which |
|
525 |
is simply a reordering of @{term "rest"}. |
|
526 |
||
527 |
The effect of this reordering needs to be |
|
528 |
understood by two cases: |
|
529 |
\begin{enumerate} |
|
530 |
\item When @{text "rest = []"}, |
|
531 |
the reordering gives rise to an empty list as well, |
|
532 |
which means there is no thread holding or waiting |
|
533 |
for resource @{term "cs"}, therefore, it is free. |
|
534 |
||
535 |
\item When @{text "rest \<noteq> []"}, the effect of |
|
536 |
this reordering is to arbitrarily |
|
537 |
switch one thread in @{term "rest"} to the |
|
538 |
head, which, by definition take over the hold |
|
539 |
of @{term "cs"} and is designated by @{text "taker"} |
|
540 |
in the following sublocale @{text "valid_trace_v_n"}. |
|
541 |
*} |
|
542 |
definition "wq' = (SOME q. distinct q \<and> set q = set rest)" |
|
543 |
||
544 |
text {* |
|
545 |
The following @{text "rest'"} is the tail of the |
|
546 |
waiting queue after the @{text "V"}-operation. |
|
547 |
It plays only auxiliary role to ease reasoning. |
|
548 |
*} |
|
549 |
definition "rest' = tl wq'" |
|
550 |
||
551 |
end |
|
552 |
||
553 |
text {* |
|
554 |
In the following, @{text "valid_trace_v"} is also |
|
555 |
divided into two |
|
556 |
sublocales: when @{text "rest"} is empty (represented |
|
557 |
by @{text "valid_trace_v_e"}), which means, there is no thread waiting |
|
558 |
for @{text "cs"}, therefore, after the @{text "V"}-operation, |
|
559 |
it will become free; otherwise (represented |
|
560 |
by @{text "valid_trace_v_n"}), one thread |
|
561 |
will be picked from those in @{text "rest"} to take |
|
562 |
over @{text "cs"}. |
|
563 |
*} |
|
564 |
||
565 |
locale valid_trace_v_e = valid_trace_v + |
|
566 |
assumes rest_nil: "rest = []" |
|
567 |
||
568 |
locale valid_trace_v_n = valid_trace_v + |
|
569 |
assumes rest_nnl: "rest \<noteq> []" |
|
570 |
begin |
|
571 |
||
572 |
text {* |
|
573 |
The following @{text "taker"} is the thread to |
|
574 |
take over @{text "cs"}. |
|
575 |
*} |
|
576 |
definition "taker = hd wq'" |
|
577 |
||
578 |
end |
|
579 |
||
580 |
||
581 |
locale valid_trace_set = valid_trace_e + |
|
582 |
fixes th prio |
|
583 |
assumes is_set: "e = Set th prio" |
|
584 |
||
585 |
context valid_trace |
|
586 |
begin |
|
587 |
||
588 |
text {* |
|
589 |
Induction rule introduced to easy the |
|
590 |
derivation of properties for valid trace @{term "s"}. |
|
591 |
One more premises, namely @{term "valid_trace_e s e"} |
|
592 |
is added, so that an interpretation of |
|
593 |
@{text "valid_trace_e"} can be instantiated |
|
594 |
so that all properties derived so far becomes |
|
595 |
available in the proof of induction step. |
|
596 |
||
597 |
You will see its use in the proofs that follows. |
|
598 |
*} |
|
599 |
lemma ind [consumes 0, case_names Nil Cons, induct type]: |
|
600 |
assumes "PP []" |
|
601 |
and "(\<And>s e. valid_trace_e s e \<Longrightarrow> |
|
602 |
PP s \<Longrightarrow> PIP s e \<Longrightarrow> PP (e # s))" |
|
603 |
shows "PP s" |
|
604 |
proof(induct rule:vt.induct[OF vt, case_names Init Step]) |
|
605 |
case Init |
|
606 |
from assms(1) show ?case . |
|
607 |
next |
|
608 |
case (Step s e) |
|
609 |
show ?case |
|
610 |
proof(rule assms(2)) |
|
611 |
show "valid_trace_e s e" using Step by (unfold_locales, auto) |
|
612 |
next |
|
613 |
show "PP s" using Step by simp |
|
614 |
next |
|
615 |
show "PIP s e" using Step by simp |
|
616 |
qed |
|
617 |
qed |
|
618 |
||
619 |
text {* |
|
620 |
The following lemma says that if @{text "s"} is a valid state, so |
|
621 |
is its any postfix. Where @{term "monent t s"} is the postfix of |
|
622 |
@{term "s"} with length @{term "t"}. |
|
623 |
*} |
|
624 |
lemma vt_moment: "\<And> t. vt (moment t s)" |
|
625 |
proof(induct rule:ind) |
|
626 |
case Nil |
|
627 |
thus ?case by (simp add:vt_nil) |
|
628 |
next |
|
629 |
case (Cons s e t) |
|
630 |
show ?case |
|
631 |
proof(cases "t \<ge> length (e#s)") |
|
632 |
case True |
|
633 |
from True have "moment t (e#s) = e#s" by simp |
|
634 |
thus ?thesis using Cons |
|
635 |
by (simp add:valid_trace_def valid_trace_e_def, auto) |
|
636 |
next |
|
637 |
case False |
|
638 |
from Cons have "vt (moment t s)" by simp |
|
639 |
moreover have "moment t (e#s) = moment t s" |
|
640 |
proof - |
|
641 |
from False have "t \<le> length s" by simp |
|
642 |
from moment_app [OF this, of "[e]"] |
|
643 |
show ?thesis by simp |
|
644 |
qed |
|
645 |
ultimately show ?thesis by simp |
|
646 |
qed |
|
647 |
qed |
|
648 |
end |
|
649 |
||
650 |
text {* |
|
651 |
The following locale @{text "valid_moment"} is to inherit the properties |
|
652 |
derived on any valid state to the prefix of it, with length @{text "i"}. |
|
653 |
*} |
|
654 |
locale valid_moment = valid_trace + |
|
655 |
fixes i :: nat |
|
656 |
||
100 | 657 |
sublocale valid_moment < vat_moment!: valid_trace "(moment i s)" |
99 | 658 |
by (unfold_locales, insert vt_moment, auto) |
659 |
||
100 | 660 |
locale valid_moment_e = valid_moment + |
661 |
assumes less_i: "i < length s" |
|
104 | 662 |
======= |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
663 |
lemma runing_ready: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
664 |
shows "runing s \<subseteq> readys s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
665 |
unfolding runing_def readys_def |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
666 |
by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
667 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
668 |
lemma readys_threads: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
669 |
shows "readys s \<subseteq> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
670 |
unfolding readys_def |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
671 |
by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
672 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
673 |
lemma wq_v_neq: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
674 |
"cs \<noteq> cs' \<Longrightarrow> wq (V thread cs#s) cs' = wq s cs'" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
675 |
by (auto simp:wq_def Let_def cp_def split:list.splits) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
676 |
|
68 | 677 |
lemma runing_head: |
678 |
assumes "th \<in> runing s" |
|
679 |
and "th \<in> set (wq_fun (schs s) cs)" |
|
680 |
shows "th = hd (wq_fun (schs s) cs)" |
|
681 |
using assms |
|
682 |
by (simp add:runing_def readys_def s_waiting_def wq_def) |
|
683 |
||
63 | 684 |
context valid_trace |
104 | 685 |
>>>>>>> other |
63 | 686 |
begin |
104 | 687 |
<<<<<<< local |
100 | 688 |
definition "next_e = hd (moment (Suc i) s)" |
689 |
||
690 |
lemma trace_e: |
|
691 |
"moment (Suc i) s = next_e#moment i s" |
|
692 |
proof - |
|
693 |
from less_i have "Suc i \<le> length s" by auto |
|
694 |
from moment_plus[OF this, folded next_e_def] |
|
695 |
show ?thesis . |
|
696 |
qed |
|
697 |
||
698 |
end |
|
699 |
||
700 |
sublocale valid_moment_e < vat_moment_e!: valid_trace_e "moment i s" "next_e" |
|
701 |
using vt_moment[of "Suc i", unfolded trace_e] |
|
702 |
by (unfold_locales, simp) |
|
703 |
||
704 |
section {* Distinctiveness of waiting queues *} |
|
99 | 705 |
|
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
706 |
context valid_trace_create |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
707 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
708 |
|
99 | 709 |
lemma wq_kept [simp]: |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
710 |
shows "wq (e#s) cs' = wq s cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
711 |
using assms unfolding is_create wq_def |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
712 |
by (auto simp:Let_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
713 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
714 |
lemma wq_distinct_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
715 |
assumes "distinct (wq s cs')" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
716 |
shows "distinct (wq (e#s) cs')" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
717 |
using assms by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
718 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
719 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
720 |
context valid_trace_exit |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
721 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
722 |
|
99 | 723 |
lemma wq_kept [simp]: |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
724 |
shows "wq (e#s) cs' = wq s cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
725 |
using assms unfolding is_exit wq_def |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
726 |
by (auto simp:Let_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
727 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
728 |
lemma wq_distinct_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
729 |
assumes "distinct (wq s cs')" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
730 |
shows "distinct (wq (e#s) cs')" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
731 |
using assms by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
732 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
733 |
|
100 | 734 |
context valid_trace_p |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
735 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
736 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
737 |
lemma wq_neq_simp [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
738 |
assumes "cs' \<noteq> cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
739 |
shows "wq (e#s) cs' = wq s cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
740 |
using assms unfolding is_p wq_def |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
741 |
by (auto simp:Let_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
742 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
743 |
lemma runing_th_s: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
744 |
shows "th \<in> runing s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
745 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
746 |
from pip_e[unfolded is_p] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
747 |
show ?thesis by (cases, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
748 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
749 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
750 |
lemma th_not_in_wq: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
751 |
shows "th \<notin> set (wq s cs)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
752 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
753 |
assume otherwise: "th \<in> set (wq s cs)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
754 |
from runing_wqE[OF runing_th_s this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
755 |
obtain rest where eq_wq: "wq s cs = th#rest" by blast |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
756 |
with otherwise |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
757 |
have "holding s th cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
758 |
by (unfold s_holding_def, fold wq_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
759 |
hence cs_th_RAG: "(Cs cs, Th th) \<in> RAG s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
760 |
by (unfold s_RAG_def, fold holding_eq, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
761 |
from pip_e[unfolded is_p] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
762 |
show False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
763 |
proof(cases) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
764 |
case (thread_P) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
765 |
with cs_th_RAG show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
766 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
767 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
768 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
769 |
lemma wq_es_cs: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
770 |
"wq (e#s) cs = wq s cs @ [th]" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
771 |
by (unfold is_p wq_def, auto simp:Let_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
772 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
773 |
lemma wq_distinct_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
774 |
assumes "distinct (wq s cs')" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
775 |
shows "distinct (wq (e#s) cs')" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
776 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
777 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
778 |
show ?thesis using True assms th_not_in_wq |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
779 |
by (unfold True wq_es_cs, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
780 |
qed (insert assms, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
781 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
782 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
783 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
784 |
context valid_trace_v |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
785 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
786 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
787 |
lemma wq_neq_simp [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
788 |
assumes "cs' \<noteq> cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
789 |
shows "wq (e#s) cs' = wq s cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
790 |
using assms unfolding is_v wq_def |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
791 |
by (auto simp:Let_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
792 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
793 |
lemma wq_s_cs: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
794 |
"wq s cs = th#rest" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
795 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
796 |
from pip_e[unfolded is_v] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
797 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
798 |
proof(cases) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
799 |
case (thread_V) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
800 |
from this(2) show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
801 |
by (unfold rest_def s_holding_def, fold wq_def, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
802 |
metis empty_iff list.collapse list.set(1)) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
803 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
804 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
805 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
806 |
lemma wq_es_cs: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
807 |
"wq (e#s) cs = wq'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
808 |
using wq_s_cs[unfolded wq_def] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
809 |
by (auto simp:Let_def wq_def rest_def wq'_def is_v, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
810 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
811 |
lemma wq_distinct_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
812 |
assumes "distinct (wq s cs')" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
813 |
shows "distinct (wq (e#s) cs')" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
814 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
815 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
816 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
817 |
proof(unfold True wq_es_cs wq'_def, rule someI2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
818 |
show "distinct rest \<and> set rest = set rest" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
819 |
using assms[unfolded True wq_s_cs] by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
820 |
qed simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
821 |
qed (insert assms, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
822 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
823 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
824 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
825 |
context valid_trace_set |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
826 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
827 |
|
99 | 828 |
lemma wq_kept [simp]: |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
829 |
shows "wq (e#s) cs' = wq s cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
830 |
using assms unfolding is_set wq_def |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
831 |
by (auto simp:Let_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
832 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
833 |
lemma wq_distinct_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
834 |
assumes "distinct (wq s cs')" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
835 |
shows "distinct (wq (e#s) cs')" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
836 |
using assms by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
837 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
838 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
839 |
context valid_trace |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
840 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
841 |
|
101 | 842 |
lemma finite_threads: |
843 |
shows "finite (threads s)" |
|
844 |
using vt by (induct) (auto elim: step.cases) |
|
845 |
||
846 |
lemma finite_readys [simp]: "finite (readys s)" |
|
847 |
using finite_threads readys_threads rev_finite_subset by blast |
|
104 | 848 |
======= |
63 | 849 |
|
67 | 850 |
lemma actor_inv: |
851 |
assumes "PIP s e" |
|
852 |
and "\<not> isCreate e" |
|
853 |
shows "actor e \<in> runing s" |
|
854 |
using assms |
|
855 |
by (induct, auto) |
|
856 |
||
63 | 857 |
lemma ind [consumes 0, case_names Nil Cons, induct type]: |
858 |
assumes "PP []" |
|
859 |
and "(\<And>s e. valid_trace s \<Longrightarrow> valid_trace (e#s) \<Longrightarrow> |
|
860 |
PP s \<Longrightarrow> PIP s e \<Longrightarrow> PP (e # s))" |
|
861 |
shows "PP s" |
|
862 |
proof(rule vt.induct[OF vt]) |
|
863 |
from assms(1) show "PP []" . |
|
864 |
next |
|
865 |
fix s e |
|
866 |
assume h: "vt s" "PP s" "PIP s e" |
|
867 |
show "PP (e # s)" |
|
868 |
proof(cases rule:assms(2)) |
|
869 |
from h(1) show v1: "valid_trace s" by (unfold_locales, simp) |
|
870 |
next |
|
871 |
from h(1,3) have "vt (e#s)" by auto |
|
872 |
thus "valid_trace (e # s)" by (unfold_locales, simp) |
|
873 |
qed (insert h, auto) |
|
874 |
qed |
|
104 | 875 |
>>>>>>> other |
63 | 876 |
|
877 |
lemma wq_distinct: "distinct (wq s cs)" |
|
68 | 878 |
proof(induct rule:ind) |
879 |
case (Cons s e) |
|
880 |
from Cons(4,3) |
|
881 |
show ?case |
|
882 |
proof(induct) |
|
883 |
case (thread_P th s cs1) |
|
884 |
show ?case |
|
885 |
proof(cases "cs = cs1") |
|
886 |
case True |
|
887 |
thus ?thesis (is "distinct ?L") |
|
888 |
proof - |
|
889 |
have "?L = wq_fun (schs s) cs1 @ [th]" using True |
|
890 |
by (simp add:wq_def wf_def Let_def split:list.splits) |
|
891 |
moreover have "distinct ..." |
|
892 |
proof - |
|
893 |
have "th \<notin> set (wq_fun (schs s) cs1)" |
|
894 |
proof |
|
895 |
assume otherwise: "th \<in> set (wq_fun (schs s) cs1)" |
|
896 |
from runing_head[OF thread_P(1) this] |
|
897 |
have "th = hd (wq_fun (schs s) cs1)" . |
|
898 |
hence "(Cs cs1, Th th) \<in> (RAG s)" using otherwise |
|
899 |
by (simp add:s_RAG_def s_holding_def wq_def cs_holding_def) |
|
900 |
with thread_P(2) show False by auto |
|
901 |
qed |
|
902 |
moreover have "distinct (wq_fun (schs s) cs1)" |
|
903 |
using True thread_P wq_def by auto |
|
904 |
ultimately show ?thesis by auto |
|
905 |
qed |
|
906 |
ultimately show ?thesis by simp |
|
907 |
qed |
|
908 |
next |
|
909 |
case False |
|
910 |
with thread_P(3) |
|
911 |
show ?thesis |
|
912 |
by (auto simp:wq_def wf_def Let_def split:list.splits) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
913 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
914 |
next |
68 | 915 |
case (thread_V th s cs1) |
916 |
thus ?case |
|
917 |
proof(cases "cs = cs1") |
|
918 |
case True |
|
919 |
show ?thesis (is "distinct ?L") |
|
920 |
proof(cases "(wq s cs)") |
|
921 |
case Nil |
|
922 |
thus ?thesis |
|
923 |
by (auto simp:wq_def wf_def Let_def split:list.splits) |
|
924 |
next |
|
925 |
case (Cons w_hd w_tl) |
|
926 |
moreover have "distinct (SOME q. distinct q \<and> set q = set w_tl)" |
|
927 |
proof(rule someI2) |
|
928 |
from thread_V(3)[unfolded Cons] |
|
929 |
show "distinct w_tl \<and> set w_tl = set w_tl" by auto |
|
930 |
qed auto |
|
931 |
ultimately show ?thesis |
|
932 |
by (auto simp:wq_def wf_def Let_def True split:list.splits) |
|
933 |
qed |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
934 |
next |
68 | 935 |
case False |
936 |
with thread_V(3) |
|
937 |
show ?thesis |
|
938 |
by (auto simp:wq_def wf_def Let_def split:list.splits) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
939 |
qed |
68 | 940 |
qed (insert Cons, auto simp: wq_def Let_def split:list.splits) |
941 |
qed (unfold wq_def Let_def, simp) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
942 |
|
63 | 943 |
end |
944 |
||
104 | 945 |
<<<<<<< local |
100 | 946 |
section {* Waiting queues and threads *} |
947 |
||
104 | 948 |
======= |
949 |
||
950 |
>>>>>>> other |
|
63 | 951 |
context valid_trace_e |
952 |
begin |
|
953 |
||
104 | 954 |
<<<<<<< local |
100 | 955 |
lemma wq_out_inv: |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
956 |
assumes s_in: "thread \<in> set (wq s cs)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
957 |
and s_hd: "thread = hd (wq s cs)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
958 |
and s_i: "thread \<noteq> hd (wq (e#s) cs)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
959 |
shows "e = V thread cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
960 |
proof(cases e) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
961 |
-- {* There are only two non-trivial cases: *} |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
962 |
case (V th cs1) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
963 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
964 |
proof(cases "cs1 = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
965 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
966 |
have "PIP s (V th cs)" using pip_e[unfolded V[unfolded True]] . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
967 |
thus ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
968 |
proof(cases) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
969 |
case (thread_V) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
970 |
moreover have "th = thread" using thread_V(2) s_hd |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
971 |
by (unfold s_holding_def wq_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
972 |
ultimately show ?thesis using V True by simp |
104 | 973 |
======= |
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
974 |
text {* |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
975 |
The following lemma shows that only the @{text "P"} |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
976 |
operation can add new thread into waiting queues. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
977 |
Such kind of lemmas are very obvious, but need to be checked formally. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
978 |
This is a kind of confirmation that our modelling is correct. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
979 |
*} |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
980 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
981 |
lemma block_pre: |
68 | 982 |
assumes s_ni: "thread \<notin> set (wq s cs)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
983 |
and s_i: "thread \<in> set (wq (e#s) cs)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
984 |
shows "e = P thread cs" |
68 | 985 |
proof(cases e) |
986 |
-- {* This is the only non-trivial case: *} |
|
987 |
case (V th cs1) |
|
988 |
have False |
|
989 |
proof(cases "cs1 = cs") |
|
990 |
case True |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
991 |
show ?thesis |
68 | 992 |
proof(cases "(wq s cs1)") |
993 |
case (Cons w_hd w_tl) |
|
994 |
have "set (wq (e#s) cs) \<subseteq> set (wq s cs)" |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
995 |
proof - |
68 | 996 |
have "(wq (e#s) cs) = (SOME q. distinct q \<and> set q = set w_tl)" |
997 |
using Cons V by (auto simp:wq_def Let_def True split:if_splits) |
|
998 |
moreover have "set ... \<subseteq> set (wq s cs)" |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
999 |
proof(rule someI2) |
68 | 1000 |
show "distinct w_tl \<and> set w_tl = set w_tl" |
1001 |
by (metis distinct.simps(2) local.Cons wq_distinct) |
|
1002 |
qed (insert Cons True, auto) |
|
1003 |
ultimately show ?thesis by simp |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1004 |
qed |
68 | 1005 |
with assms show ?thesis by auto |
1006 |
qed (insert assms V True, auto simp:wq_def Let_def split:if_splits) |
|
1007 |
qed (insert assms V, auto simp:wq_def Let_def split:if_splits) |
|
1008 |
thus ?thesis by auto |
|
1009 |
qed (insert assms, auto simp:wq_def Let_def split:if_splits) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1010 |
|
63 | 1011 |
end |
1012 |
||
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1013 |
text {* |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1014 |
The following lemmas is also obvious and shallow. It says |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1015 |
that only running thread can request for a critical resource |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1016 |
and that the requested resource must be one which is |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1017 |
not current held by the thread. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1018 |
*} |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1019 |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1020 |
lemma p_pre: "\<lbrakk>vt ((P thread cs)#s)\<rbrakk> \<Longrightarrow> |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
1021 |
thread \<in> runing s \<and> (Cs cs, Th thread) \<notin> (RAG s)^+" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1022 |
apply (ind_cases "vt ((P thread cs)#s)") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1023 |
apply (ind_cases "step s (P thread cs)") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1024 |
by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1025 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1026 |
lemma abs1: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1027 |
assumes ein: "e \<in> set es" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1028 |
and neq: "hd es \<noteq> hd (es @ [x])" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1029 |
shows "False" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1030 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1031 |
from ein have "es \<noteq> []" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1032 |
then obtain e ess where "es = e # ess" by (cases es, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1033 |
with neq show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1034 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1035 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1036 |
lemma q_head: "Q (hd es) \<Longrightarrow> hd es = hd [th\<leftarrow>es . Q th]" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1037 |
by (cases es, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1038 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1039 |
inductive_cases evt_cons: "vt (a#s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1040 |
|
63 | 1041 |
context valid_trace_e |
1042 |
begin |
|
1043 |
||
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1044 |
lemma abs2: |
63 | 1045 |
assumes inq: "thread \<in> set (wq s cs)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1046 |
and nh: "thread = hd (wq s cs)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1047 |
and qt: "thread \<noteq> hd (wq (e#s) cs)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1048 |
and inq': "thread \<in> set (wq (e#s) cs)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1049 |
shows "False" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1050 |
proof - |
63 | 1051 |
from vt_e assms show "False" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1052 |
apply (cases e) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1053 |
apply ((simp split:if_splits add:Let_def wq_def)[1])+ |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1054 |
apply (insert abs1, fast)[1] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1055 |
apply (auto simp:wq_def simp:Let_def split:if_splits list.splits) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1056 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1057 |
fix th qs |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1058 |
assume vt: "vt (V th cs # s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1059 |
and th_in: "thread \<in> set (SOME q. distinct q \<and> set q = set qs)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1060 |
and eq_wq: "wq_fun (schs s) cs = thread # qs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1061 |
show "False" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1062 |
proof - |
63 | 1063 |
from wq_distinct[of cs] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1064 |
and eq_wq[folded wq_def] have "distinct (thread#qs)" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1065 |
moreover have "thread \<in> set qs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1066 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1067 |
have "set (SOME q. distinct q \<and> set q = set qs) = set qs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1068 |
proof(rule someI2) |
63 | 1069 |
from wq_distinct [of cs] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1070 |
and eq_wq [folded wq_def] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1071 |
show "distinct qs \<and> set qs = set qs" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1072 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1073 |
fix x assume "distinct x \<and> set x = set qs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1074 |
thus "set x = set qs" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1075 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1076 |
with th_in show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1077 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1078 |
ultimately show ?thesis by auto |
104 | 1079 |
>>>>>>> other |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1080 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1081 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1082 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1083 |
|
100 | 1084 |
lemma wq_in_inv: |
1085 |
assumes s_ni: "thread \<notin> set (wq s cs)" |
|
1086 |
and s_i: "thread \<in> set (wq (e#s) cs)" |
|
1087 |
shows "e = P thread cs" |
|
1088 |
proof(cases e) |
|
1089 |
-- {* This is the only non-trivial case: *} |
|
1090 |
case (V th cs1) |
|
1091 |
have False |
|
1092 |
proof(cases "cs1 = cs") |
|
1093 |
case True |
|
1094 |
show ?thesis |
|
1095 |
proof(cases "(wq s cs1)") |
|
1096 |
case (Cons w_hd w_tl) |
|
1097 |
have "set (wq (e#s) cs) \<subseteq> set (wq s cs)" |
|
1098 |
proof - |
|
1099 |
have "(wq (e#s) cs) = (SOME q. distinct q \<and> set q = set w_tl)" |
|
1100 |
using Cons V by (auto simp:wq_def Let_def True split:if_splits) |
|
1101 |
moreover have "set ... \<subseteq> set (wq s cs)" |
|
1102 |
proof(rule someI2) |
|
1103 |
show "distinct w_tl \<and> set w_tl = set w_tl" |
|
1104 |
by (metis distinct.simps(2) local.Cons wq_distinct) |
|
1105 |
qed (insert Cons True, auto) |
|
1106 |
ultimately show ?thesis by simp |
|
1107 |
qed |
|
1108 |
with assms show ?thesis by auto |
|
1109 |
qed (insert assms V True, auto simp:wq_def Let_def split:if_splits) |
|
1110 |
qed (insert assms V, auto simp:wq_def Let_def split:if_splits) |
|
1111 |
thus ?thesis by auto |
|
1112 |
qed (insert assms, auto simp:wq_def Let_def split:if_splits) |
|
1113 |
||
63 | 1114 |
end |
1115 |
||
100 | 1116 |
lemma (in valid_trace_create) |
1117 |
th_not_in_threads: "th \<notin> threads s" |
|
1118 |
proof - |
|
1119 |
from pip_e[unfolded is_create] |
|
1120 |
show ?thesis by (cases, simp) |
|
1121 |
qed |
|
1122 |
||
1123 |
lemma (in valid_trace_create) |
|
1124 |
threads_es [simp]: "threads (e#s) = threads s \<union> {th}" |
|
1125 |
by (unfold is_create, simp) |
|
1126 |
||
1127 |
lemma (in valid_trace_exit) |
|
1128 |
threads_es [simp]: "threads (e#s) = threads s - {th}" |
|
1129 |
by (unfold is_exit, simp) |
|
1130 |
||
1131 |
lemma (in valid_trace_p) |
|
1132 |
threads_es [simp]: "threads (e#s) = threads s" |
|
1133 |
by (unfold is_p, simp) |
|
1134 |
||
1135 |
lemma (in valid_trace_v) |
|
1136 |
threads_es [simp]: "threads (e#s) = threads s" |
|
1137 |
by (unfold is_v, simp) |
|
1138 |
||
1139 |
lemma (in valid_trace_v) |
|
1140 |
th_not_in_rest[simp]: "th \<notin> set rest" |
|
1141 |
proof |
|
1142 |
assume otherwise: "th \<in> set rest" |
|
1143 |
have "distinct (wq s cs)" by (simp add: wq_distinct) |
|
1144 |
from this[unfolded wq_s_cs] and otherwise |
|
1145 |
show False by auto |
|
1146 |
qed |
|
1147 |
||
1148 |
lemma (in valid_trace_v) distinct_rest: "distinct rest" |
|
1149 |
by (simp add: distinct_tl rest_def wq_distinct) |
|
1150 |
||
1151 |
lemma (in valid_trace_v) |
|
1152 |
set_wq_es_cs [simp]: "set (wq (e#s) cs) = set (wq s cs) - {th}" |
|
1153 |
proof(unfold wq_es_cs wq'_def, rule someI2) |
|
1154 |
show "distinct rest \<and> set rest = set rest" |
|
1155 |
by (simp add: distinct_rest) |
|
1156 |
next |
|
1157 |
fix x |
|
1158 |
assume "distinct x \<and> set x = set rest" |
|
1159 |
thus "set x = set (wq s cs) - {th}" |
|
1160 |
by (unfold wq_s_cs, simp) |
|
1161 |
qed |
|
1162 |
||
1163 |
lemma (in valid_trace_exit) |
|
1164 |
th_not_in_wq: "th \<notin> set (wq s cs)" |
|
1165 |
proof - |
|
1166 |
from pip_e[unfolded is_exit] |
|
1167 |
show ?thesis |
|
1168 |
by (cases, unfold holdents_def s_holding_def, fold wq_def, |
|
1169 |
auto elim!:runing_wqE) |
|
1170 |
qed |
|
1171 |
||
1172 |
lemma (in valid_trace) wq_threads: |
|
1173 |
assumes "th \<in> set (wq s cs)" |
|
1174 |
shows "th \<in> threads s" |
|
1175 |
using assms |
|
1176 |
proof(induct rule:ind) |
|
1177 |
case (Nil) |
|
1178 |
thus ?case by (auto simp:wq_def) |
|
1179 |
next |
|
1180 |
case (Cons s e) |
|
1181 |
interpret vt_e: valid_trace_e s e using Cons by simp |
|
1182 |
show ?case |
|
1183 |
proof(cases e) |
|
1184 |
case (Create th' prio') |
|
1185 |
interpret vt: valid_trace_create s e th' prio' |
|
1186 |
using Create by (unfold_locales, simp) |
|
1187 |
show ?thesis |
|
1188 |
using Cons.hyps(2) Cons.prems by auto |
|
1189 |
next |
|
1190 |
case (Exit th') |
|
1191 |
interpret vt: valid_trace_exit s e th' |
|
1192 |
using Exit by (unfold_locales, simp) |
|
1193 |
show ?thesis |
|
1194 |
using Cons.hyps(2) Cons.prems vt.th_not_in_wq by auto |
|
1195 |
next |
|
1196 |
case (P th' cs') |
|
1197 |
interpret vt: valid_trace_p s e th' cs' |
|
1198 |
using P by (unfold_locales, simp) |
|
1199 |
show ?thesis |
|
1200 |
using Cons.hyps(2) Cons.prems readys_threads |
|
1201 |
runing_ready vt.is_p vt.runing_th_s vt_e.wq_in_inv |
|
1202 |
by fastforce |
|
1203 |
next |
|
1204 |
case (V th' cs') |
|
1205 |
interpret vt: valid_trace_v s e th' cs' |
|
1206 |
using V by (unfold_locales, simp) |
|
1207 |
show ?thesis using Cons |
|
1208 |
using vt.is_v vt.threads_es vt_e.wq_in_inv by blast |
|
1209 |
next |
|
1210 |
case (Set th' prio) |
|
1211 |
interpret vt: valid_trace_set s e th' prio |
|
1212 |
using Set by (unfold_locales, simp) |
|
1213 |
show ?thesis using Cons.hyps(2) Cons.prems vt.is_set |
|
1214 |
by (auto simp:wq_def Let_def) |
|
1215 |
qed |
|
1216 |
qed |
|
1217 |
||
1218 |
section {* RAG and threads *} |
|
68 | 1219 |
|
63 | 1220 |
context valid_trace |
1221 |
begin |
|
68 | 1222 |
lemma vt_moment: "\<And> t. vt (moment t s)" |
63 | 1223 |
proof(induct rule:ind) |
1224 |
case Nil |
|
1225 |
thus ?case by (simp add:vt_nil) |
|
1226 |
next |
|
1227 |
case (Cons s e t) |
|
1228 |
show ?case |
|
1229 |
proof(cases "t \<ge> length (e#s)") |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1230 |
case True |
63 | 1231 |
from True have "moment t (e#s) = e#s" by simp |
1232 |
thus ?thesis using Cons |
|
1233 |
by (simp add:valid_trace_def) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1234 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1235 |
case False |
63 | 1236 |
from Cons have "vt (moment t s)" by simp |
1237 |
moreover have "moment t (e#s) = moment t s" |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1238 |
proof - |
63 | 1239 |
from False have "t \<le> length s" by simp |
1240 |
from moment_app [OF this, of "[e]"] |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1241 |
show ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1242 |
qed |
63 | 1243 |
ultimately show ?thesis by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1244 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1245 |
qed |
68 | 1246 |
end |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1247 |
|
68 | 1248 |
locale valid_moment = valid_trace + |
1249 |
fixes i :: nat |
|
1250 |
||
1251 |
sublocale valid_moment < vat_moment: valid_trace "(moment i s)" |
|
1252 |
by (unfold_locales, insert vt_moment, auto) |
|
1253 |
||
1254 |
context valid_trace |
|
1255 |
begin |
|
1256 |
||
104 | 1257 |
<<<<<<< local |
100 | 1258 |
lemma dm_RAG_threads: |
1259 |
assumes in_dom: "(Th th) \<in> Domain (RAG s)" |
|
1260 |
shows "th \<in> threads s" |
|
104 | 1261 |
======= |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1262 |
|
55
b85cfbd58f59
Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents:
53
diff
changeset
|
1263 |
text {* (* ddd *) |
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1264 |
The nature of the work is like this: since it starts from a very simple and basic |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1265 |
model, even intuitively very `basic` and `obvious` properties need to derived from scratch. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1266 |
For instance, the fact |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1267 |
that one thread can not be blocked by two critical resources at the same time |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1268 |
is obvious, because only running threads can make new requests, if one is waiting for |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1269 |
a critical resource and get blocked, it can not make another resource request and get |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1270 |
blocked the second time (because it is not running). |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1271 |
|
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1272 |
To derive this fact, one needs to prove by contraction and |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1273 |
reason about time (or @{text "moement"}). The reasoning is based on a generic theorem |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1274 |
named @{text "p_split"}, which is about status changing along the time axis. It says if |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1275 |
a condition @{text "Q"} is @{text "True"} at a state @{text "s"}, |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1276 |
but it was @{text "False"} at the very beginning, then there must exits a moment @{text "t"} |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1277 |
in the history of @{text "s"} (notice that @{text "s"} itself is essentially the history |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1278 |
of events leading to it), such that @{text "Q"} switched |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1279 |
from being @{text "False"} to @{text "True"} and kept being @{text "True"} |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1280 |
till the last moment of @{text "s"}. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1281 |
|
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1282 |
Suppose a thread @{text "th"} is blocked |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1283 |
on @{text "cs1"} and @{text "cs2"} in some state @{text "s"}, |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1284 |
since no thread is blocked at the very beginning, by applying |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1285 |
@{text "p_split"} to these two blocking facts, there exist |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1286 |
two moments @{text "t1"} and @{text "t2"} in @{text "s"}, such that |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1287 |
@{text "th"} got blocked on @{text "cs1"} and @{text "cs2"} |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1288 |
and kept on blocked on them respectively ever since. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1289 |
|
68 | 1290 |
Without lost of generality, we assume @{text "t1"} is earlier than @{text "t2"}. |
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1291 |
However, since @{text "th"} was blocked ever since memonent @{text "t1"}, so it was still |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1292 |
in blocked state at moment @{text "t2"} and could not |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1293 |
make any request and get blocked the second time: Contradiction. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1294 |
*} |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1295 |
|
68 | 1296 |
lemma waiting_unique_pre: (* ccc *) |
63 | 1297 |
assumes h11: "thread \<in> set (wq s cs1)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1298 |
and h12: "thread \<noteq> hd (wq s cs1)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1299 |
assumes h21: "thread \<in> set (wq s cs2)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1300 |
and h22: "thread \<noteq> hd (wq s cs2)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1301 |
and neq12: "cs1 \<noteq> cs2" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1302 |
shows "False" |
104 | 1303 |
>>>>>>> other |
1304 |
proof - |
|
1305 |
<<<<<<< local |
|
100 | 1306 |
from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto |
1307 |
moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto |
|
1308 |
ultimately have "(Th th, Cs cs) \<in> RAG s" by simp |
|
1309 |
hence "th \<in> set (wq s cs)" |
|
1310 |
by (unfold s_RAG_def, auto simp:cs_waiting_def) |
|
1311 |
from wq_threads [OF this] show ?thesis . |
|
104 | 1312 |
======= |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1313 |
let "?Q cs s" = "thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1314 |
from h11 and h12 have q1: "?Q cs1 s" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1315 |
from h21 and h22 have q2: "?Q cs2 s" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1316 |
have nq1: "\<not> ?Q cs1 []" by (simp add:wq_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1317 |
have nq2: "\<not> ?Q cs2 []" by (simp add:wq_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1318 |
from p_split [of "?Q cs1", OF q1 nq1] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1319 |
obtain t1 where lt1: "t1 < length s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1320 |
and np1: "\<not>(thread \<in> set (wq (moment t1 s) cs1) \<and> |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1321 |
thread \<noteq> hd (wq (moment t1 s) cs1))" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1322 |
and nn1: "(\<forall>i'>t1. thread \<in> set (wq (moment i' s) cs1) \<and> |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1323 |
thread \<noteq> hd (wq (moment i' s) cs1))" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1324 |
from p_split [of "?Q cs2", OF q2 nq2] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1325 |
obtain t2 where lt2: "t2 < length s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1326 |
and np2: "\<not>(thread \<in> set (wq (moment t2 s) cs2) \<and> |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1327 |
thread \<noteq> hd (wq (moment t2 s) cs2))" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1328 |
and nn2: "(\<forall>i'>t2. thread \<in> set (wq (moment i' s) cs2) \<and> |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1329 |
thread \<noteq> hd (wq (moment i' s) cs2))" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1330 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1331 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1332 |
{ |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1333 |
assume lt12: "t1 < t2" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1334 |
let ?t3 = "Suc t2" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1335 |
from lt2 have le_t3: "?t3 \<le> length s" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1336 |
from moment_plus [OF this] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1337 |
obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1338 |
have "t2 < ?t3" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1339 |
from nn2 [rule_format, OF this] and eq_m |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1340 |
have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1341 |
h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto |
63 | 1342 |
have "vt (e#moment t2 s)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1343 |
proof - |
63 | 1344 |
from vt_moment |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1345 |
have "vt (moment ?t3 s)" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1346 |
with eq_m show ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1347 |
qed |
63 | 1348 |
then interpret vt_e: valid_trace_e "moment t2 s" "e" |
1349 |
by (unfold_locales, auto, cases, simp) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1350 |
have ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1351 |
proof(cases "thread \<in> set (wq (moment t2 s) cs2)") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1352 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1353 |
from True and np2 have eq_th: "thread = hd (wq (moment t2 s) cs2)" |
63 | 1354 |
by auto |
1355 |
from vt_e.abs2 [OF True eq_th h2 h1] |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1356 |
show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1357 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1358 |
case False |
63 | 1359 |
from vt_e.block_pre[OF False h1] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1360 |
have "e = P thread cs2" . |
63 | 1361 |
with vt_e.vt_e have "vt ((P thread cs2)# moment t2 s)" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1362 |
from p_pre [OF this] have "thread \<in> runing (moment t2 s)" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1363 |
with runing_ready have "thread \<in> readys (moment t2 s)" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1364 |
with nn1 [rule_format, OF lt12] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1365 |
show ?thesis by (simp add:readys_def wq_def s_waiting_def, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1366 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1367 |
} moreover { |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1368 |
assume lt12: "t2 < t1" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1369 |
let ?t3 = "Suc t1" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1370 |
from lt1 have le_t3: "?t3 \<le> length s" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1371 |
from moment_plus [OF this] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1372 |
obtain e where eq_m: "moment ?t3 s = e#moment t1 s" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1373 |
have lt_t3: "t1 < ?t3" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1374 |
from nn1 [rule_format, OF this] and eq_m |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1375 |
have h1: "thread \<in> set (wq (e#moment t1 s) cs1)" and |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1376 |
h2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto |
63 | 1377 |
have "vt (e#moment t1 s)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1378 |
proof - |
63 | 1379 |
from vt_moment |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1380 |
have "vt (moment ?t3 s)" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1381 |
with eq_m show ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1382 |
qed |
63 | 1383 |
then interpret vt_e: valid_trace_e "moment t1 s" e |
1384 |
by (unfold_locales, auto, cases, auto) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1385 |
have ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1386 |
proof(cases "thread \<in> set (wq (moment t1 s) cs1)") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1387 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1388 |
from True and np1 have eq_th: "thread = hd (wq (moment t1 s) cs1)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1389 |
by auto |
63 | 1390 |
from vt_e.abs2 True eq_th h2 h1 |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1391 |
show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1392 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1393 |
case False |
63 | 1394 |
from vt_e.block_pre [OF False h1] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1395 |
have "e = P thread cs1" . |
63 | 1396 |
with vt_e.vt_e have "vt ((P thread cs1)# moment t1 s)" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1397 |
from p_pre [OF this] have "thread \<in> runing (moment t1 s)" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1398 |
with runing_ready have "thread \<in> readys (moment t1 s)" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1399 |
with nn2 [rule_format, OF lt12] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1400 |
show ?thesis by (simp add:readys_def wq_def s_waiting_def, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1401 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1402 |
} moreover { |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1403 |
assume eqt12: "t1 = t2" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1404 |
let ?t3 = "Suc t1" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1405 |
from lt1 have le_t3: "?t3 \<le> length s" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1406 |
from moment_plus [OF this] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1407 |
obtain e where eq_m: "moment ?t3 s = e#moment t1 s" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1408 |
have lt_t3: "t1 < ?t3" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1409 |
from nn1 [rule_format, OF this] and eq_m |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1410 |
have h1: "thread \<in> set (wq (e#moment t1 s) cs1)" and |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1411 |
h2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1412 |
have vt_e: "vt (e#moment t1 s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1413 |
proof - |
63 | 1414 |
from vt_moment |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1415 |
have "vt (moment ?t3 s)" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1416 |
with eq_m show ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1417 |
qed |
63 | 1418 |
then interpret vt_e: valid_trace_e "moment t1 s" e |
1419 |
by (unfold_locales, auto, cases, auto) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1420 |
have ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1421 |
proof(cases "thread \<in> set (wq (moment t1 s) cs1)") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1422 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1423 |
from True and np1 have eq_th: "thread = hd (wq (moment t1 s) cs1)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1424 |
by auto |
63 | 1425 |
from vt_e.abs2 [OF True eq_th h2 h1] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1426 |
show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1427 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1428 |
case False |
63 | 1429 |
from vt_e.block_pre [OF False h1] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1430 |
have eq_e1: "e = P thread cs1" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1431 |
have lt_t3: "t1 < ?t3" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1432 |
with eqt12 have "t2 < ?t3" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1433 |
from nn2 [rule_format, OF this] and eq_m and eqt12 |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1434 |
have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1435 |
h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1436 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1437 |
proof(cases "thread \<in> set (wq (moment t2 s) cs2)") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1438 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1439 |
from True and np2 have eq_th: "thread = hd (wq (moment t2 s) cs2)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1440 |
by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1441 |
from vt_e and eqt12 have "vt (e#moment t2 s)" by simp |
63 | 1442 |
then interpret vt_e2: valid_trace_e "moment t2 s" e |
1443 |
by (unfold_locales, auto, cases, auto) |
|
1444 |
from vt_e2.abs2 [OF True eq_th h2 h1] |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1445 |
show ?thesis . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1446 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1447 |
case False |
63 | 1448 |
have "vt (e#moment t2 s)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1449 |
proof - |
63 | 1450 |
from vt_moment eqt12 |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1451 |
have "vt (moment (Suc t2) s)" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1452 |
with eq_m eqt12 show ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1453 |
qed |
63 | 1454 |
then interpret vt_e2: valid_trace_e "moment t2 s" e |
1455 |
by (unfold_locales, auto, cases, auto) |
|
1456 |
from vt_e2.block_pre [OF False h1] |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1457 |
have "e = P thread cs2" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1458 |
with eq_e1 neq12 show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1459 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1460 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1461 |
} ultimately show ?thesis by arith |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1462 |
qed |
104 | 1463 |
>>>>>>> other |
1464 |
qed |
|
1465 |
||
1466 |
<<<<<<< local |
|
100 | 1467 |
lemma rg_RAG_threads: |
1468 |
assumes "(Th th) \<in> Range (RAG s)" |
|
1469 |
shows "th \<in> threads s" |
|
1470 |
using assms |
|
1471 |
by (unfold s_RAG_def cs_waiting_def cs_holding_def, |
|
1472 |
auto intro:wq_threads) |
|
1473 |
||
1474 |
lemma RAG_threads: |
|
1475 |
assumes "(Th th) \<in> Field (RAG s)" |
|
1476 |
shows "th \<in> threads s" |
|
1477 |
using assms |
|
1478 |
by (metis Field_def UnE dm_RAG_threads rg_RAG_threads) |
|
104 | 1479 |
======= |
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1480 |
text {* |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1481 |
This lemma is a simple corrolary of @{text "waiting_unique_pre"}. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1482 |
*} |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1483 |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1484 |
lemma waiting_unique: |
63 | 1485 |
assumes "waiting s th cs1" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1486 |
and "waiting s th cs2" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1487 |
shows "cs1 = cs2" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1488 |
using waiting_unique_pre assms |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1489 |
unfolding wq_def s_waiting_def |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1490 |
by auto |
104 | 1491 |
>>>>>>> other |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1492 |
|
63 | 1493 |
end |
1494 |
||
104 | 1495 |
<<<<<<< local |
100 | 1496 |
section {* The change of @{term RAG} *} |
104 | 1497 |
======= |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1498 |
(* not used *) |
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1499 |
text {* |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1500 |
Every thread can only be blocked on one critical resource, |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1501 |
symmetrically, every critical resource can only be held by one thread. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1502 |
This fact is much more easier according to our definition. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1503 |
*} |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1504 |
lemma held_unique: |
63 | 1505 |
assumes "holding (s::event list) th1 cs" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1506 |
and "holding s th2 cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1507 |
shows "th1 = th2" |
63 | 1508 |
by (insert assms, unfold s_holding_def, auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1509 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1510 |
|
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
1511 |
lemma last_set_lt: "th \<in> threads s \<Longrightarrow> last_set th s < length s" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1512 |
apply (induct s, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1513 |
by (case_tac a, auto split:if_splits) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1514 |
|
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
1515 |
lemma last_set_unique: |
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
1516 |
"\<lbrakk>last_set th1 s = last_set th2 s; th1 \<in> threads s; th2 \<in> threads s\<rbrakk> |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1517 |
\<Longrightarrow> th1 = th2" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1518 |
apply (induct s, auto) |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
1519 |
by (case_tac a, auto split:if_splits dest:last_set_lt) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1520 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1521 |
lemma preced_unique : |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1522 |
assumes pcd_eq: "preced th1 s = preced th2 s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1523 |
and th_in1: "th1 \<in> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1524 |
and th_in2: " th2 \<in> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1525 |
shows "th1 = th2" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1526 |
proof - |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
1527 |
from pcd_eq have "last_set th1 s = last_set th2 s" by (simp add:preced_def) |
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
1528 |
from last_set_unique [OF this th_in1 th_in2] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1529 |
show ?thesis . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1530 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1531 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1532 |
lemma preced_linorder: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1533 |
assumes neq_12: "th1 \<noteq> th2" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1534 |
and th_in1: "th1 \<in> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1535 |
and th_in2: " th2 \<in> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1536 |
shows "preced th1 s < preced th2 s \<or> preced th1 s > preced th2 s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1537 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1538 |
from preced_unique [OF _ th_in1 th_in2] and neq_12 |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1539 |
have "preced th1 s \<noteq> preced th2 s" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1540 |
thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1541 |
qed |
104 | 1542 |
>>>>>>> other |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1543 |
|
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1544 |
(* An aux lemma used later *) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1545 |
lemma unique_minus: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1546 |
assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1547 |
and xy: "(x, y) \<in> r" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1548 |
and xz: "(x, z) \<in> r^+" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1549 |
and neq: "y \<noteq> z" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1550 |
shows "(y, z) \<in> r^+" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1551 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1552 |
from xz and neq show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1553 |
proof(induct) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1554 |
case (base ya) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1555 |
have "(x, ya) \<in> r" by fact |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1556 |
from unique [OF xy this] have "y = ya" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1557 |
with base show ?case by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1558 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1559 |
case (step ya z) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1560 |
show ?case |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1561 |
proof(cases "y = ya") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1562 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1563 |
from step True show ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1564 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1565 |
case False |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1566 |
from step False |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1567 |
show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1568 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1569 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1570 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1571 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1572 |
lemma unique_base: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1573 |
assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1574 |
and xy: "(x, y) \<in> r" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1575 |
and xz: "(x, z) \<in> r^+" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1576 |
and neq_yz: "y \<noteq> z" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1577 |
shows "(y, z) \<in> r^+" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1578 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1579 |
from xz neq_yz show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1580 |
proof(induct) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1581 |
case (base ya) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1582 |
from xy unique base show ?case by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1583 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1584 |
case (step ya z) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1585 |
show ?case |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1586 |
proof(cases "y = ya") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1587 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1588 |
from True step show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1589 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1590 |
case False |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1591 |
from False step |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1592 |
have "(y, ya) \<in> r\<^sup>+" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1593 |
with step show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1594 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1595 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1596 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1597 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1598 |
lemma unique_chain: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1599 |
assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1600 |
and xy: "(x, y) \<in> r^+" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1601 |
and xz: "(x, z) \<in> r^+" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1602 |
and neq_yz: "y \<noteq> z" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1603 |
shows "(y, z) \<in> r^+ \<or> (z, y) \<in> r^+" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1604 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1605 |
from xy xz neq_yz show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1606 |
proof(induct) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1607 |
case (base y) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1608 |
have h1: "(x, y) \<in> r" and h2: "(x, z) \<in> r\<^sup>+" and h3: "y \<noteq> z" using base by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1609 |
from unique_base [OF _ h1 h2 h3] and unique show ?case by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1610 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1611 |
case (step y za) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1612 |
show ?case |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1613 |
proof(cases "y = z") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1614 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1615 |
from True step show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1616 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1617 |
case False |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1618 |
from False step have "(y, z) \<in> r\<^sup>+ \<or> (z, y) \<in> r\<^sup>+" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1619 |
thus ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1620 |
proof |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1621 |
assume "(z, y) \<in> r\<^sup>+" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1622 |
with step have "(z, za) \<in> r\<^sup>+" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1623 |
thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1624 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1625 |
assume h: "(y, z) \<in> r\<^sup>+" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1626 |
from step have yza: "(y, za) \<in> r" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1627 |
from step have "za \<noteq> z" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1628 |
from unique_minus [OF _ yza h this] and unique |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1629 |
have "(za, z) \<in> r\<^sup>+" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1630 |
thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1631 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1632 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1633 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1634 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1635 |
|
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1636 |
text {* |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1637 |
The following three lemmas show that @{text "RAG"} does not change |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1638 |
by the happening of @{text "Set"}, @{text "Create"} and @{text "Exit"} |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1639 |
events, respectively. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1640 |
*} |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1641 |
|
104 | 1642 |
<<<<<<< local |
100 | 1643 |
lemma (in valid_trace_set) RAG_unchanged [simp]: "(RAG (e # s)) = RAG s" |
1644 |
by (unfold is_set s_RAG_def s_waiting_def wq_def, simp add:Let_def) |
|
1645 |
||
1646 |
lemma (in valid_trace_create) RAG_unchanged [simp]: "(RAG (e # s)) = RAG s" |
|
1647 |
by (unfold is_create s_RAG_def s_waiting_def wq_def, simp add:Let_def) |
|
1648 |
||
1649 |
lemma (in valid_trace_exit) RAG_unchanged[simp]: "(RAG (e # s)) = RAG s" |
|
1650 |
by (unfold is_exit s_RAG_def s_waiting_def wq_def, simp add:Let_def) |
|
104 | 1651 |
======= |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
1652 |
lemma RAG_set_unchanged: "(RAG (Set th prio # s)) = RAG s" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
1653 |
apply (unfold s_RAG_def s_waiting_def wq_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1654 |
by (simp add:Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1655 |
|
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
1656 |
lemma RAG_create_unchanged: "(RAG (Create th prio # s)) = RAG s" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
1657 |
apply (unfold s_RAG_def s_waiting_def wq_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1658 |
by (simp add:Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1659 |
|
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
1660 |
lemma RAG_exit_unchanged: "(RAG (Exit th # s)) = RAG s" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
1661 |
apply (unfold s_RAG_def s_waiting_def wq_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1662 |
by (simp add:Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1663 |
|
104 | 1664 |
>>>>>>> other |
1665 |
||
1666 |
<<<<<<< local |
|
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1667 |
context valid_trace_v |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1668 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1669 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1670 |
lemma holding_cs_eq_th: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1671 |
assumes "holding s t cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1672 |
shows "t = th" |
104 | 1673 |
======= |
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1674 |
text {* |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1675 |
The following lemmas are used in the proof of |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1676 |
lemma @{text "step_RAG_v"}, which characterizes how the @{text "RAG"} is changed |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1677 |
by @{text "V"}-events. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1678 |
However, since our model is very concise, such seemingly obvious lemmas need to be derived from scratch, |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1679 |
starting from the model definitions. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1680 |
*} |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1681 |
lemma step_v_hold_inv[elim_format]: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1682 |
"\<And>c t. \<lbrakk>vt (V th cs # s); |
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1683 |
\<not> holding (wq s) t c; holding (wq (V th cs # s)) t c\<rbrakk> \<Longrightarrow> |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1684 |
next_th s th cs t \<and> c = cs" |
104 | 1685 |
>>>>>>> other |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1686 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1687 |
fix c t |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1688 |
assume vt: "vt (V th cs # s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1689 |
and nhd: "\<not> holding (wq s) t c" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1690 |
and hd: "holding (wq (V th cs # s)) t c" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1691 |
show "next_th s th cs t \<and> c = cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1692 |
proof(cases "c = cs") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1693 |
case False |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1694 |
with nhd hd show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1695 |
by (unfold cs_holding_def wq_def, auto simp:Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1696 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1697 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1698 |
with step_back_step [OF vt] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1699 |
have "step s (V th c)" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1700 |
hence "next_th s th cs t" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1701 |
proof(cases) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1702 |
assume "holding s th c" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1703 |
with nhd hd show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1704 |
apply (unfold s_holding_def cs_holding_def wq_def next_th_def, |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1705 |
auto simp:Let_def split:list.splits if_splits) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1706 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1707 |
assume " hd (SOME q. distinct q \<and> q = []) \<in> set (SOME q. distinct q \<and> q = [])" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1708 |
moreover have "\<dots> = set []" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1709 |
proof(rule someI2) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1710 |
show "distinct [] \<and> [] = []" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1711 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1712 |
fix x assume "distinct x \<and> x = []" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1713 |
thus "set x = set []" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1714 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1715 |
ultimately show False by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1716 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1717 |
assume " hd (SOME q. distinct q \<and> q = []) \<in> set (SOME q. distinct q \<and> q = [])" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1718 |
moreover have "\<dots> = set []" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1719 |
proof(rule someI2) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1720 |
show "distinct [] \<and> [] = []" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1721 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1722 |
fix x assume "distinct x \<and> x = []" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1723 |
thus "set x = set []" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1724 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1725 |
ultimately show False by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1726 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1727 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1728 |
with True show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1729 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1730 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1731 |
|
104 | 1732 |
<<<<<<< local |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1733 |
lemma distinct_wq': "distinct wq'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1734 |
by (metis (mono_tags, lifting) distinct_rest some_eq_ex wq'_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1735 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1736 |
lemma set_wq': "set wq' = set rest" |
100 | 1737 |
by (metis (mono_tags, lifting) distinct_rest some_eq_ex wq'_def) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1738 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1739 |
lemma th'_in_inv: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1740 |
assumes "th' \<in> set wq'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1741 |
shows "th' \<in> set rest" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1742 |
using assms set_wq' by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1743 |
|
100 | 1744 |
lemma runing_th_s: |
1745 |
shows "th \<in> runing s" |
|
1746 |
proof - |
|
1747 |
from pip_e[unfolded is_v] |
|
1748 |
show ?thesis by (cases, simp) |
|
1749 |
qed |
|
1750 |
||
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1751 |
lemma neq_t_th: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1752 |
assumes "waiting (e#s) t c" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1753 |
shows "t \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1754 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1755 |
assume otherwise: "t = th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1756 |
show False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1757 |
proof(cases "c = cs") |
104 | 1758 |
======= |
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1759 |
text {* |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1760 |
The following @{text "step_v_wait_inv"} is also an obvious lemma, which, however, needs to be |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1761 |
derived from scratch, which confirms the correctness of the definition of @{text "next_th"}. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
1762 |
*} |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1763 |
lemma step_v_wait_inv[elim_format]: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1764 |
"\<And>t c. \<lbrakk>vt (V th cs # s); \<not> waiting (wq (V th cs # s)) t c; waiting (wq s) t c |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1765 |
\<rbrakk> |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1766 |
\<Longrightarrow> (next_th s th cs t \<and> cs = c)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1767 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1768 |
fix t c |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1769 |
assume vt: "vt (V th cs # s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1770 |
and nw: "\<not> waiting (wq (V th cs # s)) t c" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1771 |
and wt: "waiting (wq s) t c" |
63 | 1772 |
from vt interpret vt_v: valid_trace_e s "V th cs" |
1773 |
by (cases, unfold_locales, simp) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1774 |
show "next_th s th cs t \<and> cs = c" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1775 |
proof(cases "cs = c") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1776 |
case False |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1777 |
with nw wt show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1778 |
by (auto simp:cs_waiting_def wq_def Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1779 |
next |
104 | 1780 |
>>>>>>> other |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1781 |
case True |
104 | 1782 |
<<<<<<< local |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1783 |
have "t \<in> set wq'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1784 |
using assms[unfolded True s_waiting_def, folded wq_def, unfolded wq_es_cs] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1785 |
by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1786 |
from th'_in_inv[OF this] have "t \<in> set rest" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1787 |
with wq_s_cs[folded otherwise] wq_distinct[of cs] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1788 |
show ?thesis by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1789 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1790 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1791 |
have "wq (e#s) c = wq s c" using False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1792 |
by (unfold is_v, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1793 |
hence "waiting s t c" using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1794 |
by (simp add: cs_waiting_def waiting_eq) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1795 |
hence "t \<notin> readys s" by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1796 |
hence "t \<notin> runing s" using runing_ready by auto |
100 | 1797 |
with runing_th_s[folded otherwise] show ?thesis by auto |
104 | 1798 |
======= |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1799 |
from nw[folded True] wt[folded True] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1800 |
have "next_th s th cs t" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1801 |
apply (unfold next_th_def, auto simp:cs_waiting_def wq_def Let_def split:list.splits) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1802 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1803 |
fix a list |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1804 |
assume t_in: "t \<in> set list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1805 |
and t_ni: "t \<notin> set (SOME q. distinct q \<and> set q = set list)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1806 |
and eq_wq: "wq_fun (schs s) cs = a # list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1807 |
have " set (SOME q. distinct q \<and> set q = set list) = set list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1808 |
proof(rule someI2) |
63 | 1809 |
from vt_v.wq_distinct[of cs] and eq_wq[folded wq_def] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1810 |
show "distinct list \<and> set list = set list" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1811 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1812 |
show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1813 |
by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1814 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1815 |
with t_ni and t_in show "a = th" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1816 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1817 |
fix a list |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1818 |
assume t_in: "t \<in> set list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1819 |
and t_ni: "t \<notin> set (SOME q. distinct q \<and> set q = set list)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1820 |
and eq_wq: "wq_fun (schs s) cs = a # list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1821 |
have " set (SOME q. distinct q \<and> set q = set list) = set list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1822 |
proof(rule someI2) |
63 | 1823 |
from vt_v.wq_distinct[of cs] and eq_wq[folded wq_def] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1824 |
show "distinct list \<and> set list = set list" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1825 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1826 |
show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1827 |
by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1828 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1829 |
with t_ni and t_in show "t = hd (SOME q. distinct q \<and> set q = set list)" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1830 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1831 |
fix a list |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1832 |
assume eq_wq: "wq_fun (schs s) cs = a # list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1833 |
from step_back_step[OF vt] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1834 |
show "a = th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1835 |
proof(cases) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1836 |
assume "holding s th cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1837 |
with eq_wq show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1838 |
by (unfold s_holding_def wq_def, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1839 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1840 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1841 |
with True show ?thesis by simp |
104 | 1842 |
>>>>>>> other |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1843 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1844 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1845 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1846 |
lemma step_v_not_wait[consumes 3]: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1847 |
"\<lbrakk>vt (V th cs # s); next_th s th cs t; waiting (wq (V th cs # s)) t cs\<rbrakk> \<Longrightarrow> False" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1848 |
by (unfold next_th_def cs_waiting_def wq_def, auto simp:Let_def) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1849 |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1850 |
lemma step_v_release: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1851 |
"\<lbrakk>vt (V th cs # s); holding (wq (V th cs # s)) th cs\<rbrakk> \<Longrightarrow> False" |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1852 |
proof - |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1853 |
assume vt: "vt (V th cs # s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1854 |
and hd: "holding (wq (V th cs # s)) th cs" |
63 | 1855 |
from vt interpret vt_v: valid_trace_e s "V th cs" |
1856 |
by (cases, unfold_locales, simp+) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1857 |
from step_back_step [OF vt] and hd |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1858 |
show "False" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1859 |
proof(cases) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1860 |
assume "holding (wq (V th cs # s)) th cs" and "holding s th cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1861 |
thus ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1862 |
apply (unfold s_holding_def wq_def cs_holding_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1863 |
apply (auto simp:Let_def split:list.splits) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1864 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1865 |
fix list |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1866 |
assume eq_wq[folded wq_def]: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1867 |
"wq_fun (schs s) cs = hd (SOME q. distinct q \<and> set q = set list) # list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1868 |
and hd_in: "hd (SOME q. distinct q \<and> set q = set list) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1869 |
\<in> set (SOME q. distinct q \<and> set q = set list)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1870 |
have "set (SOME q. distinct q \<and> set q = set list) = set list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1871 |
proof(rule someI2) |
63 | 1872 |
from vt_v.wq_distinct[of cs] and eq_wq |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1873 |
show "distinct list \<and> set list = set list" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1874 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1875 |
show "\<And>x. distinct x \<and> set x = set list \<Longrightarrow> set x = set list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1876 |
by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1877 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1878 |
moreover have "distinct (hd (SOME q. distinct q \<and> set q = set list) # list)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1879 |
proof - |
63 | 1880 |
from vt_v.wq_distinct[of cs] and eq_wq |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1881 |
show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1882 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1883 |
moreover note eq_wq and hd_in |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1884 |
ultimately show "False" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1885 |
qed |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1886 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1887 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1888 |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1889 |
lemma step_v_get_hold: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1890 |
"\<And>th'. \<lbrakk>vt (V th cs # s); \<not> holding (wq (V th cs # s)) th' cs; next_th s th cs th'\<rbrakk> \<Longrightarrow> False" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1891 |
apply (unfold cs_holding_def next_th_def wq_def, |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1892 |
auto simp:Let_def) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1893 |
proof - |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1894 |
fix rest |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1895 |
assume vt: "vt (V th cs # s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1896 |
and eq_wq[folded wq_def]: " wq_fun (schs s) cs = th # rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1897 |
and nrest: "rest \<noteq> []" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1898 |
and ni: "hd (SOME q. distinct q \<and> set q = set rest) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1899 |
\<notin> set (SOME q. distinct q \<and> set q = set rest)" |
63 | 1900 |
from vt interpret vt_v: valid_trace_e s "V th cs" |
1901 |
by (cases, unfold_locales, simp+) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1902 |
have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1903 |
proof(rule someI2) |
63 | 1904 |
from vt_v.wq_distinct[of cs] and eq_wq |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1905 |
show "distinct rest \<and> set rest = set rest" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1906 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1907 |
fix x assume "distinct x \<and> set x = set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1908 |
hence "set x = set rest" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1909 |
with nrest |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1910 |
show "x \<noteq> []" by (case_tac x, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1911 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1912 |
with ni show "False" by auto |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1913 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1914 |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1915 |
lemma step_v_release_inv[elim_format]: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1916 |
"\<And>c t. \<lbrakk>vt (V th cs # s); \<not> holding (wq (V th cs # s)) t c; holding (wq s) t c\<rbrakk> \<Longrightarrow> |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1917 |
c = cs \<and> t = th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1918 |
apply (unfold cs_holding_def wq_def, auto simp:Let_def split:if_splits list.splits) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1919 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1920 |
fix a list |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1921 |
assume vt: "vt (V th cs # s)" and eq_wq: "wq_fun (schs s) cs = a # list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1922 |
from step_back_step [OF vt] show "a = th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1923 |
proof(cases) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1924 |
assume "holding s th cs" with eq_wq |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1925 |
show ?thesis |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1926 |
by (unfold s_holding_def wq_def, auto) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1927 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1928 |
next |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1929 |
fix a list |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1930 |
assume vt: "vt (V th cs # s)" and eq_wq: "wq_fun (schs s) cs = a # list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1931 |
from step_back_step [OF vt] show "a = th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1932 |
proof(cases) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1933 |
assume "holding s th cs" with eq_wq |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1934 |
show ?thesis |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1935 |
by (unfold s_holding_def wq_def, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1936 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1937 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1938 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1939 |
lemma step_v_waiting_mono: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1940 |
"\<And>t c. \<lbrakk>vt (V th cs # s); waiting (wq (V th cs # s)) t c\<rbrakk> \<Longrightarrow> waiting (wq s) t c" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1941 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1942 |
fix t c |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1943 |
let ?s' = "(V th cs # s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1944 |
assume vt: "vt ?s'" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1945 |
and wt: "waiting (wq ?s') t c" |
63 | 1946 |
from vt interpret vt_v: valid_trace_e s "V th cs" |
1947 |
by (cases, unfold_locales, simp+) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1948 |
show "waiting (wq s) t c" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1949 |
proof(cases "c = cs") |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1950 |
case False |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1951 |
assume neq_cs: "c \<noteq> cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1952 |
hence "waiting (wq ?s') t c = waiting (wq s) t c" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1953 |
by (unfold cs_waiting_def wq_def, auto simp:Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1954 |
with wt show ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1955 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1956 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1957 |
with wt show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1958 |
apply (unfold cs_waiting_def wq_def, auto simp:Let_def split:list.splits) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1959 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1960 |
fix a list |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1961 |
assume not_in: "t \<notin> set list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1962 |
and is_in: "t \<in> set (SOME q. distinct q \<and> set q = set list)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1963 |
and eq_wq: "wq_fun (schs s) cs = a # list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1964 |
have "set (SOME q. distinct q \<and> set q = set list) = set list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1965 |
proof(rule someI2) |
63 | 1966 |
from vt_v.wq_distinct [of cs] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1967 |
and eq_wq[folded wq_def] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1968 |
show "distinct list \<and> set list = set list" by auto |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1969 |
next |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1970 |
fix x assume "distinct x \<and> set x = set list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1971 |
thus "set x = set list" by auto |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1972 |
qed |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1973 |
with not_in is_in show "t = a" by auto |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1974 |
next |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1975 |
fix list |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1976 |
assume is_waiting: "waiting (wq (V th cs # s)) t cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1977 |
and eq_wq: "wq_fun (schs s) cs = t # list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1978 |
hence "t \<in> set list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1979 |
apply (unfold wq_def, auto simp:Let_def cs_waiting_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1980 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1981 |
assume " t \<in> set (SOME q. distinct q \<and> set q = set list)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1982 |
moreover have "\<dots> = set list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1983 |
proof(rule someI2) |
63 | 1984 |
from vt_v.wq_distinct [of cs] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1985 |
and eq_wq[folded wq_def] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1986 |
show "distinct list \<and> set list = set list" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1987 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1988 |
fix x assume "distinct x \<and> set x = set list" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1989 |
thus "set x = set list" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1990 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1991 |
ultimately show "t \<in> set list" by simp |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
1992 |
qed |
63 | 1993 |
with eq_wq and vt_v.wq_distinct [of cs, unfolded wq_def] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1994 |
show False by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1995 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1996 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1997 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1998 |
|
104 | 1999 |
<<<<<<< local |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2000 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2001 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2002 |
context valid_trace_v_n |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2003 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2004 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2005 |
lemma neq_wq': "wq' \<noteq> []" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2006 |
proof (unfold wq'_def, rule someI2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2007 |
show "distinct rest \<and> set rest = set rest" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2008 |
by (simp add: distinct_rest) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2009 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2010 |
fix x |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2011 |
assume " distinct x \<and> set x = set rest" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2012 |
thus "x \<noteq> []" using rest_nnl by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2013 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2014 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2015 |
lemma eq_wq': "wq' = taker # rest'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2016 |
by (simp add: neq_wq' rest'_def taker_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2017 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2018 |
lemma next_th_taker: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2019 |
shows "next_th s th cs taker" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2020 |
using rest_nnl taker_def wq'_def wq_s_cs |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2021 |
by (auto simp:next_th_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2022 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2023 |
lemma taker_unique: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2024 |
assumes "next_th s th cs taker'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2025 |
shows "taker' = taker" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2026 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2027 |
from assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2028 |
obtain rest' where |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2029 |
h: "wq s cs = th # rest'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2030 |
"taker' = hd (SOME q. distinct q \<and> set q = set rest')" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2031 |
by (unfold next_th_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2032 |
with wq_s_cs have "rest' = rest" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2033 |
thus ?thesis using h(2) taker_def wq'_def by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2034 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2035 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2036 |
lemma waiting_set_eq: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2037 |
"{(Th th', Cs cs) |th'. next_th s th cs th'} = {(Th taker, Cs cs)}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2038 |
by (smt all_not_in_conv bot.extremum insertI1 insert_subset |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2039 |
mem_Collect_eq next_th_taker subsetI subset_antisym taker_def taker_unique) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2040 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2041 |
lemma holding_set_eq: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2042 |
"{(Cs cs, Th th') |th'. next_th s th cs th'} = {(Cs cs, Th taker)}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2043 |
using next_th_taker taker_def waiting_set_eq |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2044 |
by fastforce |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2045 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2046 |
lemma holding_taker: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2047 |
shows "holding (e#s) taker cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2048 |
by (unfold s_holding_def, fold wq_def, unfold wq_es_cs, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2049 |
auto simp:neq_wq' taker_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2050 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2051 |
lemma waiting_esI2: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2052 |
assumes "waiting s t cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2053 |
and "t \<noteq> taker" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2054 |
shows "waiting (e#s) t cs" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2055 |
proof - |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2056 |
have "t \<in> set wq'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2057 |
proof(unfold wq'_def, rule someI2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2058 |
show "distinct rest \<and> set rest = set rest" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2059 |
by (simp add: distinct_rest) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2060 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2061 |
fix x |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2062 |
assume "distinct x \<and> set x = set rest" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2063 |
moreover have "t \<in> set rest" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2064 |
using assms(1) cs_waiting_def waiting_eq wq_s_cs by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2065 |
ultimately show "t \<in> set x" by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2066 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2067 |
moreover have "t \<noteq> hd wq'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2068 |
using assms(2) taker_def by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2069 |
ultimately show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2070 |
by (unfold s_waiting_def, fold wq_def, unfold wq_es_cs, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2071 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2072 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2073 |
lemma waiting_esE: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2074 |
assumes "waiting (e#s) t c" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2075 |
obtains "c \<noteq> cs" "waiting s t c" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2076 |
| "c = cs" "t \<noteq> taker" "waiting s t cs" "t \<in> set rest'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2077 |
proof(cases "c = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2078 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2079 |
hence "wq (e#s) c = wq s c" using is_v by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2080 |
with assms have "waiting s t c" using cs_waiting_def waiting_eq by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2081 |
from that(1)[OF False this] show ?thesis . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2082 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2083 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2084 |
from assms[unfolded s_waiting_def True, folded wq_def, unfolded wq_es_cs] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2085 |
have "t \<noteq> hd wq'" "t \<in> set wq'" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2086 |
hence "t \<noteq> taker" by (simp add: taker_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2087 |
moreover hence "t \<noteq> th" using assms neq_t_th by blast |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2088 |
moreover have "t \<in> set rest" by (simp add: `t \<in> set wq'` th'_in_inv) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2089 |
ultimately have "waiting s t cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2090 |
by (metis cs_waiting_def list.distinct(2) list.sel(1) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2091 |
list.set_sel(2) rest_def waiting_eq wq_s_cs) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2092 |
show ?thesis using that(2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2093 |
using True `t \<in> set wq'` `t \<noteq> taker` `waiting s t cs` eq_wq' by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2094 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2095 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2096 |
lemma holding_esI1: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2097 |
assumes "c = cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2098 |
and "t = taker" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2099 |
shows "holding (e#s) t c" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2100 |
by (unfold assms, simp add: holding_taker) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2101 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2102 |
lemma holding_esE: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2103 |
assumes "holding (e#s) t c" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2104 |
obtains "c = cs" "t = taker" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2105 |
| "c \<noteq> cs" "holding s t c" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2106 |
proof(cases "c = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2107 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2108 |
from assms[unfolded True, unfolded s_holding_def, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2109 |
folded wq_def, unfolded wq_es_cs] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2110 |
have "t = taker" by (simp add: taker_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2111 |
from that(1)[OF True this] show ?thesis . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2112 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2113 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2114 |
hence "wq (e#s) c = wq s c" using is_v by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2115 |
from assms[unfolded s_holding_def, folded wq_def, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2116 |
unfolded this, unfolded wq_def, folded s_holding_def] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2117 |
have "holding s t c" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2118 |
from that(2)[OF False this] show ?thesis . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2119 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2120 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2121 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2122 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2123 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2124 |
context valid_trace_v_e |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2125 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2126 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2127 |
lemma nil_wq': "wq' = []" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2128 |
proof (unfold wq'_def, rule someI2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2129 |
show "distinct rest \<and> set rest = set rest" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2130 |
by (simp add: distinct_rest) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2131 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2132 |
fix x |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2133 |
assume " distinct x \<and> set x = set rest" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2134 |
thus "x = []" using rest_nil by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2135 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2136 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2137 |
lemma no_taker: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2138 |
assumes "next_th s th cs taker" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2139 |
shows "False" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2140 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2141 |
from assms[unfolded next_th_def] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2142 |
obtain rest' where "wq s cs = th # rest'" "rest' \<noteq> []" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2143 |
by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2144 |
thus ?thesis using rest_def rest_nil by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2145 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2146 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2147 |
lemma waiting_set_eq: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2148 |
"{(Th th', Cs cs) |th'. next_th s th cs th'} = {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2149 |
using no_taker by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2150 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2151 |
lemma holding_set_eq: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2152 |
"{(Cs cs, Th th') |th'. next_th s th cs th'} = {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2153 |
using no_taker by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2154 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2155 |
lemma no_holding: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2156 |
assumes "holding (e#s) taker cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2157 |
shows False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2158 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2159 |
from wq_es_cs[unfolded nil_wq'] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2160 |
have " wq (e # s) cs = []" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2161 |
from assms[unfolded s_holding_def, folded wq_def, unfolded this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2162 |
show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2163 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2164 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2165 |
lemma no_waiting: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2166 |
assumes "waiting (e#s) t cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2167 |
shows False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2168 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2169 |
from wq_es_cs[unfolded nil_wq'] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2170 |
have " wq (e # s) cs = []" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2171 |
from assms[unfolded s_waiting_def, folded wq_def, unfolded this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2172 |
show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2173 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2174 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2175 |
lemma waiting_esI2: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2176 |
assumes "waiting s t c" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2177 |
shows "waiting (e#s) t c" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2178 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2179 |
have "c \<noteq> cs" using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2180 |
using cs_waiting_def rest_nil waiting_eq wq_s_cs by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2181 |
from waiting_esI1[OF assms this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2182 |
show ?thesis . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2183 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2184 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2185 |
lemma waiting_esE: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2186 |
assumes "waiting (e#s) t c" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2187 |
obtains "c \<noteq> cs" "waiting s t c" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2188 |
proof(cases "c = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2189 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2190 |
hence "wq (e#s) c = wq s c" using is_v by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2191 |
with assms have "waiting s t c" using cs_waiting_def waiting_eq by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2192 |
from that(1)[OF False this] show ?thesis . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2193 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2194 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2195 |
from no_waiting[OF assms[unfolded True]] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2196 |
show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2197 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2198 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2199 |
lemma holding_esE: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2200 |
assumes "holding (e#s) t c" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2201 |
obtains "c \<noteq> cs" "holding s t c" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2202 |
proof(cases "c = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2203 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2204 |
from no_holding[OF assms[unfolded True]] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2205 |
show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2206 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2207 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2208 |
hence "wq (e#s) c = wq s c" using is_v by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2209 |
from assms[unfolded s_holding_def, folded wq_def, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2210 |
unfolded this, unfolded wq_def, folded s_holding_def] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2211 |
have "holding s t c" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2212 |
from that[OF False this] show ?thesis . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2213 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2214 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2215 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2216 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2217 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2218 |
context valid_trace_v |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2219 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2220 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2221 |
lemma RAG_es: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2222 |
"RAG (e # s) = |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2223 |
RAG s - {(Cs cs, Th th)} - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2224 |
{(Th th', Cs cs) |th'. next_th s th cs th'} \<union> |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2225 |
{(Cs cs, Th th') |th'. next_th s th cs th'}" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2226 |
proof(rule rel_eqI) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2227 |
fix n1 n2 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2228 |
assume "(n1, n2) \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2229 |
thus "(n1, n2) \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2230 |
proof(cases rule:in_RAG_E) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2231 |
case (waiting th' cs') |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2232 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2233 |
proof(cases "rest = []") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2234 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2235 |
interpret h_n: valid_trace_v_n s e th cs |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2236 |
by (unfold_locales, insert False, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2237 |
from waiting(3) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2238 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2239 |
proof(cases rule:h_n.waiting_esE) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2240 |
case 1 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2241 |
with waiting(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2242 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2243 |
by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2244 |
fold waiting_eq, auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2245 |
next |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2246 |
case 2 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2247 |
with waiting(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2248 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2249 |
by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2250 |
fold waiting_eq, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2251 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2252 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2253 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2254 |
interpret h_e: valid_trace_v_e s e th cs |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2255 |
by (unfold_locales, insert True, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2256 |
from waiting(3) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2257 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2258 |
proof(cases rule:h_e.waiting_esE) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2259 |
case 1 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2260 |
with waiting(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2261 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2262 |
by (unfold h_e.waiting_set_eq h_e.holding_set_eq s_RAG_def, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2263 |
fold waiting_eq, auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2264 |
qed |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2265 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2266 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2267 |
case (holding th' cs') |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2268 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2269 |
proof(cases "rest = []") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2270 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2271 |
interpret h_n: valid_trace_v_n s e th cs |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2272 |
by (unfold_locales, insert False, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2273 |
from holding(3) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2274 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2275 |
proof(cases rule:h_n.holding_esE) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2276 |
case 1 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2277 |
with holding(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2278 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2279 |
by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2280 |
fold waiting_eq, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2281 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2282 |
case 2 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2283 |
with holding(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2284 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2285 |
by (unfold h_n.waiting_set_eq h_n.holding_set_eq s_RAG_def, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2286 |
fold holding_eq, auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2287 |
qed |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2288 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2289 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2290 |
interpret h_e: valid_trace_v_e s e th cs |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2291 |
by (unfold_locales, insert True, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2292 |
from holding(3) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2293 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2294 |
proof(cases rule:h_e.holding_esE) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2295 |
case 1 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2296 |
with holding(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2297 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2298 |
by (unfold h_e.waiting_set_eq h_e.holding_set_eq s_RAG_def, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2299 |
fold holding_eq, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2300 |
qed |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2301 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2302 |
qed |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2303 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2304 |
fix n1 n2 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2305 |
assume h: "(n1, n2) \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2306 |
show "(n1, n2) \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2307 |
proof(cases "rest = []") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2308 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2309 |
interpret h_n: valid_trace_v_n s e th cs |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2310 |
by (unfold_locales, insert False, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2311 |
from h[unfolded h_n.waiting_set_eq h_n.holding_set_eq] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2312 |
have "((n1, n2) \<in> RAG s \<and> (n1 \<noteq> Cs cs \<or> n2 \<noteq> Th th) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2313 |
\<and> (n1 \<noteq> Th h_n.taker \<or> n2 \<noteq> Cs cs)) \<or> |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2314 |
(n2 = Th h_n.taker \<and> n1 = Cs cs)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2315 |
by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2316 |
thus ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2317 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2318 |
assume "n2 = Th h_n.taker \<and> n1 = Cs cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2319 |
with h_n.holding_taker |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2320 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2321 |
by (unfold s_RAG_def, fold holding_eq, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2322 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2323 |
assume h: "(n1, n2) \<in> RAG s \<and> |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2324 |
(n1 \<noteq> Cs cs \<or> n2 \<noteq> Th th) \<and> (n1 \<noteq> Th h_n.taker \<or> n2 \<noteq> Cs cs)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2325 |
hence "(n1, n2) \<in> RAG s" by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2326 |
thus ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2327 |
proof(cases rule:in_RAG_E) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2328 |
case (waiting th' cs') |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2329 |
from h and this(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2330 |
have "th' \<noteq> h_n.taker \<or> cs' \<noteq> cs" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2331 |
hence "waiting (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2332 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2333 |
assume "cs' \<noteq> cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2334 |
from waiting_esI1[OF waiting(3) this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2335 |
show ?thesis . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2336 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2337 |
assume neq_th': "th' \<noteq> h_n.taker" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2338 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2339 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2340 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2341 |
from waiting_esI1[OF waiting(3) this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2342 |
show ?thesis . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2343 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2344 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2345 |
from h_n.waiting_esI2[OF waiting(3)[unfolded True] neq_th', folded True] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2346 |
show ?thesis . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2347 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2348 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2349 |
thus ?thesis using waiting(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2350 |
by (unfold s_RAG_def, fold waiting_eq, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2351 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2352 |
case (holding th' cs') |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2353 |
from h this(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2354 |
have "cs' \<noteq> cs \<or> th' \<noteq> th" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2355 |
hence "holding (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2356 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2357 |
assume "cs' \<noteq> cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2358 |
from holding_esI2[OF this holding(3)] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2359 |
show ?thesis . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2360 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2361 |
assume "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2362 |
from holding_esI1[OF holding(3) this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2363 |
show ?thesis . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2364 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2365 |
thus ?thesis using holding(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2366 |
by (unfold s_RAG_def, fold holding_eq, auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2367 |
qed |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2368 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2369 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2370 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2371 |
interpret h_e: valid_trace_v_e s e th cs |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2372 |
by (unfold_locales, insert True, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2373 |
from h[unfolded h_e.waiting_set_eq h_e.holding_set_eq] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2374 |
have h_s: "(n1, n2) \<in> RAG s" "(n1, n2) \<noteq> (Cs cs, Th th)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2375 |
by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2376 |
from h_s(1) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2377 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2378 |
proof(cases rule:in_RAG_E) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2379 |
case (waiting th' cs') |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2380 |
from h_e.waiting_esI2[OF this(3)] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2381 |
show ?thesis using waiting(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2382 |
by (unfold s_RAG_def, fold waiting_eq, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2383 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2384 |
case (holding th' cs') |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2385 |
with h_s(2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2386 |
have "cs' \<noteq> cs \<or> th' \<noteq> th" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2387 |
thus ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2388 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2389 |
assume neq_cs: "cs' \<noteq> cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2390 |
from holding_esI2[OF this holding(3)] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2391 |
show ?thesis using holding(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2392 |
by (unfold s_RAG_def, fold holding_eq, auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2393 |
next |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2394 |
assume "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2395 |
from holding_esI1[OF holding(3) this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2396 |
show ?thesis using holding(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2397 |
by (unfold s_RAG_def, fold holding_eq, auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2398 |
qed |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2399 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2400 |
qed |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2401 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2402 |
|
100 | 2403 |
lemma |
2404 |
finite_RAG_kept: |
|
2405 |
assumes "finite (RAG s)" |
|
2406 |
shows "finite (RAG (e#s))" |
|
2407 |
proof(cases "rest = []") |
|
2408 |
case True |
|
2409 |
interpret vt: valid_trace_v_e using True |
|
2410 |
by (unfold_locales, simp) |
|
2411 |
show ?thesis using assms |
|
2412 |
by (unfold RAG_es vt.waiting_set_eq vt.holding_set_eq, simp) |
|
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2413 |
next |
100 | 2414 |
case False |
2415 |
interpret vt: valid_trace_v_n using False |
|
2416 |
by (unfold_locales, simp) |
|
2417 |
show ?thesis using assms |
|
2418 |
by (unfold RAG_es vt.waiting_set_eq vt.holding_set_eq, simp) |
|
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2419 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2420 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2421 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2422 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2423 |
context valid_trace_p |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2424 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2425 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2426 |
lemma waiting_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2427 |
assumes "waiting s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2428 |
shows "waiting (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2429 |
using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2430 |
by (metis cs_waiting_def hd_append2 list.sel(1) list.set_intros(2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2431 |
rotate1.simps(2) self_append_conv2 set_rotate1 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2432 |
th_not_in_wq waiting_eq wq_es_cs wq_neq_simp) |
100 | 2433 |
|
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2434 |
lemma holding_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2435 |
assumes "holding s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2436 |
shows "holding (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2437 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2438 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2439 |
hence "wq (e#s) cs' = wq s cs'" by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2440 |
with assms show ?thesis using cs_holding_def holding_eq by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2441 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2442 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2443 |
from assms[unfolded s_holding_def, folded wq_def] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2444 |
obtain rest where eq_wq: "wq s cs' = th'#rest" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2445 |
by (metis empty_iff list.collapse list.set(1)) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2446 |
hence "wq (e#s) cs' = th'#(rest@[th])" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2447 |
by (simp add: True wq_es_cs) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2448 |
thus ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2449 |
by (simp add: cs_holding_def holding_eq) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2450 |
qed |
100 | 2451 |
end |
2452 |
||
2453 |
lemma (in valid_trace_p) th_not_waiting: "\<not> waiting s th c" |
|
2454 |
proof - |
|
2455 |
have "th \<in> readys s" |
|
2456 |
using runing_ready runing_th_s by blast |
|
2457 |
thus ?thesis |
|
2458 |
by (unfold readys_def, auto) |
|
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2459 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2460 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2461 |
context valid_trace_p_h |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2462 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2463 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2464 |
lemma wq_es_cs': "wq (e#s) cs = [th]" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2465 |
using wq_es_cs[unfolded we] by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2466 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2467 |
lemma holding_es_th_cs: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2468 |
shows "holding (e#s) th cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2469 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2470 |
from wq_es_cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2471 |
have "th \<in> set (wq (e#s) cs)" "th = hd (wq (e#s) cs)" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2472 |
thus ?thesis using cs_holding_def holding_eq by blast |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2473 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2474 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2475 |
lemma RAG_edge: "(Cs cs, Th th) \<in> RAG (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2476 |
by (unfold s_RAG_def, fold holding_eq, insert holding_es_th_cs, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2477 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2478 |
lemma waiting_esE: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2479 |
assumes "waiting (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2480 |
obtains "waiting s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2481 |
using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2482 |
by (metis cs_waiting_def event.distinct(15) is_p list.sel(1) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2483 |
set_ConsD waiting_eq we wq_es_cs' wq_neq_simp wq_out_inv) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2484 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2485 |
lemma holding_esE: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2486 |
assumes "holding (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2487 |
obtains "cs' \<noteq> cs" "holding s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2488 |
| "cs' = cs" "th' = th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2489 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2490 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2491 |
from held_unique[OF holding_es_th_cs assms[unfolded True]] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2492 |
have "th' = th" by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2493 |
from that(2)[OF True this] show ?thesis . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2494 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2495 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2496 |
have "holding s th' cs'" using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2497 |
using False cs_holding_def holding_eq by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2498 |
from that(1)[OF False this] show ?thesis . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2499 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2500 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2501 |
lemma RAG_es: "RAG (e # s) = RAG s \<union> {(Cs cs, Th th)}" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2502 |
proof(rule rel_eqI) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2503 |
fix n1 n2 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2504 |
assume "(n1, n2) \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2505 |
thus "(n1, n2) \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2506 |
proof(cases rule:in_RAG_E) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2507 |
case (waiting th' cs') |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2508 |
from this(3) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2509 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2510 |
proof(cases rule:waiting_esE) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2511 |
case 1 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2512 |
thus ?thesis using waiting(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2513 |
by (unfold s_RAG_def, fold waiting_eq, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2514 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2515 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2516 |
case (holding th' cs') |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2517 |
from this(3) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2518 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2519 |
proof(cases rule:holding_esE) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2520 |
case 1 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2521 |
with holding(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2522 |
show ?thesis by (unfold s_RAG_def, fold holding_eq, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2523 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2524 |
case 2 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2525 |
with holding(1,2) show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2526 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2527 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2528 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2529 |
fix n1 n2 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2530 |
assume "(n1, n2) \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2531 |
hence "(n1, n2) \<in> RAG s \<or> (n1 = Cs cs \<and> n2 = Th th)" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2532 |
thus "(n1, n2) \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2533 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2534 |
assume "(n1, n2) \<in> RAG s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2535 |
thus ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2536 |
proof(cases rule:in_RAG_E) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2537 |
case (waiting th' cs') |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2538 |
from waiting_kept[OF this(3)] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2539 |
show ?thesis using waiting(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2540 |
by (unfold s_RAG_def, fold waiting_eq, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2541 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2542 |
case (holding th' cs') |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2543 |
from holding_kept[OF this(3)] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2544 |
show ?thesis using holding(1,2) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2545 |
by (unfold s_RAG_def, fold holding_eq, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2546 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2547 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2548 |
assume "n1 = Cs cs \<and> n2 = Th th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2549 |
with holding_es_th_cs |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2550 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2551 |
by (unfold s_RAG_def, fold holding_eq, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2552 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2553 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2554 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2555 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2556 |
|
100 | 2557 |
context valid_trace_p_w |
2558 |
begin |
|
2559 |
||
2560 |
lemma wq_s_cs: "wq s cs = holder#waiters" |
|
2561 |
by (simp add: holder_def waiters_def wne) |
|
2562 |
||
2563 |
lemma wq_es_cs': "wq (e#s) cs = holder#waiters@[th]" |
|
2564 |
by (simp add: wq_es_cs wq_s_cs) |
|
2565 |
||
2566 |
lemma waiting_es_th_cs: "waiting (e#s) th cs" |
|
2567 |
using cs_waiting_def th_not_in_wq waiting_eq wq_es_cs' wq_s_cs by auto |
|
2568 |
||
2569 |
lemma RAG_edge: "(Th th, Cs cs) \<in> RAG (e#s)" |
|
2570 |
by (unfold s_RAG_def, fold waiting_eq, insert waiting_es_th_cs, auto) |
|
2571 |
||
2572 |
lemma holding_esE: |
|
2573 |
assumes "holding (e#s) th' cs'" |
|
2574 |
obtains "holding s th' cs'" |
|
2575 |
using assms |
|
2576 |
proof(cases "cs' = cs") |
|
2577 |
case False |
|
2578 |
hence "wq (e#s) cs' = wq s cs'" by simp |
|
2579 |
with assms show ?thesis |
|
2580 |
using cs_holding_def holding_eq that by auto |
|
2581 |
next |
|
2582 |
case True |
|
2583 |
with assms show ?thesis |
|
2584 |
by (metis cs_holding_def holding_eq list.sel(1) list.set_intros(1) that |
|
2585 |
wq_es_cs' wq_s_cs) |
|
2586 |
qed |
|
2587 |
||
2588 |
lemma waiting_esE: |
|
2589 |
assumes "waiting (e#s) th' cs'" |
|
2590 |
obtains "th' \<noteq> th" "waiting s th' cs'" |
|
2591 |
| "th' = th" "cs' = cs" |
|
2592 |
proof(cases "waiting s th' cs'") |
|
2593 |
case True |
|
2594 |
have "th' \<noteq> th" |
|
2595 |
proof |
|
2596 |
assume otherwise: "th' = th" |
|
2597 |
from True[unfolded this] |
|
2598 |
show False by (simp add: th_not_waiting) |
|
2599 |
qed |
|
2600 |
from that(1)[OF this True] show ?thesis . |
|
2601 |
next |
|
2602 |
case False |
|
2603 |
hence "th' = th \<and> cs' = cs" |
|
2604 |
by (metis assms cs_waiting_def holder_def list.sel(1) rotate1.simps(2) |
|
2605 |
set_ConsD set_rotate1 waiting_eq wq_es_cs wq_es_cs' wq_neq_simp) |
|
2606 |
with that(2) show ?thesis by metis |
|
2607 |
qed |
|
2608 |
||
2609 |
lemma RAG_es: "RAG (e # s) = RAG s \<union> {(Th th, Cs cs)}" (is "?L = ?R") |
|
2610 |
proof(rule rel_eqI) |
|
2611 |
fix n1 n2 |
|
2612 |
assume "(n1, n2) \<in> ?L" |
|
2613 |
thus "(n1, n2) \<in> ?R" |
|
2614 |
proof(cases rule:in_RAG_E) |
|
2615 |
case (waiting th' cs') |
|
2616 |
from this(3) |
|
2617 |
show ?thesis |
|
2618 |
proof(cases rule:waiting_esE) |
|
2619 |
case 1 |
|
2620 |
thus ?thesis using waiting(1,2) |
|
2621 |
by (unfold s_RAG_def, fold waiting_eq, auto) |
|
2622 |
next |
|
2623 |
case 2 |
|
2624 |
thus ?thesis using waiting(1,2) by auto |
|
2625 |
qed |
|
2626 |
next |
|
2627 |
case (holding th' cs') |
|
2628 |
from this(3) |
|
2629 |
show ?thesis |
|
2630 |
proof(cases rule:holding_esE) |
|
2631 |
case 1 |
|
2632 |
with holding(1,2) |
|
2633 |
show ?thesis by (unfold s_RAG_def, fold holding_eq, auto) |
|
2634 |
qed |
|
2635 |
qed |
|
2636 |
next |
|
2637 |
fix n1 n2 |
|
2638 |
assume "(n1, n2) \<in> ?R" |
|
2639 |
hence "(n1, n2) \<in> RAG s \<or> (n1 = Th th \<and> n2 = Cs cs)" by auto |
|
2640 |
thus "(n1, n2) \<in> ?L" |
|
2641 |
proof |
|
2642 |
assume "(n1, n2) \<in> RAG s" |
|
2643 |
thus ?thesis |
|
2644 |
proof(cases rule:in_RAG_E) |
|
2645 |
case (waiting th' cs') |
|
2646 |
from waiting_kept[OF this(3)] |
|
2647 |
show ?thesis using waiting(1,2) |
|
2648 |
by (unfold s_RAG_def, fold waiting_eq, auto) |
|
2649 |
next |
|
2650 |
case (holding th' cs') |
|
2651 |
from holding_kept[OF this(3)] |
|
2652 |
show ?thesis using holding(1,2) |
|
2653 |
by (unfold s_RAG_def, fold holding_eq, auto) |
|
2654 |
qed |
|
2655 |
next |
|
2656 |
assume "n1 = Th th \<and> n2 = Cs cs" |
|
2657 |
thus ?thesis using RAG_edge by auto |
|
2658 |
qed |
|
2659 |
qed |
|
2660 |
||
2661 |
end |
|
2662 |
||
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2663 |
context valid_trace_p |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2664 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2665 |
|
100 | 2666 |
lemma RAG_es: "RAG (e # s) = (if (wq s cs = []) then RAG s \<union> {(Cs cs, Th th)} |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2667 |
else RAG s \<union> {(Th th, Cs cs)})" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2668 |
proof(cases "wq s cs = []") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2669 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2670 |
interpret vt_p: valid_trace_p_h using True |
63 | 2671 |
by (unfold_locales, simp) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2672 |
show ?thesis by (simp add: vt_p.RAG_es vt_p.we) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2673 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2674 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2675 |
interpret vt_p: valid_trace_p_w using False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2676 |
by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2677 |
show ?thesis by (simp add: vt_p.RAG_es vt_p.wne) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2678 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2679 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2680 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2681 |
|
100 | 2682 |
section {* Finiteness of RAG *} |
2683 |
||
2684 |
context valid_trace |
|
2685 |
begin |
|
2686 |
||
2687 |
lemma finite_RAG: |
|
2688 |
shows "finite (RAG s)" |
|
2689 |
proof(induct rule:ind) |
|
2690 |
case Nil |
|
2691 |
show ?case |
|
2692 |
by (auto simp: s_RAG_def cs_waiting_def |
|
2693 |
cs_holding_def wq_def acyclic_def) |
|
2694 |
next |
|
2695 |
case (Cons s e) |
|
2696 |
interpret vt_e: valid_trace_e s e using Cons by simp |
|
2697 |
show ?case |
|
2698 |
proof(cases e) |
|
2699 |
case (Create th prio) |
|
2700 |
interpret vt: valid_trace_create s e th prio using Create |
|
2701 |
by (unfold_locales, simp) |
|
2702 |
show ?thesis using Cons by simp |
|
2703 |
next |
|
2704 |
case (Exit th) |
|
2705 |
interpret vt: valid_trace_exit s e th using Exit |
|
2706 |
by (unfold_locales, simp) |
|
2707 |
show ?thesis using Cons by simp |
|
2708 |
next |
|
2709 |
case (P th cs) |
|
2710 |
interpret vt: valid_trace_p s e th cs using P |
|
2711 |
by (unfold_locales, simp) |
|
2712 |
show ?thesis using Cons using vt.RAG_es by auto |
|
2713 |
next |
|
2714 |
case (V th cs) |
|
2715 |
interpret vt: valid_trace_v s e th cs using V |
|
2716 |
by (unfold_locales, simp) |
|
2717 |
show ?thesis using Cons by (simp add: vt.finite_RAG_kept) |
|
2718 |
next |
|
2719 |
case (Set th prio) |
|
2720 |
interpret vt: valid_trace_set s e th prio using Set |
|
2721 |
by (unfold_locales, simp) |
|
2722 |
show ?thesis using Cons by simp |
|
2723 |
qed |
|
2724 |
qed |
|
2725 |
end |
|
2726 |
||
2727 |
section {* RAG is acyclic *} |
|
2728 |
||
2729 |
text {* (* ddd *) |
|
2730 |
The nature of the work is like this: since it starts from a very simple and basic |
|
2731 |
model, even intuitively very `basic` and `obvious` properties need to derived from scratch. |
|
2732 |
For instance, the fact |
|
2733 |
that one thread can not be blocked by two critical resources at the same time |
|
2734 |
is obvious, because only running threads can make new requests, if one is waiting for |
|
2735 |
a critical resource and get blocked, it can not make another resource request and get |
|
2736 |
blocked the second time (because it is not running). |
|
2737 |
||
2738 |
To derive this fact, one needs to prove by contraction and |
|
2739 |
reason about time (or @{text "moement"}). The reasoning is based on a generic theorem |
|
2740 |
named @{text "p_split"}, which is about status changing along the time axis. It says if |
|
2741 |
a condition @{text "Q"} is @{text "True"} at a state @{text "s"}, |
|
2742 |
but it was @{text "False"} at the very beginning, then there must exits a moment @{text "t"} |
|
2743 |
in the history of @{text "s"} (notice that @{text "s"} itself is essentially the history |
|
2744 |
of events leading to it), such that @{text "Q"} switched |
|
2745 |
from being @{text "False"} to @{text "True"} and kept being @{text "True"} |
|
2746 |
till the last moment of @{text "s"}. |
|
2747 |
||
2748 |
Suppose a thread @{text "th"} is blocked |
|
2749 |
on @{text "cs1"} and @{text "cs2"} in some state @{text "s"}, |
|
2750 |
since no thread is blocked at the very beginning, by applying |
|
2751 |
@{text "p_split"} to these two blocking facts, there exist |
|
2752 |
two moments @{text "t1"} and @{text "t2"} in @{text "s"}, such that |
|
2753 |
@{text "th"} got blocked on @{text "cs1"} and @{text "cs2"} |
|
2754 |
and kept on blocked on them respectively ever since. |
|
2755 |
||
2756 |
Without lost of generality, we assume @{text "t1"} is earlier than @{text "t2"}. |
|
2757 |
However, since @{text "th"} was blocked ever since memonent @{text "t1"}, so it was still |
|
2758 |
in blocked state at moment @{text "t2"} and could not |
|
2759 |
make any request and get blocked the second time: Contradiction. |
|
2760 |
*} |
|
2761 |
||
2762 |
||
2763 |
context valid_trace |
|
2764 |
begin |
|
2765 |
||
2766 |
lemma waiting_unique_pre: (* ddd *) |
|
2767 |
assumes h11: "thread \<in> set (wq s cs1)" |
|
2768 |
and h12: "thread \<noteq> hd (wq s cs1)" |
|
2769 |
assumes h21: "thread \<in> set (wq s cs2)" |
|
2770 |
and h22: "thread \<noteq> hd (wq s cs2)" |
|
2771 |
and neq12: "cs1 \<noteq> cs2" |
|
2772 |
shows "False" |
|
2773 |
proof - |
|
2774 |
let "?Q" = "\<lambda> cs s. thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs)" |
|
2775 |
from h11 and h12 have q1: "?Q cs1 s" by simp |
|
2776 |
from h21 and h22 have q2: "?Q cs2 s" by simp |
|
2777 |
have nq1: "\<not> ?Q cs1 []" by (simp add:wq_def) |
|
2778 |
have nq2: "\<not> ?Q cs2 []" by (simp add:wq_def) |
|
2779 |
from p_split [of "?Q cs1", OF q1 nq1] |
|
2780 |
obtain t1 where lt1: "t1 < length s" |
|
2781 |
and np1: "\<not> ?Q cs1 (moment t1 s)" |
|
2782 |
and nn1: "(\<forall>i'>t1. ?Q cs1 (moment i' s))" by auto |
|
2783 |
from p_split [of "?Q cs2", OF q2 nq2] |
|
2784 |
obtain t2 where lt2: "t2 < length s" |
|
2785 |
and np2: "\<not> ?Q cs2 (moment t2 s)" |
|
2786 |
and nn2: "(\<forall>i'>t2. ?Q cs2 (moment i' s))" by auto |
|
2787 |
{ fix s cs |
|
2788 |
assume q: "?Q cs s" |
|
2789 |
have "thread \<notin> runing s" |
|
2790 |
proof |
|
2791 |
assume "thread \<in> runing s" |
|
2792 |
hence " \<forall>cs. \<not> (thread \<in> set (wq_fun (schs s) cs) \<and> |
|
2793 |
thread \<noteq> hd (wq_fun (schs s) cs))" |
|
2794 |
by (unfold runing_def s_waiting_def readys_def, auto) |
|
2795 |
from this[rule_format, of cs] q |
|
2796 |
show False by (simp add: wq_def) |
|
2797 |
qed |
|
2798 |
} note q_not_runing = this |
|
2799 |
{ fix t1 t2 cs1 cs2 |
|
2800 |
assume lt1: "t1 < length s" |
|
2801 |
and np1: "\<not> ?Q cs1 (moment t1 s)" |
|
2802 |
and nn1: "(\<forall>i'>t1. ?Q cs1 (moment i' s))" |
|
2803 |
and lt2: "t2 < length s" |
|
2804 |
and np2: "\<not> ?Q cs2 (moment t2 s)" |
|
2805 |
and nn2: "(\<forall>i'>t2. ?Q cs2 (moment i' s))" |
|
2806 |
and lt12: "t1 < t2" |
|
2807 |
let ?t3 = "Suc t2" |
|
2808 |
interpret ve2: valid_moment_e _ t2 using lt2 |
|
2809 |
by (unfold_locales, simp) |
|
2810 |
let ?e = ve2.next_e |
|
2811 |
have "t2 < ?t3" by simp |
|
2812 |
from nn2 [rule_format, OF this] and ve2.trace_e |
|
2813 |
have h1: "thread \<in> set (wq (?e#moment t2 s) cs2)" and |
|
2814 |
h2: "thread \<noteq> hd (wq (?e#moment t2 s) cs2)" by auto |
|
2815 |
have ?thesis |
|
2816 |
proof - |
|
2817 |
have "thread \<in> runing (moment t2 s)" |
|
2818 |
proof(cases "thread \<in> set (wq (moment t2 s) cs2)") |
|
2819 |
case True |
|
2820 |
have "?e = V thread cs2" |
|
2821 |
proof - |
|
2822 |
have eq_th: "thread = hd (wq (moment t2 s) cs2)" |
|
2823 |
using True and np2 by auto |
|
2824 |
thus ?thesis |
|
2825 |
using True h2 ve2.vat_moment_e.wq_out_inv by blast |
|
2826 |
qed |
|
2827 |
thus ?thesis |
|
2828 |
using step.cases ve2.vat_moment_e.pip_e by auto |
|
2829 |
next |
|
2830 |
case False |
|
2831 |
hence "?e = P thread cs2" |
|
2832 |
using h1 ve2.vat_moment_e.wq_in_inv by blast |
|
2833 |
thus ?thesis |
|
2834 |
using step.cases ve2.vat_moment_e.pip_e by auto |
|
2835 |
qed |
|
2836 |
moreover have "thread \<notin> runing (moment t2 s)" |
|
2837 |
by (rule q_not_runing[OF nn1[rule_format, OF lt12]]) |
|
2838 |
ultimately show ?thesis by simp |
|
2839 |
qed |
|
2840 |
} note lt_case = this |
|
2841 |
show ?thesis |
|
2842 |
proof - |
|
2843 |
{ assume "t1 < t2" |
|
2844 |
from lt_case[OF lt1 np1 nn1 lt2 np2 nn2 this] |
|
2845 |
have ?thesis . |
|
2846 |
} moreover { |
|
2847 |
assume "t2 < t1" |
|
2848 |
from lt_case[OF lt2 np2 nn2 lt1 np1 nn1 this] |
|
2849 |
have ?thesis . |
|
2850 |
} moreover { |
|
2851 |
assume eq_12: "t1 = t2" |
|
2852 |
let ?t3 = "Suc t2" |
|
2853 |
interpret ve2: valid_moment_e _ t2 using lt2 |
|
2854 |
by (unfold_locales, simp) |
|
2855 |
let ?e = ve2.next_e |
|
2856 |
have "t2 < ?t3" by simp |
|
2857 |
from nn2 [rule_format, OF this] and ve2.trace_e |
|
2858 |
have h1: "thread \<in> set (wq (?e#moment t2 s) cs2)" by auto |
|
2859 |
have lt_2: "t2 < ?t3" by simp |
|
2860 |
from nn2 [rule_format, OF this] and ve2.trace_e |
|
2861 |
have h1: "thread \<in> set (wq (?e#moment t2 s) cs2)" and |
|
2862 |
h2: "thread \<noteq> hd (wq (?e#moment t2 s) cs2)" by auto |
|
2863 |
from nn1[rule_format, OF lt_2[folded eq_12], unfolded ve2.trace_e[folded eq_12]] |
|
2864 |
eq_12[symmetric] |
|
2865 |
have g1: "thread \<in> set (wq (?e#moment t1 s) cs1)" and |
|
2866 |
g2: "thread \<noteq> hd (wq (?e#moment t1 s) cs1)" by auto |
|
2867 |
have "?e = V thread cs2 \<or> ?e = P thread cs2" |
|
2868 |
using h1 h2 np2 ve2.vat_moment_e.wq_in_inv |
|
2869 |
ve2.vat_moment_e.wq_out_inv by blast |
|
2870 |
moreover have "?e = V thread cs1 \<or> ?e = P thread cs1" |
|
2871 |
using eq_12 g1 g2 np1 ve2.vat_moment_e.wq_in_inv |
|
2872 |
ve2.vat_moment_e.wq_out_inv by blast |
|
2873 |
ultimately have ?thesis using neq12 by auto |
|
2874 |
} ultimately show ?thesis using nat_neq_iff by blast |
|
2875 |
qed |
|
2876 |
qed |
|
2877 |
||
2878 |
text {* |
|
2879 |
This lemma is a simple corrolary of @{text "waiting_unique_pre"}. |
|
2880 |
*} |
|
2881 |
||
2882 |
lemma waiting_unique: |
|
2883 |
assumes "waiting s th cs1" |
|
2884 |
and "waiting s th cs2" |
|
2885 |
shows "cs1 = cs2" |
|
2886 |
using waiting_unique_pre assms |
|
2887 |
unfolding wq_def s_waiting_def |
|
2888 |
by auto |
|
2889 |
||
2890 |
end |
|
2891 |
||
2892 |
lemma (in valid_trace_v) |
|
2893 |
preced_es [simp]: "preced th (e#s) = preced th s" |
|
2894 |
by (unfold is_v preced_def, simp) |
|
2895 |
||
2896 |
lemma the_preced_v[simp]: "the_preced (V th cs#s) = the_preced s" |
|
2897 |
proof |
|
2898 |
fix th' |
|
2899 |
show "the_preced (V th cs # s) th' = the_preced s th'" |
|
2900 |
by (unfold the_preced_def preced_def, simp) |
|
2901 |
qed |
|
2902 |
||
2903 |
||
2904 |
lemma (in valid_trace_v) |
|
2905 |
the_preced_es: "the_preced (e#s) = the_preced s" |
|
2906 |
by (unfold is_v preced_def, simp) |
|
2907 |
||
2908 |
context valid_trace_p |
|
2909 |
begin |
|
2910 |
||
2911 |
lemma not_holding_s_th_cs: "\<not> holding s th cs" |
|
2912 |
proof |
|
2913 |
assume otherwise: "holding s th cs" |
|
2914 |
from pip_e[unfolded is_p] |
|
2915 |
show False |
|
2916 |
proof(cases) |
|
2917 |
case (thread_P) |
|
2918 |
moreover have "(Cs cs, Th th) \<in> RAG s" |
|
2919 |
using otherwise cs_holding_def |
|
2920 |
holding_eq th_not_in_wq by auto |
|
2921 |
ultimately show ?thesis by auto |
|
2922 |
qed |
|
2923 |
qed |
|
2924 |
||
2925 |
end |
|
2926 |
||
2927 |
||
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2928 |
lemma (in valid_trace_v_n) finite_waiting_set: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2929 |
"finite {(Th th', Cs cs) |th'. next_th s th cs th'}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2930 |
by (simp add: waiting_set_eq) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2931 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2932 |
lemma (in valid_trace_v_n) finite_holding_set: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2933 |
"finite {(Cs cs, Th th') |th'. next_th s th cs th'}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2934 |
by (simp add: holding_set_eq) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2935 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2936 |
lemma (in valid_trace_v_e) finite_waiting_set: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2937 |
"finite {(Th th', Cs cs) |th'. next_th s th cs th'}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2938 |
by (simp add: waiting_set_eq) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2939 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2940 |
lemma (in valid_trace_v_e) finite_holding_set: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2941 |
"finite {(Cs cs, Th th') |th'. next_th s th cs th'}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2942 |
by (simp add: holding_set_eq) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2943 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2944 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2945 |
context valid_trace_v_e |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2946 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2947 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2948 |
lemma |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2949 |
acylic_RAG_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2950 |
assumes "acyclic (RAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2951 |
shows "acyclic (RAG (e#s))" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2952 |
proof(rule acyclic_subset[OF assms]) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2953 |
show "RAG (e # s) \<subseteq> RAG s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2954 |
by (unfold RAG_es waiting_set_eq holding_set_eq, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2955 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2956 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2957 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2958 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2959 |
context valid_trace_v_n |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2960 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2961 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2962 |
lemma waiting_taker: "waiting s taker cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2963 |
apply (unfold s_waiting_def, fold wq_def, unfold wq_s_cs taker_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2964 |
using eq_wq' th'_in_inv wq'_def by fastforce |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2965 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2966 |
lemma |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2967 |
acylic_RAG_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2968 |
assumes "acyclic (RAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2969 |
shows "acyclic (RAG (e#s))" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2970 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2971 |
have "acyclic ((RAG s - {(Cs cs, Th th)} - {(Th taker, Cs cs)}) \<union> |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2972 |
{(Cs cs, Th taker)})" (is "acyclic (?A \<union> _)") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2973 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2974 |
from assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2975 |
have "acyclic ?A" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2976 |
by (rule acyclic_subset, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2977 |
moreover have "(Th taker, Cs cs) \<notin> ?A^*" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2978 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2979 |
assume otherwise: "(Th taker, Cs cs) \<in> ?A^*" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2980 |
hence "(Th taker, Cs cs) \<in> ?A^+" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2981 |
by (unfold rtrancl_eq_or_trancl, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2982 |
from tranclD[OF this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2983 |
obtain cs' where h: "(Th taker, Cs cs') \<in> ?A" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2984 |
"(Th taker, Cs cs') \<in> RAG s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2985 |
by (unfold s_RAG_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2986 |
from this(2) have "waiting s taker cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2987 |
by (unfold s_RAG_def, fold waiting_eq, auto) |
100 | 2988 |
from waiting_unique[OF this waiting_taker] |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2989 |
have "cs' = cs" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2990 |
from h(1)[unfolded this] show False by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2991 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2992 |
ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2993 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2994 |
thus ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2995 |
by (unfold RAG_es waiting_set_eq holding_set_eq, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2996 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2997 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2998 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
2999 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3000 |
context valid_trace_p_h |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3001 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3002 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3003 |
lemma |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3004 |
acylic_RAG_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3005 |
assumes "acyclic (RAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3006 |
shows "acyclic (RAG (e#s))" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3007 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3008 |
have "acyclic (RAG s \<union> {(Cs cs, Th th)})" (is "acyclic (?A \<union> _)") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3009 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3010 |
from assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3011 |
have "acyclic ?A" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3012 |
by (rule acyclic_subset, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3013 |
moreover have "(Th th, Cs cs) \<notin> ?A^*" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3014 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3015 |
assume otherwise: "(Th th, Cs cs) \<in> ?A^*" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3016 |
hence "(Th th, Cs cs) \<in> ?A^+" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3017 |
by (unfold rtrancl_eq_or_trancl, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3018 |
from tranclD[OF this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3019 |
obtain cs' where h: "(Th th, Cs cs') \<in> RAG s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3020 |
by (unfold s_RAG_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3021 |
hence "waiting s th cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3022 |
by (unfold s_RAG_def, fold waiting_eq, auto) |
101 | 3023 |
with th_not_waiting show False by auto |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3024 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3025 |
ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3026 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3027 |
thus ?thesis by (unfold RAG_es, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3028 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3029 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3030 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3031 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3032 |
context valid_trace_p_w |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3033 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3034 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3035 |
lemma |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3036 |
acylic_RAG_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3037 |
assumes "acyclic (RAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3038 |
shows "acyclic (RAG (e#s))" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3039 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3040 |
have "acyclic (RAG s \<union> {(Th th, Cs cs)})" (is "acyclic (?A \<union> _)") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3041 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3042 |
from assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3043 |
have "acyclic ?A" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3044 |
by (rule acyclic_subset, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3045 |
moreover have "(Cs cs, Th th) \<notin> ?A^*" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3046 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3047 |
assume otherwise: "(Cs cs, Th th) \<in> ?A^*" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3048 |
from pip_e[unfolded is_p] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3049 |
show False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3050 |
proof(cases) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3051 |
case (thread_P) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3052 |
moreover from otherwise have "(Cs cs, Th th) \<in> ?A^+" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3053 |
by (unfold rtrancl_eq_or_trancl, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3054 |
ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3055 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3056 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3057 |
ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3058 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3059 |
thus ?thesis by (unfold RAG_es, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3060 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3061 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3062 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3063 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3064 |
context valid_trace |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3065 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3066 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3067 |
lemma acyclic_RAG: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3068 |
shows "acyclic (RAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3069 |
proof(induct rule:ind) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3070 |
case Nil |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3071 |
show ?case |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3072 |
by (auto simp: s_RAG_def cs_waiting_def |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3073 |
cs_holding_def wq_def acyclic_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3074 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3075 |
case (Cons s e) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3076 |
interpret vt_e: valid_trace_e s e using Cons by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3077 |
show ?case |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3078 |
proof(cases e) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3079 |
case (Create th prio) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3080 |
interpret vt: valid_trace_create s e th prio using Create |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3081 |
by (unfold_locales, simp) |
100 | 3082 |
show ?thesis using Cons by simp |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3083 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3084 |
case (Exit th) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3085 |
interpret vt: valid_trace_exit s e th using Exit |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3086 |
by (unfold_locales, simp) |
100 | 3087 |
show ?thesis using Cons by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3088 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3089 |
case (P th cs) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3090 |
interpret vt: valid_trace_p s e th cs using P |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3091 |
by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3092 |
show ?thesis |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3093 |
proof(cases "wq s cs = []") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3094 |
case True |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3095 |
then interpret vt_h: valid_trace_p_h s e th cs |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3096 |
by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3097 |
show ?thesis using Cons by (simp add: vt_h.acylic_RAG_kept) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3098 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3099 |
case False |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3100 |
then interpret vt_w: valid_trace_p_w s e th cs |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3101 |
by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3102 |
show ?thesis using Cons by (simp add: vt_w.acylic_RAG_kept) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3103 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3104 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3105 |
case (V th cs) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3106 |
interpret vt: valid_trace_v s e th cs using V |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3107 |
by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3108 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3109 |
proof(cases "vt.rest = []") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3110 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3111 |
then interpret vt_e: valid_trace_v_e s e th cs |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3112 |
by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3113 |
show ?thesis by (simp add: Cons.hyps(2) vt_e.acylic_RAG_kept) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3114 |
next |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3115 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3116 |
then interpret vt_n: valid_trace_v_n s e th cs |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3117 |
by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3118 |
show ?thesis by (simp add: Cons.hyps(2) vt_n.acylic_RAG_kept) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3119 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3120 |
next |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3121 |
case (Set th prio) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3122 |
interpret vt: valid_trace_set s e th prio using Set |
63 | 3123 |
by (unfold_locales, simp) |
100 | 3124 |
show ?thesis using Cons by simp |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3125 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3126 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3127 |
|
100 | 3128 |
end |
3129 |
||
3130 |
section {* RAG is single-valued *} |
|
3131 |
||
3132 |
context valid_trace |
|
3133 |
begin |
|
3134 |
||
3135 |
lemma unique_RAG: "\<lbrakk>(n, n1) \<in> RAG s; (n, n2) \<in> RAG s\<rbrakk> \<Longrightarrow> n1 = n2" |
|
3136 |
apply(unfold s_RAG_def, auto, fold waiting_eq holding_eq) |
|
3137 |
by(auto elim:waiting_unique held_unique) |
|
3138 |
||
3139 |
lemma sgv_RAG: "single_valued (RAG s)" |
|
3140 |
using unique_RAG by (auto simp:single_valued_def) |
|
3141 |
||
3142 |
end |
|
3143 |
||
3144 |
section {* RAG is well-founded *} |
|
3145 |
||
3146 |
context valid_trace |
|
3147 |
begin |
|
3148 |
||
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3149 |
lemma wf_RAG: "wf (RAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3150 |
proof(rule finite_acyclic_wf) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3151 |
from finite_RAG show "finite (RAG s)" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3152 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3153 |
from acyclic_RAG show "acyclic (RAG s)" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3154 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3155 |
|
100 | 3156 |
lemma wf_RAG_converse: |
3157 |
shows "wf ((RAG s)^-1)" |
|
3158 |
proof(rule finite_acyclic_wf_converse) |
|
3159 |
from finite_RAG |
|
3160 |
show "finite (RAG s)" . |
|
3161 |
next |
|
3162 |
from acyclic_RAG |
|
3163 |
show "acyclic (RAG s)" . |
|
3164 |
qed |
|
3165 |
||
3166 |
end |
|
3167 |
||
3168 |
section {* RAG forms a forest (or tree) *} |
|
3169 |
||
3170 |
context valid_trace |
|
3171 |
begin |
|
3172 |
||
3173 |
lemma rtree_RAG: "rtree (RAG s)" |
|
3174 |
using sgv_RAG acyclic_RAG |
|
3175 |
by (unfold rtree_def rtree_axioms_def sgv_def, auto) |
|
3176 |
||
3177 |
end |
|
3178 |
||
3179 |
sublocale valid_trace < rtree_RAG: rtree "RAG s" |
|
3180 |
using rtree_RAG . |
|
3181 |
||
3182 |
sublocale valid_trace < fsbtRAGs : fsubtree "RAG s" |
|
3183 |
proof - |
|
3184 |
show "fsubtree (RAG s)" |
|
3185 |
proof(intro_locales) |
|
3186 |
show "fbranch (RAG s)" using finite_fbranchI[OF finite_RAG] . |
|
3187 |
next |
|
3188 |
show "fsubtree_axioms (RAG s)" |
|
3189 |
proof(unfold fsubtree_axioms_def) |
|
3190 |
from wf_RAG show "wf (RAG s)" . |
|
3191 |
qed |
|
3192 |
qed |
|
3193 |
qed |
|
3194 |
||
101 | 3195 |
|
100 | 3196 |
section {* Derived properties for parts of RAG *} |
3197 |
||
3198 |
context valid_trace |
|
3199 |
begin |
|
3200 |
||
3201 |
lemma acyclic_tRAG: "acyclic (tRAG s)" |
|
3202 |
proof(unfold tRAG_def, rule acyclic_compose) |
|
3203 |
show "acyclic (RAG s)" using acyclic_RAG . |
|
3204 |
next |
|
3205 |
show "wRAG s \<subseteq> RAG s" unfolding RAG_split by auto |
|
3206 |
next |
|
3207 |
show "hRAG s \<subseteq> RAG s" unfolding RAG_split by auto |
|
3208 |
qed |
|
3209 |
||
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3210 |
lemma sgv_wRAG: "single_valued (wRAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3211 |
using waiting_unique |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3212 |
by (unfold single_valued_def wRAG_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3213 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3214 |
lemma sgv_hRAG: "single_valued (hRAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3215 |
using held_unique |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3216 |
by (unfold single_valued_def hRAG_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3217 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3218 |
lemma sgv_tRAG: "single_valued (tRAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3219 |
by (unfold tRAG_def, rule single_valued_relcomp, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3220 |
insert sgv_wRAG sgv_hRAG, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3221 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3222 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3223 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3224 |
sublocale valid_trace < rtree_s: rtree "tRAG s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3225 |
proof(unfold_locales) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3226 |
from sgv_tRAG show "single_valued (tRAG s)" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3227 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3228 |
from acyclic_tRAG show "acyclic (tRAG s)" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3229 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3230 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3231 |
sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3232 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3233 |
have "fsubtree (tRAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3234 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3235 |
have "fbranch (tRAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3236 |
proof(unfold tRAG_def, rule fbranch_compose) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3237 |
show "fbranch (wRAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3238 |
proof(rule finite_fbranchI) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3239 |
from finite_RAG show "finite (wRAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3240 |
by (unfold RAG_split, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3241 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3242 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3243 |
show "fbranch (hRAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3244 |
proof(rule finite_fbranchI) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3245 |
from finite_RAG |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3246 |
show "finite (hRAG s)" by (unfold RAG_split, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3247 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3248 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3249 |
moreover have "wf (tRAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3250 |
proof(rule wf_subset) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3251 |
show "wf (RAG s O RAG s)" using wf_RAG |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3252 |
by (fold wf_comp_self, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3253 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3254 |
show "tRAG s \<subseteq> (RAG s O RAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3255 |
by (unfold tRAG_alt_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3256 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3257 |
ultimately show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3258 |
by (unfold fsubtree_def fsubtree_axioms_def,auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3259 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3260 |
from this[folded tRAG_def] show "fsubtree (tRAG s)" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3261 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3262 |
|
101 | 3263 |
lemma tRAG_nodeE: |
3264 |
assumes "(n1, n2) \<in> tRAG s" |
|
3265 |
obtains th1 th2 where "n1 = Th th1" "n2 = Th th2" |
|
3266 |
using assms |
|
3267 |
by (auto simp: tRAG_def wRAG_def hRAG_def) |
|
3268 |
||
3269 |
lemma tRAG_ancestorsE: |
|
3270 |
assumes "x \<in> ancestors (tRAG s) u" |
|
3271 |
obtains th where "x = Th th" |
|
3272 |
proof - |
|
3273 |
from assms have "(u, x) \<in> (tRAG s)^+" |
|
3274 |
by (unfold ancestors_def, auto) |
|
3275 |
from tranclE[OF this] obtain c where "(c, x) \<in> tRAG s" by auto |
|
3276 |
then obtain th where "x = Th th" |
|
3277 |
by (unfold tRAG_alt_def, auto) |
|
3278 |
from that[OF this] show ?thesis . |
|
3279 |
qed |
|
3280 |
||
3281 |
lemma subtree_nodeE: |
|
3282 |
assumes "n \<in> subtree (tRAG s) (Th th)" |
|
3283 |
obtains th1 where "n = Th th1" |
|
3284 |
proof - |
|
3285 |
show ?thesis |
|
3286 |
proof(rule subtreeE[OF assms]) |
|
3287 |
assume "n = Th th" |
|
3288 |
from that[OF this] show ?thesis . |
|
3289 |
next |
|
3290 |
assume "Th th \<in> ancestors (tRAG s) n" |
|
3291 |
hence "(n, Th th) \<in> (tRAG s)^+" by (auto simp:ancestors_def) |
|
3292 |
hence "\<exists> th1. n = Th th1" |
|
3293 |
proof(induct) |
|
3294 |
case (base y) |
|
3295 |
from tRAG_nodeE[OF this] show ?case by metis |
|
3296 |
next |
|
3297 |
case (step y z) |
|
3298 |
thus ?case by auto |
|
3299 |
qed |
|
3300 |
with that show ?thesis by auto |
|
3301 |
qed |
|
3302 |
qed |
|
3303 |
||
3304 |
lemma tRAG_star_RAG: "(tRAG s)^* \<subseteq> (RAG s)^*" |
|
3305 |
proof - |
|
3306 |
have "(wRAG s O hRAG s)^* \<subseteq> (RAG s O RAG s)^*" |
|
3307 |
by (rule rtrancl_mono, auto simp:RAG_split) |
|
3308 |
also have "... \<subseteq> ((RAG s)^*)^*" |
|
3309 |
by (rule rtrancl_mono, auto) |
|
3310 |
also have "... = (RAG s)^*" by simp |
|
3311 |
finally show ?thesis by (unfold tRAG_def, simp) |
|
3312 |
qed |
|
3313 |
||
3314 |
lemma tRAG_subtree_RAG: "subtree (tRAG s) x \<subseteq> subtree (RAG s) x" |
|
3315 |
proof - |
|
3316 |
{ fix a |
|
3317 |
assume "a \<in> subtree (tRAG s) x" |
|
3318 |
hence "(a, x) \<in> (tRAG s)^*" by (auto simp:subtree_def) |
|
3319 |
with tRAG_star_RAG |
|
3320 |
have "(a, x) \<in> (RAG s)^*" by auto |
|
3321 |
hence "a \<in> subtree (RAG s) x" by (auto simp:subtree_def) |
|
3322 |
} thus ?thesis by auto |
|
3323 |
qed |
|
3324 |
||
3325 |
lemma tRAG_trancl_eq: |
|
3326 |
"{th'. (Th th', Th th) \<in> (tRAG s)^+} = |
|
3327 |
{th'. (Th th', Th th) \<in> (RAG s)^+}" |
|
3328 |
(is "?L = ?R") |
|
100 | 3329 |
proof - |
101 | 3330 |
{ fix th' |
3331 |
assume "th' \<in> ?L" |
|
3332 |
hence "(Th th', Th th) \<in> (tRAG s)^+" by auto |
|
3333 |
from tranclD[OF this] |
|
3334 |
obtain z where h: "(Th th', z) \<in> tRAG s" "(z, Th th) \<in> (tRAG s)\<^sup>*" by auto |
|
3335 |
from tRAG_subtree_RAG and this(2) |
|
3336 |
have "(z, Th th) \<in> (RAG s)^*" by (meson subsetCE tRAG_star_RAG) |
|
3337 |
moreover from h(1) have "(Th th', z) \<in> (RAG s)^+" using tRAG_alt_def by auto |
|
3338 |
ultimately have "th' \<in> ?R" by auto |
|
3339 |
} moreover |
|
3340 |
{ fix th' |
|
3341 |
assume "th' \<in> ?R" |
|
3342 |
hence "(Th th', Th th) \<in> (RAG s)^+" by (auto) |
|
3343 |
from plus_rpath[OF this] |
|
3344 |
obtain xs where rp: "rpath (RAG s) (Th th') xs (Th th)" "xs \<noteq> []" by auto |
|
3345 |
hence "(Th th', Th th) \<in> (tRAG s)^+" |
|
3346 |
proof(induct xs arbitrary:th' th rule:length_induct) |
|
3347 |
case (1 xs th' th) |
|
3348 |
then obtain x1 xs1 where Cons1: "xs = x1#xs1" by (cases xs, auto) |
|
3349 |
show ?case |
|
3350 |
proof(cases "xs1") |
|
3351 |
case Nil |
|
3352 |
from 1(2)[unfolded Cons1 Nil] |
|
3353 |
have rp: "rpath (RAG s) (Th th') [x1] (Th th)" . |
|
3354 |
hence "(Th th', x1) \<in> (RAG s)" |
|
3355 |
by (cases, auto) |
|
3356 |
then obtain cs where "x1 = Cs cs" |
|
3357 |
by (unfold s_RAG_def, auto) |
|
3358 |
from rpath_nnl_lastE[OF rp[unfolded this]] |
|
3359 |
show ?thesis by auto |
|
3360 |
next |
|
3361 |
case (Cons x2 xs2) |
|
3362 |
from 1(2)[unfolded Cons1[unfolded this]] |
|
3363 |
have rp: "rpath (RAG s) (Th th') (x1 # x2 # xs2) (Th th)" . |
|
3364 |
from rpath_edges_on[OF this] |
|
3365 |
have eds: "edges_on (Th th' # x1 # x2 # xs2) \<subseteq> RAG s" . |
|
3366 |
have "(Th th', x1) \<in> edges_on (Th th' # x1 # x2 # xs2)" |
|
3367 |
by (simp add: edges_on_unfold) |
|
3368 |
with eds have rg1: "(Th th', x1) \<in> RAG s" by auto |
|
3369 |
then obtain cs1 where eq_x1: "x1 = Cs cs1" by (unfold s_RAG_def, auto) |
|
3370 |
have "(x1, x2) \<in> edges_on (Th th' # x1 # x2 # xs2)" |
|
3371 |
by (simp add: edges_on_unfold) |
|
3372 |
from this eds |
|
3373 |
have rg2: "(x1, x2) \<in> RAG s" by auto |
|
3374 |
from this[unfolded eq_x1] |
|
3375 |
obtain th1 where eq_x2: "x2 = Th th1" by (unfold s_RAG_def, auto) |
|
3376 |
from rg1[unfolded eq_x1] rg2[unfolded eq_x1 eq_x2] |
|
3377 |
have rt1: "(Th th', Th th1) \<in> tRAG s" by (unfold tRAG_alt_def, auto) |
|
3378 |
from rp have "rpath (RAG s) x2 xs2 (Th th)" |
|
3379 |
by (elim rpath_ConsE, simp) |
|
3380 |
from this[unfolded eq_x2] have rp': "rpath (RAG s) (Th th1) xs2 (Th th)" . |
|
3381 |
show ?thesis |
|
3382 |
proof(cases "xs2 = []") |
|
3383 |
case True |
|
3384 |
from rpath_nilE[OF rp'[unfolded this]] |
|
3385 |
have "th1 = th" by auto |
|
3386 |
from rt1[unfolded this] show ?thesis by auto |
|
3387 |
next |
|
3388 |
case False |
|
3389 |
from 1(1)[rule_format, OF _ rp' this, unfolded Cons1 Cons] |
|
3390 |
have "(Th th1, Th th) \<in> (tRAG s)\<^sup>+" by simp |
|
3391 |
with rt1 show ?thesis by auto |
|
3392 |
qed |
|
3393 |
qed |
|
3394 |
qed |
|
3395 |
hence "th' \<in> ?L" by auto |
|
3396 |
} ultimately show ?thesis by blast |
|
100 | 3397 |
qed |
3398 |
||
101 | 3399 |
lemma tRAG_trancl_eq_Th: |
3400 |
"{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = |
|
3401 |
{Th th' | th'. (Th th', Th th) \<in> (RAG s)^+}" |
|
3402 |
using tRAG_trancl_eq by auto |
|
3403 |
||
3404 |
||
3405 |
lemma tRAG_Field: |
|
3406 |
"Field (tRAG s) \<subseteq> Field (RAG s)" |
|
3407 |
by (unfold tRAG_alt_def Field_def, auto) |
|
3408 |
||
3409 |
lemma tRAG_mono: |
|
3410 |
assumes "RAG s' \<subseteq> RAG s" |
|
3411 |
shows "tRAG s' \<subseteq> tRAG s" |
|
3412 |
using assms |
|
3413 |
by (unfold tRAG_alt_def, auto) |
|
3414 |
||
3415 |
lemma tRAG_subtree_eq: |
|
3416 |
"(subtree (tRAG s) (Th th)) = {Th th' | th'. Th th' \<in> (subtree (RAG s) (Th th))}" |
|
3417 |
(is "?L = ?R") |
|
3418 |
proof - |
|
3419 |
{ fix n |
|
3420 |
assume h: "n \<in> ?L" |
|
3421 |
hence "n \<in> ?R" |
|
3422 |
by (smt mem_Collect_eq subsetCE subtree_def subtree_nodeE tRAG_subtree_RAG) |
|
3423 |
} moreover { |
|
3424 |
fix n |
|
3425 |
assume "n \<in> ?R" |
|
3426 |
then obtain th' where h: "n = Th th'" "(Th th', Th th) \<in> (RAG s)^*" |
|
3427 |
by (auto simp:subtree_def) |
|
3428 |
from rtranclD[OF this(2)] |
|
3429 |
have "n \<in> ?L" |
|
3430 |
proof |
|
3431 |
assume "Th th' \<noteq> Th th \<and> (Th th', Th th) \<in> (RAG s)\<^sup>+" |
|
3432 |
with h have "n \<in> {Th th' | th'. (Th th', Th th) \<in> (RAG s)^+}" by auto |
|
3433 |
thus ?thesis using subtree_def tRAG_trancl_eq by fastforce |
|
3434 |
qed (insert h, auto simp:subtree_def) |
|
3435 |
} ultimately show ?thesis by auto |
|
3436 |
qed |
|
3437 |
||
3438 |
lemma threads_set_eq: |
|
3439 |
"the_thread ` (subtree (tRAG s) (Th th)) = |
|
3440 |
{th'. Th th' \<in> (subtree (RAG s) (Th th))}" (is "?L = ?R") |
|
3441 |
by (auto intro:rev_image_eqI simp:tRAG_subtree_eq) |
|
3442 |
||
102
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3443 |
context valid_trace |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3444 |
begin |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3445 |
|
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3446 |
lemma RAG_tRAG_transfer: |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3447 |
assumes "RAG s' = RAG s \<union> {(Th th, Cs cs)}" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3448 |
and "(Cs cs, Th th'') \<in> RAG s" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3449 |
shows "tRAG s' = tRAG s \<union> {(Th th, Th th'')}" (is "?L = ?R") |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3450 |
proof - |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3451 |
{ fix n1 n2 |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3452 |
assume "(n1, n2) \<in> ?L" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3453 |
from this[unfolded tRAG_alt_def] |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3454 |
obtain th1 th2 cs' where |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3455 |
h: "n1 = Th th1" "n2 = Th th2" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3456 |
"(Th th1, Cs cs') \<in> RAG s'" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3457 |
"(Cs cs', Th th2) \<in> RAG s'" by auto |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3458 |
from h(4) and assms(1) have cs_in: "(Cs cs', Th th2) \<in> RAG s" by auto |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3459 |
from h(3) and assms(1) |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3460 |
have "(Th th1, Cs cs') = (Th th, Cs cs) \<or> |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3461 |
(Th th1, Cs cs') \<in> RAG s" by auto |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3462 |
hence "(n1, n2) \<in> ?R" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3463 |
proof |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3464 |
assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3465 |
hence eq_th1: "th1 = th" by simp |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3466 |
moreover have "th2 = th''" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3467 |
proof - |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3468 |
from h1 have "cs' = cs" by simp |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3469 |
from assms(2) cs_in[unfolded this] |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3470 |
show ?thesis using unique_RAG by auto |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3471 |
qed |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3472 |
ultimately show ?thesis using h(1,2) by auto |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3473 |
next |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3474 |
assume "(Th th1, Cs cs') \<in> RAG s" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3475 |
with cs_in have "(Th th1, Th th2) \<in> tRAG s" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3476 |
by (unfold tRAG_alt_def, auto) |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3477 |
from this[folded h(1, 2)] show ?thesis by auto |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3478 |
qed |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3479 |
} moreover { |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3480 |
fix n1 n2 |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3481 |
assume "(n1, n2) \<in> ?R" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3482 |
hence "(n1, n2) \<in>tRAG s \<or> (n1, n2) = (Th th, Th th'')" by auto |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3483 |
hence "(n1, n2) \<in> ?L" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3484 |
proof |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3485 |
assume "(n1, n2) \<in> tRAG s" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3486 |
moreover have "... \<subseteq> ?L" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3487 |
proof(rule tRAG_mono) |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3488 |
show "RAG s \<subseteq> RAG s'" by (unfold assms(1), auto) |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3489 |
qed |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3490 |
ultimately show ?thesis by auto |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3491 |
next |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3492 |
assume eq_n: "(n1, n2) = (Th th, Th th'')" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3493 |
from assms(1, 2) have "(Cs cs, Th th'') \<in> RAG s'" by auto |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3494 |
moreover have "(Th th, Cs cs) \<in> RAG s'" using assms(1) by auto |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3495 |
ultimately show ?thesis |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3496 |
by (unfold eq_n tRAG_alt_def, auto) |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3497 |
qed |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3498 |
} ultimately show ?thesis by auto |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3499 |
qed |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3500 |
|
103
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3501 |
lemma subtree_tRAG_thread: |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3502 |
assumes "th \<in> threads s" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3503 |
shows "subtree (tRAG s) (Th th) \<subseteq> Th ` threads s" (is "?L \<subseteq> ?R") |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3504 |
proof - |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3505 |
have "?L = {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3506 |
by (unfold tRAG_subtree_eq, simp) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3507 |
also have "... \<subseteq> ?R" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3508 |
proof |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3509 |
fix x |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3510 |
assume "x \<in> {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3511 |
then obtain th' where h: "x = Th th'" "Th th' \<in> subtree (RAG s) (Th th)" by auto |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3512 |
from this(2) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3513 |
show "x \<in> ?R" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3514 |
proof(cases rule:subtreeE) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3515 |
case 1 |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3516 |
thus ?thesis by (simp add: assms h(1)) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3517 |
next |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3518 |
case 2 |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3519 |
thus ?thesis by (metis ancestors_Field dm_RAG_threads h(1) image_eqI) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3520 |
qed |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3521 |
qed |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3522 |
finally show ?thesis . |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3523 |
qed |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
3524 |
|
101 | 3525 |
lemma dependants_alt_def: |
3526 |
"dependants s th = {th'. (Th th', Th th) \<in> (tRAG s)^+}" |
|
3527 |
by (metis eq_RAG s_dependants_def tRAG_trancl_eq) |
|
3528 |
||
3529 |
lemma dependants_alt_def1: |
|
3530 |
"dependants (s::state) th = {th'. (Th th', Th th) \<in> (RAG s)^+}" |
|
3531 |
using dependants_alt_def tRAG_trancl_eq by auto |
|
3532 |
||
102
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3533 |
end |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
3534 |
|
101 | 3535 |
section {* Chain to readys *} |
100 | 3536 |
|
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3537 |
context valid_trace |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3538 |
begin |
104 | 3539 |
======= |
55
b85cfbd58f59
Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents:
53
diff
changeset
|
3540 |
text {* (* ddd *) |
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
3541 |
The following @{text "step_RAG_v"} lemma charaterizes how @{text "RAG"} is changed |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
3542 |
with the happening of @{text "V"}-events: |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
3543 |
*} |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3544 |
lemma step_RAG_v: |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3545 |
assumes vt: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3546 |
"vt (V th cs#s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3547 |
shows " |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3548 |
RAG (V th cs # s) = |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3549 |
RAG s - {(Cs cs, Th th)} - |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3550 |
{(Th th', Cs cs) |th'. next_th s th cs th'} \<union> |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3551 |
{(Cs cs, Th th') |th'. next_th s th cs th'}" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3552 |
apply (insert vt, unfold s_RAG_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3553 |
apply (auto split:if_splits list.splits simp:Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3554 |
apply (auto elim: step_v_waiting_mono step_v_hold_inv |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3555 |
step_v_release step_v_wait_inv |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3556 |
step_v_get_hold step_v_release_inv) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3557 |
apply (erule_tac step_v_not_wait, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3558 |
done |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3559 |
|
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
3560 |
text {* |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
3561 |
The following @{text "step_RAG_p"} lemma charaterizes how @{text "RAG"} is changed |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
3562 |
with the happening of @{text "P"}-events: |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
3563 |
*} |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3564 |
lemma step_RAG_p: |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3565 |
"vt (P th cs#s) \<Longrightarrow> |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3566 |
RAG (P th cs # s) = (if (wq s cs = []) then RAG s \<union> {(Cs cs, Th th)} |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3567 |
else RAG s \<union> {(Th th, Cs cs)})" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3568 |
apply(simp only: s_RAG_def wq_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3569 |
apply (auto split:list.splits prod.splits simp:Let_def wq_def cs_waiting_def cs_holding_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3570 |
apply(case_tac "csa = cs", auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3571 |
apply(fold wq_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3572 |
apply(drule_tac step_back_step) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3573 |
apply(ind_cases " step s (P (hd (wq s cs)) cs)") |
36
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3574 |
apply(simp add:s_RAG_def wq_def cs_holding_def) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3575 |
apply(auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3576 |
done |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3577 |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3578 |
|
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3579 |
lemma RAG_target_th: "(Th th, x) \<in> RAG (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3580 |
by (unfold s_RAG_def, auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3581 |
|
63 | 3582 |
context valid_trace |
3583 |
begin |
|
3584 |
||
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
3585 |
text {* |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
3586 |
The following lemma shows that @{text "RAG"} is acyclic. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
3587 |
The overall structure is by induction on the formation of @{text "vt s"} |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
3588 |
and then case analysis on event @{text "e"}, where the non-trivial cases |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
3589 |
for those for @{text "V"} and @{text "P"} events. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
3590 |
*} |
63 | 3591 |
lemma acyclic_RAG: |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3592 |
shows "acyclic (RAG s)" |
63 | 3593 |
using vt |
36
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3594 |
proof(induct) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3595 |
case (vt_cons s e) |
63 | 3596 |
interpret vt_s: valid_trace s using vt_cons(1) |
3597 |
by (unfold_locales, simp) |
|
36
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3598 |
assume ih: "acyclic (RAG s)" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3599 |
and stp: "step s e" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3600 |
and vt: "vt s" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3601 |
show ?case |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3602 |
proof(cases e) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3603 |
case (Create th prio) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3604 |
with ih |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3605 |
show ?thesis by (simp add:RAG_create_unchanged) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3606 |
next |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3607 |
case (Exit th) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3608 |
with ih show ?thesis by (simp add:RAG_exit_unchanged) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3609 |
next |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3610 |
case (V th cs) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3611 |
from V vt stp have vtt: "vt (V th cs#s)" by auto |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3612 |
from step_RAG_v [OF this] |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3613 |
have eq_de: |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3614 |
"RAG (e # s) = |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3615 |
RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union> |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3616 |
{(Cs cs, Th th') |th'. next_th s th cs th'}" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3617 |
(is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3618 |
from ih have ac: "acyclic (?A - ?B - ?C)" by (auto elim:acyclic_subset) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3619 |
from step_back_step [OF vtt] |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3620 |
have "step s (V th cs)" . |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3621 |
thus ?thesis |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3622 |
proof(cases) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3623 |
assume "holding s th cs" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3624 |
hence th_in: "th \<in> set (wq s cs)" and |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3625 |
eq_hd: "th = hd (wq s cs)" unfolding s_holding_def wq_def by auto |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3626 |
then obtain rest where |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3627 |
eq_wq: "wq s cs = th#rest" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3628 |
by (cases "wq s cs", auto) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3629 |
show ?thesis |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3630 |
proof(cases "rest = []") |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3631 |
case False |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3632 |
let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3633 |
from eq_wq False have eq_D: "?D = {(Cs cs, Th ?th')}" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3634 |
by (unfold next_th_def, auto) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3635 |
let ?E = "(?A - ?B - ?C)" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3636 |
have "(Th ?th', Cs cs) \<notin> ?E\<^sup>*" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3637 |
proof |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3638 |
assume "(Th ?th', Cs cs) \<in> ?E\<^sup>*" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3639 |
hence " (Th ?th', Cs cs) \<in> ?E\<^sup>+" by (simp add: rtrancl_eq_or_trancl) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3640 |
from tranclD [OF this] |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3641 |
obtain x where th'_e: "(Th ?th', x) \<in> ?E" by blast |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3642 |
hence th_d: "(Th ?th', x) \<in> ?A" by simp |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3643 |
from RAG_target_th [OF this] |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3644 |
obtain cs' where eq_x: "x = Cs cs'" by auto |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3645 |
with th_d have "(Th ?th', Cs cs') \<in> ?A" by simp |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3646 |
hence wt_th': "waiting s ?th' cs'" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3647 |
unfolding s_RAG_def s_waiting_def cs_waiting_def wq_def by simp |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3648 |
hence "cs' = cs" |
63 | 3649 |
proof(rule vt_s.waiting_unique) |
3650 |
from eq_wq vt_s.wq_distinct[of cs] |
|
36
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3651 |
show "waiting s ?th' cs" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3652 |
apply (unfold s_waiting_def wq_def, auto) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3653 |
proof - |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3654 |
assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3655 |
and eq_wq: "wq_fun (schs s) cs = th # rest" |
36
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3656 |
have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3657 |
proof(rule someI2) |
63 | 3658 |
from vt_s.wq_distinct[of cs] and eq_wq |
36
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3659 |
show "distinct rest \<and> set rest = set rest" unfolding wq_def by auto |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3660 |
next |
36
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3661 |
fix x assume "distinct x \<and> set x = set rest" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3662 |
with False show "x \<noteq> []" by auto |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3663 |
qed |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3664 |
hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3665 |
set (SOME q. distinct q \<and> set q = set rest)" by auto |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3666 |
moreover have "\<dots> = set rest" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3667 |
proof(rule someI2) |
63 | 3668 |
from vt_s.wq_distinct[of cs] and eq_wq |
36
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3669 |
show "distinct rest \<and> set rest = set rest" unfolding wq_def by auto |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3670 |
next |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3671 |
show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3672 |
qed |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3673 |
moreover note hd_in |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3674 |
ultimately show "hd (SOME q. distinct q \<and> set q = set rest) = th" by auto |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3675 |
next |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3676 |
assume hd_in: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3677 |
and eq_wq: "wq s cs = hd (SOME q. distinct q \<and> set q = set rest) # rest" |
36
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3678 |
have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3679 |
proof(rule someI2) |
63 | 3680 |
from vt_s.wq_distinct[of cs] and eq_wq |
36
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3681 |
show "distinct rest \<and> set rest = set rest" by auto |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3682 |
next |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3683 |
fix x assume "distinct x \<and> set x = set rest" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3684 |
with False show "x \<noteq> []" by auto |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3685 |
qed |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3686 |
hence "hd (SOME q. distinct q \<and> set q = set rest) \<in> |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3687 |
set (SOME q. distinct q \<and> set q = set rest)" by auto |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3688 |
moreover have "\<dots> = set rest" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3689 |
proof(rule someI2) |
63 | 3690 |
from vt_s.wq_distinct[of cs] and eq_wq |
36
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3691 |
show "distinct rest \<and> set rest = set rest" by auto |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3692 |
next |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3693 |
show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3694 |
qed |
36
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3695 |
moreover note hd_in |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3696 |
ultimately show False by auto |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3697 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3698 |
qed |
36
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3699 |
with th'_e eq_x have "(Th ?th', Cs cs) \<in> ?E" by simp |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3700 |
with False |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3701 |
show "False" by (auto simp: next_th_def eq_wq) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3702 |
qed |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3703 |
with acyclic_insert[symmetric] and ac |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3704 |
and eq_de eq_D show ?thesis by auto |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3705 |
next |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3706 |
case True |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3707 |
with eq_wq |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3708 |
have eq_D: "?D = {}" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3709 |
by (unfold next_th_def, auto) |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3710 |
with eq_de ac |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3711 |
show ?thesis by auto |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3712 |
qed |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3713 |
qed |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3714 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3715 |
case (P th cs) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3716 |
from P vt stp have vtt: "vt (P th cs#s)" by auto |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3717 |
from step_RAG_p [OF this] P |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3718 |
have "RAG (e # s) = |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3719 |
(if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3720 |
RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R") |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3721 |
by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3722 |
moreover have "acyclic ?R" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3723 |
proof(cases "wq s cs = []") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3724 |
case True |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3725 |
hence eq_r: "?R = RAG s \<union> {(Cs cs, Th th)}" by simp |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3726 |
have "(Th th, Cs cs) \<notin> (RAG s)\<^sup>*" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3727 |
proof |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3728 |
assume "(Th th, Cs cs) \<in> (RAG s)\<^sup>*" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3729 |
hence "(Th th, Cs cs) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3730 |
from tranclD2 [OF this] |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3731 |
obtain x where "(x, Cs cs) \<in> RAG s" by auto |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3732 |
with True show False by (auto simp:s_RAG_def cs_waiting_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3733 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3734 |
with acyclic_insert ih eq_r show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3735 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3736 |
case False |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3737 |
hence eq_r: "?R = RAG s \<union> {(Th th, Cs cs)}" by simp |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3738 |
have "(Cs cs, Th th) \<notin> (RAG s)\<^sup>*" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3739 |
proof |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3740 |
assume "(Cs cs, Th th) \<in> (RAG s)\<^sup>*" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3741 |
hence "(Cs cs, Th th) \<in> (RAG s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl) |
36
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3742 |
moreover from step_back_step [OF vtt] have "step s (P th cs)" . |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3743 |
ultimately show False |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3744 |
proof - |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3745 |
show " \<lbrakk>(Cs cs, Th th) \<in> (RAG s)\<^sup>+; step s (P th cs)\<rbrakk> \<Longrightarrow> False" |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3746 |
by (ind_cases "step s (P th cs)", simp) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3747 |
qed |
36
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3748 |
qed |
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3749 |
with acyclic_insert ih eq_r show ?thesis by auto |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3750 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3751 |
ultimately show ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3752 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3753 |
case (Set thread prio) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3754 |
with ih |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3755 |
thm RAG_set_unchanged |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3756 |
show ?thesis by (simp add:RAG_set_unchanged) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3757 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3758 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3759 |
case vt_nil |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3760 |
show "acyclic (RAG ([]::state))" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3761 |
by (auto simp: s_RAG_def cs_waiting_def |
36
af38526275f8
added a test theory for polishing teh proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
35
diff
changeset
|
3762 |
cs_holding_def wq_def acyclic_def) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3763 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3764 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3765 |
|
63 | 3766 |
lemma finite_RAG: |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3767 |
shows "finite (RAG s)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3768 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3769 |
from vt show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3770 |
proof(induct) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3771 |
case (vt_cons s e) |
63 | 3772 |
interpret vt_s: valid_trace s using vt_cons(1) |
3773 |
by (unfold_locales, simp) |
|
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3774 |
assume ih: "finite (RAG s)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3775 |
and stp: "step s e" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3776 |
and vt: "vt s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3777 |
show ?case |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3778 |
proof(cases e) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3779 |
case (Create th prio) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3780 |
with ih |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3781 |
show ?thesis by (simp add:RAG_create_unchanged) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3782 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3783 |
case (Exit th) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3784 |
with ih show ?thesis by (simp add:RAG_exit_unchanged) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3785 |
next |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3786 |
case (V th cs) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3787 |
from V vt stp have vtt: "vt (V th cs#s)" by auto |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3788 |
from step_RAG_v [OF this] |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3789 |
have eq_de: "RAG (e # s) = |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3790 |
RAG s - {(Cs cs, Th th)} - {(Th th', Cs cs) |th'. next_th s th cs th'} \<union> |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3791 |
{(Cs cs, Th th') |th'. next_th s th cs th'} |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3792 |
" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3793 |
(is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3794 |
moreover from ih have ac: "finite (?A - ?B - ?C)" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3795 |
moreover have "finite ?D" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3796 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3797 |
have "?D = {} \<or> (\<exists> a. ?D = {a})" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3798 |
by (unfold next_th_def, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3799 |
thus ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3800 |
proof |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3801 |
assume h: "?D = {}" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3802 |
show ?thesis by (unfold h, simp) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3803 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3804 |
assume "\<exists> a. ?D = {a}" |
3
51019d035a79
made everything working
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
0
diff
changeset
|
3805 |
thus ?thesis |
51019d035a79
made everything working
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
0
diff
changeset
|
3806 |
by (metis finite.simps) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3807 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3808 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3809 |
ultimately show ?thesis by simp |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3810 |
next |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3811 |
case (P th cs) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3812 |
from P vt stp have vtt: "vt (P th cs#s)" by auto |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3813 |
from step_RAG_p [OF this] P |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3814 |
have "RAG (e # s) = |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3815 |
(if wq s cs = [] then RAG s \<union> {(Cs cs, Th th)} else |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3816 |
RAG s \<union> {(Th th, Cs cs)})" (is "?L = ?R") |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3817 |
by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3818 |
moreover have "finite ?R" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3819 |
proof(cases "wq s cs = []") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3820 |
case True |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3821 |
hence eq_r: "?R = RAG s \<union> {(Cs cs, Th th)}" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3822 |
with True and ih show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3823 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3824 |
case False |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3825 |
hence "?R = RAG s \<union> {(Th th, Cs cs)}" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3826 |
with False and ih show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3827 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3828 |
ultimately show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3829 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3830 |
case (Set thread prio) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3831 |
with ih |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3832 |
show ?thesis by (simp add:RAG_set_unchanged) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3833 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3834 |
next |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3835 |
case vt_nil |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3836 |
show "finite (RAG ([]::state))" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3837 |
by (auto simp: s_RAG_def cs_waiting_def |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3838 |
cs_holding_def wq_def acyclic_def) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3839 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3840 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3841 |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3842 |
text {* Several useful lemmas *} |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
3843 |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3844 |
lemma wf_dep_converse: |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3845 |
shows "wf ((RAG s)^-1)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3846 |
proof(rule finite_acyclic_wf_converse) |
63 | 3847 |
from finite_RAG |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3848 |
show "finite (RAG s)" . |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3849 |
next |
63 | 3850 |
from acyclic_RAG |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3851 |
show "acyclic (RAG s)" . |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3852 |
qed |
104 | 3853 |
>>>>>>> other |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3854 |
|
63 | 3855 |
end |
3856 |
||
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3857 |
lemma hd_np_in: "x \<in> set l \<Longrightarrow> hd l \<in> set l" |
63 | 3858 |
by (induct l, auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3859 |
|
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3860 |
lemma th_chasing: "(Th th, Cs cs) \<in> RAG (s::state) \<Longrightarrow> \<exists> th'. (Cs cs, Th th') \<in> RAG s" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3861 |
by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3862 |
|
63 | 3863 |
context valid_trace |
3864 |
begin |
|
3865 |
||
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3866 |
lemma wq_threads: |
63 | 3867 |
assumes h: "th \<in> set (wq s cs)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3868 |
shows "th \<in> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3869 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3870 |
from vt and h show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3871 |
proof(induct arbitrary: th cs) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3872 |
case (vt_cons s e) |
63 | 3873 |
interpret vt_s: valid_trace s |
3874 |
using vt_cons(1) by (unfold_locales, auto) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3875 |
assume ih: "\<And>th cs. th \<in> set (wq s cs) \<Longrightarrow> th \<in> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3876 |
and stp: "step s e" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3877 |
and vt: "vt s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3878 |
and h: "th \<in> set (wq (e # s) cs)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3879 |
show ?case |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3880 |
proof(cases e) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3881 |
case (Create th' prio) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3882 |
with ih h show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3883 |
by (auto simp:wq_def Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3884 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3885 |
case (Exit th') |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3886 |
with stp ih h show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3887 |
apply (auto simp:wq_def Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3888 |
apply (ind_cases "step s (Exit th')") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3889 |
apply (auto simp:runing_def readys_def s_holding_def s_waiting_def holdents_def |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3890 |
s_RAG_def s_holding_def cs_holding_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3891 |
done |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3892 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3893 |
case (V th' cs') |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3894 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3895 |
proof(cases "cs' = cs") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3896 |
case False |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3897 |
with h |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3898 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3899 |
apply(unfold wq_def V, auto simp:Let_def V split:prod.splits, fold wq_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3900 |
by (drule_tac ih, simp) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3901 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3902 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3903 |
from h |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3904 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3905 |
proof(unfold V wq_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3906 |
assume th_in: "th \<in> set (wq_fun (schs (V th' cs' # s)) cs)" (is "th \<in> set ?l") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3907 |
show "th \<in> threads (V th' cs' # s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3908 |
proof(cases "cs = cs'") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3909 |
case False |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3910 |
hence "?l = wq_fun (schs s) cs" by (simp add:Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3911 |
with th_in have " th \<in> set (wq s cs)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3912 |
by (fold wq_def, simp) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3913 |
from ih [OF this] show ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3914 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3915 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3916 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3917 |
proof(cases "wq_fun (schs s) cs'") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3918 |
case Nil |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3919 |
with h V show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3920 |
apply (auto simp:wq_def Let_def split:if_splits) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3921 |
by (fold wq_def, drule_tac ih, simp) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3922 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3923 |
case (Cons a rest) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3924 |
assume eq_wq: "wq_fun (schs s) cs' = a # rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3925 |
with h V show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3926 |
apply (auto simp:Let_def wq_def split:if_splits) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3927 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3928 |
assume th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3929 |
have "set (SOME q. distinct q \<and> set q = set rest) = set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3930 |
proof(rule someI2) |
63 | 3931 |
from vt_s.wq_distinct[of cs'] and eq_wq[folded wq_def] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3932 |
show "distinct rest \<and> set rest = set rest" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3933 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3934 |
show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3935 |
by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3936 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3937 |
with eq_wq th_in have "th \<in> set (wq_fun (schs s) cs')" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3938 |
from ih[OF this[folded wq_def]] show "th \<in> threads s" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3939 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3940 |
assume th_in: "th \<in> set (wq_fun (schs s) cs)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3941 |
from ih[OF this[folded wq_def]] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3942 |
show "th \<in> threads s" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3943 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3944 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3945 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3946 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3947 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3948 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3949 |
case (P th' cs') |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3950 |
from h stp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3951 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3952 |
apply (unfold P wq_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3953 |
apply (auto simp:Let_def split:if_splits, fold wq_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3954 |
apply (auto intro:ih) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3955 |
apply(ind_cases "step s (P th' cs')") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3956 |
by (unfold runing_def readys_def, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3957 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3958 |
case (Set thread prio) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3959 |
with ih h show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3960 |
by (auto simp:wq_def Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3961 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3962 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3963 |
case vt_nil |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3964 |
thus ?case by (auto simp:wq_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3965 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3966 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3967 |
|
63 | 3968 |
lemma range_in: "\<lbrakk>(Th th) \<in> Range (RAG (s::state))\<rbrakk> \<Longrightarrow> th \<in> threads s" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
3969 |
apply(unfold s_RAG_def cs_waiting_def cs_holding_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3970 |
by (auto intro:wq_threads) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3971 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3972 |
lemma readys_v_eq: |
63 | 3973 |
assumes neq_th: "th \<noteq> thread" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3974 |
and eq_wq: "wq s cs = thread#rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3975 |
and not_in: "th \<notin> set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3976 |
shows "(th \<in> readys (V thread cs#s)) = (th \<in> readys s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3977 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3978 |
from assms show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3979 |
apply (auto simp:readys_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3980 |
apply(simp add:s_waiting_def[folded wq_def]) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3981 |
apply (erule_tac x = csa in allE) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3982 |
apply (simp add:s_waiting_def wq_def Let_def split:if_splits) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3983 |
apply (case_tac "csa = cs", simp) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3984 |
apply (erule_tac x = cs in allE) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3985 |
apply(auto simp add: s_waiting_def[folded wq_def] Let_def split: list.splits) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3986 |
apply(auto simp add: wq_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3987 |
apply (auto simp:s_waiting_def wq_def Let_def split:list.splits) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3988 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3989 |
assume th_nin: "th \<notin> set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3990 |
and th_in: "th \<in> set (SOME q. distinct q \<and> set q = set rest)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3991 |
and eq_wq: "wq_fun (schs s) cs = thread # rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3992 |
have "set (SOME q. distinct q \<and> set q = set rest) = set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3993 |
proof(rule someI2) |
63 | 3994 |
from wq_distinct[of cs, unfolded wq_def] and eq_wq[unfolded wq_def] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3995 |
show "distinct rest \<and> set rest = set rest" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3996 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3997 |
show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3998 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3999 |
with th_nin th_in show False by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4000 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4001 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4002 |
|
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
4003 |
text {* \noindent |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
4004 |
The following lemmas shows that: starting from any node in @{text "RAG"}, |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
4005 |
by chasing out-going edges, it is always possible to reach a node representing a ready |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
4006 |
thread. In this lemma, it is the @{text "th'"}. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
4007 |
*} |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
4008 |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4009 |
lemma chain_building: |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4010 |
shows "node \<in> Domain (RAG s) \<longrightarrow> (\<exists> th'. th' \<in> readys s \<and> (node, Th th') \<in> (RAG s)^+)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4011 |
proof - |
63 | 4012 |
from wf_dep_converse |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4013 |
have h: "wf ((RAG s)\<inverse>)" . |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4014 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4015 |
proof(induct rule:wf_induct [OF h]) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4016 |
fix x |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4017 |
assume ih [rule_format]: |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4018 |
"\<forall>y. (y, x) \<in> (RAG s)\<inverse> \<longrightarrow> |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4019 |
y \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (y, Th th') \<in> (RAG s)\<^sup>+)" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4020 |
show "x \<in> Domain (RAG s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4021 |
proof |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4022 |
assume x_d: "x \<in> Domain (RAG s)" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4023 |
show "\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (RAG s)\<^sup>+" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4024 |
proof(cases x) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4025 |
case (Th th) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4026 |
from x_d Th obtain cs where x_in: "(Th th, Cs cs) \<in> RAG s" by (auto simp:s_RAG_def) |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4027 |
with Th have x_in_r: "(Cs cs, x) \<in> (RAG s)^-1" by simp |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4028 |
from th_chasing [OF x_in] obtain th' where "(Cs cs, Th th') \<in> RAG s" by blast |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4029 |
hence "Cs cs \<in> Domain (RAG s)" by auto |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4030 |
from ih [OF x_in_r this] obtain th' |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4031 |
where th'_ready: " th' \<in> readys s" and cs_in: "(Cs cs, Th th') \<in> (RAG s)\<^sup>+" by auto |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4032 |
have "(x, Th th') \<in> (RAG s)\<^sup>+" using Th x_in cs_in by auto |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4033 |
with th'_ready show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4034 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4035 |
case (Cs cs) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4036 |
from x_d Cs obtain th' where th'_d: "(Th th', x) \<in> (RAG s)^-1" by (auto simp:s_RAG_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4037 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4038 |
proof(cases "th' \<in> readys s") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4039 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4040 |
from True and th'_d show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4041 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4042 |
case False |
63 | 4043 |
from th'_d and range_in have "th' \<in> threads s" by auto |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4044 |
with False have "Th th' \<in> Domain (RAG s)" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4045 |
by (auto simp:readys_def wq_def s_waiting_def s_RAG_def cs_waiting_def Domain_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4046 |
from ih [OF th'_d this] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4047 |
obtain th'' where |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4048 |
th''_r: "th'' \<in> readys s" and |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4049 |
th''_in: "(Th th', Th th'') \<in> (RAG s)\<^sup>+" by auto |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4050 |
from th'_d and th''_in |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4051 |
have "(x, Th th'') \<in> (RAG s)\<^sup>+" by auto |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4052 |
with th''_r show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4053 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4054 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4055 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4056 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4057 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4058 |
|
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
4059 |
text {* \noindent |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
4060 |
The following is just an instance of @{text "chain_building"}. |
100 | 4061 |
*} |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4062 |
lemma th_chain_to_ready: |
63 | 4063 |
assumes th_in: "th \<in> threads s" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4064 |
shows "th \<in> readys s \<or> (\<exists> th'. th' \<in> readys s \<and> (Th th, Th th') \<in> (RAG s)^+)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4065 |
proof(cases "th \<in> readys s") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4066 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4067 |
thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4068 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4069 |
case False |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4070 |
from False and th_in have "Th th \<in> Domain (RAG s)" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4071 |
by (auto simp:readys_def s_waiting_def s_RAG_def wq_def cs_waiting_def Domain_def) |
63 | 4072 |
from chain_building [rule_format, OF this] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4073 |
show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4074 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4075 |
|
104 | 4076 |
<<<<<<< local |
101 | 4077 |
lemma finite_subtree_threads: |
4078 |
"finite {th'. Th th' \<in> subtree (RAG s) (Th th)}" (is "finite ?A") |
|
104 | 4079 |
======= |
63 | 4080 |
end |
4081 |
||
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4082 |
lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4083 |
by (unfold s_waiting_def cs_waiting_def wq_def, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4084 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4085 |
lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4086 |
by (unfold s_holding_def wq_def cs_holding_def, simp) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4087 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4088 |
lemma holding_unique: "\<lbrakk>holding (s::state) th1 cs; holding s th2 cs\<rbrakk> \<Longrightarrow> th1 = th2" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4089 |
by (unfold s_holding_def cs_holding_def, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4090 |
|
63 | 4091 |
context valid_trace |
4092 |
begin |
|
4093 |
||
4094 |
lemma unique_RAG: "\<lbrakk>(n, n1) \<in> RAG s; (n, n2) \<in> RAG s\<rbrakk> \<Longrightarrow> n1 = n2" |
|
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4095 |
apply(unfold s_RAG_def, auto, fold waiting_eq holding_eq) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4096 |
by(auto elim:waiting_unique holding_unique) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4097 |
|
63 | 4098 |
end |
4099 |
||
4100 |
||
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4101 |
lemma trancl_split: "(a, b) \<in> r^+ \<Longrightarrow> \<exists> c. (a, c) \<in> r" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4102 |
by (induct rule:trancl_induct, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4103 |
|
63 | 4104 |
context valid_trace |
4105 |
begin |
|
4106 |
||
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4107 |
lemma dchain_unique: |
63 | 4108 |
assumes th1_d: "(n, Th th1) \<in> (RAG s)^+" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4109 |
and th1_r: "th1 \<in> readys s" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4110 |
and th2_d: "(n, Th th2) \<in> (RAG s)^+" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4111 |
and th2_r: "th2 \<in> readys s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4112 |
shows "th1 = th2" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4113 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4114 |
{ assume neq: "th1 \<noteq> th2" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4115 |
hence "Th th1 \<noteq> Th th2" by simp |
63 | 4116 |
from unique_chain [OF _ th1_d th2_d this] and unique_RAG |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4117 |
have "(Th th1, Th th2) \<in> (RAG s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (RAG s)\<^sup>+" by auto |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4118 |
hence "False" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4119 |
proof |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4120 |
assume "(Th th1, Th th2) \<in> (RAG s)\<^sup>+" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4121 |
from trancl_split [OF this] |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4122 |
obtain n where dd: "(Th th1, n) \<in> RAG s" by auto |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4123 |
then obtain cs where eq_n: "n = Cs cs" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4124 |
by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4125 |
from dd eq_n have "th1 \<notin> readys s" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4126 |
by (auto simp:readys_def s_RAG_def wq_def s_waiting_def cs_waiting_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4127 |
with th1_r show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4128 |
next |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4129 |
assume "(Th th2, Th th1) \<in> (RAG s)\<^sup>+" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4130 |
from trancl_split [OF this] |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4131 |
obtain n where dd: "(Th th2, n) \<in> RAG s" by auto |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4132 |
then obtain cs where eq_n: "n = Cs cs" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4133 |
by (auto simp:s_RAG_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4134 |
from dd eq_n have "th2 \<notin> readys s" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4135 |
by (auto simp:readys_def wq_def s_RAG_def s_waiting_def cs_waiting_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4136 |
with th2_r show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4137 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4138 |
} thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4139 |
qed |
63 | 4140 |
|
4141 |
end |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4142 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4143 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4144 |
lemma step_holdents_p_add: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4145 |
assumes vt: "vt (P th cs#s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4146 |
and "wq s cs = []" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4147 |
shows "holdents (P th cs#s) th = holdents s th \<union> {cs}" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4148 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4149 |
from assms show ?thesis |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4150 |
unfolding holdents_test step_RAG_p[OF vt] by (auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4151 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4152 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4153 |
lemma step_holdents_p_eq: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4154 |
assumes vt: "vt (P th cs#s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4155 |
and "wq s cs \<noteq> []" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4156 |
shows "holdents (P th cs#s) th = holdents s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4157 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4158 |
from assms show ?thesis |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4159 |
unfolding holdents_test step_RAG_p[OF vt] by auto |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4160 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4161 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4162 |
|
63 | 4163 |
lemma (in valid_trace) finite_holding : |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4164 |
shows "finite (holdents s th)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4165 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4166 |
let ?F = "\<lambda> (x, y). the_cs x" |
63 | 4167 |
from finite_RAG |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4168 |
have "finite (RAG s)" . |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4169 |
hence "finite (?F `(RAG s))" by simp |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4170 |
moreover have "{cs . (Cs cs, Th th) \<in> RAG s} \<subseteq> \<dots>" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4171 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4172 |
{ have h: "\<And> a A f. a \<in> A \<Longrightarrow> f a \<in> f ` A" by auto |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4173 |
fix x assume "(Cs x, Th th) \<in> RAG s" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4174 |
hence "?F (Cs x, Th th) \<in> ?F `(RAG s)" by (rule h) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4175 |
moreover have "?F (Cs x, Th th) = x" by simp |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4176 |
ultimately have "x \<in> (\<lambda>(x, y). the_cs x) ` RAG s" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4177 |
} thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4178 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4179 |
ultimately show ?thesis by (unfold holdents_test, auto intro:finite_subset) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4180 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4181 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4182 |
lemma cntCS_v_dec: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4183 |
assumes vtv: "vt (V thread cs#s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4184 |
shows "(cntCS (V thread cs#s) thread + 1) = cntCS s thread" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4185 |
proof - |
63 | 4186 |
from vtv interpret vt_s: valid_trace s |
4187 |
by (cases, unfold_locales, simp) |
|
4188 |
from vtv interpret vt_v: valid_trace "V thread cs#s" |
|
4189 |
by (unfold_locales, simp) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4190 |
from step_back_step[OF vtv] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4191 |
have cs_in: "cs \<in> holdents s thread" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4192 |
apply (cases, unfold holdents_test s_RAG_def, simp) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4193 |
by (unfold cs_holding_def s_holding_def wq_def, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4194 |
moreover have cs_not_in: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4195 |
"(holdents (V thread cs#s) thread) = holdents s thread - {cs}" |
63 | 4196 |
apply (insert vt_s.wq_distinct[of cs]) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4197 |
apply (unfold holdents_test, unfold step_RAG_v[OF vtv], |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4198 |
auto simp:next_th_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4199 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4200 |
fix rest |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4201 |
assume dst: "distinct (rest::thread list)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4202 |
and ne: "rest \<noteq> []" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4203 |
and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4204 |
moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4205 |
proof(rule someI2) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4206 |
from dst show "distinct rest \<and> set rest = set rest" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4207 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4208 |
show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4209 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4210 |
ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4211 |
set (SOME q. distinct q \<and> set q = set rest)" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4212 |
moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4213 |
proof(rule someI2) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4214 |
from dst show "distinct rest \<and> set rest = set rest" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4215 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4216 |
fix x assume " distinct x \<and> set x = set rest" with ne |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4217 |
show "x \<noteq> []" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4218 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4219 |
ultimately |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4220 |
show "(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4221 |
by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4222 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4223 |
fix rest |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4224 |
assume dst: "distinct (rest::thread list)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4225 |
and ne: "rest \<noteq> []" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4226 |
and hd_ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4227 |
moreover have "set (SOME q. distinct q \<and> set q = set rest) = set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4228 |
proof(rule someI2) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4229 |
from dst show "distinct rest \<and> set rest = set rest" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4230 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4231 |
show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4232 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4233 |
ultimately have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4234 |
set (SOME q. distinct q \<and> set q = set rest)" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4235 |
moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4236 |
proof(rule someI2) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4237 |
from dst show "distinct rest \<and> set rest = set rest" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4238 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4239 |
fix x assume " distinct x \<and> set x = set rest" with ne |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4240 |
show "x \<noteq> []" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4241 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4242 |
ultimately show "False" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4243 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4244 |
ultimately |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4245 |
have "holdents s thread = insert cs (holdents (V thread cs#s) thread)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4246 |
by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4247 |
moreover have "card \<dots> = |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4248 |
Suc (card ((holdents (V thread cs#s) thread) - {cs}))" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4249 |
proof(rule card_insert) |
63 | 4250 |
from vt_v.finite_holding |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4251 |
show " finite (holdents (V thread cs # s) thread)" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4252 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4253 |
moreover from cs_not_in |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4254 |
have "cs \<notin> (holdents (V thread cs#s) thread)" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4255 |
ultimately show ?thesis by (simp add:cntCS_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4256 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4257 |
|
67 | 4258 |
lemma count_rec1 [simp]: |
4259 |
assumes "Q e" |
|
4260 |
shows "count Q (e#es) = Suc (count Q es)" |
|
4261 |
using assms |
|
4262 |
by (unfold count_def, auto) |
|
4263 |
||
4264 |
lemma count_rec2 [simp]: |
|
4265 |
assumes "\<not>Q e" |
|
4266 |
shows "count Q (e#es) = (count Q es)" |
|
4267 |
using assms |
|
4268 |
by (unfold count_def, auto) |
|
4269 |
||
4270 |
lemma count_rec3 [simp]: |
|
4271 |
shows "count Q [] = 0" |
|
4272 |
by (unfold count_def, auto) |
|
4273 |
||
4274 |
lemma cntP_diff_inv: |
|
4275 |
assumes "cntP (e#s) th \<noteq> cntP s th" |
|
4276 |
shows "isP e \<and> actor e = th" |
|
4277 |
proof(cases e) |
|
4278 |
case (P th' pty) |
|
4279 |
show ?thesis |
|
4280 |
by (cases "(\<lambda>e. \<exists>cs. e = P th cs) (P th' pty)", |
|
4281 |
insert assms P, auto simp:cntP_def) |
|
4282 |
qed (insert assms, auto simp:cntP_def) |
|
4283 |
||
4284 |
lemma isP_E: |
|
4285 |
assumes "isP e" |
|
4286 |
obtains cs where "e = P (actor e) cs" |
|
4287 |
using assms by (cases e, auto) |
|
4288 |
||
4289 |
lemma isV_E: |
|
4290 |
assumes "isV e" |
|
4291 |
obtains cs where "e = V (actor e) cs" |
|
4292 |
using assms by (cases e, auto) (* ccc *) |
|
4293 |
||
4294 |
lemma cntV_diff_inv: |
|
4295 |
assumes "cntV (e#s) th \<noteq> cntV s th" |
|
4296 |
shows "isV e \<and> actor e = th" |
|
4297 |
proof(cases e) |
|
4298 |
case (V th' pty) |
|
4299 |
show ?thesis |
|
4300 |
by (cases "(\<lambda>e. \<exists>cs. e = V th cs) (V th' pty)", |
|
4301 |
insert assms V, auto simp:cntV_def) |
|
4302 |
qed (insert assms, auto simp:cntV_def) |
|
4303 |
||
63 | 4304 |
context valid_trace |
4305 |
begin |
|
4306 |
||
55
b85cfbd58f59
Comments for Set-operation finished
xingyuan zhang <xingyuanzhang@126.com>
parents:
53
diff
changeset
|
4307 |
text {* (* ddd *) \noindent |
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
4308 |
The relationship between @{text "cntP"}, @{text "cntV"} and @{text "cntCS"} |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
4309 |
of one particular thread. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
4310 |
*} |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4311 |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4312 |
lemma cnp_cnv_cncs: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4313 |
shows "cntP s th = cntV s th + (if (th \<in> readys s \<or> th \<notin> threads s) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4314 |
then cntCS s th else cntCS s th + 1)" |
104 | 4315 |
>>>>>>> other |
4316 |
proof - |
|
4317 |
<<<<<<< local |
|
101 | 4318 |
have "?A = the_thread ` {Th th' | th' . Th th' \<in> subtree (RAG s) (Th th)}" |
4319 |
by (auto, insert image_iff, fastforce) |
|
4320 |
moreover have "finite {Th th' | th' . Th th' \<in> subtree (RAG s) (Th th)}" |
|
4321 |
(is "finite ?B") |
|
4322 |
proof - |
|
4323 |
have "?B = (subtree (RAG s) (Th th)) \<inter> {Th th' | th'. True}" |
|
4324 |
by auto |
|
4325 |
moreover have "... \<subseteq> (subtree (RAG s) (Th th))" by auto |
|
4326 |
moreover have "finite ..." by (simp add: fsbtRAGs.finite_subtree) |
|
4327 |
ultimately show ?thesis by auto |
|
4328 |
qed |
|
4329 |
ultimately show ?thesis by auto |
|
104 | 4330 |
======= |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4331 |
from vt show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4332 |
proof(induct arbitrary:th) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4333 |
case (vt_cons s e) |
63 | 4334 |
interpret vt_s: valid_trace s using vt_cons(1) by (unfold_locales, simp) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4335 |
assume vt: "vt s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4336 |
and ih: "\<And>th. cntP s th = cntV s th + |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4337 |
(if (th \<in> readys s \<or> th \<notin> threads s) then cntCS s th else cntCS s th + 1)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4338 |
and stp: "step s e" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4339 |
from stp show ?case |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4340 |
proof(cases) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4341 |
case (thread_create thread prio) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4342 |
assume eq_e: "e = Create thread prio" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4343 |
and not_in: "thread \<notin> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4344 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4345 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4346 |
{ fix cs |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4347 |
assume "thread \<in> set (wq s cs)" |
63 | 4348 |
from vt_s.wq_threads [OF this] have "thread \<in> threads s" . |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4349 |
with not_in have "False" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4350 |
} with eq_e have eq_readys: "readys (e#s) = readys s \<union> {thread}" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4351 |
by (auto simp:readys_def threads.simps s_waiting_def |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4352 |
wq_def cs_waiting_def Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4353 |
from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4354 |
from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4355 |
have eq_cncs: "cntCS (e#s) th = cntCS s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4356 |
unfolding cntCS_def holdents_test |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4357 |
by (simp add:RAG_create_unchanged eq_e) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4358 |
{ assume "th \<noteq> thread" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4359 |
with eq_readys eq_e |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4360 |
have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4361 |
(th \<in> readys (s) \<or> th \<notin> threads (s))" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4362 |
by (simp add:threads.simps) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4363 |
with eq_cnp eq_cnv eq_cncs ih not_in |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4364 |
have ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4365 |
} moreover { |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4366 |
assume eq_th: "th = thread" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4367 |
with not_in ih have " cntP s th = cntV s th + cntCS s th" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4368 |
moreover from eq_th and eq_readys have "th \<in> readys (e#s)" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4369 |
moreover note eq_cnp eq_cnv eq_cncs |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4370 |
ultimately have ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4371 |
} ultimately show ?thesis by blast |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4372 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4373 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4374 |
case (thread_exit thread) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4375 |
assume eq_e: "e = Exit thread" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4376 |
and is_runing: "thread \<in> runing s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4377 |
and no_hold: "holdents s thread = {}" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4378 |
from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4379 |
from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4380 |
have eq_cncs: "cntCS (e#s) th = cntCS s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4381 |
unfolding cntCS_def holdents_test |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4382 |
by (simp add:RAG_exit_unchanged eq_e) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4383 |
{ assume "th \<noteq> thread" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4384 |
with eq_e |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4385 |
have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4386 |
(th \<in> readys (s) \<or> th \<notin> threads (s))" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4387 |
apply (simp add:threads.simps readys_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4388 |
apply (subst s_waiting_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4389 |
apply (simp add:Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4390 |
apply (subst s_waiting_def, simp) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4391 |
done |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4392 |
with eq_cnp eq_cnv eq_cncs ih |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4393 |
have ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4394 |
} moreover { |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4395 |
assume eq_th: "th = thread" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4396 |
with ih is_runing have " cntP s th = cntV s th + cntCS s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4397 |
by (simp add:runing_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4398 |
moreover from eq_th eq_e have "th \<notin> threads (e#s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4399 |
by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4400 |
moreover note eq_cnp eq_cnv eq_cncs |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4401 |
ultimately have ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4402 |
} ultimately show ?thesis by blast |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4403 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4404 |
case (thread_P thread cs) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4405 |
assume eq_e: "e = P thread cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4406 |
and is_runing: "thread \<in> runing s" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4407 |
and no_dep: "(Cs cs, Th thread) \<notin> (RAG s)\<^sup>+" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4408 |
from thread_P vt stp ih have vtp: "vt (P thread cs#s)" by auto |
63 | 4409 |
then interpret vt_p: valid_trace "(P thread cs#s)" |
4410 |
by (unfold_locales, simp) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4411 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4412 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4413 |
{ have hh: "\<And> A B C. (B = C) \<Longrightarrow> (A \<and> B) = (A \<and> C)" by blast |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4414 |
assume neq_th: "th \<noteq> thread" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4415 |
with eq_e |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4416 |
have eq_readys: "(th \<in> readys (e#s)) = (th \<in> readys (s))" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4417 |
apply (simp add:readys_def s_waiting_def wq_def Let_def) |
44
f676a68935a0
updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
4418 |
apply (rule_tac hh) |
f676a68935a0
updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
4419 |
apply (intro iffI allI, clarify) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4420 |
apply (erule_tac x = csa in allE, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4421 |
apply (subgoal_tac "wq_fun (schs s) cs \<noteq> []", auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4422 |
apply (erule_tac x = cs in allE, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4423 |
by (case_tac "(wq_fun (schs s) cs)", auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4424 |
moreover from neq_th eq_e have "cntCS (e # s) th = cntCS s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4425 |
apply (simp add:cntCS_def holdents_test) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4426 |
by (unfold step_RAG_p [OF vtp], auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4427 |
moreover from eq_e neq_th have "cntP (e # s) th = cntP s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4428 |
by (simp add:cntP_def count_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4429 |
moreover from eq_e neq_th have "cntV (e#s) th = cntV s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4430 |
by (simp add:cntV_def count_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4431 |
moreover from eq_e neq_th have "threads (e#s) = threads s" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4432 |
moreover note ih [of th] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4433 |
ultimately have ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4434 |
} moreover { |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4435 |
assume eq_th: "th = thread" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4436 |
have ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4437 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4438 |
from eq_e eq_th have eq_cnp: "cntP (e # s) th = 1 + (cntP s th)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4439 |
by (simp add:cntP_def count_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4440 |
from eq_e eq_th have eq_cnv: "cntV (e#s) th = cntV s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4441 |
by (simp add:cntV_def count_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4442 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4443 |
proof (cases "wq s cs = []") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4444 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4445 |
with is_runing |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4446 |
have "th \<in> readys (e#s)" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4447 |
apply (unfold eq_e wq_def, unfold readys_def s_RAG_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4448 |
apply (simp add: wq_def[symmetric] runing_def eq_th s_waiting_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4449 |
by (auto simp:readys_def wq_def Let_def s_waiting_def wq_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4450 |
moreover have "cntCS (e # s) th = 1 + cntCS s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4451 |
proof - |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4452 |
have "card {csa. csa = cs \<or> (Cs csa, Th thread) \<in> RAG s} = |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4453 |
Suc (card {cs. (Cs cs, Th thread) \<in> RAG s})" (is "card ?L = Suc (card ?R)") |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4454 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4455 |
have "?L = insert cs ?R" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4456 |
moreover have "card \<dots> = Suc (card (?R - {cs}))" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4457 |
proof(rule card_insert) |
63 | 4458 |
from vt_s.finite_holding [of thread] |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4459 |
show " finite {cs. (Cs cs, Th thread) \<in> RAG s}" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4460 |
by (unfold holdents_test, simp) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4461 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4462 |
moreover have "?R - {cs} = ?R" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4463 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4464 |
have "cs \<notin> ?R" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4465 |
proof |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4466 |
assume "cs \<in> {cs. (Cs cs, Th thread) \<in> RAG s}" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4467 |
with no_dep show False by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4468 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4469 |
thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4470 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4471 |
ultimately show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4472 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4473 |
thus ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4474 |
apply (unfold eq_e eq_th cntCS_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4475 |
apply (simp add: holdents_test) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4476 |
by (unfold step_RAG_p [OF vtp], auto simp:True) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4477 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4478 |
moreover from is_runing have "th \<in> readys s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4479 |
by (simp add:runing_def eq_th) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4480 |
moreover note eq_cnp eq_cnv ih [of th] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4481 |
ultimately show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4482 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4483 |
case False |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4484 |
have eq_wq: "wq (e#s) cs = wq s cs @ [th]" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4485 |
by (unfold eq_th eq_e wq_def, auto simp:Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4486 |
have "th \<notin> readys (e#s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4487 |
proof |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4488 |
assume "th \<in> readys (e#s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4489 |
hence "\<forall>cs. \<not> waiting (e # s) th cs" by (simp add:readys_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4490 |
from this[rule_format, of cs] have " \<not> waiting (e # s) th cs" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4491 |
hence "th \<in> set (wq (e#s) cs) \<Longrightarrow> th = hd (wq (e#s) cs)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4492 |
by (simp add:s_waiting_def wq_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4493 |
moreover from eq_wq have "th \<in> set (wq (e#s) cs)" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4494 |
ultimately have "th = hd (wq (e#s) cs)" by blast |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4495 |
with eq_wq have "th = hd (wq s cs @ [th])" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4496 |
hence "th = hd (wq s cs)" using False by auto |
63 | 4497 |
with False eq_wq vt_p.wq_distinct [of cs] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4498 |
show False by (fold eq_e, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4499 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4500 |
moreover from is_runing have "th \<in> threads (e#s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4501 |
by (unfold eq_e, auto simp:runing_def readys_def eq_th) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4502 |
moreover have "cntCS (e # s) th = cntCS s th" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4503 |
apply (unfold cntCS_def holdents_test eq_e step_RAG_p[OF vtp]) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4504 |
by (auto simp:False) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4505 |
moreover note eq_cnp eq_cnv ih[of th] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4506 |
moreover from is_runing have "th \<in> readys s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4507 |
by (simp add:runing_def eq_th) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4508 |
ultimately show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4509 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4510 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4511 |
} ultimately show ?thesis by blast |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4512 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4513 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4514 |
case (thread_V thread cs) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4515 |
from assms vt stp ih thread_V have vtv: "vt (V thread cs # s)" by auto |
63 | 4516 |
then interpret vt_v: valid_trace "(V thread cs # s)" by (unfold_locales, simp) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4517 |
assume eq_e: "e = V thread cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4518 |
and is_runing: "thread \<in> runing s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4519 |
and hold: "holding s thread cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4520 |
from hold obtain rest |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4521 |
where eq_wq: "wq s cs = thread # rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4522 |
by (case_tac "wq s cs", auto simp: wq_def s_holding_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4523 |
have eq_threads: "threads (e#s) = threads s" by (simp add: eq_e) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4524 |
have eq_set: "set (SOME q. distinct q \<and> set q = set rest) = set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4525 |
proof(rule someI2) |
63 | 4526 |
from vt_v.wq_distinct[of cs] and eq_wq |
4527 |
show "distinct rest \<and> set rest = set rest" |
|
4528 |
by (metis distinct.simps(2) vt_s.wq_distinct) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4529 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4530 |
show "\<And>x. distinct x \<and> set x = set rest \<Longrightarrow> set x = set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4531 |
by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4532 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4533 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4534 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4535 |
{ assume eq_th: "th = thread" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4536 |
from eq_th have eq_cnp: "cntP (e # s) th = cntP s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4537 |
by (unfold eq_e, simp add:cntP_def count_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4538 |
moreover from eq_th have eq_cnv: "cntV (e#s) th = 1 + cntV s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4539 |
by (unfold eq_e, simp add:cntV_def count_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4540 |
moreover from cntCS_v_dec [OF vtv] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4541 |
have "cntCS (e # s) thread + 1 = cntCS s thread" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4542 |
by (simp add:eq_e) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4543 |
moreover from is_runing have rd_before: "thread \<in> readys s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4544 |
by (unfold runing_def, simp) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4545 |
moreover have "thread \<in> readys (e # s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4546 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4547 |
from is_runing |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4548 |
have "thread \<in> threads (e#s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4549 |
by (unfold eq_e, auto simp:runing_def readys_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4550 |
moreover have "\<forall> cs1. \<not> waiting (e#s) thread cs1" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4551 |
proof |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4552 |
fix cs1 |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4553 |
{ assume eq_cs: "cs1 = cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4554 |
have "\<not> waiting (e # s) thread cs1" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4555 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4556 |
from eq_wq |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4557 |
have "thread \<notin> set (wq (e#s) cs1)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4558 |
apply(unfold eq_e wq_def eq_cs s_holding_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4559 |
apply (auto simp:Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4560 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4561 |
assume "thread \<in> set (SOME q. distinct q \<and> set q = set rest)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4562 |
with eq_set have "thread \<in> set rest" by simp |
63 | 4563 |
with vt_v.wq_distinct[of cs] |
4564 |
and eq_wq show False |
|
4565 |
by (metis distinct.simps(2) vt_s.wq_distinct) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4566 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4567 |
thus ?thesis by (simp add:wq_def s_waiting_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4568 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4569 |
} moreover { |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4570 |
assume neq_cs: "cs1 \<noteq> cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4571 |
have "\<not> waiting (e # s) thread cs1" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4572 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4573 |
from wq_v_neq [OF neq_cs[symmetric]] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4574 |
have "wq (V thread cs # s) cs1 = wq s cs1" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4575 |
moreover have "\<not> waiting s thread cs1" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4576 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4577 |
from runing_ready and is_runing |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4578 |
have "thread \<in> readys s" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4579 |
thus ?thesis by (simp add:readys_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4580 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4581 |
ultimately show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4582 |
by (auto simp:wq_def s_waiting_def eq_e) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4583 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4584 |
} ultimately show "\<not> waiting (e # s) thread cs1" by blast |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4585 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4586 |
ultimately show ?thesis by (simp add:readys_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4587 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4588 |
moreover note eq_th ih |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4589 |
ultimately have ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4590 |
} moreover { |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4591 |
assume neq_th: "th \<noteq> thread" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4592 |
from neq_th eq_e have eq_cnp: "cntP (e # s) th = cntP s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4593 |
by (simp add:cntP_def count_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4594 |
from neq_th eq_e have eq_cnv: "cntV (e # s) th = cntV s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4595 |
by (simp add:cntV_def count_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4596 |
have ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4597 |
proof(cases "th \<in> set rest") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4598 |
case False |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4599 |
have "(th \<in> readys (e # s)) = (th \<in> readys s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4600 |
apply (insert step_back_vt[OF vtv]) |
63 | 4601 |
by (simp add: False eq_e eq_wq neq_th vt_s.readys_v_eq) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4602 |
moreover have "cntCS (e#s) th = cntCS s th" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4603 |
apply (insert neq_th, unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4604 |
proof - |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4605 |
have "{csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} = |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4606 |
{cs. (Cs cs, Th th) \<in> RAG s}" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4607 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4608 |
from False eq_wq |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4609 |
have " next_th s thread cs th \<Longrightarrow> (Cs cs, Th th) \<in> RAG s" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4610 |
apply (unfold next_th_def, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4611 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4612 |
assume ne: "rest \<noteq> []" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4613 |
and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4614 |
and eq_wq: "wq s cs = thread # rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4615 |
from eq_set ni have "hd (SOME q. distinct q \<and> set q = set rest) \<notin> |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4616 |
set (SOME q. distinct q \<and> set q = set rest) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4617 |
" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4618 |
moreover have "(SOME q. distinct q \<and> set q = set rest) \<noteq> []" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4619 |
proof(rule someI2) |
63 | 4620 |
from vt_s.wq_distinct[ of cs] and eq_wq |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4621 |
show "distinct rest \<and> set rest = set rest" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4622 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4623 |
fix x assume "distinct x \<and> set x = set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4624 |
with ne show "x \<noteq> []" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4625 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4626 |
ultimately show |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4627 |
"(Cs cs, Th (hd (SOME q. distinct q \<and> set q = set rest))) \<in> RAG s" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4628 |
by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4629 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4630 |
thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4631 |
qed |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4632 |
thus "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs \<and> next_th s thread cs th} = |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4633 |
card {cs. (Cs cs, Th th) \<in> RAG s}" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4634 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4635 |
moreover note ih eq_cnp eq_cnv eq_threads |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4636 |
ultimately show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4637 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4638 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4639 |
assume th_in: "th \<in> set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4640 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4641 |
proof(cases "next_th s thread cs th") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4642 |
case False |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4643 |
with eq_wq and th_in have |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4644 |
neq_hd: "th \<noteq> hd (SOME q. distinct q \<and> set q = set rest)" (is "th \<noteq> hd ?rest") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4645 |
by (auto simp:next_th_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4646 |
have "(th \<in> readys (e # s)) = (th \<in> readys s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4647 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4648 |
from eq_wq and th_in |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4649 |
have "\<not> th \<in> readys s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4650 |
apply (auto simp:readys_def s_waiting_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4651 |
apply (rule_tac x = cs in exI, auto) |
63 | 4652 |
by (insert vt_s.wq_distinct[of cs], auto simp add: wq_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4653 |
moreover |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4654 |
from eq_wq and th_in and neq_hd |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4655 |
have "\<not> (th \<in> readys (e # s))" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4656 |
apply (auto simp:readys_def s_waiting_def eq_e wq_def Let_def split:list.splits) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4657 |
by (rule_tac x = cs in exI, auto simp:eq_set) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4658 |
ultimately show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4659 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4660 |
moreover have "cntCS (e#s) th = cntCS s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4661 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4662 |
from eq_wq and th_in and neq_hd |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4663 |
have "(holdents (e # s) th) = (holdents s th)" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4664 |
apply (unfold eq_e step_RAG_v[OF vtv], |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4665 |
auto simp:next_th_def eq_set s_RAG_def holdents_test wq_def |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4666 |
Let_def cs_holding_def) |
63 | 4667 |
by (insert vt_s.wq_distinct[of cs], auto simp:wq_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4668 |
thus ?thesis by (simp add:cntCS_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4669 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4670 |
moreover note ih eq_cnp eq_cnv eq_threads |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4671 |
ultimately show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4672 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4673 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4674 |
let ?rest = " (SOME q. distinct q \<and> set q = set rest)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4675 |
let ?t = "hd ?rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4676 |
from True eq_wq th_in neq_th |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4677 |
have "th \<in> readys (e # s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4678 |
apply (auto simp:eq_e readys_def s_waiting_def wq_def |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4679 |
Let_def next_th_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4680 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4681 |
assume eq_wq: "wq_fun (schs s) cs = thread # rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4682 |
and t_in: "?t \<in> set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4683 |
show "?t \<in> threads s" |
63 | 4684 |
proof(rule vt_s.wq_threads) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4685 |
from eq_wq and t_in |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4686 |
show "?t \<in> set (wq s cs)" by (auto simp:wq_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4687 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4688 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4689 |
fix csa |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4690 |
assume eq_wq: "wq_fun (schs s) cs = thread # rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4691 |
and t_in: "?t \<in> set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4692 |
and neq_cs: "csa \<noteq> cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4693 |
and t_in': "?t \<in> set (wq_fun (schs s) csa)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4694 |
show "?t = hd (wq_fun (schs s) csa)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4695 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4696 |
{ assume neq_hd': "?t \<noteq> hd (wq_fun (schs s) csa)" |
63 | 4697 |
from vt_s.wq_distinct[of cs] and |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4698 |
eq_wq[folded wq_def] and t_in eq_wq |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4699 |
have "?t \<noteq> thread" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4700 |
with eq_wq and t_in |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4701 |
have w1: "waiting s ?t cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4702 |
by (auto simp:s_waiting_def wq_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4703 |
from t_in' neq_hd' |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4704 |
have w2: "waiting s ?t csa" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4705 |
by (auto simp:s_waiting_def wq_def) |
63 | 4706 |
from vt_s.waiting_unique[OF w1 w2] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4707 |
and neq_cs have "False" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4708 |
} thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4709 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4710 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4711 |
moreover have "cntP s th = cntV s th + cntCS s th + 1" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4712 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4713 |
have "th \<notin> readys s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4714 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4715 |
from True eq_wq neq_th th_in |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4716 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4717 |
apply (unfold readys_def s_waiting_def, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4718 |
by (rule_tac x = cs in exI, auto simp add: wq_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4719 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4720 |
moreover have "th \<in> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4721 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4722 |
from th_in eq_wq |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4723 |
have "th \<in> set (wq s cs)" by simp |
63 | 4724 |
from vt_s.wq_threads [OF this] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4725 |
show ?thesis . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4726 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4727 |
ultimately show ?thesis using ih by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4728 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4729 |
moreover from True neq_th have "cntCS (e # s) th = 1 + cntCS s th" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4730 |
apply (unfold cntCS_def holdents_test eq_e step_RAG_v[OF vtv], auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4731 |
proof - |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4732 |
show "card {csa. (Cs csa, Th th) \<in> RAG s \<or> csa = cs} = |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4733 |
Suc (card {cs. (Cs cs, Th th) \<in> RAG s})" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4734 |
(is "card ?A = Suc (card ?B)") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4735 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4736 |
have "?A = insert cs ?B" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4737 |
hence "card ?A = card (insert cs ?B)" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4738 |
also have "\<dots> = Suc (card ?B)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4739 |
proof(rule card_insert_disjoint) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4740 |
have "?B \<subseteq> ((\<lambda> (x, y). the_cs x) ` RAG s)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4741 |
apply (auto simp:image_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4742 |
by (rule_tac x = "(Cs x, Th th)" in bexI, auto) |
63 | 4743 |
with vt_s.finite_RAG |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4744 |
show "finite {cs. (Cs cs, Th th) \<in> RAG s}" by (auto intro:finite_subset) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4745 |
next |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4746 |
show "cs \<notin> {cs. (Cs cs, Th th) \<in> RAG s}" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4747 |
proof |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4748 |
assume "cs \<in> {cs. (Cs cs, Th th) \<in> RAG s}" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4749 |
hence "(Cs cs, Th th) \<in> RAG s" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4750 |
with True neq_th eq_wq show False |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4751 |
by (auto simp:next_th_def s_RAG_def cs_holding_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4752 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4753 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4754 |
finally show ?thesis . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4755 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4756 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4757 |
moreover note eq_cnp eq_cnv |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4758 |
ultimately show ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4759 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4760 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4761 |
} ultimately show ?thesis by blast |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4762 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4763 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4764 |
case (thread_set thread prio) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4765 |
assume eq_e: "e = Set thread prio" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4766 |
and is_runing: "thread \<in> runing s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4767 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4768 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4769 |
from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4770 |
from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4771 |
have eq_cncs: "cntCS (e#s) th = cntCS s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4772 |
unfolding cntCS_def holdents_test |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4773 |
by (simp add:RAG_set_unchanged eq_e) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4774 |
from eq_e have eq_readys: "readys (e#s) = readys s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4775 |
by (simp add:readys_def cs_waiting_def s_waiting_def wq_def, |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4776 |
auto simp:Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4777 |
{ assume "th \<noteq> thread" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4778 |
with eq_readys eq_e |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4779 |
have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4780 |
(th \<in> readys (s) \<or> th \<notin> threads (s))" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4781 |
by (simp add:threads.simps) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4782 |
with eq_cnp eq_cnv eq_cncs ih is_runing |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4783 |
have ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4784 |
} moreover { |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4785 |
assume eq_th: "th = thread" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4786 |
with is_runing ih have " cntP s th = cntV s th + cntCS s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4787 |
by (unfold runing_def, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4788 |
moreover from eq_th and eq_readys is_runing have "th \<in> readys (e#s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4789 |
by (simp add:runing_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4790 |
moreover note eq_cnp eq_cnv eq_cncs |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4791 |
ultimately have ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4792 |
} ultimately show ?thesis by blast |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4793 |
qed |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4794 |
qed |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4795 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4796 |
case vt_nil |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4797 |
show ?case |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4798 |
by (unfold cntP_def cntV_def cntCS_def, |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4799 |
auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4800 |
qed |
104 | 4801 |
>>>>>>> other |
4802 |
qed |
|
4803 |
||
4804 |
<<<<<<< local |
|
4805 |
======= |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4806 |
lemma not_thread_cncs: |
63 | 4807 |
assumes not_in: "th \<notin> threads s" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4808 |
shows "cntCS s th = 0" |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4809 |
proof - |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4810 |
from vt not_in show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4811 |
proof(induct arbitrary:th) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4812 |
case (vt_cons s e th) |
63 | 4813 |
interpret vt_s: valid_trace s using vt_cons(1) |
4814 |
by (unfold_locales, simp) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4815 |
assume vt: "vt s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4816 |
and ih: "\<And>th. th \<notin> threads s \<Longrightarrow> cntCS s th = 0" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4817 |
and stp: "step s e" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4818 |
and not_in: "th \<notin> threads (e # s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4819 |
from stp show ?case |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4820 |
proof(cases) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4821 |
case (thread_create thread prio) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4822 |
assume eq_e: "e = Create thread prio" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4823 |
and not_in': "thread \<notin> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4824 |
have "cntCS (e # s) th = cntCS s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4825 |
apply (unfold eq_e cntCS_def holdents_test) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4826 |
by (simp add:RAG_create_unchanged) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4827 |
moreover have "th \<notin> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4828 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4829 |
from not_in eq_e show ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4830 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4831 |
moreover note ih ultimately show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4832 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4833 |
case (thread_exit thread) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4834 |
assume eq_e: "e = Exit thread" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4835 |
and nh: "holdents s thread = {}" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4836 |
have eq_cns: "cntCS (e # s) th = cntCS s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4837 |
apply (unfold eq_e cntCS_def holdents_test) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4838 |
by (simp add:RAG_exit_unchanged) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4839 |
show ?thesis |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4840 |
proof(cases "th = thread") |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4841 |
case True |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4842 |
have "cntCS s th = 0" by (unfold cntCS_def, auto simp:nh True) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4843 |
with eq_cns show ?thesis by simp |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4844 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4845 |
case False |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4846 |
with not_in and eq_e |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4847 |
have "th \<notin> threads s" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4848 |
from ih[OF this] and eq_cns show ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4849 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4850 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4851 |
case (thread_P thread cs) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4852 |
assume eq_e: "e = P thread cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4853 |
and is_runing: "thread \<in> runing s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4854 |
from assms thread_P ih vt stp thread_P have vtp: "vt (P thread cs#s)" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4855 |
have neq_th: "th \<noteq> thread" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4856 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4857 |
from not_in eq_e have "th \<notin> threads s" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4858 |
moreover from is_runing have "thread \<in> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4859 |
by (simp add:runing_def readys_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4860 |
ultimately show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4861 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4862 |
hence "cntCS (e # s) th = cntCS s th " |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4863 |
apply (unfold cntCS_def holdents_test eq_e) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4864 |
by (unfold step_RAG_p[OF vtp], auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4865 |
moreover have "cntCS s th = 0" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4866 |
proof(rule ih) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4867 |
from not_in eq_e show "th \<notin> threads s" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4868 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4869 |
ultimately show ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4870 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4871 |
case (thread_V thread cs) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4872 |
assume eq_e: "e = V thread cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4873 |
and is_runing: "thread \<in> runing s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4874 |
and hold: "holding s thread cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4875 |
have neq_th: "th \<noteq> thread" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4876 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4877 |
from not_in eq_e have "th \<notin> threads s" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4878 |
moreover from is_runing have "thread \<in> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4879 |
by (simp add:runing_def readys_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4880 |
ultimately show ?thesis by auto |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4881 |
qed |
63 | 4882 |
from assms thread_V vt stp ih |
4883 |
have vtv: "vt (V thread cs#s)" by auto |
|
4884 |
then interpret vt_v: valid_trace "(V thread cs#s)" |
|
4885 |
by (unfold_locales, simp) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4886 |
from hold obtain rest |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4887 |
where eq_wq: "wq s cs = thread # rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4888 |
by (case_tac "wq s cs", auto simp: wq_def s_holding_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4889 |
from not_in eq_e eq_wq |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4890 |
have "\<not> next_th s thread cs th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4891 |
apply (auto simp:next_th_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4892 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4893 |
assume ne: "rest \<noteq> []" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4894 |
and ni: "hd (SOME q. distinct q \<and> set q = set rest) \<notin> threads s" (is "?t \<notin> threads s") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4895 |
have "?t \<in> set rest" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4896 |
proof(rule someI2) |
63 | 4897 |
from vt_v.wq_distinct[of cs] and eq_wq |
4898 |
show "distinct rest \<and> set rest = set rest" |
|
4899 |
by (metis distinct.simps(2) vt_s.wq_distinct) |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4900 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4901 |
fix x assume "distinct x \<and> set x = set rest" with ne |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4902 |
show "hd x \<in> set rest" by (cases x, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4903 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4904 |
with eq_wq have "?t \<in> set (wq s cs)" by simp |
63 | 4905 |
from vt_s.wq_threads[OF this] and ni |
4906 |
show False |
|
4907 |
using `hd (SOME q. distinct q \<and> set q = set rest) \<in> set (wq s cs)` |
|
4908 |
ni vt_s.wq_threads by blast |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4909 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4910 |
moreover note neq_th eq_wq |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4911 |
ultimately have "cntCS (e # s) th = cntCS s th" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4912 |
by (unfold eq_e cntCS_def holdents_test step_RAG_v[OF vtv], auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4913 |
moreover have "cntCS s th = 0" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4914 |
proof(rule ih) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4915 |
from not_in eq_e show "th \<notin> threads s" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4916 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4917 |
ultimately show ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4918 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4919 |
case (thread_set thread prio) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4920 |
print_facts |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4921 |
assume eq_e: "e = Set thread prio" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4922 |
and is_runing: "thread \<in> runing s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4923 |
from not_in and eq_e have "th \<notin> threads s" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4924 |
from ih [OF this] and eq_e |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4925 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4926 |
apply (unfold eq_e cntCS_def holdents_test) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4927 |
by (simp add:RAG_set_unchanged) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4928 |
qed |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4929 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4930 |
case vt_nil |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4931 |
show ?case |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4932 |
by (unfold cntCS_def, |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4933 |
auto simp:count_def holdents_test s_RAG_def wq_def cs_holding_def) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4934 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4935 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4936 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4937 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4938 |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4939 |
lemma eq_waiting: "waiting (wq (s::state)) th cs = waiting s th cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4940 |
by (auto simp:s_waiting_def cs_waiting_def wq_def) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4941 |
|
63 | 4942 |
context valid_trace |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4943 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4944 |
|
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4945 |
lemma dm_RAG_threads: |
63 | 4946 |
assumes in_dom: "(Th th) \<in> Domain (RAG s)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4947 |
shows "th \<in> threads s" |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4948 |
proof - |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4949 |
from in_dom obtain n where "(Th th, n) \<in> RAG s" by auto |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4950 |
moreover from RAG_target_th[OF this] obtain cs where "n = Cs cs" by auto |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4951 |
ultimately have "(Th th, Cs cs) \<in> RAG s" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4952 |
hence "th \<in> set (wq s cs)" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
4953 |
by (unfold s_RAG_def, auto simp:cs_waiting_def) |
63 | 4954 |
from wq_threads [OF this] show ?thesis . |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4955 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4956 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4957 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4958 |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4959 |
lemma cp_eq_cpreced: "cp s th = cpreced (wq s) s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4960 |
unfolding cp_def wq_def |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4961 |
apply(induct s rule: schs.induct) |
44
f676a68935a0
updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
4962 |
thm cpreced_initial |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4963 |
apply(simp add: Let_def cpreced_initial) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4964 |
apply(simp add: Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4965 |
apply(simp add: Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4966 |
apply(simp add: Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4967 |
apply(subst (2) schs.simps) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4968 |
apply(simp add: Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4969 |
apply(subst (2) schs.simps) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4970 |
apply(simp add: Let_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4971 |
done |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4972 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4973 |
context valid_trace |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4974 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
4975 |
|
104 | 4976 |
>>>>>>> other |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4977 |
lemma runing_unique: |
63 | 4978 |
assumes runing_1: "th1 \<in> runing s" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4979 |
and runing_2: "th2 \<in> runing s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4980 |
shows "th1 = th2" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4981 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4982 |
from runing_1 and runing_2 have "cp s th1 = cp s th2" |
104 | 4983 |
<<<<<<< local |
101 | 4984 |
unfolding runing_def by auto |
4985 |
from this[unfolded cp_alt_def] |
|
4986 |
have eq_max: |
|
4987 |
"Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th1)}) = |
|
4988 |
Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th2)})" |
|
4989 |
(is "Max ?L = Max ?R") . |
|
4990 |
have "Max ?L \<in> ?L" |
|
4991 |
proof(rule Max_in) |
|
4992 |
show "finite ?L" by (simp add: finite_subtree_threads) |
|
4993 |
next |
|
4994 |
show "?L \<noteq> {}" using subtree_def by fastforce |
|
104 | 4995 |
======= |
44
f676a68935a0
updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
4996 |
unfolding runing_def |
f676a68935a0
updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
4997 |
apply(simp) |
f676a68935a0
updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
4998 |
done |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
4999 |
hence eq_max: "Max ((\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1)) = |
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5000 |
Max ((\<lambda>th. preced th s) ` ({th2} \<union> dependants (wq s) th2))" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5001 |
(is "Max (?f ` ?A) = Max (?f ` ?B)") |
44
f676a68935a0
updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
5002 |
unfolding cp_eq_cpreced |
f676a68935a0
updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
5003 |
unfolding cpreced_def . |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5004 |
obtain th1' where th1_in: "th1' \<in> ?A" and eq_f_th1: "?f th1' = Max (?f ` ?A)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5005 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5006 |
have h1: "finite (?f ` ?A)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5007 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5008 |
have "finite ?A" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5009 |
proof - |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5010 |
have "finite (dependants (wq s) th1)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5011 |
proof- |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5012 |
have "finite {th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+}" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5013 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5014 |
let ?F = "\<lambda> (x, y). the_th x" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5015 |
have "{th'. (Th th', Th th1) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5016 |
apply (auto simp:image_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5017 |
by (rule_tac x = "(Th x, Th th1)" in bexI, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5018 |
moreover have "finite \<dots>" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5019 |
proof - |
63 | 5020 |
from finite_RAG have "finite (RAG s)" . |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5021 |
hence "finite ((RAG (wq s))\<^sup>+)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5022 |
apply (unfold finite_trancl) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5023 |
by (auto simp: s_RAG_def cs_RAG_def wq_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5024 |
thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5025 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5026 |
ultimately show ?thesis by (auto intro:finite_subset) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5027 |
qed |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5028 |
thus ?thesis by (simp add:cs_dependants_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5029 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5030 |
thus ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5031 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5032 |
thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5033 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5034 |
moreover have h2: "(?f ` ?A) \<noteq> {}" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5035 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5036 |
have "?A \<noteq> {}" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5037 |
thus ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5038 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5039 |
from Max_in [OF h1 h2] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5040 |
have "Max (?f ` ?A) \<in> (?f ` ?A)" . |
44
f676a68935a0
updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
5041 |
thus ?thesis |
f676a68935a0
updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
5042 |
thm cpreced_def |
f676a68935a0
updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
5043 |
unfolding cpreced_def[symmetric] |
f676a68935a0
updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
5044 |
unfolding cp_eq_cpreced[symmetric] |
f676a68935a0
updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
5045 |
unfolding cpreced_def |
f676a68935a0
updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
5046 |
using that[intro] by (auto) |
104 | 5047 |
>>>>>>> other |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5048 |
qed |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5049 |
obtain th2' where th2_in: "th2' \<in> ?B" and eq_f_th2: "?f th2' = Max (?f ` ?B)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5050 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5051 |
have h1: "finite (?f ` ?B)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5052 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5053 |
have "finite ?B" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5054 |
proof - |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5055 |
have "finite (dependants (wq s) th2)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5056 |
proof- |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5057 |
have "finite {th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+}" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5058 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5059 |
let ?F = "\<lambda> (x, y). the_th x" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5060 |
have "{th'. (Th th', Th th2) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5061 |
apply (auto simp:image_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5062 |
by (rule_tac x = "(Th x, Th th2)" in bexI, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5063 |
moreover have "finite \<dots>" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5064 |
proof - |
63 | 5065 |
from finite_RAG have "finite (RAG s)" . |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5066 |
hence "finite ((RAG (wq s))\<^sup>+)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5067 |
apply (unfold finite_trancl) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5068 |
by (auto simp: s_RAG_def cs_RAG_def wq_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5069 |
thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5070 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5071 |
ultimately show ?thesis by (auto intro:finite_subset) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5072 |
qed |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5073 |
thus ?thesis by (simp add:cs_dependants_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5074 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5075 |
thus ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5076 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5077 |
thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5078 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5079 |
moreover have h2: "(?f ` ?B) \<noteq> {}" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5080 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5081 |
have "?B \<noteq> {}" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5082 |
thus ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5083 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5084 |
from Max_in [OF h1 h2] |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5085 |
have "Max (?f ` ?B) \<in> (?f ` ?B)" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5086 |
thus ?thesis by (auto intro:that) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5087 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5088 |
from eq_f_th1 eq_f_th2 eq_max |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5089 |
have eq_preced: "preced th1' s = preced th2' s" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5090 |
hence eq_th12: "th1' = th2'" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5091 |
proof (rule preced_unique) |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5092 |
from th1_in have "th1' = th1 \<or> (th1' \<in> dependants (wq s) th1)" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5093 |
thus "th1' \<in> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5094 |
proof |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5095 |
assume "th1' \<in> dependants (wq s) th1" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5096 |
hence "(Th th1') \<in> Domain ((RAG s)^+)" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5097 |
apply (unfold cs_dependants_def cs_RAG_def s_RAG_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5098 |
by (auto simp:Domain_def) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5099 |
hence "(Th th1') \<in> Domain (RAG s)" by (simp add:trancl_domain) |
63 | 5100 |
from dm_RAG_threads[OF this] show ?thesis . |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5101 |
next |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5102 |
assume "th1' = th1" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5103 |
with runing_1 show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5104 |
by (unfold runing_def readys_def, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5105 |
qed |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5106 |
next |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5107 |
from th2_in have "th2' = th2 \<or> (th2' \<in> dependants (wq s) th2)" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5108 |
thus "th2' \<in> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5109 |
proof |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5110 |
assume "th2' \<in> dependants (wq s) th2" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5111 |
hence "(Th th2') \<in> Domain ((RAG s)^+)" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5112 |
apply (unfold cs_dependants_def cs_RAG_def s_RAG_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5113 |
by (auto simp:Domain_def) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5114 |
hence "(Th th2') \<in> Domain (RAG s)" by (simp add:trancl_domain) |
63 | 5115 |
from dm_RAG_threads[OF this] show ?thesis . |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5116 |
next |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5117 |
assume "th2' = th2" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5118 |
with runing_2 show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5119 |
by (unfold runing_def readys_def, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5120 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5121 |
qed |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5122 |
from th1_in have "th1' = th1 \<or> th1' \<in> dependants (wq s) th1" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5123 |
thus ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5124 |
proof |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5125 |
assume eq_th': "th1' = th1" |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5126 |
from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5127 |
thus ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5128 |
proof |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5129 |
assume "th2' = th2" thus ?thesis using eq_th' eq_th12 by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5130 |
next |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5131 |
assume "th2' \<in> dependants (wq s) th2" |
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5132 |
with eq_th12 eq_th' have "th1 \<in> dependants (wq s) th2" by simp |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5133 |
hence "(Th th1, Th th2) \<in> (RAG s)^+" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5134 |
by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp) |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5135 |
hence "Th th1 \<in> Domain ((RAG s)^+)" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5136 |
apply (unfold cs_dependants_def cs_RAG_def s_RAG_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5137 |
by (auto simp:Domain_def) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5138 |
hence "Th th1 \<in> Domain (RAG s)" by (simp add:trancl_domain) |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5139 |
then obtain n where d: "(Th th1, n) \<in> RAG s" by (auto simp:Domain_def) |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5140 |
from RAG_target_th [OF this] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5141 |
obtain cs' where "n = Cs cs'" by auto |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5142 |
with d have "(Th th1, Cs cs') \<in> RAG s" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5143 |
with runing_1 have "False" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5144 |
apply (unfold runing_def readys_def s_RAG_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5145 |
by (auto simp:eq_waiting) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5146 |
thus ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5147 |
qed |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5148 |
next |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5149 |
assume th1'_in: "th1' \<in> dependants (wq s) th1" |
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5150 |
from th2_in have "th2' = th2 \<or> th2' \<in> dependants (wq s) th2" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5151 |
thus ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5152 |
proof |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5153 |
assume "th2' = th2" |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5154 |
with th1'_in eq_th12 have "th2 \<in> dependants (wq s) th1" by simp |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5155 |
hence "(Th th2, Th th1) \<in> (RAG s)^+" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5156 |
by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp) |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5157 |
hence "Th th2 \<in> Domain ((RAG s)^+)" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5158 |
apply (unfold cs_dependants_def cs_RAG_def s_RAG_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5159 |
by (auto simp:Domain_def) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5160 |
hence "Th th2 \<in> Domain (RAG s)" by (simp add:trancl_domain) |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5161 |
then obtain n where d: "(Th th2, n) \<in> RAG s" by (auto simp:Domain_def) |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5162 |
from RAG_target_th [OF this] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5163 |
obtain cs' where "n = Cs cs'" by auto |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5164 |
with d have "(Th th2, Cs cs') \<in> RAG s" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5165 |
with runing_2 have "False" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5166 |
apply (unfold runing_def readys_def s_RAG_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5167 |
by (auto simp:eq_waiting) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5168 |
thus ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5169 |
next |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5170 |
assume "th2' \<in> dependants (wq s) th2" |
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
5171 |
with eq_th12 have "th1' \<in> dependants (wq s) th2" by simp |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5172 |
hence h1: "(Th th1', Th th2) \<in> (RAG s)^+" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5173 |
by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp) |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5174 |
from th1'_in have h2: "(Th th1', Th th1) \<in> (RAG s)^+" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
5175 |
by (unfold cs_dependants_def s_RAG_def cs_RAG_def, simp) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5176 |
show ?thesis |
63 | 5177 |
proof(rule dchain_unique[OF h1 _ h2, symmetric]) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5178 |
from runing_1 show "th1 \<in> readys s" by (simp add:runing_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5179 |
from runing_2 show "th2 \<in> readys s" by (simp add:runing_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5180 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5181 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5182 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5183 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5184 |
|
39
7ea6b019ce24
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
38
diff
changeset
|
5185 |
|
63 | 5186 |
lemma "card (runing s) \<le> 1" |
41
66ed924aaa5c
added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
39
diff
changeset
|
5187 |
apply(subgoal_tac "finite (runing s)") |
66ed924aaa5c
added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
39
diff
changeset
|
5188 |
prefer 2 |
66ed924aaa5c
added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
39
diff
changeset
|
5189 |
apply (metis finite_nat_set_iff_bounded lessI runing_unique) |
44
f676a68935a0
updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
5190 |
apply(rule ccontr) |
f676a68935a0
updated teh theories to newer Isabelle version
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
5191 |
apply(simp) |
41
66ed924aaa5c
added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
39
diff
changeset
|
5192 |
apply(case_tac "Suc (Suc 0) \<le> card (runing s)") |
66ed924aaa5c
added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
39
diff
changeset
|
5193 |
apply(subst (asm) card_le_Suc_iff) |
66ed924aaa5c
added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
39
diff
changeset
|
5194 |
apply(simp) |
66ed924aaa5c
added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
39
diff
changeset
|
5195 |
apply(auto)[1] |
66ed924aaa5c
added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
39
diff
changeset
|
5196 |
apply (metis insertCI runing_unique) |
66ed924aaa5c
added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
39
diff
changeset
|
5197 |
apply(auto) |
66ed924aaa5c
added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
39
diff
changeset
|
5198 |
done |
66ed924aaa5c
added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
39
diff
changeset
|
5199 |
|
63 | 5200 |
end |
5201 |
||
5202 |
||
5203 |
end |
|
5204 |
||
101 | 5205 |
|
5206 |
section {* Relating @{term cp} and @{term the_preced} and @{term preced} *} |
|
5207 |
||
5208 |
context valid_trace |
|
5209 |
begin |
|
5210 |
||
5211 |
lemma le_cp: |
|
5212 |
shows "preced th s \<le> cp s th" |
|
5213 |
proof(unfold cp_alt_def, rule Max_ge) |
|
5214 |
show "finite (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})" |
|
5215 |
by (simp add: finite_subtree_threads) |
|
5216 |
next |
|
5217 |
show "preced th s \<in> the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)}" |
|
5218 |
by (simp add: subtree_def the_preced_def) |
|
5219 |
qed |
|
5220 |
||
5221 |
||
5222 |
lemma cp_le: |
|
5223 |
assumes th_in: "th \<in> threads s" |
|
5224 |
shows "cp s th \<le> Max (the_preced s ` threads s)" |
|
5225 |
proof(unfold cp_alt_def, rule Max_f_mono) |
|
5226 |
show "finite (threads s)" by (simp add: finite_threads) |
|
5227 |
next |
|
5228 |
show " {th'. Th th' \<in> subtree (RAG s) (Th th)} \<noteq> {}" |
|
5229 |
using subtree_def by fastforce |
|
5230 |
next |
|
5231 |
show "{th'. Th th' \<in> subtree (RAG s) (Th th)} \<subseteq> threads s" |
|
5232 |
using assms |
|
5233 |
by (smt Domain.DomainI dm_RAG_threads mem_Collect_eq |
|
5234 |
node.inject(1) rtranclD subsetI subtree_def trancl_domain) |
|
5235 |
qed |
|
5236 |
||
5237 |
lemma max_cp_eq: |
|
5238 |
shows "Max ((cp s) ` threads s) = Max (the_preced s ` threads s)" |
|
5239 |
(is "?L = ?R") |
|
5240 |
proof - |
|
5241 |
have "?L \<le> ?R" |
|
5242 |
proof(cases "threads s = {}") |
|
5243 |
case False |
|
5244 |
show ?thesis |
|
5245 |
by (rule Max.boundedI, |
|
5246 |
insert cp_le, |
|
5247 |
auto simp:finite_threads False) |
|
5248 |
qed auto |
|
5249 |
moreover have "?R \<le> ?L" |
|
5250 |
by (rule Max_fg_mono, |
|
5251 |
simp add: finite_threads, |
|
5252 |
simp add: le_cp the_preced_def) |
|
5253 |
ultimately show ?thesis by auto |
|
5254 |
qed |
|
5255 |
||
103
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5256 |
lemma threads_alt_def: |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5257 |
"(threads s) = (\<Union> th \<in> readys s. {th'. Th th' \<in> subtree (RAG s) (Th th)})" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5258 |
(is "?L = ?R") |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5259 |
proof - |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5260 |
{ fix th1 |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5261 |
assume "th1 \<in> ?L" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5262 |
from th_chain_to_ready[OF this] |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5263 |
have "th1 \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th th1, Th th') \<in> (RAG s)\<^sup>+)" . |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5264 |
hence "th1 \<in> ?R" by (auto simp:subtree_def) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5265 |
} moreover |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5266 |
{ fix th' |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5267 |
assume "th' \<in> ?R" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5268 |
then obtain th where h: "th \<in> readys s" " Th th' \<in> subtree (RAG s) (Th th)" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5269 |
by auto |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5270 |
from this(2) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5271 |
have "th' \<in> ?L" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5272 |
proof(cases rule:subtreeE) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5273 |
case 1 |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5274 |
with h(1) show ?thesis by (auto simp:readys_def) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5275 |
next |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5276 |
case 2 |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5277 |
from tranclD[OF this(2)[unfolded ancestors_def, simplified]] |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5278 |
have "Th th' \<in> Domain (RAG s)" by auto |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5279 |
from dm_RAG_threads[OF this] |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5280 |
show ?thesis . |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5281 |
qed |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5282 |
} ultimately show ?thesis by auto |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5283 |
qed |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5284 |
|
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5285 |
|
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5286 |
text {* (* ccc *) \noindent |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5287 |
Since the current precedence of the threads in ready queue will always be boosted, |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5288 |
there must be one inside it has the maximum precedence of the whole system. |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5289 |
*} |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5290 |
lemma max_cp_readys_threads: |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5291 |
shows "Max (cp s ` readys s) = Max (cp s ` threads s)" (is "?L = ?R") |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5292 |
proof(cases "readys s = {}") |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5293 |
case False |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5294 |
have "?R = Max (the_preced s ` threads s)" by (unfold max_cp_eq, simp) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5295 |
also have "... = |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5296 |
Max (the_preced s ` (\<Union>th\<in>readys s. {th'. Th th' \<in> subtree (RAG s) (Th th)}))" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5297 |
by (unfold threads_alt_def, simp) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5298 |
also have "... = |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5299 |
Max ((\<Union>th\<in>readys s. the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)}))" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5300 |
by (unfold image_UN, simp) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5301 |
also have "... = |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5302 |
Max (Max ` (\<lambda>th. the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)}) ` readys s)" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5303 |
proof(rule Max_UNION) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5304 |
show "\<forall>M\<in>(\<lambda>x. the_preced s ` |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5305 |
{th'. Th th' \<in> subtree (RAG s) (Th x)}) ` readys s. finite M" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5306 |
using finite_subtree_threads by auto |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5307 |
qed (auto simp:False subtree_def) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5308 |
also have "... = |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5309 |
Max ((Max \<circ> (\<lambda>th. the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})) ` readys s)" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5310 |
by (unfold image_comp, simp) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5311 |
also have "... = ?L" (is "Max (?f ` ?A) = Max (?g ` ?A)") |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5312 |
proof - |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5313 |
have "(?f ` ?A) = (?g ` ?A)" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5314 |
proof(rule f_image_eq) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5315 |
fix th1 |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5316 |
assume "th1 \<in> ?A" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5317 |
thus "?f th1 = ?g th1" |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5318 |
by (unfold cp_alt_def, simp) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5319 |
qed |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5320 |
thus ?thesis by simp |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5321 |
qed |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5322 |
finally show ?thesis by simp |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5323 |
qed (auto simp:threads_alt_def) |
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
5324 |
|
101 | 5325 |
end |
5326 |
||
5327 |
section {* Relating @{term cntP}, @{term cntV}, @{term cntCS} and @{term pvD} *} |
|
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5328 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5329 |
context valid_trace_p_w |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5330 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5331 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5332 |
lemma holding_s_holder: "holding s holder cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5333 |
by (unfold s_holding_def, fold wq_def, unfold wq_s_cs, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5334 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5335 |
lemma holding_es_holder: "holding (e#s) holder cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5336 |
by (unfold s_holding_def, fold wq_def, unfold wq_es_cs wq_s_cs, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5337 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5338 |
lemma holdents_es: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5339 |
shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R") |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5340 |
proof - |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5341 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5342 |
assume "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5343 |
hence h: "holding (e#s) th' cs'" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5344 |
have "holding s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5345 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5346 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5347 |
from held_unique[OF h[unfolded True] holding_es_holder] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5348 |
have "th' = holder" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5349 |
thus ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5350 |
by (unfold True holdents_def, insert holding_s_holder, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5351 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5352 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5353 |
hence "wq (e#s) cs' = wq s cs'" by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5354 |
from h[unfolded s_holding_def, folded wq_def, unfolded this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5355 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5356 |
by (unfold s_holding_def, fold wq_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5357 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5358 |
hence "cs' \<in> ?R" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5359 |
} moreover { |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5360 |
fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5361 |
assume "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5362 |
hence h: "holding s th' cs'" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5363 |
have "holding (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5364 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5365 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5366 |
from held_unique[OF h[unfolded True] holding_s_holder] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5367 |
have "th' = holder" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5368 |
thus ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5369 |
by (unfold True holdents_def, insert holding_es_holder, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5370 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5371 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5372 |
hence "wq s cs' = wq (e#s) cs'" by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5373 |
from h[unfolded s_holding_def, folded wq_def, unfolded this] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5374 |
show ?thesis |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5375 |
by (unfold s_holding_def, fold wq_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5376 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5377 |
hence "cs' \<in> ?L" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5378 |
} ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5379 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5380 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5381 |
lemma cntCS_es_th[simp]: "cntCS (e#s) th' = cntCS s th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5382 |
by (unfold cntCS_def holdents_es, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5383 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5384 |
lemma th_not_ready_es: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5385 |
shows "th \<notin> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5386 |
using waiting_es_th_cs |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5387 |
by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5388 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5389 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5390 |
|
100 | 5391 |
lemma (in valid_trace) finite_holdents: "finite (holdents s th)" |
5392 |
by (unfold holdents_alt_def, insert fsbtRAGs.finite_children, auto) |
|
5393 |
||
101 | 5394 |
context valid_trace_p |
5395 |
begin |
|
5396 |
||
5397 |
lemma ready_th_s: "th \<in> readys s" |
|
5398 |
using runing_th_s |
|
5399 |
by (unfold runing_def, auto) |
|
5400 |
||
5401 |
lemma live_th_s: "th \<in> threads s" |
|
5402 |
using readys_threads ready_th_s by auto |
|
5403 |
||
5404 |
lemma live_th_es: "th \<in> threads (e#s)" |
|
5405 |
using live_th_s |
|
5406 |
by (unfold is_p, simp) |
|
5407 |
||
5408 |
lemma waiting_neq_th: |
|
5409 |
assumes "waiting s t c" |
|
5410 |
shows "t \<noteq> th" |
|
5411 |
using assms using th_not_waiting by blast |
|
5412 |
||
5413 |
end |
|
5414 |
||
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5415 |
context valid_trace_p_h |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5416 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5417 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5418 |
lemma th_not_waiting': |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5419 |
"\<not> waiting (e#s) th cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5420 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5421 |
case True |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5422 |
show ?thesis |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5423 |
by (unfold True s_waiting_def, fold wq_def, unfold wq_es_cs', auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5424 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5425 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5426 |
from th_not_waiting[of cs', unfolded s_waiting_def, folded wq_def] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5427 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5428 |
by (unfold s_waiting_def, fold wq_def, insert False, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5429 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5430 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5431 |
lemma ready_th_es: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5432 |
shows "th \<in> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5433 |
using th_not_waiting' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5434 |
by (unfold readys_def, insert live_th_es, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5435 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5436 |
lemma holdents_es_th: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5437 |
"holdents (e#s) th = (holdents s th) \<union> {cs}" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5438 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5439 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5440 |
assume "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5441 |
hence "holding (e#s) th cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5442 |
by (unfold holdents_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5443 |
hence "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5444 |
by (cases rule:holding_esE, auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5445 |
} moreover { |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5446 |
fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5447 |
assume "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5448 |
hence "holding s th cs' \<or> cs' = cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5449 |
by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5450 |
hence "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5451 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5452 |
assume "holding s th cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5453 |
from holding_kept[OF this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5454 |
show ?thesis by (auto simp:holdents_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5455 |
next |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5456 |
assume "cs' = cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5457 |
thus ?thesis using holding_es_th_cs |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5458 |
by (unfold holdents_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5459 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5460 |
} ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5461 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5462 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5463 |
lemma cntCS_es_th: "cntCS (e#s) th = cntCS s th + 1" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5464 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5465 |
have "card (holdents s th \<union> {cs}) = card (holdents s th) + 1" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5466 |
proof(subst card_Un_disjoint) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5467 |
show "holdents s th \<inter> {cs} = {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5468 |
using not_holding_s_th_cs by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5469 |
qed (auto simp:finite_holdents) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5470 |
thus ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5471 |
by (unfold cntCS_def holdents_es_th, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5472 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5473 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5474 |
lemma no_holder: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5475 |
"\<not> holding s th' cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5476 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5477 |
assume otherwise: "holding s th' cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5478 |
from this[unfolded s_holding_def, folded wq_def, unfolded we] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5479 |
show False by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5480 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5481 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5482 |
lemma holdents_es_th': |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5483 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5484 |
shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5485 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5486 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5487 |
assume "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5488 |
hence h_e: "holding (e#s) th' cs'" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5489 |
have "cs' \<noteq> cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5490 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5491 |
assume "cs' = cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5492 |
from held_unique[OF h_e[unfolded this] holding_es_th_cs] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5493 |
have "th' = th" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5494 |
with assms show False by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5495 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5496 |
from h_e[unfolded s_holding_def, folded wq_def, unfolded wq_neq_simp[OF this]] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5497 |
have "th' \<in> set (wq s cs') \<and> th' = hd (wq s cs')" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5498 |
hence "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5499 |
by (unfold holdents_def s_holding_def, fold wq_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5500 |
} moreover { |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5501 |
fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5502 |
assume "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5503 |
hence "holding s th' cs'" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5504 |
from holding_kept[OF this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5505 |
have "holding (e # s) th' cs'" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5506 |
hence "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5507 |
by (unfold holdents_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5508 |
} ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5509 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5510 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5511 |
lemma cntCS_es_th'[simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5512 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5513 |
shows "cntCS (e#s) th' = cntCS s th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5514 |
by (unfold cntCS_def holdents_es_th'[OF assms], simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5515 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5516 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5517 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5518 |
context valid_trace_p |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5519 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5520 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5521 |
lemma readys_kept1: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5522 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5523 |
and "th' \<in> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5524 |
shows "th' \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5525 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5526 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5527 |
assume wait: "waiting s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5528 |
have n_wait: "\<not> waiting (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5529 |
using assms(2)[unfolded readys_def] by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5530 |
have False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5531 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5532 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5533 |
with n_wait wait |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5534 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5535 |
by (unfold s_waiting_def, fold wq_def, auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5536 |
next |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5537 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5538 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5539 |
proof(cases "wq s cs = []") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5540 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5541 |
then interpret vt: valid_trace_p_h |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5542 |
by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5543 |
show ?thesis using n_wait wait waiting_kept by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5544 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5545 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5546 |
then interpret vt: valid_trace_p_w by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5547 |
show ?thesis using n_wait wait waiting_kept by blast |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5548 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5549 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5550 |
} with assms(2) show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5551 |
by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5552 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5553 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5554 |
lemma readys_kept2: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5555 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5556 |
and "th' \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5557 |
shows "th' \<in> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5558 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5559 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5560 |
assume wait: "waiting (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5561 |
have n_wait: "\<not> waiting s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5562 |
using assms(2)[unfolded readys_def] by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5563 |
have False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5564 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5565 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5566 |
with n_wait wait |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5567 |
show ?thesis |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5568 |
by (unfold s_waiting_def, fold wq_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5569 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5570 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5571 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5572 |
proof(cases "wq s cs = []") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5573 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5574 |
then interpret vt: valid_trace_p_h |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5575 |
by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5576 |
show ?thesis using n_wait vt.waiting_esE wait by blast |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5577 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5578 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5579 |
then interpret vt: valid_trace_p_w by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5580 |
show ?thesis using assms(1) n_wait vt.waiting_esE wait by auto |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5581 |
qed |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5582 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5583 |
} with assms(2) show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5584 |
by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5585 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5586 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5587 |
lemma readys_simp [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5588 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5589 |
shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5590 |
using readys_kept1[OF assms] readys_kept2[OF assms] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5591 |
by metis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5592 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5593 |
lemma cnp_cnv_cncs_kept: (* ddd *) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5594 |
assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5595 |
shows "cntP (e#s) th' = cntV (e#s) th' + cntCS (e#s) th' + pvD (e#s) th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5596 |
proof(cases "th' = th") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5597 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5598 |
note eq_th' = this |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5599 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5600 |
proof(cases "wq s cs = []") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5601 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5602 |
then interpret vt: valid_trace_p_h by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5603 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5604 |
using assms eq_th' is_p ready_th_s vt.cntCS_es_th vt.ready_th_es pvD_def by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5605 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5606 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5607 |
then interpret vt: valid_trace_p_w by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5608 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5609 |
using add.commute add.left_commute assms eq_th' is_p live_th_s |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5610 |
ready_th_s vt.th_not_ready_es pvD_def |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5611 |
apply (auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5612 |
by (fold is_p, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5613 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5614 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5615 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5616 |
note h_False = False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5617 |
thus ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5618 |
proof(cases "wq s cs = []") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5619 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5620 |
then interpret vt: valid_trace_p_h by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5621 |
show ?thesis using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5622 |
by (insert True h_False pvD_def, auto split:if_splits,unfold is_p, auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5623 |
next |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5624 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5625 |
then interpret vt: valid_trace_p_w by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5626 |
show ?thesis using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5627 |
by (insert False h_False pvD_def, auto split:if_splits,unfold is_p, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5628 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5629 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5630 |
|
104 | 5631 |
<<<<<<< local |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5632 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5633 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5634 |
|
100 | 5635 |
context valid_trace_v |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5636 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5637 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5638 |
lemma holding_th_cs_s: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5639 |
"holding s th cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5640 |
by (unfold s_holding_def, fold wq_def, unfold wq_s_cs, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5641 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5642 |
lemma th_ready_s [simp]: "th \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5643 |
using runing_th_s |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5644 |
by (unfold runing_def readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5645 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5646 |
lemma th_live_s [simp]: "th \<in> threads s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5647 |
using th_ready_s by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5648 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5649 |
lemma th_ready_es [simp]: "th \<in> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5650 |
using runing_th_s neq_t_th |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5651 |
by (unfold is_v runing_def readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5652 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5653 |
lemma th_live_es [simp]: "th \<in> threads (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5654 |
using th_ready_es by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5655 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5656 |
lemma pvD_th_s[simp]: "pvD s th = 0" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5657 |
by (unfold pvD_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5658 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5659 |
lemma pvD_th_es[simp]: "pvD (e#s) th = 0" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5660 |
by (unfold pvD_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5661 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5662 |
lemma cntCS_s_th [simp]: "cntCS s th > 0" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5663 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5664 |
have "cs \<in> holdents s th" using holding_th_cs_s |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5665 |
by (unfold holdents_def, simp) |
101 | 5666 |
moreover have "finite (holdents s th)" using finite_holdents |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5667 |
by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5668 |
ultimately show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5669 |
by (unfold cntCS_def, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5670 |
auto intro!:card_gt_0_iff[symmetric, THEN iffD1]) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5671 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5672 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5673 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5674 |
|
101 | 5675 |
context valid_trace_v |
5676 |
begin |
|
5677 |
||
5678 |
lemma th_not_waiting: |
|
5679 |
"\<not> waiting s th c" |
|
5680 |
proof - |
|
5681 |
have "th \<in> readys s" |
|
5682 |
using runing_ready runing_th_s by blast |
|
5683 |
thus ?thesis |
|
5684 |
by (unfold readys_def, auto) |
|
5685 |
qed |
|
5686 |
||
5687 |
lemma waiting_neq_th: |
|
5688 |
assumes "waiting s t c" |
|
5689 |
shows "t \<noteq> th" |
|
5690 |
using assms using th_not_waiting by blast |
|
5691 |
||
5692 |
end |
|
5693 |
||
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5694 |
context valid_trace_v_n |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5695 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5696 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5697 |
lemma not_ready_taker_s[simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5698 |
"taker \<notin> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5699 |
using waiting_taker |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5700 |
by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5701 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5702 |
lemma taker_live_s [simp]: "taker \<in> threads s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5703 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5704 |
have "taker \<in> set wq'" by (simp add: eq_wq') |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5705 |
from th'_in_inv[OF this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5706 |
have "taker \<in> set rest" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5707 |
hence "taker \<in> set (wq s cs)" by (simp add: wq_s_cs) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5708 |
thus ?thesis using wq_threads by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5709 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5710 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5711 |
lemma taker_live_es [simp]: "taker \<in> threads (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5712 |
using taker_live_s threads_es by blast |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5713 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5714 |
lemma taker_ready_es [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5715 |
shows "taker \<in> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5716 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5717 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5718 |
assume "waiting (e#s) taker cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5719 |
hence False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5720 |
proof(cases rule:waiting_esE) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5721 |
case 1 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5722 |
thus ?thesis using waiting_taker waiting_unique by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5723 |
qed simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5724 |
} thus ?thesis by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5725 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5726 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5727 |
lemma neq_taker_th: "taker \<noteq> th" |
101 | 5728 |
using th_not_waiting waiting_taker by blast |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5729 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5730 |
lemma not_holding_taker_s_cs: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5731 |
shows "\<not> holding s taker cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5732 |
using holding_cs_eq_th neq_taker_th by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5733 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5734 |
lemma holdents_es_taker: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5735 |
"holdents (e#s) taker = holdents s taker \<union> {cs}" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5736 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5737 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5738 |
assume "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5739 |
hence "holding (e#s) taker cs'" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5740 |
hence "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5741 |
proof(cases rule:holding_esE) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5742 |
case 2 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5743 |
thus ?thesis by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5744 |
qed auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5745 |
} moreover { |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5746 |
fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5747 |
assume "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5748 |
hence "holding s taker cs' \<or> cs' = cs" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5749 |
hence "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5750 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5751 |
assume "holding s taker cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5752 |
hence "holding (e#s) taker cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5753 |
using holding_esI2 holding_taker by fastforce |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5754 |
thus ?thesis by (auto simp:holdents_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5755 |
next |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5756 |
assume "cs' = cs" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5757 |
with holding_taker |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5758 |
show ?thesis by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5759 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5760 |
} ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5761 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5762 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5763 |
lemma cntCS_es_taker [simp]: "cntCS (e#s) taker = cntCS s taker + 1" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5764 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5765 |
have "card (holdents s taker \<union> {cs}) = card (holdents s taker) + 1" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5766 |
proof(subst card_Un_disjoint) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5767 |
show "holdents s taker \<inter> {cs} = {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5768 |
using not_holding_taker_s_cs by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5769 |
qed (auto simp:finite_holdents) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5770 |
thus ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5771 |
by (unfold cntCS_def, insert holdents_es_taker, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5772 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5773 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5774 |
lemma pvD_taker_s[simp]: "pvD s taker = 1" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5775 |
by (unfold pvD_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5776 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5777 |
lemma pvD_taker_es[simp]: "pvD (e#s) taker = 0" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5778 |
by (unfold pvD_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5779 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5780 |
lemma pvD_th_s[simp]: "pvD s th = 0" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5781 |
by (unfold pvD_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5782 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5783 |
lemma pvD_th_es[simp]: "pvD (e#s) th = 0" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5784 |
by (unfold pvD_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5785 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5786 |
lemma holdents_es_th: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5787 |
"holdents (e#s) th = holdents s th - {cs}" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5788 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5789 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5790 |
assume "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5791 |
hence "holding (e#s) th cs'" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5792 |
hence "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5793 |
proof(cases rule:holding_esE) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5794 |
case 2 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5795 |
thus ?thesis by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5796 |
qed (insert neq_taker_th, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5797 |
} moreover { |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5798 |
fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5799 |
assume "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5800 |
hence "cs' \<noteq> cs" "holding s th cs'" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5801 |
from holding_esI2[OF this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5802 |
have "cs' \<in> ?L" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5803 |
} ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5804 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5805 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5806 |
lemma cntCS_es_th [simp]: "cntCS (e#s) th = cntCS s th - 1" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5807 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5808 |
have "card (holdents s th - {cs}) = card (holdents s th) - 1" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5809 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5810 |
have "cs \<in> holdents s th" using holding_th_cs_s |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5811 |
by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5812 |
moreover have "finite (holdents s th)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5813 |
by (simp add: finite_holdents) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5814 |
ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5815 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5816 |
thus ?thesis by (unfold cntCS_def holdents_es_th) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5817 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5818 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5819 |
lemma holdents_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5820 |
assumes "th' \<noteq> taker" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5821 |
and "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5822 |
shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5823 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5824 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5825 |
assume h: "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5826 |
have "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5827 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5828 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5829 |
hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5830 |
from h have "holding (e#s) th' cs'" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5831 |
from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5832 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5833 |
by (unfold holdents_def s_holding_def, fold wq_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5834 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5835 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5836 |
from h[unfolded this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5837 |
have "holding (e#s) th' cs" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5838 |
from held_unique[OF this holding_taker] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5839 |
have "th' = taker" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5840 |
with assms show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5841 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5842 |
} moreover { |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5843 |
fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5844 |
assume h: "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5845 |
have "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5846 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5847 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5848 |
hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5849 |
from h have "holding s th' cs'" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5850 |
from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5851 |
show ?thesis |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5852 |
by (unfold holdents_def s_holding_def, fold wq_def, insert eq_wq, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5853 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5854 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5855 |
from h[unfolded this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5856 |
have "holding s th' cs" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5857 |
from held_unique[OF this holding_th_cs_s] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5858 |
have "th' = th" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5859 |
with assms show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5860 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5861 |
} ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5862 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5863 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5864 |
lemma cntCS_kept [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5865 |
assumes "th' \<noteq> taker" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5866 |
and "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5867 |
shows "cntCS (e#s) th' = cntCS s th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5868 |
by (unfold cntCS_def holdents_kept[OF assms], simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5869 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5870 |
lemma readys_kept1: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5871 |
assumes "th' \<noteq> taker" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5872 |
and "th' \<in> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5873 |
shows "th' \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5874 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5875 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5876 |
assume wait: "waiting s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5877 |
have n_wait: "\<not> waiting (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5878 |
using assms(2)[unfolded readys_def] by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5879 |
have False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5880 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5881 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5882 |
with n_wait wait |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5883 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5884 |
by (unfold s_waiting_def, fold wq_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5885 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5886 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5887 |
have "th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5888 |
using wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5889 |
moreover have "\<not> (th' \<in> set rest \<and> th' \<noteq> hd (taker # rest'))" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5890 |
using n_wait[unfolded True s_waiting_def, folded wq_def, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5891 |
unfolded wq_es_cs set_wq', unfolded eq_wq'] . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5892 |
ultimately have "th' = taker" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5893 |
with assms(1) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5894 |
show ?thesis by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5895 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5896 |
} with assms(2) show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5897 |
by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5898 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5899 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5900 |
lemma readys_kept2: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5901 |
assumes "th' \<noteq> taker" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5902 |
and "th' \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5903 |
shows "th' \<in> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5904 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5905 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5906 |
assume wait: "waiting (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5907 |
have n_wait: "\<not> waiting s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5908 |
using assms(2)[unfolded readys_def] by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5909 |
have False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5910 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5911 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5912 |
with n_wait wait |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5913 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5914 |
by (unfold s_waiting_def, fold wq_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5915 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5916 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5917 |
have "th' \<in> set rest \<and> th' \<noteq> hd (taker # rest')" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5918 |
using wait [unfolded True s_waiting_def, folded wq_def, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5919 |
unfolded wq_es_cs set_wq', unfolded eq_wq'] . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5920 |
moreover have "\<not> (th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest))" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5921 |
using n_wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5922 |
ultimately have "th' = taker" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5923 |
with assms(1) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5924 |
show ?thesis by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5925 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5926 |
} with assms(2) show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5927 |
by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5928 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5929 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5930 |
lemma readys_simp [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5931 |
assumes "th' \<noteq> taker" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5932 |
shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5933 |
using readys_kept1[OF assms] readys_kept2[OF assms] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5934 |
by metis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5935 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5936 |
lemma cnp_cnv_cncs_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5937 |
assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5938 |
shows "cntP (e#s) th' = cntV (e#s) th' + cntCS (e#s) th' + pvD (e#s) th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5939 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5940 |
{ assume eq_th': "th' = taker" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5941 |
have ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5942 |
apply (unfold eq_th' pvD_taker_es cntCS_es_taker) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5943 |
by (insert neq_taker_th assms[unfolded eq_th'], unfold is_v, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5944 |
} moreover { |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5945 |
assume eq_th': "th' = th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5946 |
have ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5947 |
apply (unfold eq_th' pvD_th_es cntCS_es_th) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5948 |
by (insert assms[unfolded eq_th'], unfold is_v, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5949 |
} moreover { |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5950 |
assume h: "th' \<noteq> taker" "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5951 |
have ?thesis using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5952 |
apply (unfold cntCS_kept[OF h], insert h, unfold is_v, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5953 |
by (fold is_v, unfold pvD_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5954 |
} ultimately show ?thesis by metis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5955 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5956 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5957 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5958 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5959 |
context valid_trace_v_e |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5960 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5961 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5962 |
lemma holdents_es_th: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5963 |
"holdents (e#s) th = holdents s th - {cs}" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5964 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5965 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5966 |
assume "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5967 |
hence "holding (e#s) th cs'" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5968 |
hence "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5969 |
proof(cases rule:holding_esE) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5970 |
case 1 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5971 |
thus ?thesis by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5972 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5973 |
} moreover { |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5974 |
fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5975 |
assume "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5976 |
hence "cs' \<noteq> cs" "holding s th cs'" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5977 |
from holding_esI2[OF this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5978 |
have "cs' \<in> ?L" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5979 |
} ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5980 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5981 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5982 |
lemma cntCS_es_th [simp]: "cntCS (e#s) th = cntCS s th - 1" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5983 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5984 |
have "card (holdents s th - {cs}) = card (holdents s th) - 1" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5985 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5986 |
have "cs \<in> holdents s th" using holding_th_cs_s |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5987 |
by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5988 |
moreover have "finite (holdents s th)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5989 |
by (simp add: finite_holdents) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5990 |
ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5991 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5992 |
thus ?thesis by (unfold cntCS_def holdents_es_th) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5993 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5994 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5995 |
lemma holdents_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5996 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5997 |
shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5998 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
5999 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6000 |
assume h: "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6001 |
have "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6002 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6003 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6004 |
hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6005 |
from h have "holding (e#s) th' cs'" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6006 |
from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6007 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6008 |
by (unfold holdents_def s_holding_def, fold wq_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6009 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6010 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6011 |
from h[unfolded this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6012 |
have "holding (e#s) th' cs" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6013 |
from this[unfolded s_holding_def, folded wq_def, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6014 |
unfolded wq_es_cs nil_wq'] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6015 |
show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6016 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6017 |
} moreover { |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6018 |
fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6019 |
assume h: "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6020 |
have "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6021 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6022 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6023 |
hence eq_wq: "wq (e#s) cs' = wq s cs'" by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6024 |
from h have "holding s th' cs'" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6025 |
from this[unfolded s_holding_def, folded wq_def, unfolded eq_wq] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6026 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6027 |
by (unfold holdents_def s_holding_def, fold wq_def, insert eq_wq, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6028 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6029 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6030 |
from h[unfolded this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6031 |
have "holding s th' cs" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6032 |
from held_unique[OF this holding_th_cs_s] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6033 |
have "th' = th" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6034 |
with assms show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6035 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6036 |
} ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6037 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6038 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6039 |
lemma cntCS_kept [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6040 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6041 |
shows "cntCS (e#s) th' = cntCS s th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6042 |
by (unfold cntCS_def holdents_kept[OF assms], simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6043 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6044 |
lemma readys_kept1: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6045 |
assumes "th' \<in> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6046 |
shows "th' \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6047 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6048 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6049 |
assume wait: "waiting s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6050 |
have n_wait: "\<not> waiting (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6051 |
using assms(1)[unfolded readys_def] by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6052 |
have False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6053 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6054 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6055 |
with n_wait wait |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6056 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6057 |
by (unfold s_waiting_def, fold wq_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6058 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6059 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6060 |
have "th' \<in> set (th # rest) \<and> th' \<noteq> hd (th # rest)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6061 |
using wait[unfolded True s_waiting_def, folded wq_def, unfolded wq_s_cs] . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6062 |
hence "th' \<in> set rest" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6063 |
with set_wq' have "th' \<in> set wq'" by metis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6064 |
with nil_wq' show ?thesis by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6065 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6066 |
} thus ?thesis using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6067 |
by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6068 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6069 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6070 |
lemma readys_kept2: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6071 |
assumes "th' \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6072 |
shows "th' \<in> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6073 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6074 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6075 |
assume wait: "waiting (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6076 |
have n_wait: "\<not> waiting s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6077 |
using assms[unfolded readys_def] by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6078 |
have False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6079 |
proof(cases "cs' = cs") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6080 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6081 |
with n_wait wait |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6082 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6083 |
by (unfold s_waiting_def, fold wq_def, auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6084 |
next |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6085 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6086 |
have "th' \<in> set [] \<and> th' \<noteq> hd []" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6087 |
using wait[unfolded True s_waiting_def, folded wq_def, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6088 |
unfolded wq_es_cs nil_wq'] . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6089 |
thus ?thesis by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6090 |
qed |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6091 |
} with assms show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6092 |
by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6093 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6094 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6095 |
lemma readys_simp [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6096 |
shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6097 |
using readys_kept1[OF assms] readys_kept2[OF assms] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6098 |
by metis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6099 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6100 |
lemma cnp_cnv_cncs_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6101 |
assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6102 |
shows "cntP (e#s) th' = cntV (e#s) th' + cntCS (e#s) th' + pvD (e#s) th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6103 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6104 |
{ |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6105 |
assume eq_th': "th' = th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6106 |
have ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6107 |
apply (unfold eq_th' pvD_th_es cntCS_es_th) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6108 |
by (insert assms[unfolded eq_th'], unfold is_v, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6109 |
} moreover { |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6110 |
assume h: "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6111 |
have ?thesis using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6112 |
apply (unfold cntCS_kept[OF h], insert h, unfold is_v, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6113 |
by (fold is_v, unfold pvD_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6114 |
} ultimately show ?thesis by metis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6115 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6116 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6117 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6118 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6119 |
context valid_trace_v |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6120 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6121 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6122 |
lemma cnp_cnv_cncs_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6123 |
assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6124 |
shows "cntP (e#s) th' = cntV (e#s) th' + cntCS (e#s) th' + pvD (e#s) th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6125 |
proof(cases "rest = []") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6126 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6127 |
then interpret vt: valid_trace_v_e by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6128 |
show ?thesis using assms using vt.cnp_cnv_cncs_kept by blast |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6129 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6130 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6131 |
then interpret vt: valid_trace_v_n by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6132 |
show ?thesis using assms using vt.cnp_cnv_cncs_kept by blast |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6133 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6134 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6135 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6136 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6137 |
context valid_trace_create |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6138 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6139 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6140 |
lemma th_not_live_s [simp]: "th \<notin> threads s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6141 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6142 |
from pip_e[unfolded is_create] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6143 |
show ?thesis by (cases, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6144 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6145 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6146 |
lemma th_not_ready_s [simp]: "th \<notin> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6147 |
using th_not_live_s by (unfold readys_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6148 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6149 |
lemma th_live_es [simp]: "th \<in> threads (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6150 |
by (unfold is_create, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6151 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6152 |
lemma not_waiting_th_s [simp]: "\<not> waiting s th cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6153 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6154 |
assume "waiting s th cs'" |
99 | 6155 |
from this[unfolded s_waiting_def, folded wq_def, unfolded wq_kept] |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6156 |
have "th \<in> set (wq s cs')" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6157 |
from wq_threads[OF this] have "th \<in> threads s" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6158 |
with th_not_live_s show False by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6159 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6160 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6161 |
lemma not_holding_th_s [simp]: "\<not> holding s th cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6162 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6163 |
assume "holding s th cs'" |
99 | 6164 |
from this[unfolded s_holding_def, folded wq_def, unfolded wq_kept] |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6165 |
have "th \<in> set (wq s cs')" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6166 |
from wq_threads[OF this] have "th \<in> threads s" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6167 |
with th_not_live_s show False by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6168 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6169 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6170 |
lemma not_waiting_th_es [simp]: "\<not> waiting (e#s) th cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6171 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6172 |
assume "waiting (e # s) th cs'" |
99 | 6173 |
from this[unfolded s_waiting_def, folded wq_def, unfolded wq_kept] |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6174 |
have "th \<in> set (wq s cs')" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6175 |
from wq_threads[OF this] have "th \<in> threads s" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6176 |
with th_not_live_s show False by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6177 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6178 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6179 |
lemma not_holding_th_es [simp]: "\<not> holding (e#s) th cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6180 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6181 |
assume "holding (e # s) th cs'" |
99 | 6182 |
from this[unfolded s_holding_def, folded wq_def, unfolded wq_kept] |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6183 |
have "th \<in> set (wq s cs')" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6184 |
from wq_threads[OF this] have "th \<in> threads s" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6185 |
with th_not_live_s show False by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6186 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6187 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6188 |
lemma ready_th_es [simp]: "th \<in> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6189 |
by (simp add:readys_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6190 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6191 |
lemma holdents_th_s: "holdents s th = {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6192 |
by (unfold holdents_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6193 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6194 |
lemma holdents_th_es: "holdents (e#s) th = {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6195 |
by (unfold holdents_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6196 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6197 |
lemma cntCS_th_s [simp]: "cntCS s th = 0" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6198 |
by (unfold cntCS_def, simp add:holdents_th_s) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6199 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6200 |
lemma cntCS_th_es [simp]: "cntCS (e#s) th = 0" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6201 |
by (unfold cntCS_def, simp add:holdents_th_es) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6202 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6203 |
lemma pvD_th_s [simp]: "pvD s th = 0" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6204 |
by (unfold pvD_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6205 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6206 |
lemma pvD_th_es [simp]: "pvD (e#s) th = 0" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6207 |
by (unfold pvD_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6208 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6209 |
lemma holdents_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6210 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6211 |
shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6212 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6213 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6214 |
assume h: "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6215 |
hence "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6216 |
by (unfold holdents_def s_holding_def, fold wq_def, |
99 | 6217 |
unfold wq_kept, auto) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6218 |
} moreover { |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6219 |
fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6220 |
assume h: "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6221 |
hence "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6222 |
by (unfold holdents_def s_holding_def, fold wq_def, |
99 | 6223 |
unfold wq_kept, auto) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6224 |
} ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6225 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6226 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6227 |
lemma cntCS_kept [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6228 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6229 |
shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6230 |
using holdents_kept[OF assms] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6231 |
by (unfold cntCS_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6232 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6233 |
lemma readys_kept1: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6234 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6235 |
and "th' \<in> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6236 |
shows "th' \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6237 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6238 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6239 |
assume wait: "waiting s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6240 |
have n_wait: "\<not> waiting (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6241 |
using assms by (auto simp:readys_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6242 |
from wait[unfolded s_waiting_def, folded wq_def] |
99 | 6243 |
n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_kept] |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6244 |
have False by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6245 |
} thus ?thesis using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6246 |
by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6247 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6248 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6249 |
lemma readys_kept2: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6250 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6251 |
and "th' \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6252 |
shows "th' \<in> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6253 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6254 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6255 |
assume wait: "waiting (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6256 |
have n_wait: "\<not> waiting s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6257 |
using assms(2) by (auto simp:readys_def) |
99 | 6258 |
from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_kept] |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6259 |
n_wait[unfolded s_waiting_def, folded wq_def] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6260 |
have False by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6261 |
} with assms show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6262 |
by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6263 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6264 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6265 |
lemma readys_simp [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6266 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6267 |
shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6268 |
using readys_kept1[OF assms] readys_kept2[OF assms] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6269 |
by metis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6270 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6271 |
lemma pvD_kept [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6272 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6273 |
shows "pvD (e#s) th' = pvD s th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6274 |
using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6275 |
by (unfold pvD_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6276 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6277 |
lemma cnp_cnv_cncs_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6278 |
assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6279 |
shows "cntP (e#s) th' = cntV (e#s) th' + cntCS (e#s) th' + pvD (e#s) th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6280 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6281 |
{ |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6282 |
assume eq_th': "th' = th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6283 |
have ?thesis using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6284 |
by (unfold eq_th', simp, unfold is_create, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6285 |
} moreover { |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6286 |
assume h: "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6287 |
hence ?thesis using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6288 |
by (simp, simp add:is_create) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6289 |
} ultimately show ?thesis by metis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6290 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6291 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6292 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6293 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6294 |
context valid_trace_exit |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6295 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6296 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6297 |
lemma th_live_s [simp]: "th \<in> threads s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6298 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6299 |
from pip_e[unfolded is_exit] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6300 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6301 |
by (cases, unfold runing_def readys_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6302 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6303 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6304 |
lemma th_ready_s [simp]: "th \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6305 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6306 |
from pip_e[unfolded is_exit] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6307 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6308 |
by (cases, unfold runing_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6309 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6310 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6311 |
lemma th_not_live_es [simp]: "th \<notin> threads (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6312 |
by (unfold is_exit, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6313 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6314 |
lemma not_holding_th_s [simp]: "\<not> holding s th cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6315 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6316 |
from pip_e[unfolded is_exit] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6317 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6318 |
by (cases, unfold holdents_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6319 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6320 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6321 |
lemma cntCS_th_s [simp]: "cntCS s th = 0" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6322 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6323 |
from pip_e[unfolded is_exit] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6324 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6325 |
by (cases, unfold cntCS_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6326 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6327 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6328 |
lemma not_holding_th_es [simp]: "\<not> holding (e#s) th cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6329 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6330 |
assume "holding (e # s) th cs'" |
99 | 6331 |
from this[unfolded s_holding_def, folded wq_def, unfolded wq_kept] |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6332 |
have "holding s th cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6333 |
by (unfold s_holding_def, fold wq_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6334 |
with not_holding_th_s |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6335 |
show False by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6336 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6337 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6338 |
lemma ready_th_es [simp]: "th \<notin> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6339 |
by (simp add:readys_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6340 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6341 |
lemma holdents_th_s: "holdents s th = {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6342 |
by (unfold holdents_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6343 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6344 |
lemma holdents_th_es: "holdents (e#s) th = {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6345 |
by (unfold holdents_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6346 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6347 |
lemma cntCS_th_es [simp]: "cntCS (e#s) th = 0" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6348 |
by (unfold cntCS_def, simp add:holdents_th_es) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6349 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6350 |
lemma pvD_th_s [simp]: "pvD s th = 0" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6351 |
by (unfold pvD_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6352 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6353 |
lemma pvD_th_es [simp]: "pvD (e#s) th = 0" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6354 |
by (unfold pvD_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6355 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6356 |
lemma holdents_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6357 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6358 |
shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6359 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6360 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6361 |
assume h: "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6362 |
hence "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6363 |
by (unfold holdents_def s_holding_def, fold wq_def, |
99 | 6364 |
unfold wq_kept, auto) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6365 |
} moreover { |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6366 |
fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6367 |
assume h: "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6368 |
hence "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6369 |
by (unfold holdents_def s_holding_def, fold wq_def, |
99 | 6370 |
unfold wq_kept, auto) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6371 |
} ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6372 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6373 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6374 |
lemma cntCS_kept [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6375 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6376 |
shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6377 |
using holdents_kept[OF assms] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6378 |
by (unfold cntCS_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6379 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6380 |
lemma readys_kept1: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6381 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6382 |
and "th' \<in> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6383 |
shows "th' \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6384 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6385 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6386 |
assume wait: "waiting s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6387 |
have n_wait: "\<not> waiting (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6388 |
using assms by (auto simp:readys_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6389 |
from wait[unfolded s_waiting_def, folded wq_def] |
99 | 6390 |
n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_kept] |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6391 |
have False by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6392 |
} thus ?thesis using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6393 |
by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6394 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6395 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6396 |
lemma readys_kept2: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6397 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6398 |
and "th' \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6399 |
shows "th' \<in> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6400 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6401 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6402 |
assume wait: "waiting (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6403 |
have n_wait: "\<not> waiting s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6404 |
using assms(2) by (auto simp:readys_def) |
99 | 6405 |
from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_kept] |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6406 |
n_wait[unfolded s_waiting_def, folded wq_def] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6407 |
have False by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6408 |
} with assms show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6409 |
by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6410 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6411 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6412 |
lemma readys_simp [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6413 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6414 |
shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6415 |
using readys_kept1[OF assms] readys_kept2[OF assms] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6416 |
by metis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6417 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6418 |
lemma pvD_kept [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6419 |
assumes "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6420 |
shows "pvD (e#s) th' = pvD s th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6421 |
using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6422 |
by (unfold pvD_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6423 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6424 |
lemma cnp_cnv_cncs_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6425 |
assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6426 |
shows "cntP (e#s) th' = cntV (e#s) th' + cntCS (e#s) th' + pvD (e#s) th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6427 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6428 |
{ |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6429 |
assume eq_th': "th' = th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6430 |
have ?thesis using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6431 |
by (unfold eq_th', simp, unfold is_exit, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6432 |
} moreover { |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6433 |
assume h: "th' \<noteq> th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6434 |
hence ?thesis using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6435 |
by (simp, simp add:is_exit) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6436 |
} ultimately show ?thesis by metis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6437 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6438 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6439 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6440 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6441 |
context valid_trace_set |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6442 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6443 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6444 |
lemma th_live_s [simp]: "th \<in> threads s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6445 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6446 |
from pip_e[unfolded is_set] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6447 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6448 |
by (cases, unfold runing_def readys_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6449 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6450 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6451 |
lemma th_ready_s [simp]: "th \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6452 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6453 |
from pip_e[unfolded is_set] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6454 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6455 |
by (cases, unfold runing_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6456 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6457 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6458 |
lemma th_not_live_es [simp]: "th \<in> threads (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6459 |
by (unfold is_set, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6460 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6461 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6462 |
lemma holdents_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6463 |
shows "holdents (e#s) th' = holdents s th'" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6464 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6465 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6466 |
assume h: "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6467 |
hence "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6468 |
by (unfold holdents_def s_holding_def, fold wq_def, |
99 | 6469 |
unfold wq_kept, auto) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6470 |
} moreover { |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6471 |
fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6472 |
assume h: "cs' \<in> ?R" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6473 |
hence "cs' \<in> ?L" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6474 |
by (unfold holdents_def s_holding_def, fold wq_def, |
99 | 6475 |
unfold wq_kept, auto) |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6476 |
} ultimately show ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6477 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6478 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6479 |
lemma cntCS_kept [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6480 |
shows "cntCS (e#s) th' = cntCS s th'" (is "?L = ?R") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6481 |
using holdents_kept |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6482 |
by (unfold cntCS_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6483 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6484 |
lemma threads_kept[simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6485 |
"threads (e#s) = threads s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6486 |
by (unfold is_set, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6487 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6488 |
lemma readys_kept1: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6489 |
assumes "th' \<in> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6490 |
shows "th' \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6491 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6492 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6493 |
assume wait: "waiting s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6494 |
have n_wait: "\<not> waiting (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6495 |
using assms by (auto simp:readys_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6496 |
from wait[unfolded s_waiting_def, folded wq_def] |
99 | 6497 |
n_wait[unfolded s_waiting_def, folded wq_def, unfolded wq_kept] |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6498 |
have False by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6499 |
} moreover have "th' \<in> threads s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6500 |
using assms[unfolded readys_def] by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6501 |
ultimately show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6502 |
by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6503 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6504 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6505 |
lemma readys_kept2: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6506 |
assumes "th' \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6507 |
shows "th' \<in> readys (e#s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6508 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6509 |
{ fix cs' |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6510 |
assume wait: "waiting (e#s) th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6511 |
have n_wait: "\<not> waiting s th' cs'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6512 |
using assms by (auto simp:readys_def) |
99 | 6513 |
from wait[unfolded s_waiting_def, folded wq_def, unfolded wq_kept] |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6514 |
n_wait[unfolded s_waiting_def, folded wq_def] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6515 |
have False by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6516 |
} with assms show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6517 |
by (unfold readys_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6518 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6519 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6520 |
lemma readys_simp [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6521 |
shows "(th' \<in> readys (e#s)) = (th' \<in> readys s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6522 |
using readys_kept1 readys_kept2 |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6523 |
by metis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6524 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6525 |
lemma pvD_kept [simp]: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6526 |
shows "pvD (e#s) th' = pvD s th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6527 |
by (unfold pvD_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6528 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6529 |
lemma cnp_cnv_cncs_kept: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6530 |
assumes "cntP s th' = cntV s th' + cntCS s th' + pvD s th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6531 |
shows "cntP (e#s) th' = cntV (e#s) th' + cntCS (e#s) th' + pvD (e#s) th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6532 |
using assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6533 |
by (unfold is_set, simp, fold is_set, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6534 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6535 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6536 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6537 |
context valid_trace |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6538 |
begin |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6539 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6540 |
lemma cnp_cnv_cncs: "cntP s th' = cntV s th' + cntCS s th' + pvD s th'" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6541 |
proof(induct rule:ind) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6542 |
case Nil |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6543 |
thus ?case |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6544 |
by (unfold cntP_def cntV_def pvD_def cntCS_def holdents_def |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6545 |
s_holding_def, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6546 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6547 |
case (Cons s e) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6548 |
interpret vt_e: valid_trace_e s e using Cons by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6549 |
show ?case |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6550 |
proof(cases e) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6551 |
case (Create th prio) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6552 |
interpret vt_create: valid_trace_create s e th prio |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6553 |
using Create by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6554 |
show ?thesis using Cons by (simp add: vt_create.cnp_cnv_cncs_kept) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6555 |
next |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6556 |
case (Exit th) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6557 |
interpret vt_exit: valid_trace_exit s e th |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6558 |
using Exit by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6559 |
show ?thesis using Cons by (simp add: vt_exit.cnp_cnv_cncs_kept) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6560 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6561 |
case (P th cs) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6562 |
interpret vt_p: valid_trace_p s e th cs using P by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6563 |
show ?thesis using Cons by (simp add: vt_p.cnp_cnv_cncs_kept) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6564 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6565 |
case (V th cs) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6566 |
interpret vt_v: valid_trace_v s e th cs using V by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6567 |
show ?thesis using Cons by (simp add: vt_v.cnp_cnv_cncs_kept) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6568 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6569 |
case (Set th prio) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6570 |
interpret vt_set: valid_trace_set s e th prio |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6571 |
using Set by (unfold_locales, simp) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6572 |
show ?thesis using Cons by (simp add: vt_set.cnp_cnv_cncs_kept) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6573 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6574 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6575 |
|
101 | 6576 |
end |
6577 |
||
6578 |
section {* Corollaries of @{thm valid_trace.cnp_cnv_cncs} *} |
|
6579 |
||
6580 |
context valid_trace |
|
6581 |
begin |
|
6582 |
||
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6583 |
lemma not_thread_holdents: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6584 |
assumes not_in: "th \<notin> threads s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6585 |
shows "holdents s th = {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6586 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6587 |
{ fix cs |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6588 |
assume "cs \<in> holdents s th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6589 |
hence "holding s th cs" by (auto simp:holdents_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6590 |
from this[unfolded s_holding_def, folded wq_def] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6591 |
have "th \<in> set (wq s cs)" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6592 |
with wq_threads have "th \<in> threads s" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6593 |
with assms |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6594 |
have False by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6595 |
} thus ?thesis by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6596 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6597 |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6598 |
lemma not_thread_cncs: |
63 | 6599 |
assumes not_in: "th \<notin> threads s" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6600 |
shows "cntCS s th = 0" |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6601 |
using not_thread_holdents[OF assms] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6602 |
by (simp add:cntCS_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6603 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6604 |
lemma cnp_cnv_eq: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6605 |
assumes "th \<notin> threads s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6606 |
shows "cntP s th = cntV s th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6607 |
using assms cnp_cnv_cncs not_thread_cncs pvD_def |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6608 |
by (auto) |
63 | 6609 |
|
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6610 |
lemma eq_pv_children: |
104 | 6611 |
======= |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6612 |
|
63 | 6613 |
context valid_trace |
6614 |
begin |
|
6615 |
||
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6616 |
lemma cnp_cnv_eq: |
63 | 6617 |
assumes "th \<notin> threads s" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6618 |
shows "cntP s th = cntV s th" |
63 | 6619 |
using assms |
6620 |
using cnp_cnv_cncs not_thread_cncs by auto |
|
6621 |
||
6622 |
end |
|
6623 |
||
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6624 |
lemma eq_RAG: |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6625 |
"RAG (wq s) = RAG s" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6626 |
by (unfold cs_RAG_def s_RAG_def, auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6627 |
|
63 | 6628 |
context valid_trace |
6629 |
begin |
|
6630 |
||
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6631 |
lemma count_eq_dependants: |
104 | 6632 |
>>>>>>> other |
63 | 6633 |
assumes eq_pv: "cntP s th = cntV s th" |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6634 |
shows "dependants (wq s) th = {}" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6635 |
proof - |
63 | 6636 |
from cnp_cnv_cncs and eq_pv |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6637 |
have "cntCS s th = 0" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6638 |
by (auto split:if_splits) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6639 |
moreover have "finite {cs. (Cs cs, Th th) \<in> RAG s}" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6640 |
proof - |
63 | 6641 |
from finite_holding[of th] show ?thesis |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6642 |
by (simp add:holdents_test) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6643 |
qed |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6644 |
ultimately have h: "{cs. (Cs cs, Th th) \<in> RAG s} = {}" |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6645 |
by (unfold cntCS_def holdents_test cs_dependants_def, auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6646 |
show ?thesis |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6647 |
proof(unfold cs_dependants_def) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6648 |
{ assume "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6649 |
then obtain th' where "(Th th', Th th) \<in> (RAG (wq s))\<^sup>+" by auto |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6650 |
hence "False" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6651 |
proof(cases) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6652 |
assume "(Th th', Th th) \<in> RAG (wq s)" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6653 |
thus "False" by (auto simp:cs_RAG_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6654 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6655 |
fix c |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6656 |
assume "(c, Th th) \<in> RAG (wq s)" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6657 |
with h and eq_RAG show "False" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6658 |
by (cases c, auto simp:cs_RAG_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6659 |
qed |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6660 |
} thus "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} = {}" by auto |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6661 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6662 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6663 |
|
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6664 |
lemma dependants_threads: |
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6665 |
shows "dependants (wq s) th \<subseteq> threads s" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6666 |
proof |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6667 |
{ fix th th' |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6668 |
assume h: "th \<in> {th'a. (Th th'a, Th th') \<in> (RAG (wq s))\<^sup>+}" |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6669 |
have "Th th \<in> Domain (RAG s)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6670 |
proof - |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6671 |
from h obtain th' where "(Th th, Th th') \<in> (RAG (wq s))\<^sup>+" by auto |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6672 |
hence "(Th th) \<in> Domain ( (RAG (wq s))\<^sup>+)" by (auto simp:Domain_def) |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6673 |
with trancl_domain have "(Th th) \<in> Domain (RAG (wq s))" by simp |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6674 |
thus ?thesis using eq_RAG by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6675 |
qed |
63 | 6676 |
from dm_RAG_threads[OF this] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6677 |
have "th \<in> threads s" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6678 |
} note hh = this |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6679 |
fix th1 |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6680 |
assume "th1 \<in> dependants (wq s) th" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6681 |
hence "th1 \<in> {th'a. (Th th'a, Th th) \<in> (RAG (wq s))\<^sup>+}" |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6682 |
by (unfold cs_dependants_def, simp) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6683 |
from hh [OF this] show "th1 \<in> threads s" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6684 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6685 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6686 |
lemma finite_threads: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6687 |
shows "finite (threads s)" |
63 | 6688 |
using vt by (induct) (auto elim: step.cases) |
6689 |
||
104 | 6690 |
<<<<<<< local |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6691 |
lemma count_eq_RAG_plus: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6692 |
assumes "cntP s th = cntV s th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6693 |
shows "{th'. (Th th', Th th) \<in> (RAG s)^+} = {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6694 |
proof(rule ccontr) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6695 |
assume otherwise: "{th'. (Th th', Th th) \<in> (RAG s)\<^sup>+} \<noteq> {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6696 |
then obtain th' where "(Th th', Th th) \<in> (RAG s)^+" by auto |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6697 |
from tranclD2[OF this] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6698 |
obtain z where "z \<in> children (RAG s) (Th th)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6699 |
by (auto simp:children_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6700 |
with eq_pv_children[OF assms] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6701 |
show False by simp |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6702 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6703 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6704 |
lemma eq_pv_dependants: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6705 |
assumes eq_pv: "cntP s th = cntV s th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6706 |
shows "dependants s th = {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6707 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6708 |
from count_eq_RAG_plus[OF assms, folded dependants_alt_def1] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6709 |
show ?thesis . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6710 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6711 |
|
104 | 6712 |
======= |
63 | 6713 |
end |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6714 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6715 |
lemma Max_f_mono: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6716 |
assumes seq: "A \<subseteq> B" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6717 |
and np: "A \<noteq> {}" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6718 |
and fnt: "finite B" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6719 |
shows "Max (f ` A) \<le> Max (f ` B)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6720 |
proof(rule Max_mono) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6721 |
from seq show "f ` A \<subseteq> f ` B" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6722 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6723 |
from np show "f ` A \<noteq> {}" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6724 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6725 |
from fnt and seq show "finite (f ` B)" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6726 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6727 |
|
63 | 6728 |
context valid_trace |
6729 |
begin |
|
6730 |
||
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6731 |
lemma cp_le: |
63 | 6732 |
assumes th_in: "th \<in> threads s" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6733 |
shows "cp s th \<le> Max ((\<lambda> th. (preced th s)) ` threads s)" |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6734 |
proof(unfold cp_eq_cpreced cpreced_def cs_dependants_def) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6735 |
show "Max ((\<lambda>th. preced th s) ` ({th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+})) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6736 |
\<le> Max ((\<lambda>th. preced th s) ` threads s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6737 |
(is "Max (?f ` ?A) \<le> Max (?f ` ?B)") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6738 |
proof(rule Max_f_mono) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6739 |
show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<noteq> {}" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6740 |
next |
63 | 6741 |
from finite_threads |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6742 |
show "finite (threads s)" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6743 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6744 |
from th_in |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6745 |
show "{th} \<union> {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> threads s" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6746 |
apply (auto simp:Domain_def) |
63 | 6747 |
apply (rule_tac dm_RAG_threads) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6748 |
apply (unfold trancl_domain [of "RAG s", symmetric]) |
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6749 |
by (unfold cs_RAG_def s_RAG_def, auto simp:Domain_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6750 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6751 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6752 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6753 |
lemma le_cp: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6754 |
shows "preced th s \<le> cp s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6755 |
proof(unfold cp_eq_cpreced preced_def cpreced_def, simp) |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6756 |
show "Prc (priority th s) (last_set th s) |
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6757 |
\<le> Max (insert (Prc (priority th s) (last_set th s)) |
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6758 |
((\<lambda>th. Prc (priority th s) (last_set th s)) ` dependants (wq s) th))" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6759 |
(is "?l \<le> Max (insert ?l ?A)") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6760 |
proof(cases "?A = {}") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6761 |
case False |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6762 |
have "finite ?A" (is "finite (?f ` ?B)") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6763 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6764 |
have "finite ?B" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6765 |
proof- |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6766 |
have "finite {th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+}" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6767 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6768 |
let ?F = "\<lambda> (x, y). the_th x" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6769 |
have "{th'. (Th th', Th th) \<in> (RAG (wq s))\<^sup>+} \<subseteq> ?F ` ((RAG (wq s))\<^sup>+)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6770 |
apply (auto simp:image_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6771 |
by (rule_tac x = "(Th x, Th th)" in bexI, auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6772 |
moreover have "finite \<dots>" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6773 |
proof - |
63 | 6774 |
from finite_RAG have "finite (RAG s)" . |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6775 |
hence "finite ((RAG (wq s))\<^sup>+)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6776 |
apply (unfold finite_trancl) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6777 |
by (auto simp: s_RAG_def cs_RAG_def wq_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6778 |
thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6779 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6780 |
ultimately show ?thesis by (auto intro:finite_subset) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6781 |
qed |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6782 |
thus ?thesis by (simp add:cs_dependants_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6783 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6784 |
thus ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6785 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6786 |
from Max_insert [OF this False, of ?l] show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6787 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6788 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6789 |
thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6790 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6791 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6792 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6793 |
lemma max_cp_eq: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6794 |
shows "Max ((cp s) ` threads s) = Max ((\<lambda> th. (preced th s)) ` threads s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6795 |
(is "?l = ?r") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6796 |
proof(cases "threads s = {}") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6797 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6798 |
thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6799 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6800 |
case False |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6801 |
have "?l \<in> ((cp s) ` threads s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6802 |
proof(rule Max_in) |
63 | 6803 |
from finite_threads |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6804 |
show "finite (cp s ` threads s)" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6805 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6806 |
from False show "cp s ` threads s \<noteq> {}" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6807 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6808 |
then obtain th |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6809 |
where th_in: "th \<in> threads s" and eq_l: "?l = cp s th" by auto |
63 | 6810 |
have "\<dots> \<le> ?r" by (rule cp_le[OF th_in]) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6811 |
moreover have "?r \<le> cp s th" (is "Max (?f ` ?A) \<le> cp s th") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6812 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6813 |
have "?r \<in> (?f ` ?A)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6814 |
proof(rule Max_in) |
63 | 6815 |
from finite_threads |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6816 |
show " finite ((\<lambda>th. preced th s) ` threads s)" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6817 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6818 |
from False show " (\<lambda>th. preced th s) ` threads s \<noteq> {}" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6819 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6820 |
then obtain th' where |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6821 |
th_in': "th' \<in> ?A " and eq_r: "?r = ?f th'" by auto |
63 | 6822 |
from le_cp [of th'] eq_r |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6823 |
have "?r \<le> cp s th'" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6824 |
moreover have "\<dots> \<le> cp s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6825 |
proof(fold eq_l) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6826 |
show " cp s th' \<le> Max (cp s ` threads s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6827 |
proof(rule Max_ge) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6828 |
from th_in' show "cp s th' \<in> cp s ` threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6829 |
by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6830 |
next |
63 | 6831 |
from finite_threads |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6832 |
show "finite (cp s ` threads s)" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6833 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6834 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6835 |
ultimately show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6836 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6837 |
ultimately show ?thesis using eq_l by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6838 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6839 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6840 |
lemma max_cp_readys_threads_pre: |
63 | 6841 |
assumes np: "threads s \<noteq> {}" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6842 |
shows "Max (cp s ` readys s) = Max (cp s ` threads s)" |
63 | 6843 |
proof(unfold max_cp_eq) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6844 |
show "Max (cp s ` readys s) = Max ((\<lambda>th. preced th s) ` threads s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6845 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6846 |
let ?p = "Max ((\<lambda>th. preced th s) ` threads s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6847 |
let ?f = "(\<lambda>th. preced th s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6848 |
have "?p \<in> ((\<lambda>th. preced th s) ` threads s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6849 |
proof(rule Max_in) |
63 | 6850 |
from finite_threads show "finite (?f ` threads s)" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6851 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6852 |
from np show "?f ` threads s \<noteq> {}" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6853 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6854 |
then obtain tm where tm_max: "?f tm = ?p" and tm_in: "tm \<in> threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6855 |
by (auto simp:Image_def) |
63 | 6856 |
from th_chain_to_ready [OF tm_in] |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6857 |
have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+)" . |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6858 |
thus ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6859 |
proof |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6860 |
assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (RAG s)\<^sup>+ " |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6861 |
then obtain th' where th'_in: "th' \<in> readys s" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6862 |
and tm_chain:"(Th tm, Th th') \<in> (RAG s)\<^sup>+" by auto |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6863 |
have "cp s th' = ?f tm" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6864 |
proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI) |
63 | 6865 |
from dependants_threads finite_threads |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6866 |
show "finite ((\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th'))" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6867 |
by (auto intro:finite_subset) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6868 |
next |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6869 |
fix p assume p_in: "p \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6870 |
from tm_max have " preced tm s = Max ((\<lambda>th. preced th s) ` threads s)" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6871 |
moreover have "p \<le> \<dots>" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6872 |
proof(rule Max_ge) |
63 | 6873 |
from finite_threads |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6874 |
show "finite ((\<lambda>th. preced th s) ` threads s)" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6875 |
next |
63 | 6876 |
from p_in and th'_in and dependants_threads[of th'] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6877 |
show "p \<in> (\<lambda>th. preced th s) ` threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6878 |
by (auto simp:readys_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6879 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6880 |
ultimately show "p \<le> preced tm s" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6881 |
next |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6882 |
show "preced tm s \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependants (wq s) th')" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6883 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6884 |
from tm_chain |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6885 |
have "tm \<in> dependants (wq s) th'" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
6886 |
by (unfold cs_dependants_def s_RAG_def cs_RAG_def, auto) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6887 |
thus ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6888 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6889 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6890 |
with tm_max |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6891 |
have h: "cp s th' = Max ((\<lambda>th. preced th s) ` threads s)" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6892 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6893 |
proof (fold h, rule Max_eqI) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6894 |
fix q |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6895 |
assume "q \<in> cp s ` readys s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6896 |
then obtain th1 where th1_in: "th1 \<in> readys s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6897 |
and eq_q: "q = cp s th1" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6898 |
show "q \<le> cp s th'" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6899 |
apply (unfold h eq_q) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6900 |
apply (unfold cp_eq_cpreced cpreced_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6901 |
apply (rule Max_mono) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6902 |
proof - |
63 | 6903 |
from dependants_threads [of th1] th1_in |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6904 |
show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<subseteq> |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6905 |
(\<lambda>th. preced th s) ` threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6906 |
by (auto simp:readys_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6907 |
next |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6908 |
show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}" by simp |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6909 |
next |
63 | 6910 |
from finite_threads |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6911 |
show " finite ((\<lambda>th. preced th s) ` threads s)" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6912 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6913 |
next |
63 | 6914 |
from finite_threads |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6915 |
show "finite (cp s ` readys s)" by (auto simp:readys_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6916 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6917 |
from th'_in |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6918 |
show "cp s th' \<in> cp s ` readys s" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6919 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6920 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6921 |
assume tm_ready: "tm \<in> readys s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6922 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6923 |
proof(fold tm_max) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6924 |
have cp_eq_p: "cp s tm = preced tm s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6925 |
proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6926 |
fix y |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6927 |
assume hy: "y \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6928 |
show "y \<le> preced tm s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6929 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6930 |
{ fix y' |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6931 |
assume hy' : "y' \<in> ((\<lambda>th. preced th s) ` dependants (wq s) tm)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6932 |
have "y' \<le> preced tm s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6933 |
proof(unfold tm_max, rule Max_ge) |
63 | 6934 |
from hy' dependants_threads[of tm] |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6935 |
show "y' \<in> (\<lambda>th. preced th s) ` threads s" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6936 |
next |
63 | 6937 |
from finite_threads |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6938 |
show "finite ((\<lambda>th. preced th s) ` threads s)" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6939 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6940 |
} with hy show ?thesis by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6941 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6942 |
next |
63 | 6943 |
from dependants_threads[of tm] finite_threads |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6944 |
show "finite ((\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm))" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6945 |
by (auto intro:finite_subset) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6946 |
next |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6947 |
show "preced tm s \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependants (wq s) tm)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6948 |
by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6949 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6950 |
moreover have "Max (cp s ` readys s) = cp s tm" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6951 |
proof(rule Max_eqI) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6952 |
from tm_ready show "cp s tm \<in> cp s ` readys s" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6953 |
next |
63 | 6954 |
from finite_threads |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6955 |
show "finite (cp s ` readys s)" by (auto simp:readys_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6956 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6957 |
fix y assume "y \<in> cp s ` readys s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6958 |
then obtain th1 where th1_readys: "th1 \<in> readys s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6959 |
and h: "y = cp s th1" by auto |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6960 |
show "y \<le> cp s tm" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6961 |
apply(unfold cp_eq_p h) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6962 |
apply(unfold cp_eq_cpreced cpreced_def tm_max, rule Max_mono) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6963 |
proof - |
63 | 6964 |
from finite_threads |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6965 |
show "finite ((\<lambda>th. preced th s) ` threads s)" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6966 |
next |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6967 |
show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) \<noteq> {}" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6968 |
by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6969 |
next |
63 | 6970 |
from dependants_threads[of th1] th1_readys |
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
3
diff
changeset
|
6971 |
show "(\<lambda>th. preced th s) ` ({th1} \<union> dependants (wq s) th1) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6972 |
\<subseteq> (\<lambda>th. preced th s) ` threads s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6973 |
by (auto simp:readys_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6974 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6975 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6976 |
ultimately show " Max (cp s ` readys s) = preced tm s" by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6977 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6978 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6979 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6980 |
qed |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
6981 |
|
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
6982 |
text {* (* ccc *) \noindent |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
6983 |
Since the current precedence of the threads in ready queue will always be boosted, |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
6984 |
there must be one inside it has the maximum precedence of the whole system. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
6985 |
*} |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6986 |
lemma max_cp_readys_threads: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6987 |
shows "Max (cp s ` readys s) = Max (cp s ` threads s)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6988 |
proof(cases "threads s = {}") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6989 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6990 |
thus ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6991 |
by (auto simp:readys_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6992 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6993 |
case False |
63 | 6994 |
show ?thesis by (rule max_cp_readys_threads_pre[OF False]) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6995 |
qed |
63 | 6996 |
|
6997 |
end |
|
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6998 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6999 |
lemma eq_holding: "holding (wq s) th cs = holding s th cs" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7000 |
apply (unfold s_holding_def cs_holding_def wq_def, simp) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7001 |
done |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7002 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7003 |
lemma f_image_eq: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7004 |
assumes h: "\<And> a. a \<in> A \<Longrightarrow> f a = g a" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7005 |
shows "f ` A = g ` A" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7006 |
proof |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7007 |
show "f ` A \<subseteq> g ` A" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7008 |
by(rule image_subsetI, auto intro:h) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7009 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7010 |
show "g ` A \<subseteq> f ` A" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7011 |
by (rule image_subsetI, auto intro:h[symmetric]) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7012 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7013 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7014 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7015 |
definition detached :: "state \<Rightarrow> thread \<Rightarrow> bool" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7016 |
where "detached s th \<equiv> (\<not>(\<exists> cs. holding s th cs)) \<and> (\<not>(\<exists>cs. waiting s th cs))" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7017 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7018 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7019 |
lemma detached_test: |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
7020 |
shows "detached s th = (Th th \<notin> Field (RAG s))" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7021 |
apply(simp add: detached_def Field_def) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
7022 |
apply(simp add: s_RAG_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7023 |
apply(simp add: s_holding_abv s_waiting_abv) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7024 |
apply(simp add: Domain_iff Range_iff) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7025 |
apply(simp add: wq_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7026 |
apply(auto) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7027 |
done |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7028 |
|
63 | 7029 |
context valid_trace |
7030 |
begin |
|
7031 |
||
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7032 |
lemma detached_intro: |
63 | 7033 |
assumes eq_pv: "cntP s th = cntV s th" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7034 |
shows "detached s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7035 |
proof - |
63 | 7036 |
from cnp_cnv_cncs |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7037 |
have eq_cnt: "cntP s th = |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7038 |
cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7039 |
hence cncs_zero: "cntCS s th = 0" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7040 |
by (auto simp:eq_pv split:if_splits) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7041 |
with eq_cnt |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7042 |
have "th \<in> readys s \<or> th \<notin> threads s" by (auto simp:eq_pv) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7043 |
thus ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7044 |
proof |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7045 |
assume "th \<notin> threads s" |
63 | 7046 |
with range_in dm_RAG_threads |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7047 |
show ?thesis |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
7048 |
by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def Domain_iff Range_iff) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7049 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7050 |
assume "th \<in> readys s" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
7051 |
moreover have "Th th \<notin> Range (RAG s)" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7052 |
proof - |
63 | 7053 |
from card_0_eq [OF finite_holding] and cncs_zero |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7054 |
have "holdents s th = {}" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7055 |
by (simp add:cntCS_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7056 |
thus ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7057 |
apply(auto simp:holdents_test) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7058 |
apply(case_tac a) |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
7059 |
apply(auto simp:holdents_test s_RAG_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7060 |
done |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7061 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7062 |
ultimately show ?thesis |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
7063 |
by (auto simp add: detached_def s_RAG_def s_waiting_abv s_holding_abv wq_def readys_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7064 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7065 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7066 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7067 |
lemma detached_elim: |
63 | 7068 |
assumes dtc: "detached s th" |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7069 |
shows "cntP s th = cntV s th" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7070 |
proof - |
63 | 7071 |
from cnp_cnv_cncs |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7072 |
have eq_pv: " cntP s th = |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7073 |
cntV s th + (if th \<in> readys s \<or> th \<notin> threads s then cntCS s th else cntCS s th + 1)" . |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7074 |
have cncs_z: "cntCS s th = 0" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7075 |
proof - |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7076 |
from dtc have "holdents s th = {}" |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
7077 |
unfolding detached_def holdents_test s_RAG_def |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7078 |
by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7079 |
thus ?thesis by (auto simp:cntCS_def) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7080 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7081 |
show ?thesis |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7082 |
proof(cases "th \<in> threads s") |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7083 |
case True |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7084 |
with dtc |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7085 |
have "th \<in> readys s" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7086 |
by (unfold readys_def detached_def Field_def Domain_def Range_def, |
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
7087 |
auto simp:eq_waiting s_RAG_def) |
0
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7088 |
with cncs_z and eq_pv show ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7089 |
next |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7090 |
case False |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7091 |
with cncs_z and eq_pv show ?thesis by simp |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7092 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7093 |
qed |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7094 |
|
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7095 |
lemma detached_eq: |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7096 |
shows "(detached s th) = (cntP s th = cntV s th)" |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7097 |
by (insert vt, auto intro:detached_intro detached_elim) |
110247f9d47e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7098 |
|
63 | 7099 |
end |
7100 |
||
53
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
7101 |
text {* |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
7102 |
The lemmas in this .thy file are all obvious lemmas, however, they still needs to be derived |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
7103 |
from the concise and miniature model of PIP given in PrioGDef.thy. |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
7104 |
*} |
8142e80f5d58
Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents:
44
diff
changeset
|
7105 |
|
61 | 7106 |
lemma eq_dependants: "dependants (wq s) = dependants s" |
7107 |
by (simp add: s_dependants_abv wq_def) |
|
7108 |
||
7109 |
lemma next_th_unique: |
|
7110 |
assumes nt1: "next_th s th cs th1" |
|
7111 |
and nt2: "next_th s th cs th2" |
|
7112 |
shows "th1 = th2" |
|
7113 |
using assms by (unfold next_th_def, auto) |
|
7114 |
||
63 | 7115 |
lemma birth_time_lt: "s \<noteq> [] \<Longrightarrow> last_set th s < length s" |
7116 |
apply (induct s, simp) |
|
7117 |
proof - |
|
7118 |
fix a s |
|
7119 |
assume ih: "s \<noteq> [] \<Longrightarrow> last_set th s < length s" |
|
7120 |
and eq_as: "a # s \<noteq> []" |
|
7121 |
show "last_set th (a # s) < length (a # s)" |
|
7122 |
proof(cases "s \<noteq> []") |
|
7123 |
case False |
|
7124 |
from False show ?thesis |
|
7125 |
by (cases a, auto simp:last_set.simps) |
|
7126 |
next |
|
7127 |
case True |
|
7128 |
from ih [OF True] show ?thesis |
|
7129 |
by (cases a, auto simp:last_set.simps) |
|
7130 |
qed |
|
7131 |
qed |
|
7132 |
||
7133 |
lemma th_in_ne: "th \<in> threads s \<Longrightarrow> s \<noteq> []" |
|
7134 |
by (induct s, auto simp:threads.simps) |
|
7135 |
||
7136 |
lemma preced_tm_lt: "th \<in> threads s \<Longrightarrow> preced th s = Prc x y \<Longrightarrow> y < length s" |
|
7137 |
apply (drule_tac th_in_ne) |
|
7138 |
by (unfold preced_def, auto intro: birth_time_lt) |
|
7139 |
||
65
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7140 |
lemma inj_the_preced: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7141 |
"inj_on (the_preced s) (threads s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7142 |
by (metis inj_onI preced_unique the_preced_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7143 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7144 |
lemma tRAG_alt_def: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7145 |
"tRAG s = {(Th th1, Th th2) | th1 th2. |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7146 |
\<exists> cs. (Th th1, Cs cs) \<in> RAG s \<and> (Cs cs, Th th2) \<in> RAG s}" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7147 |
by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7148 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7149 |
lemma tRAG_Field: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7150 |
"Field (tRAG s) \<subseteq> Field (RAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7151 |
by (unfold tRAG_alt_def Field_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7152 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7153 |
lemma tRAG_ancestorsE: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7154 |
assumes "x \<in> ancestors (tRAG s) u" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7155 |
obtains th where "x = Th th" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7156 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7157 |
from assms have "(u, x) \<in> (tRAG s)^+" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7158 |
by (unfold ancestors_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7159 |
from tranclE[OF this] obtain c where "(c, x) \<in> tRAG s" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7160 |
then obtain th where "x = Th th" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7161 |
by (unfold tRAG_alt_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7162 |
from that[OF this] show ?thesis . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7163 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7164 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7165 |
lemma tRAG_mono: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7166 |
assumes "RAG s' \<subseteq> RAG s" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7167 |
shows "tRAG s' \<subseteq> tRAG s" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7168 |
using assms |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7169 |
by (unfold tRAG_alt_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7170 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7171 |
lemma holding_next_thI: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7172 |
assumes "holding s th cs" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7173 |
and "length (wq s cs) > 1" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7174 |
obtains th' where "next_th s th cs th'" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7175 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7176 |
from assms(1)[folded eq_holding, unfolded cs_holding_def] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7177 |
have " th \<in> set (wq s cs) \<and> th = hd (wq s cs)" . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7178 |
then obtain rest where h1: "wq s cs = th#rest" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7179 |
by (cases "wq s cs", auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7180 |
with assms(2) have h2: "rest \<noteq> []" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7181 |
let ?th' = "hd (SOME q. distinct q \<and> set q = set rest)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7182 |
have "next_th s th cs ?th'" using h1(1) h2 |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7183 |
by (unfold next_th_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7184 |
from that[OF this] show ?thesis . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7185 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7186 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7187 |
lemma RAG_tRAG_transfer: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7188 |
assumes "vt s'" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7189 |
assumes "RAG s = RAG s' \<union> {(Th th, Cs cs)}" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7190 |
and "(Cs cs, Th th'') \<in> RAG s'" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7191 |
shows "tRAG s = tRAG s' \<union> {(Th th, Th th'')}" (is "?L = ?R") |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7192 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7193 |
interpret vt_s': valid_trace "s'" using assms(1) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7194 |
by (unfold_locales, simp) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7195 |
interpret rtree: rtree "RAG s'" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7196 |
proof |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7197 |
show "single_valued (RAG s')" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7198 |
apply (intro_locales) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7199 |
by (unfold single_valued_def, |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7200 |
auto intro:vt_s'.unique_RAG) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7201 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7202 |
show "acyclic (RAG s')" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7203 |
by (rule vt_s'.acyclic_RAG) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7204 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7205 |
{ fix n1 n2 |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7206 |
assume "(n1, n2) \<in> ?L" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7207 |
from this[unfolded tRAG_alt_def] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7208 |
obtain th1 th2 cs' where |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7209 |
h: "n1 = Th th1" "n2 = Th th2" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7210 |
"(Th th1, Cs cs') \<in> RAG s" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7211 |
"(Cs cs', Th th2) \<in> RAG s" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7212 |
from h(4) and assms(2) have cs_in: "(Cs cs', Th th2) \<in> RAG s'" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7213 |
from h(3) and assms(2) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7214 |
have "(Th th1, Cs cs') = (Th th, Cs cs) \<or> |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7215 |
(Th th1, Cs cs') \<in> RAG s'" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7216 |
hence "(n1, n2) \<in> ?R" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7217 |
proof |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7218 |
assume h1: "(Th th1, Cs cs') = (Th th, Cs cs)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7219 |
hence eq_th1: "th1 = th" by simp |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7220 |
moreover have "th2 = th''" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7221 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7222 |
from h1 have "cs' = cs" by simp |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7223 |
from assms(3) cs_in[unfolded this] rtree.sgv |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7224 |
show ?thesis |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7225 |
by (unfold single_valued_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7226 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7227 |
ultimately show ?thesis using h(1,2) by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7228 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7229 |
assume "(Th th1, Cs cs') \<in> RAG s'" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7230 |
with cs_in have "(Th th1, Th th2) \<in> tRAG s'" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7231 |
by (unfold tRAG_alt_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7232 |
from this[folded h(1, 2)] show ?thesis by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7233 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7234 |
} moreover { |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7235 |
fix n1 n2 |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7236 |
assume "(n1, n2) \<in> ?R" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7237 |
hence "(n1, n2) \<in>tRAG s' \<or> (n1, n2) = (Th th, Th th'')" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7238 |
hence "(n1, n2) \<in> ?L" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7239 |
proof |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7240 |
assume "(n1, n2) \<in> tRAG s'" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7241 |
moreover have "... \<subseteq> ?L" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7242 |
proof(rule tRAG_mono) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7243 |
show "RAG s' \<subseteq> RAG s" by (unfold assms(2), auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7244 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7245 |
ultimately show ?thesis by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7246 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7247 |
assume eq_n: "(n1, n2) = (Th th, Th th'')" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7248 |
from assms(2, 3) have "(Cs cs, Th th'') \<in> RAG s" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7249 |
moreover have "(Th th, Cs cs) \<in> RAG s" using assms(2) by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7250 |
ultimately show ?thesis |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7251 |
by (unfold eq_n tRAG_alt_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7252 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7253 |
} ultimately show ?thesis by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7254 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7255 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7256 |
context valid_trace |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7257 |
begin |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7258 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7259 |
lemmas RAG_tRAG_transfer = RAG_tRAG_transfer[OF vt] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7260 |
|
35
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
32
diff
changeset
|
7261 |
end |
65
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7262 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7263 |
lemma cp_alt_def: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7264 |
"cp s th = |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7265 |
Max ((the_preced s) ` {th'. Th th' \<in> (subtree (RAG s) (Th th))})" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7266 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7267 |
have "Max (the_preced s ` ({th} \<union> dependants (wq s) th)) = |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7268 |
Max (the_preced s ` {th'. Th th' \<in> subtree (RAG s) (Th th)})" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7269 |
(is "Max (_ ` ?L) = Max (_ ` ?R)") |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7270 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7271 |
have "?L = ?R" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7272 |
by (auto dest:rtranclD simp:cs_dependants_def cs_RAG_def s_RAG_def subtree_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7273 |
thus ?thesis by simp |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7274 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7275 |
thus ?thesis by (unfold cp_eq_cpreced cpreced_def, fold the_preced_def, simp) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7276 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7277 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7278 |
lemma cp_gen_alt_def: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7279 |
"cp_gen s = (Max \<circ> (\<lambda>x. (the_preced s \<circ> the_thread) ` subtree (tRAG s) x))" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7280 |
by (auto simp:cp_gen_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7281 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7282 |
lemma tRAG_nodeE: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7283 |
assumes "(n1, n2) \<in> tRAG s" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7284 |
obtains th1 th2 where "n1 = Th th1" "n2 = Th th2" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7285 |
using assms |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7286 |
by (auto simp: tRAG_def wRAG_def hRAG_def tRAG_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7287 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7288 |
lemma subtree_nodeE: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7289 |
assumes "n \<in> subtree (tRAG s) (Th th)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7290 |
obtains th1 where "n = Th th1" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7291 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7292 |
show ?thesis |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7293 |
proof(rule subtreeE[OF assms]) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7294 |
assume "n = Th th" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7295 |
from that[OF this] show ?thesis . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7296 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7297 |
assume "Th th \<in> ancestors (tRAG s) n" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7298 |
hence "(n, Th th) \<in> (tRAG s)^+" by (auto simp:ancestors_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7299 |
hence "\<exists> th1. n = Th th1" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7300 |
proof(induct) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7301 |
case (base y) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7302 |
from tRAG_nodeE[OF this] show ?case by metis |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7303 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7304 |
case (step y z) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7305 |
thus ?case by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7306 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7307 |
with that show ?thesis by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7308 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7309 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7310 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7311 |
lemma tRAG_star_RAG: "(tRAG s)^* \<subseteq> (RAG s)^*" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7312 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7313 |
have "(wRAG s O hRAG s)^* \<subseteq> (RAG s O RAG s)^*" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7314 |
by (rule rtrancl_mono, auto simp:RAG_split) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7315 |
also have "... \<subseteq> ((RAG s)^*)^*" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7316 |
by (rule rtrancl_mono, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7317 |
also have "... = (RAG s)^*" by simp |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7318 |
finally show ?thesis by (unfold tRAG_def, simp) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7319 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7320 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7321 |
lemma tRAG_subtree_RAG: "subtree (tRAG s) x \<subseteq> subtree (RAG s) x" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7322 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7323 |
{ fix a |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7324 |
assume "a \<in> subtree (tRAG s) x" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7325 |
hence "(a, x) \<in> (tRAG s)^*" by (auto simp:subtree_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7326 |
with tRAG_star_RAG[of s] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7327 |
have "(a, x) \<in> (RAG s)^*" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7328 |
hence "a \<in> subtree (RAG s) x" by (auto simp:subtree_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7329 |
} thus ?thesis by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7330 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7331 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7332 |
lemma tRAG_trancl_eq: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7333 |
"{th'. (Th th', Th th) \<in> (tRAG s)^+} = |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7334 |
{th'. (Th th', Th th) \<in> (RAG s)^+}" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7335 |
(is "?L = ?R") |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7336 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7337 |
{ fix th' |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7338 |
assume "th' \<in> ?L" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7339 |
hence "(Th th', Th th) \<in> (tRAG s)^+" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7340 |
from tranclD[OF this] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7341 |
obtain z where h: "(Th th', z) \<in> tRAG s" "(z, Th th) \<in> (tRAG s)\<^sup>*" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7342 |
from tRAG_subtree_RAG[of s] and this(2) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7343 |
have "(z, Th th) \<in> (RAG s)^*" by (meson subsetCE tRAG_star_RAG) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7344 |
moreover from h(1) have "(Th th', z) \<in> (RAG s)^+" using tRAG_alt_def by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7345 |
ultimately have "th' \<in> ?R" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7346 |
} moreover |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7347 |
{ fix th' |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7348 |
assume "th' \<in> ?R" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7349 |
hence "(Th th', Th th) \<in> (RAG s)^+" by (auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7350 |
from plus_rpath[OF this] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7351 |
obtain xs where rp: "rpath (RAG s) (Th th') xs (Th th)" "xs \<noteq> []" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7352 |
hence "(Th th', Th th) \<in> (tRAG s)^+" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7353 |
proof(induct xs arbitrary:th' th rule:length_induct) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7354 |
case (1 xs th' th) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7355 |
then obtain x1 xs1 where Cons1: "xs = x1#xs1" by (cases xs, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7356 |
show ?case |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7357 |
proof(cases "xs1") |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7358 |
case Nil |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7359 |
from 1(2)[unfolded Cons1 Nil] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7360 |
have rp: "rpath (RAG s) (Th th') [x1] (Th th)" . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7361 |
hence "(Th th', x1) \<in> (RAG s)" by (cases, simp) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7362 |
then obtain cs where "x1 = Cs cs" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7363 |
by (unfold s_RAG_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7364 |
from rpath_nnl_lastE[OF rp[unfolded this]] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7365 |
show ?thesis by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7366 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7367 |
case (Cons x2 xs2) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7368 |
from 1(2)[unfolded Cons1[unfolded this]] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7369 |
have rp: "rpath (RAG s) (Th th') (x1 # x2 # xs2) (Th th)" . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7370 |
from rpath_edges_on[OF this] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7371 |
have eds: "edges_on (Th th' # x1 # x2 # xs2) \<subseteq> RAG s" . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7372 |
have "(Th th', x1) \<in> edges_on (Th th' # x1 # x2 # xs2)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7373 |
by (simp add: edges_on_unfold) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7374 |
with eds have rg1: "(Th th', x1) \<in> RAG s" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7375 |
then obtain cs1 where eq_x1: "x1 = Cs cs1" by (unfold s_RAG_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7376 |
have "(x1, x2) \<in> edges_on (Th th' # x1 # x2 # xs2)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7377 |
by (simp add: edges_on_unfold) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7378 |
from this eds |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7379 |
have rg2: "(x1, x2) \<in> RAG s" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7380 |
from this[unfolded eq_x1] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7381 |
obtain th1 where eq_x2: "x2 = Th th1" by (unfold s_RAG_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7382 |
from rg1[unfolded eq_x1] rg2[unfolded eq_x1 eq_x2] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7383 |
have rt1: "(Th th', Th th1) \<in> tRAG s" by (unfold tRAG_alt_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7384 |
from rp have "rpath (RAG s) x2 xs2 (Th th)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7385 |
by (elim rpath_ConsE, simp) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7386 |
from this[unfolded eq_x2] have rp': "rpath (RAG s) (Th th1) xs2 (Th th)" . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7387 |
show ?thesis |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7388 |
proof(cases "xs2 = []") |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7389 |
case True |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7390 |
from rpath_nilE[OF rp'[unfolded this]] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7391 |
have "th1 = th" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7392 |
from rt1[unfolded this] show ?thesis by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7393 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7394 |
case False |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7395 |
from 1(1)[rule_format, OF _ rp' this, unfolded Cons1 Cons] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7396 |
have "(Th th1, Th th) \<in> (tRAG s)\<^sup>+" by simp |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7397 |
with rt1 show ?thesis by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7398 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7399 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7400 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7401 |
hence "th' \<in> ?L" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7402 |
} ultimately show ?thesis by blast |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7403 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7404 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7405 |
lemma tRAG_trancl_eq_Th: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7406 |
"{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7407 |
{Th th' | th'. (Th th', Th th) \<in> (RAG s)^+}" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7408 |
using tRAG_trancl_eq by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7409 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7410 |
lemma dependants_alt_def: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7411 |
"dependants s th = {th'. (Th th', Th th) \<in> (tRAG s)^+}" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7412 |
by (metis eq_RAG s_dependants_def tRAG_trancl_eq) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7413 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7414 |
context valid_trace |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7415 |
begin |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7416 |
|
104 | 7417 |
>>>>>>> other |
65
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7418 |
lemma count_eq_tRAG_plus: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7419 |
assumes "cntP s th = cntV s th" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7420 |
shows "{th'. (Th th', Th th) \<in> (tRAG s)^+} = {}" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7421 |
using assms count_eq_dependants dependants_alt_def eq_dependants by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7422 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7423 |
lemma count_eq_RAG_plus: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7424 |
assumes "cntP s th = cntV s th" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7425 |
shows "{th'. (Th th', Th th) \<in> (RAG s)^+} = {}" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7426 |
using assms count_eq_dependants cs_dependants_def eq_RAG by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7427 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7428 |
lemma count_eq_RAG_plus_Th: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7429 |
assumes "cntP s th = cntV s th" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7430 |
shows "{Th th' | th'. (Th th', Th th) \<in> (RAG s)^+} = {}" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7431 |
using count_eq_RAG_plus[OF assms] by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7432 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7433 |
lemma count_eq_tRAG_plus_Th: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7434 |
assumes "cntP s th = cntV s th" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7435 |
shows "{Th th' | th'. (Th th', Th th) \<in> (tRAG s)^+} = {}" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7436 |
using count_eq_tRAG_plus[OF assms] by auto |
104 | 7437 |
<<<<<<< local |
7438 |
======= |
|
65
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7439 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7440 |
end |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7441 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7442 |
lemma tRAG_subtree_eq: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7443 |
"(subtree (tRAG s) (Th th)) = {Th th' | th'. Th th' \<in> (subtree (RAG s) (Th th))}" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7444 |
(is "?L = ?R") |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7445 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7446 |
{ fix n |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7447 |
assume h: "n \<in> ?L" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7448 |
hence "n \<in> ?R" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7449 |
by (smt mem_Collect_eq subsetCE subtree_def subtree_nodeE tRAG_subtree_RAG) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7450 |
} moreover { |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7451 |
fix n |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7452 |
assume "n \<in> ?R" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7453 |
then obtain th' where h: "n = Th th'" "(Th th', Th th) \<in> (RAG s)^*" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7454 |
by (auto simp:subtree_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7455 |
from rtranclD[OF this(2)] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7456 |
have "n \<in> ?L" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7457 |
proof |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7458 |
assume "Th th' \<noteq> Th th \<and> (Th th', Th th) \<in> (RAG s)\<^sup>+" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7459 |
with h have "n \<in> {Th th' | th'. (Th th', Th th) \<in> (RAG s)^+}" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7460 |
thus ?thesis using subtree_def tRAG_trancl_eq by fastforce |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7461 |
qed (insert h, auto simp:subtree_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7462 |
} ultimately show ?thesis by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7463 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7464 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7465 |
lemma threads_set_eq: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7466 |
"the_thread ` (subtree (tRAG s) (Th th)) = |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7467 |
{th'. Th th' \<in> (subtree (RAG s) (Th th))}" (is "?L = ?R") |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7468 |
by (auto intro:rev_image_eqI simp:tRAG_subtree_eq) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7469 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7470 |
lemma cp_alt_def1: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7471 |
"cp s th = Max ((the_preced s o the_thread) ` (subtree (tRAG s) (Th th)))" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7472 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7473 |
have "(the_preced s ` the_thread ` subtree (tRAG s) (Th th)) = |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7474 |
((the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th))" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7475 |
by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7476 |
thus ?thesis by (unfold cp_alt_def, fold threads_set_eq, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7477 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7478 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7479 |
lemma cp_gen_def_cond: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7480 |
assumes "x = Th th" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7481 |
shows "cp s th = cp_gen s (Th th)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7482 |
by (unfold cp_alt_def1 cp_gen_def, simp) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7483 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7484 |
lemma cp_gen_over_set: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7485 |
assumes "\<forall> x \<in> A. \<exists> th. x = Th th" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7486 |
shows "cp_gen s ` A = (cp s \<circ> the_thread) ` A" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7487 |
proof(rule f_image_eq) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7488 |
fix a |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7489 |
assume "a \<in> A" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7490 |
from assms[rule_format, OF this] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7491 |
obtain th where eq_a: "a = Th th" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7492 |
show "cp_gen s a = (cp s \<circ> the_thread) a" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7493 |
by (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7494 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7495 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7496 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7497 |
context valid_trace |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7498 |
begin |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7499 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7500 |
lemma RAG_threads: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7501 |
assumes "(Th th) \<in> Field (RAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7502 |
shows "th \<in> threads s" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7503 |
using assms |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7504 |
by (metis Field_def UnE dm_RAG_threads range_in vt) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7505 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7506 |
lemma subtree_tRAG_thread: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7507 |
assumes "th \<in> threads s" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7508 |
shows "subtree (tRAG s) (Th th) \<subseteq> Th ` threads s" (is "?L \<subseteq> ?R") |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7509 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7510 |
have "?L = {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7511 |
by (unfold tRAG_subtree_eq, simp) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7512 |
also have "... \<subseteq> ?R" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7513 |
proof |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7514 |
fix x |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7515 |
assume "x \<in> {Th th' |th'. Th th' \<in> subtree (RAG s) (Th th)}" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7516 |
then obtain th' where h: "x = Th th'" "Th th' \<in> subtree (RAG s) (Th th)" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7517 |
from this(2) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7518 |
show "x \<in> ?R" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7519 |
proof(cases rule:subtreeE) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7520 |
case 1 |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7521 |
thus ?thesis by (simp add: assms h(1)) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7522 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7523 |
case 2 |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7524 |
thus ?thesis by (metis ancestors_Field dm_RAG_threads h(1) image_eqI) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7525 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7526 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7527 |
finally show ?thesis . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7528 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7529 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7530 |
lemma readys_root: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7531 |
assumes "th \<in> readys s" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7532 |
shows "root (RAG s) (Th th)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7533 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7534 |
{ fix x |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7535 |
assume "x \<in> ancestors (RAG s) (Th th)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7536 |
hence h: "(Th th, x) \<in> (RAG s)^+" by (auto simp:ancestors_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7537 |
from tranclD[OF this] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7538 |
obtain z where "(Th th, z) \<in> RAG s" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7539 |
with assms(1) have False |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7540 |
apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7541 |
by (fold wq_def, blast) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7542 |
} thus ?thesis by (unfold root_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7543 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7544 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7545 |
lemma readys_in_no_subtree: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7546 |
assumes "th \<in> readys s" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7547 |
and "th' \<noteq> th" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7548 |
shows "Th th \<notin> subtree (RAG s) (Th th')" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7549 |
proof |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7550 |
assume "Th th \<in> subtree (RAG s) (Th th')" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7551 |
thus False |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7552 |
proof(cases rule:subtreeE) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7553 |
case 1 |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7554 |
with assms show ?thesis by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7555 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7556 |
case 2 |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7557 |
with readys_root[OF assms(1)] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7558 |
show ?thesis by (auto simp:root_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7559 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7560 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7561 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7562 |
lemma not_in_thread_isolated: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7563 |
assumes "th \<notin> threads s" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7564 |
shows "(Th th) \<notin> Field (RAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7565 |
proof |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7566 |
assume "(Th th) \<in> Field (RAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7567 |
with dm_RAG_threads and range_in assms |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7568 |
show False by (unfold Field_def, blast) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7569 |
qed |
104 | 7570 |
>>>>>>> other |
65
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7571 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7572 |
lemma wf_RAG: "wf (RAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7573 |
proof(rule finite_acyclic_wf) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7574 |
from finite_RAG show "finite (RAG s)" . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7575 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7576 |
from acyclic_RAG show "acyclic (RAG s)" . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7577 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7578 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7579 |
lemma sgv_wRAG: "single_valued (wRAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7580 |
using waiting_unique |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7581 |
by (unfold single_valued_def wRAG_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7582 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7583 |
lemma sgv_hRAG: "single_valued (hRAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7584 |
using holding_unique |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7585 |
by (unfold single_valued_def hRAG_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7586 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7587 |
lemma sgv_tRAG: "single_valued (tRAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7588 |
by (unfold tRAG_def, rule single_valued_relcomp, |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7589 |
insert sgv_wRAG sgv_hRAG, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7590 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7591 |
lemma acyclic_tRAG: "acyclic (tRAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7592 |
proof(unfold tRAG_def, rule acyclic_compose) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7593 |
show "acyclic (RAG s)" using acyclic_RAG . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7594 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7595 |
show "wRAG s \<subseteq> RAG s" unfolding RAG_split by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7596 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7597 |
show "hRAG s \<subseteq> RAG s" unfolding RAG_split by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7598 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7599 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7600 |
lemma sgv_RAG: "single_valued (RAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7601 |
using unique_RAG by (auto simp:single_valued_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7602 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7603 |
lemma rtree_RAG: "rtree (RAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7604 |
using sgv_RAG acyclic_RAG |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7605 |
by (unfold rtree_def rtree_axioms_def sgv_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7606 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7607 |
end |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7608 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7609 |
sublocale valid_trace < rtree_RAG: rtree "RAG s" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7610 |
proof |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7611 |
show "single_valued (RAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7612 |
apply (intro_locales) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7613 |
by (unfold single_valued_def, |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7614 |
auto intro:unique_RAG) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7615 |
|
104 | 7616 |
<<<<<<< local |
7617 |
lemma detached_test: |
|
7618 |
shows "detached s th = (Th th \<notin> Field (RAG s))" |
|
7619 |
apply(simp add: detached_def Field_def) |
|
7620 |
apply(simp add: s_RAG_def) |
|
7621 |
apply(simp add: s_holding_abv s_waiting_abv) |
|
7622 |
apply(simp add: Domain_iff Range_iff) |
|
7623 |
apply(simp add: wq_def) |
|
7624 |
apply(auto) |
|
7625 |
done |
|
7626 |
======= |
|
65
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7627 |
show "acyclic (RAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7628 |
by (rule acyclic_RAG) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7629 |
qed |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7630 |
|
65
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7631 |
sublocale valid_trace < rtree_s: rtree "tRAG s" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7632 |
proof(unfold_locales) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7633 |
from sgv_tRAG show "single_valued (tRAG s)" . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7634 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7635 |
from acyclic_tRAG show "acyclic (tRAG s)" . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7636 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7637 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7638 |
sublocale valid_trace < fsbtRAGs : fsubtree "RAG s" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7639 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7640 |
show "fsubtree (RAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7641 |
proof(intro_locales) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7642 |
show "fbranch (RAG s)" using finite_fbranchI[OF finite_RAG] . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7643 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7644 |
show "fsubtree_axioms (RAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7645 |
proof(unfold fsubtree_axioms_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7646 |
from wf_RAG show "wf (RAG s)" . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7647 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7648 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7649 |
qed |
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7650 |
|
65
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7651 |
sublocale valid_trace < fsbttRAGs: fsubtree "tRAG s" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7652 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7653 |
have "fsubtree (tRAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7654 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7655 |
have "fbranch (tRAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7656 |
proof(unfold tRAG_def, rule fbranch_compose) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7657 |
show "fbranch (wRAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7658 |
proof(rule finite_fbranchI) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7659 |
from finite_RAG show "finite (wRAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7660 |
by (unfold RAG_split, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7661 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7662 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7663 |
show "fbranch (hRAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7664 |
proof(rule finite_fbranchI) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7665 |
from finite_RAG |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7666 |
show "finite (hRAG s)" by (unfold RAG_split, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7667 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7668 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7669 |
moreover have "wf (tRAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7670 |
proof(rule wf_subset) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7671 |
show "wf (RAG s O RAG s)" using wf_RAG |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7672 |
by (fold wf_comp_self, simp) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7673 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7674 |
show "tRAG s \<subseteq> (RAG s O RAG s)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7675 |
by (unfold tRAG_alt_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7676 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7677 |
ultimately show ?thesis |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7678 |
by (unfold fsubtree_def fsubtree_axioms_def,auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7679 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7680 |
from this[folded tRAG_def] show "fsubtree (tRAG s)" . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7681 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7682 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7683 |
lemma Max_UNION: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7684 |
assumes "finite A" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7685 |
and "A \<noteq> {}" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7686 |
and "\<forall> M \<in> f ` A. finite M" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7687 |
and "\<forall> M \<in> f ` A. M \<noteq> {}" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7688 |
shows "Max (\<Union>x\<in> A. f x) = Max (Max ` f ` A)" (is "?L = ?R") |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7689 |
using assms[simp] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7690 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7691 |
have "?L = Max (\<Union>(f ` A))" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7692 |
by (fold Union_image_eq, simp) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7693 |
also have "... = ?R" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7694 |
by (subst Max_Union, simp+) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7695 |
finally show ?thesis . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7696 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7697 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7698 |
lemma max_Max_eq: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7699 |
assumes "finite A" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7700 |
and "A \<noteq> {}" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7701 |
and "x = y" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7702 |
shows "max x (Max A) = Max ({y} \<union> A)" (is "?L = ?R") |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7703 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7704 |
have "?R = Max (insert y A)" by simp |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7705 |
also from assms have "... = ?L" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7706 |
by (subst Max.insert, simp+) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7707 |
finally show ?thesis by simp |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7708 |
qed |
104 | 7709 |
>>>>>>> other |
65
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7710 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7711 |
context valid_trace |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7712 |
begin |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7713 |
|
104 | 7714 |
<<<<<<< local |
7715 |
lemma detached_intro: |
|
7716 |
assumes eq_pv: "cntP s th = cntV s th" |
|
7717 |
shows "detached s th" |
|
7718 |
proof - |
|
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7719 |
from eq_pv cnp_cnv_cncs |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7720 |
have "th \<in> readys s \<or> th \<notin> threads s" by (auto simp:pvD_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7721 |
thus ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7722 |
proof |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7723 |
assume "th \<notin> threads s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7724 |
with rg_RAG_threads dm_RAG_threads |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7725 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7726 |
by (auto simp add: detached_def s_RAG_def s_waiting_abv |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7727 |
s_holding_abv wq_def Domain_iff Range_iff) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7728 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7729 |
assume "th \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7730 |
moreover have "Th th \<notin> Range (RAG s)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7731 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7732 |
from eq_pv_children[OF assms] |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7733 |
have "children (RAG s) (Th th) = {}" . |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7734 |
thus ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7735 |
by (unfold children_def, auto) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7736 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7737 |
ultimately show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7738 |
by (auto simp add: detached_def s_RAG_def s_waiting_abv |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7739 |
s_holding_abv wq_def readys_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7740 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7741 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7742 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7743 |
lemma detached_elim: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7744 |
assumes dtc: "detached s th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7745 |
shows "cntP s th = cntV s th" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7746 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7747 |
have cncs_z: "cntCS s th = 0" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7748 |
proof - |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7749 |
from dtc have "holdents s th = {}" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7750 |
unfolding detached_def holdents_test s_RAG_def |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7751 |
by (simp add: s_waiting_abv wq_def s_holding_abv Domain_iff Range_iff) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7752 |
thus ?thesis by (auto simp:cntCS_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7753 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7754 |
show ?thesis |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7755 |
proof(cases "th \<in> threads s") |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7756 |
case True |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7757 |
with dtc |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7758 |
have "th \<in> readys s" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7759 |
by (unfold readys_def detached_def Field_def Domain_def Range_def, |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7760 |
auto simp:waiting_eq s_RAG_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7761 |
with cncs_z show ?thesis using cnp_cnv_cncs by (simp add:pvD_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7762 |
next |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7763 |
case False |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7764 |
with cncs_z and cnp_cnv_cncs show ?thesis by (simp add:pvD_def) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7765 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7766 |
qed |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7767 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7768 |
lemma detached_eq: |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7769 |
shows "(detached s th) = (cntP s th = cntV s th)" |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7770 |
by (insert vt, auto intro:detached_intro detached_elim) |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7771 |
|
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7772 |
end |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7773 |
|
103
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
7774 |
section {* Recursive definition of @{term "cp"} *} |
102
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7775 |
|
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7776 |
lemma cp_alt_def1: |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7777 |
"cp s th = Max ((the_preced s o the_thread) ` (subtree (tRAG s) (Th th)))" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7778 |
proof - |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7779 |
have "(the_preced s ` the_thread ` subtree (tRAG s) (Th th)) = |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7780 |
((the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th))" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7781 |
by auto |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7782 |
thus ?thesis by (unfold cp_alt_def, fold threads_set_eq, auto) |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7783 |
qed |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7784 |
|
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7785 |
lemma cp_gen_def_cond: |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7786 |
assumes "x = Th th" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7787 |
shows "cp s th = cp_gen s (Th th)" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7788 |
by (unfold cp_alt_def1 cp_gen_def, simp) |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7789 |
|
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7790 |
lemma cp_gen_over_set: |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7791 |
assumes "\<forall> x \<in> A. \<exists> th. x = Th th" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7792 |
shows "cp_gen s ` A = (cp s \<circ> the_thread) ` A" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7793 |
proof(rule f_image_eq) |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7794 |
fix a |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7795 |
assume "a \<in> A" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7796 |
from assms[rule_format, OF this] |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7797 |
obtain th where eq_a: "a = Th th" by auto |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7798 |
show "cp_gen s a = (cp s \<circ> the_thread) a" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7799 |
by (unfold eq_a, simp, unfold cp_gen_def_cond[OF refl[of "Th th"]], simp) |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7800 |
qed |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7801 |
|
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7802 |
|
92
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7803 |
context valid_trace |
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
zhangx
parents:
90
diff
changeset
|
7804 |
begin |
104 | 7805 |
======= |
7806 |
>>>>>>> other |
|
65
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7807 |
(* ddd *) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7808 |
lemma cp_gen_rec: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7809 |
assumes "x = Th th" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7810 |
shows "cp_gen s x = Max ({the_preced s th} \<union> (cp_gen s) ` children (tRAG s) x)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7811 |
proof(cases "children (tRAG s) x = {}") |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7812 |
case True |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7813 |
show ?thesis |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7814 |
by (unfold True cp_gen_def subtree_children, simp add:assms) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7815 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7816 |
case False |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7817 |
hence [simp]: "children (tRAG s) x \<noteq> {}" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7818 |
note fsbttRAGs.finite_subtree[simp] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7819 |
have [simp]: "finite (children (tRAG s) x)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7820 |
by (intro rev_finite_subset[OF fsbttRAGs.finite_subtree], |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7821 |
rule children_subtree) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7822 |
{ fix r x |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7823 |
have "subtree r x \<noteq> {}" by (auto simp:subtree_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7824 |
} note this[simp] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7825 |
have [simp]: "\<exists>x\<in>children (tRAG s) x. subtree (tRAG s) x \<noteq> {}" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7826 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7827 |
from False obtain q where "q \<in> children (tRAG s) x" by blast |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7828 |
moreover have "subtree (tRAG s) q \<noteq> {}" by simp |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7829 |
ultimately show ?thesis by blast |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7830 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7831 |
have h: "Max ((the_preced s \<circ> the_thread) ` |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7832 |
({x} \<union> \<Union>(subtree (tRAG s) ` children (tRAG s) x))) = |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7833 |
Max ({the_preced s th} \<union> cp_gen s ` children (tRAG s) x)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7834 |
(is "?L = ?R") |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7835 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7836 |
let "Max (?f ` (?A \<union> \<Union> (?g ` ?B)))" = ?L |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7837 |
let "Max (_ \<union> (?h ` ?B))" = ?R |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7838 |
let ?L1 = "?f ` \<Union>(?g ` ?B)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7839 |
have eq_Max_L1: "Max ?L1 = Max (?h ` ?B)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7840 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7841 |
have "?L1 = ?f ` (\<Union> x \<in> ?B.(?g x))" by simp |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7842 |
also have "... = (\<Union> x \<in> ?B. ?f ` (?g x))" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7843 |
finally have "Max ?L1 = Max ..." by simp |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7844 |
also have "... = Max (Max ` (\<lambda>x. ?f ` subtree (tRAG s) x) ` ?B)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7845 |
by (subst Max_UNION, simp+) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7846 |
also have "... = Max (cp_gen s ` children (tRAG s) x)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7847 |
by (unfold image_comp cp_gen_alt_def, simp) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7848 |
finally show ?thesis . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7849 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7850 |
show ?thesis |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7851 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7852 |
have "?L = Max (?f ` ?A \<union> ?L1)" by simp |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7853 |
also have "... = max (the_preced s (the_thread x)) (Max ?L1)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7854 |
by (subst Max_Un, simp+) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7855 |
also have "... = max (?f x) (Max (?h ` ?B))" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7856 |
by (unfold eq_Max_L1, simp) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7857 |
also have "... =?R" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7858 |
by (rule max_Max_eq, (simp)+, unfold assms, simp) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7859 |
finally show ?thesis . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7860 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7861 |
qed thus ?thesis |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7862 |
by (fold h subtree_children, unfold cp_gen_def, simp) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7863 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7864 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7865 |
lemma cp_rec: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7866 |
"cp s th = Max ({the_preced s th} \<union> |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7867 |
(cp s o the_thread) ` children (tRAG s) (Th th))" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7868 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7869 |
have "Th th = Th th" by simp |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7870 |
note h = cp_gen_def_cond[OF this] cp_gen_rec[OF this] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7871 |
show ?thesis |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7872 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7873 |
have "cp_gen s ` children (tRAG s) (Th th) = |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7874 |
(cp s \<circ> the_thread) ` children (tRAG s) (Th th)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7875 |
proof(rule cp_gen_over_set) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7876 |
show " \<forall>x\<in>children (tRAG s) (Th th). \<exists>th. x = Th th" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7877 |
by (unfold tRAG_alt_def, auto simp:children_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7878 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7879 |
thus ?thesis by (subst (1) h(1), unfold h(2), simp) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7880 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7881 |
qed |
101 | 7882 |
end |
7883 |
||
103
d5e9653fbf19
PIPBasics.thy reorganized into sections. The structure is now stable. Next step is to put in more comments.
zhangx
parents:
102
diff
changeset
|
7884 |
section {* Other properties useful in Implementation.thy or Correctness.thy *} |
101 | 7885 |
|
7886 |
context valid_trace_e |
|
7887 |
begin |
|
7888 |
||
7889 |
lemma actor_inv: |
|
7890 |
assumes "\<not> isCreate e" |
|
7891 |
shows "actor e \<in> runing s" |
|
7892 |
using pip_e assms |
|
7893 |
by (induct, auto) |
|
65
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7894 |
end |
80
17305a85493d
CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents:
77
diff
changeset
|
7895 |
|
102
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7896 |
context valid_trace |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7897 |
begin |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7898 |
|
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7899 |
lemma readys_root: |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7900 |
assumes "th \<in> readys s" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7901 |
shows "root (RAG s) (Th th)" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7902 |
proof - |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7903 |
{ fix x |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7904 |
assume "x \<in> ancestors (RAG s) (Th th)" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7905 |
hence h: "(Th th, x) \<in> (RAG s)^+" by (auto simp:ancestors_def) |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7906 |
from tranclD[OF this] |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7907 |
obtain z where "(Th th, z) \<in> RAG s" by auto |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7908 |
with assms(1) have False |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7909 |
apply (case_tac z, auto simp:readys_def s_RAG_def s_waiting_def cs_waiting_def) |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7910 |
by (fold wq_def, blast) |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7911 |
} thus ?thesis by (unfold root_def, auto) |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7912 |
qed |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7913 |
|
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7914 |
lemma readys_in_no_subtree: |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7915 |
assumes "th \<in> readys s" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7916 |
and "th' \<noteq> th" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7917 |
shows "Th th \<notin> subtree (RAG s) (Th th')" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7918 |
proof |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7919 |
assume "Th th \<in> subtree (RAG s) (Th th')" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7920 |
thus False |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7921 |
proof(cases rule:subtreeE) |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7922 |
case 1 |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7923 |
with assms show ?thesis by auto |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7924 |
next |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7925 |
case 2 |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7926 |
with readys_root[OF assms(1)] |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7927 |
show ?thesis by (auto simp:root_def) |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7928 |
qed |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7929 |
qed |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7930 |
|
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7931 |
lemma not_in_thread_isolated: |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7932 |
assumes "th \<notin> threads s" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7933 |
shows "(Th th) \<notin> Field (RAG s)" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7934 |
proof |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7935 |
assume "(Th th) \<in> Field (RAG s)" |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7936 |
with dm_RAG_threads and rg_RAG_threads assms |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7937 |
show False by (unfold Field_def, blast) |
3a801bbd2687
Reorganizing PIPBasics.thy and making small changes to Implementation.thy and Correctness.thy.
zhangx
parents:
101
diff
changeset
|
7938 |
qed |
65
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7939 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7940 |
end |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7941 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7942 |
(* keep *) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7943 |
lemma next_th_holding: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7944 |
assumes vt: "vt s" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7945 |
and nxt: "next_th s th cs th'" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7946 |
shows "holding (wq s) th cs" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7947 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7948 |
from nxt[unfolded next_th_def] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7949 |
obtain rest where h: "wq s cs = th # rest" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7950 |
"rest \<noteq> []" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7951 |
"th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7952 |
thus ?thesis |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7953 |
by (unfold cs_holding_def, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7954 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7955 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7956 |
context valid_trace |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7957 |
begin |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7958 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7959 |
lemma next_th_waiting: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7960 |
assumes nxt: "next_th s th cs th'" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7961 |
shows "waiting (wq s) th' cs" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7962 |
proof - |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7963 |
from nxt[unfolded next_th_def] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7964 |
obtain rest where h: "wq s cs = th # rest" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7965 |
"rest \<noteq> []" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7966 |
"th' = hd (SOME q. distinct q \<and> set q = set rest)" by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7967 |
from wq_distinct[of cs, unfolded h] |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7968 |
have dst: "distinct (th # rest)" . |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7969 |
have in_rest: "th' \<in> set rest" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7970 |
proof(unfold h, rule someI2) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7971 |
show "distinct rest \<and> set rest = set rest" using dst by auto |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7972 |
next |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7973 |
fix x assume "distinct x \<and> set x = set rest" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7974 |
with h(2) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7975 |
show "hd x \<in> set (rest)" by (cases x, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7976 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7977 |
hence "th' \<in> set (wq s cs)" by (unfold h(1), auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7978 |
moreover have "th' \<noteq> hd (wq s cs)" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7979 |
by (unfold h(1), insert in_rest dst, auto) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7980 |
ultimately show ?thesis by (auto simp:cs_waiting_def) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7981 |
qed |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7982 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7983 |
lemma next_th_RAG: |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7984 |
assumes nxt: "next_th (s::event list) th cs th'" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7985 |
shows "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7986 |
using vt assms next_th_holding next_th_waiting |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7987 |
by (unfold s_RAG_def, simp) |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7988 |
|
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7989 |
end |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7990 |
|
104 | 7991 |
<<<<<<< local |
7992 |
end======= |
|
65
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7993 |
-- {* A useless definition *} |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7994 |
definition cps:: "state \<Rightarrow> (thread \<times> precedence) set" |
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7995 |
where "cps s = {(th, cp s th) | th . th \<in> threads s}" |
77
d37703e0c5c4
CpsG.thy updated. It is a copy of PIPBasics.thy under drastic improvement.
zhangx
parents:
69
diff
changeset
|
7996 |
|
65
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents:
64
diff
changeset
|
7997 |
end |
104 | 7998 |
>>>>>>> other |