| author | Christian Urban <urbanc@in.tum.de> |
| Fri, 23 Jun 2017 00:27:16 +0100 | |
| changeset 178 | da27bece9575 |
| parent 177 | abe117821c32 |
| child 179 | f9e6c4166476 |
| permissions | -rw-r--r-- |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1 |
(*<*) |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2 |
theory Paper |
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
3 |
imports "../Implementation" |
|
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
4 |
"../Correctness" |
|
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
5 |
"~~/src/HOL/Library/LaTeXsugar" |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6 |
begin |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7 |
|
|
126
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
8 |
ML {* Scan.succeed *}
|
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
9 |
|
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
10 |
ML {*
|
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
11 |
fun strip_quants ctxt trm = |
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
12 |
case trm of |
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
13 |
Const("HOL.Trueprop", _) $ t => strip_quants ctxt t
|
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
14 |
| Const("Pure.imp", _) $ _ $ t => strip_quants ctxt t
|
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
15 |
| Const("Pure.all", _) $ Abs(n, T, t) =>
|
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
16 |
strip_quants ctxt (subst_bound (Free (n, T), t)) |
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
17 |
| Const("HOL.All", _) $ Abs(n, T, t) =>
|
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
18 |
strip_quants ctxt (subst_bound (Free (n, T), t)) |
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
19 |
| Const("HOL.Ex", _) $ Abs(n, T, t) =>
|
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
20 |
strip_quants ctxt (subst_bound (Free (n, T), t)) |
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
21 |
| _ => trm |
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
22 |
*} |
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
23 |
|
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
24 |
|
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
25 |
setup {* Term_Style.setup @{binding "no_quants"} (Scan.succeed strip_quants) *}
|
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
26 |
|
|
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
27 |
|
|
20
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
28 |
declare [[show_question_marks = false]] |
|
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
29 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
30 |
notation (latex output) |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
31 |
Cons ("_::_" [78,77] 73) 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:
33
diff
changeset
|
32 |
If ("(\<^raw:\textrm{>if\<^raw:}> (_)/ \<^raw:\textrm{>then\<^raw:}> (_)/ \<^raw:\textrm{>else\<^raw:}> (_))" 10) and
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
33 |
vt ("valid'_state") and
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
34 |
Prc ("'(_, _')") and
|
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
35 |
holding_raw ("holds") and
|
|
130
0f124691c191
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
127
diff
changeset
|
36 |
holding ("holds") and
|
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
37 |
waiting_raw ("waits") and
|
|
130
0f124691c191
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
127
diff
changeset
|
38 |
waiting ("waits") and
|
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
39 |
dependants_raw ("dependants") and
|
|
130
0f124691c191
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
127
diff
changeset
|
40 |
dependants ("dependants") and
|
|
0f124691c191
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
127
diff
changeset
|
41 |
RAG_raw ("RAG") and
|
|
0f124691c191
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
127
diff
changeset
|
42 |
RAG ("RAG") and
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
43 |
Th ("T") and
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
44 |
Cs ("C") and
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
45 |
readys ("ready") and
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
46 |
preced ("prec") and
|
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
47 |
preceds ("precs") and
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
48 |
cpreced ("cprec") and
|
| 166 | 49 |
cpreceds ("cprecs") and
|
| 178 | 50 |
(*wq_fun ("wq") and*)
|
| 166 | 51 |
cp ("cprec") and
|
52 |
(*cprec_fun ("cp_fun") and*)
|
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
53 |
holdents ("resources") and
|
|
76
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
54 |
DUMMY ("\<^raw:\mbox{$\_\!\_$}>") and
|
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
55 |
cntP ("c\<^bsub>P\<^esub>") and
|
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
56 |
cntV ("c\<^bsub>V\<^esub>")
|
|
126
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
57 |
|
|
28
7fa738a9615a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
27
diff
changeset
|
58 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
59 |
(*>*) |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
60 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
61 |
section {* Introduction *}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
62 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
63 |
text {*
|
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
64 |
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
65 |
Many real-time systems need to support threads involving priorities |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
66 |
and locking of resources. Locking of resources ensures mutual |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
67 |
exclusion when accessing shared data or devices that cannot be |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
68 |
preempted. Priorities allow scheduling of threads that need to |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
69 |
finish their work within deadlines. Unfortunately, both features |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
70 |
can interact in subtle ways leading to a problem, called |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
71 |
\emph{Priority Inversion}. Suppose three threads having priorities
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
72 |
$H$(igh), $M$(edium) and $L$(ow). We would expect that the thread |
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
73 |
$H$ blocks any other thread with lower priority and the thread |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
74 |
itself cannot be blocked indefinitely by threads with lower |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
75 |
priority. Alas, in a naive implementation of resource locking and |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
76 |
priorities this property can be violated. For this let $L$ be in the |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
77 |
possession of a lock for a resource that $H$ also needs. $H$ must |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
78 |
therefore wait for $L$ to exit the critical section and release this |
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
79 |
lock. The problem is that $L$ might in turn be blocked by any thread |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
80 |
with priority $M$, and so $H$ sits there potentially waiting |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
81 |
indefinitely. Since $H$ is blocked by threads with lower priorities, |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
82 |
the problem is called Priority Inversion. It was first described in |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
83 |
\cite{Lampson80} in the context of the Mesa programming language
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
84 |
designed for concurrent programming. |
|
22
9f0b78fcc894
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
20
diff
changeset
|
85 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
86 |
If the problem of Priority Inversion is ignored, real-time systems |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
87 |
can become unpredictable and resulting bugs can be hard to diagnose. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
88 |
The classic example where this happened is the software that |
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
89 |
controlled the Mars Pathfinder mission in 1997 \cite{Reeves98}. On
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
90 |
Earth the software run mostly without any problem, but once the |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
91 |
spacecraft landed on Mars, it shut down at irregular, but frequent, |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
92 |
intervals leading to loss of project time as normal operation of the |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
93 |
craft could only resume the next day (the mission and data already |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
94 |
collected were fortunately not lost, because of a clever system |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
95 |
design). The reason for the shutdowns was that the scheduling |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
96 |
software fell victim to Priority Inversion: a low priority thread |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
97 |
locking a resource prevented a high priority thread from running in |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
98 |
time, leading to a system reset. Once the problem was found, it was |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
99 |
rectified by enabling the \emph{Priority Inheritance Protocol} (PIP)
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
100 |
\cite{Sha90}\footnote{Sha et al.~call it the \emph{Basic Priority
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
101 |
Inheritance Protocol} \cite{Sha90} and others sometimes also call it
|
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
102 |
\emph{Priority Boosting}, \emph{Priority Donation} or \emph{Priority
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
103 |
Lending}.} in the scheduling software. |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
104 |
|
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
105 |
The idea behind PIP is to let the thread $L$ temporarily inherit the |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
106 |
high priority from $H$ until $L$ leaves the critical section |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
107 |
unlocking the resource. This solves the problem of $H$ having to |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
108 |
wait indefinitely, because $L$ cannot be blocked by threads having |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
109 |
priority $M$. While a few other solutions exist for the Priority |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
110 |
Inversion problem, PIP is one that is widely deployed and |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
111 |
implemented. This includes VxWorks (a proprietary real-time OS used |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
112 |
in the Mars Pathfinder mission, in Boeing's 787 Dreamliner, Honda's |
|
22
9f0b78fcc894
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
20
diff
changeset
|
113 |
ASIMO robot, etc.) and ThreadX (another proprietary real-time OS |
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
114 |
used in nearly all HP inkjet printers \cite{ThreadX}), but also the
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
115 |
POSIX 1003.1c Standard realised for example in libraries for |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
116 |
FreeBSD, Solaris and Linux. |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
117 |
|
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
118 |
Two advantages of PIP are that it is deterministic and that |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
119 |
increasing the priority of a thread can be performed dynamically by |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
120 |
the scheduler. This is in contrast to \emph{Priority Ceiling}
|
|
20
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
121 |
\cite{Sha90}, another solution to the Priority Inversion problem,
|
|
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
122 |
which requires static analysis of the program in order to prevent |
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
123 |
Priority Inversion, and also in contrast to the approach taken in |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
124 |
the Windows NT scheduler, which avoids this problem by randomly |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
125 |
boosting the priority of ready low-priority threads (see for |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
126 |
instance~\cite{WINDOWSNT}). However, there has also been strong
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
127 |
criticism against PIP. For instance, PIP cannot prevent deadlocks |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
128 |
when lock dependencies are circular, and also blocking times can be |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
129 |
substantial (more than just the duration of a critical section). |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
130 |
Though, most criticism against PIP centres around unreliable |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
131 |
implementations and PIP being too complicated and too inefficient. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
132 |
For example, Yodaiken writes in \cite{Yodaiken02}:
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
133 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
134 |
\begin{quote}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
135 |
\it{}``Priority inheritance is neither efficient nor reliable. Implementations
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
136 |
are either incomplete (and unreliable) or surprisingly complex and intrusive.'' |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
137 |
\end{quote}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
138 |
|
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
139 |
\noindent He suggests avoiding PIP altogether by designing the |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
140 |
system so that no priority inversion may happen in the first |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
141 |
place. However, such ideal designs may not always be achievable in |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
142 |
practice. |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
143 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
144 |
In our opinion, there is clearly a need for investigating correct |
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
145 |
algorithms for PIP. A few specifications for PIP exist (in informal |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
146 |
English) and also a few high-level descriptions of implementations |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
147 |
(e.g.~in the textbooks \cite[Section 12.3.1]{Liu00} and
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
148 |
\cite[Section 5.6.5]{Vahalia96}), but they help little with actual
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
149 |
implementations. That this is a problem in practice is proved by an |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
150 |
email by Baker, who wrote on 13 July 2009 on the Linux Kernel |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
151 |
mailing list: |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
152 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
153 |
\begin{quote}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
154 |
\it{}``I observed in the kernel code (to my disgust), the Linux PIP
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
155 |
implementation is a nightmare: extremely heavy weight, involving |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
156 |
maintenance of a full wait-for graph, and requiring updates for a |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
157 |
range of events, including priority changes and interruptions of |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
158 |
wait operations.'' |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
159 |
\end{quote}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
160 |
|
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
161 |
\noindent The criticism by Yodaiken, Baker and others suggests |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
162 |
another look at PIP from a more abstract level (but still concrete |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
163 |
enough to inform an implementation), and makes PIP a good candidate |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
164 |
for a formal verification. An additional reason is that the original |
|
43
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
165 |
specification of PIP~\cite{Sha90}, despite being informally
|
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
166 |
``proved'' correct, is actually \emph{flawed}.
|
|
43
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
167 |
|
|
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
168 |
Yodaiken \cite{Yodaiken02} and also Moylan et
|
|
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
169 |
al.~\cite{deinheritance} point to a subtlety that had been
|
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
170 |
overlooked in the informal proof by Sha et al. They specify PIP in |
| 164 | 171 |
\cite[Section III]{Sha90} so that after the thread (whose priority has been
|
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
172 |
raised) completes its critical section and releases the lock, it |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
173 |
``{\it returns to its original priority level}''. This leads them to
|
|
43
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
174 |
believe that an implementation of PIP is ``{\it rather
|
|
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
175 |
straightforward}''~\cite{Sha90}. Unfortunately, as Yodaiken and
|
|
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
176 |
Moylan et al.~point out, this behaviour is too simplistic. Moylan et |
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
177 |
al.~write that there are ``{\it some hidden
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
178 |
traps}''~\cite{deinheritance}. Consider the case where the low
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
179 |
priority thread $L$ locks \emph{two} resources, and two
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
180 |
high-priority threads $H$ and $H'$ each wait for one of them. If |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
181 |
$L$ releases one resource so that $H$, say, can proceed, then we |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
182 |
still have Priority Inversion with $H'$ (which waits for the other |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
183 |
resource). The correct behaviour for $L$ is to switch to the highest |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
184 |
remaining priority of the threads that it blocks. A similar error |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
185 |
is made in the textbook \cite[Section 2.3.1]{book} which specifies
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
186 |
for a process that inherited a higher priority and exits a critical |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
187 |
section ``{\it it resumes the priority it had at the point of entry
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
188 |
into the critical section}''. This error can also be found in the |
|
134
8a13b37b4d95
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
130
diff
changeset
|
189 |
textbook \cite[Section 16.4.1]{LiYao03} where the authors write
|
|
135
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
190 |
about this process: ``{\it its priority is immediately lowered to the level originally assigned}'';
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
191 |
and also in the |
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
192 |
more recent textbook \cite[Page 119]{Laplante11} where the authors
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
193 |
state: ``{\it when [the task] exits the critical section that caused
|
|
43
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
194 |
the block, it reverts to the priority it had when it entered that |
|
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
195 |
section}''. The textbook \cite[Page 286]{Liu00} contains a simlar
|
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
196 |
flawed specification and even goes on to develop pseudo-code based |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
197 |
on this flawed specification. Accordingly, the operating system |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
198 |
primitives for inheritance and restoration of priorities in |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
199 |
\cite{Liu00} depend on maintaining a data structure called
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
200 |
\emph{inheritance log}. This log is maintained for every thread and
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
201 |
broadly specified as containing ``{\it [h]istorical information on
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
202 |
how the thread inherited its current priority}'' \cite[Page |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
203 |
527]{Liu00}. Unfortunately, the important information about actually
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
204 |
computing the priority to be restored solely from this log is not |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
205 |
explained in \cite{Liu00} but left as an ``{\it excercise}'' to the
|
|
135
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
206 |
reader. As we shall see, a correct version of PIP does not need to |
|
75
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
207 |
maintain this (potentially expensive) data structure at |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
208 |
all. Surprisingly also the widely read and frequently updated |
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
209 |
textbook \cite{Silberschatz13} gives the wrong specification. For
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
210 |
example on Page 254 the authors write: ``{\it Upon releasing the
|
|
2aa37de77f31
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
70
diff
changeset
|
211 |
lock, the [low-priority] thread will revert to its original |
|
135
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
212 |
priority.}'' The same error is also repeated later in this popular textbook. |
|
43
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
213 |
|
|
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
214 |
|
|
134
8a13b37b4d95
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
130
diff
changeset
|
215 |
While \cite{Laplante11,LiYao03,Liu00,book,Sha90,Silberschatz13} are the only
|
|
76
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
216 |
formal publications we have found that specify the incorrect |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
217 |
behaviour, it seems also many informal descriptions of PIP overlook |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
218 |
the possibility that another high-priority might wait for a |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
219 |
low-priority process to finish. A notable exception is the texbook |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
220 |
\cite{buttazzo}, which gives the correct behaviour of resetting the
|
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
221 |
priority of a thread to the highest remaining priority of the |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
222 |
threads it blocks. This textbook also gives an informal proof for |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
223 |
the correctness of PIP in the style of Sha et al. Unfortunately, |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
224 |
this informal proof is too vague to be useful for formalising the |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
225 |
correctness of PIP and the specification leaves out nearly all |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
226 |
details in order to implement PIP efficiently.\medskip\smallskip |
|
43
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
227 |
% |
|
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
228 |
%The advantage of formalising the |
|
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
229 |
%correctness of a high-level specification of PIP in a theorem prover |
|
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
230 |
%is that such issues clearly show up and cannot be overlooked as in |
|
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
231 |
%informal reasoning (since we have to analyse all possible behaviours |
|
45e1d324c493
a few additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
41
diff
changeset
|
232 |
%of threads, i.e.~\emph{traces}, that could possibly happen).
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
233 |
|
|
76
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
234 |
\noindent {\bf Contributions:} There have been earlier formal
|
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
235 |
investigations into PIP \cite{Faria08,Jahier09,Wellings07}, but they
|
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
236 |
employ model checking techniques. This paper presents a formalised |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
237 |
and mechanically checked proof for the correctness of PIP. For this |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
238 |
we needed to design a new correctness criterion for PIP. In contrast |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
239 |
to model checking, our formalisation provides insight into why PIP |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
240 |
is correct and allows us to prove stronger properties that, as we |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
241 |
will show, can help with an efficient implementation of PIP. We |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
242 |
illustrate this with an implementation of PIP in the educational |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
243 |
PINTOS operating system \cite{PINTOS}. For example, we found by
|
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
244 |
``playing'' with the formalisation that the choice of the next |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
245 |
thread to take over a lock when a resource is released is irrelevant |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
246 |
for PIP being correct---a fact that has not been mentioned in the |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
247 |
literature and not been used in the reference implementation of PIP |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
248 |
in PINTOS. This fact, however, is important for an efficient |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
249 |
implementation of PIP, because we can give the lock to the thread |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
250 |
with the highest priority so that it terminates more quickly. We |
|
135
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
251 |
are also being able to generalise the scheduler of Sha et |
|
76
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
252 |
al.~\cite{Sha90} to the practically relevant case where critical
|
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
253 |
sections can overlap; see Figure~\ref{overlap} \emph{a)} below for
|
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
254 |
an example of this restriction. In the existing literature there is |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
255 |
no proof and also no proving method that cover this generalised |
|
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
256 |
case. |
|
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
257 |
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
258 |
\begin{figure}
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
259 |
\begin{center}
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
260 |
\begin{tikzpicture}[scale=1]
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
261 |
%%\draw[step=2mm] (0,0) grid (10,2); |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
262 |
\draw [->,line width=0.6mm] (0,0) -- (10,0); |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
263 |
\draw [->,line width=0.6mm] (0,1.5) -- (10,1.5); |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
264 |
\draw [line width=0.6mm, pattern=horizontal lines] (0.8,0) rectangle (4,0.5); |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
265 |
\draw [line width=0.6mm, pattern=north east lines] (3.0,0) rectangle (6,0.5); |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
266 |
\draw [line width=0.6mm, pattern=vertical lines] (5.0,0) rectangle (9,0.5); |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
267 |
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
268 |
\draw [line width=0.6mm, pattern=horizontal lines] (0.6,1.5) rectangle (4.0,2); |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
269 |
\draw [line width=0.6mm, pattern=north east lines] (1.0,1.5) rectangle (3.4,2); |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
270 |
\draw [line width=0.6mm, pattern=vertical lines] (5.0,1.5) rectangle (8.8,2); |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
271 |
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
272 |
\node at (0.8,-0.3) {@{term "P\<^sub>1"}};
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
273 |
\node at (3.0,-0.3) {@{term "P\<^sub>2"}};
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
274 |
\node at (4.0,-0.3) {@{term "V\<^sub>1"}};
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
275 |
\node at (5.0,-0.3) {@{term "P\<^sub>3"}};
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
276 |
\node at (6.0,-0.3) {@{term "V\<^sub>2"}};
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
277 |
\node at (9.0,-0.3) {@{term "V\<^sub>3"}};
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
278 |
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
279 |
\node at (0.6,1.2) {@{term "P\<^sub>1"}};
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
280 |
\node at (1.0,1.2) {@{term "P\<^sub>2"}};
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
281 |
\node at (3.4,1.2) {@{term "V\<^sub>2"}};
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
282 |
\node at (4.0,1.2) {@{term "V\<^sub>1"}};
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
283 |
\node at (5.0,1.2) {@{term "P\<^sub>3"}};
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
284 |
\node at (8.8,1.2) {@{term "V\<^sub>3"}};
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
285 |
\node at (10.3,0) {$t$};
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
286 |
\node at (10.3,1.5) {$t$};
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
287 |
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
288 |
\node at (-0.3,0.2) {$b)$};
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
289 |
\node at (-0.3,1.7) {$a)$};
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
290 |
\end{tikzpicture}\mbox{}\\[-10mm]\mbox{}
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
291 |
\end{center}
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
292 |
\caption{Assume a process is over time locking and unlocking, say, three resources.
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
293 |
The locking requests are labelled @{term "P\<^sub>1"}, @{term "P\<^sub>2"}, and @{term "P\<^sub>3"}
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
294 |
respectively, and the corresponding unlocking operations are labelled |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
295 |
@{term "V\<^sub>1"}, @{term "V\<^sub>2"}, and @{term "V\<^sub>3"}.
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
296 |
Then graph $a)$ shows \emph{properly nested} critical sections as required
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
297 |
by Sha et al.~\cite{Sha90} in their proof---the sections must either be contained within
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
298 |
each other |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
299 |
(the section @{term "P\<^sub>2"}--@{term "V\<^sub>2"} is contained in @{term "P\<^sub>1"}--@{term "V\<^sub>1"}) or
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
300 |
be independent (@{term "P\<^sub>3"}--@{term "V\<^sub>3"} is independent from the other
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
301 |
two). Graph $b)$ shows the general case where |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
302 |
the locking and unlocking of different critical sections can |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
303 |
overlap.\label{overlap}}
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
304 |
\end{figure}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
305 |
*} |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
306 |
|
|
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
307 |
section {* Formal Model of the Priority Inheritance Protocol\label{model} *}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
308 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
309 |
text {*
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
310 |
The Priority Inheritance Protocol, short PIP, is a scheduling |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
311 |
algorithm for a single-processor system.\footnote{We shall come back
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
312 |
later to the case of PIP on multi-processor systems.} |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
313 |
Following good experience in earlier work \cite{Wang09},
|
|
143
c5a65d98191a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
314 |
our model of PIP is based on Paulson's inductive approach for protocol |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
315 |
verification \cite{Paulson98}. In this approach a \emph{state} of a system is
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
316 |
given by a list of events that happened so far (with new events prepended to the list). |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
317 |
\emph{Events} of PIP fall
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
318 |
into five categories defined as the datatype: |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
319 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
320 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
321 |
\mbox{\begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{7mm}}l}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
322 |
\isacommand{datatype} event
|
|
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
323 |
& @{text "="} & @{term "Create thread priority\<iota>"}\\
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
324 |
& @{text "|"} & @{term "Exit thread"} \\
|
|
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
325 |
& @{text "|"} & @{term "Set thread priority\<iota>"} & {\rm reset of the priority for} @{text thread}\\
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
326 |
& @{text "|"} & @{term "P thread cs"} & {\rm request of resource} @{text "cs"} {\rm by} @{text "thread"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
327 |
& @{text "|"} & @{term "V thread cs"} & {\rm release of resource} @{text "cs"} {\rm by} @{text "thread"}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
328 |
\end{tabular}}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
329 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
330 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
331 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
332 |
whereby threads, priorities and (critical) resources are represented |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
333 |
as natural numbers. The event @{term Set} models the situation that
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
334 |
a thread obtains a new priority given by the programmer or |
|
135
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
335 |
user (for example via the {\tt nice} utility under UNIX). For states
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
336 |
we define the following type-synonym: |
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
337 |
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
338 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
339 |
\isacommand{type\_synonym} @{text "state = event list"}
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
340 |
\end{isabelle}
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
341 |
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
342 |
\noindent As in Paulson's work, we need to define functions that |
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
343 |
allow us to make some observations about states. One function, |
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
344 |
called @{term threads}, calculates the set of ``live'' threads that
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
345 |
we have seen so far in a state: |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
346 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
347 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
348 |
\mbox{\begin{tabular}{lcl}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
349 |
@{thm (lhs) threads.simps(1)} & @{text "\<equiv>"} &
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
350 |
@{thm (rhs) threads.simps(1)}\\
|
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
351 |
@{thm (lhs) threads.simps(2)} & @{text "\<equiv>"} &
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
352 |
@{thm (rhs) threads.simps(2)}\\
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
353 |
@{thm (lhs) threads.simps(3)} & @{text "\<equiv>"} &
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
354 |
@{thm (rhs) threads.simps(3)}\\
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
355 |
@{term "threads (DUMMY#s)"} & @{text "\<equiv>"} & @{term "threads s"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
356 |
\end{tabular}}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
357 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
358 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
359 |
\noindent |
|
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
360 |
In this definition @{term "DUMMY # DUMMY"} stands for list-cons and @{term "[]"} for the empty list.
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
361 |
Another function calculates the priority for a thread @{text "th"}, which is
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
362 |
defined as |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
363 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
364 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
365 |
\mbox{\begin{tabular}{lcl}
|
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
366 |
@{thm (lhs) priority.simps(1)} & @{text "\<equiv>"} &
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
367 |
@{thm (rhs) priority.simps(1)}\\
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
368 |
@{thm (lhs) priority.simps(2)} & @{text "\<equiv>"} &
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
369 |
@{thm (rhs) priority.simps(2)}\\
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
370 |
@{thm (lhs) priority.simps(3)} & @{text "\<equiv>"} &
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
371 |
@{thm (rhs) priority.simps(3)}\\
|
|
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
372 |
@{term "priority th (DUMMY#s)"} & @{text "\<equiv>"} & @{term "priority th s"}\\
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
373 |
\end{tabular}}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
374 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
375 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
376 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
377 |
In this definition we set @{text 0} as the default priority for
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
378 |
threads that have not (yet) been created. The last function we need |
|
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
379 |
calculates the ``time'', or index, at which time a thread had its |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
380 |
priority last set. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
381 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
382 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
383 |
\mbox{\begin{tabular}{lcl}
|
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
384 |
@{thm (lhs) last_set.simps(1)} & @{text "\<equiv>"} &
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
385 |
@{thm (rhs) last_set.simps(1)}\\
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
386 |
@{thm (lhs) last_set.simps(2)} & @{text "\<equiv>"} &
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
387 |
@{thm (rhs) last_set.simps(2)}\\
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
388 |
@{thm (lhs) last_set.simps(3)} & @{text "\<equiv>"} &
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
389 |
@{thm (rhs) last_set.simps(3)}\\
|
|
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
390 |
@{term "last_set th (DUMMY#s)"} & @{text "\<equiv>"} & @{term "last_set th s"}\\
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
391 |
\end{tabular}}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
392 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
393 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
394 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
395 |
In this definition @{term "length s"} stands for the length of the list
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
396 |
of events @{text s}. Again the default value in this function is @{text 0}
|
| 161 | 397 |
for threads that have not been created yet. An \emph{actor} of an event is
|
| 158 | 398 |
defined as |
399 |
||
400 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
401 |
\mbox{\begin{tabular}{lcl}
|
|
402 |
@{thm (lhs) actor.simps(5)} & @{text "\<equiv>"} &
|
|
403 |
@{thm (rhs) actor.simps(5)}\\
|
|
404 |
@{thm (lhs) actor.simps(1)} & @{text "\<equiv>"} &
|
|
405 |
@{thm (rhs) actor.simps(1)}\\
|
|
| 161 | 406 |
@{thm (lhs) actor.simps(4)} & @{text "\<equiv>"} &
|
407 |
@{thm (rhs) actor.simps(4)}\\
|
|
| 158 | 408 |
@{thm (lhs) actor.simps(2)} & @{text "\<equiv>"} &
|
409 |
@{thm (rhs) actor.simps(2)}\\
|
|
410 |
@{thm (lhs) actor.simps(3)} & @{text "\<equiv>"} &
|
|
411 |
@{thm (rhs) actor.simps(3)}\\
|
|
412 |
\end{tabular}}
|
|
413 |
\end{isabelle}
|
|
414 |
||
| 161 | 415 |
\noindent |
416 |
This allows us to define what actions a set of threads @{text ths} might
|
|
417 |
perform in a list of events @{text s}, namely
|
|
| 163 | 418 |
|
419 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
420 |
@{thm actions_of_def[where ?s="s" and ?ths="ths", THEN eq_reflection]}.
|
|
421 |
\end{isabelle}
|
|
422 |
||
423 |
where we use Isabelle's notation for list-comprehensions. This |
|
424 |
notation is very similar to notation used in Haskell for list |
|
425 |
comprehensions. A \emph{precedence} of a thread @{text th} in a
|
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
426 |
state @{text s} is the pair of natural numbers defined as
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
427 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
428 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
429 |
@{thm preced_def}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
430 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
431 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
432 |
\noindent |
|
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:
33
diff
changeset
|
433 |
We also use the abbreviation |
|
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:
33
diff
changeset
|
434 |
|
|
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:
33
diff
changeset
|
435 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
140
389ef8b1959c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
139
diff
changeset
|
436 |
@{abbrev "preceds ths 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:
33
diff
changeset
|
437 |
\end{isabelle}
|
|
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:
33
diff
changeset
|
438 |
|
|
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:
33
diff
changeset
|
439 |
\noindent |
|
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:
33
diff
changeset
|
440 |
for the set of precedences of threads @{text ths} in state @{text s}.
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
441 |
The point of precedences is to schedule threads not according to priorities (because what should |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
442 |
we do in case two threads have the same priority), but according to precedences. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
443 |
Precedences allow us to always discriminate between two threads with equal priority by |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
444 |
taking into account the time when the priority was last set. We order precedences so |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
445 |
that threads with the same priority get a higher precedence if their priority has been |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
446 |
set earlier, since for such threads it is more urgent to finish their work. In an implementation |
|
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
447 |
this choice would translate to a quite natural FIFO-scheduling of threads with |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
448 |
the same priority. |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
449 |
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
450 |
Moylan et al.~\cite{deinheritance} considered the alternative of
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
451 |
``time-slicing'' threads with equal priority, but found that it does not lead to |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
452 |
advantages in practice. On the contrary, according to their work having a policy |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
453 |
like our FIFO-scheduling of threads with equal priority reduces the number of |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
454 |
tasks involved in the inheritance process and thus minimises the number |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
455 |
of potentially expensive thread-switches. |
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
456 |
|
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
457 |
%\endnote{{\bf NEEDED?} We will also need counters for @{term P} and @{term V} events of a thread @{term th}
|
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
458 |
%in a state @{term s}. This can be straightforwardly defined in Isabelle as
|
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
459 |
% |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
460 |
%\begin{isabelle}\ \ \ \ \ %%%
|
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
461 |
%\mbox{\begin{tabular}{@ {}l}
|
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
462 |
%@{thm cntP_def}\\
|
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
463 |
%@{thm cntV_def}
|
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
464 |
%\end{tabular}}
|
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
465 |
%\end{isabelle}
|
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
466 |
% |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
467 |
%\noindent using the predefined function @{const count} for lists.}
|
|
76
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
468 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
469 |
Next, we introduce the concept of \emph{waiting queues}. They are
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
470 |
lists of threads associated with every resource. The first thread in |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
471 |
this list (i.e.~the head, or short @{term hd}) is chosen to be the one
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
472 |
that is in possession of the |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
473 |
``lock'' of the corresponding resource. We model waiting queues as |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
474 |
functions, below abbreviated as @{text wq}. They take a resource as
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
475 |
argument and return a list of threads. This allows us to define |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
476 |
when a thread \emph{holds}, respectively \emph{waits} for, a
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
477 |
resource @{text cs} given a waiting queue function @{text wq}.
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
478 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
479 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
480 |
\begin{tabular}{@ {}l}
|
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
481 |
@{thm holding_raw_def[where thread="th"]}\\
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
482 |
@{thm waiting_raw_def[where thread="th"]}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
483 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
484 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
485 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
486 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
487 |
In this definition we assume @{text "set"} converts a list into a set.
|
|
17
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
488 |
Note that in the first definition the condition about @{text "th \<in> set (wq cs)"} does not follow
|
|
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
489 |
from @{text "th = hd (set (wq cs))"}, since the head of an empty list is undefined in Isabelle/HOL.
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
490 |
At the beginning, that is in the state where no thread is created yet, |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
491 |
the waiting queue function will be the function that returns the |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
492 |
empty list for every resource. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
493 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
494 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
495 |
@{abbrev all_unlocked}\hfill\numbered{allunlocked}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
496 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
497 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
498 |
\noindent |
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
499 |
Using @{term "holding_raw"} and @{term waiting_raw}, we can introduce \emph{Resource Allocation Graphs}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
500 |
(RAG), which represent the dependencies between threads and resources. |
|
130
0f124691c191
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
127
diff
changeset
|
501 |
We choose to represent RAGs as relations using pairs of the form |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
502 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
503 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
504 |
@{term "(Th th, Cs cs)"} \hspace{5mm}{\rm and}\hspace{5mm}
|
|
136
fb3f52fe99d1
updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
135
diff
changeset
|
505 |
@{term "(Cs cs, Th th)"}\hfill\numbered{pairs}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
506 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
507 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
508 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
509 |
where the first stands for a \emph{waiting edge} and the second for a
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
510 |
\emph{holding edge} (@{term Cs} and @{term Th} are constructors of a
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
511 |
datatype for vertices). Given a waiting queue function, a RAG is defined |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
512 |
as the union of the sets of waiting and holding edges, namely |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
513 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
514 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
515 |
@{thm RAG_raw_def}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
516 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
517 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
518 |
|
|
22
9f0b78fcc894
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
20
diff
changeset
|
519 |
\begin{figure}[t]
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
520 |
\begin{center}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
521 |
\newcommand{\fnt}{\fontsize{7}{8}\selectfont}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
522 |
\begin{tikzpicture}[scale=1]
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
523 |
%%\draw[step=2mm] (-3,2) grid (1,-1); |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
524 |
|
|
20
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
525 |
\node (A) at (0,0) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^sub>0"}};
|
|
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
526 |
\node (B) at (2,0) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^sub>1"}};
|
|
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
527 |
\node (C) at (4,0.7) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^sub>1"}};
|
|
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
528 |
\node (D) at (4,-0.7) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^sub>2"}};
|
|
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
529 |
\node (E) at (6,-0.7) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^sub>2"}};
|
|
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
530 |
\node (E1) at (6, 0.2) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^sub>3"}};
|
|
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
531 |
\node (F) at (8,-0.7) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^sub>3"}};
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
532 |
|
|
20
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
533 |
\node (X) at (0,-2) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^sub>4"}};
|
|
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
534 |
\node (Y) at (2,-2) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^sub>4"}};
|
|
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
535 |
\node (Z) at (2,-2.9) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^sub>5"}};
|
|
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
536 |
\node (U1) at (4,-2) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^sub>5"}};
|
|
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
537 |
\node (U2) at (4,-2.9) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^sub>6"}};
|
|
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
538 |
\node (R) at (6,-2.9) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^sub>6"}};
|
|
12
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
539 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
540 |
\draw [<-,line width=0.6mm] (A) to node [pos=0.54,sloped,above=-0.5mm] {\fnt{}holding} (B);
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
541 |
\draw [->,line width=0.6mm] (C) to node [pos=0.4,sloped,above=-0.5mm] {\fnt{}waiting} (B);
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
542 |
\draw [->,line width=0.6mm] (D) to node [pos=0.4,sloped,below=-0.5mm] {\fnt{}waiting} (B);
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
543 |
\draw [<-,line width=0.6mm] (D) to node [pos=0.54,sloped,below=-0.5mm] {\fnt{}holding} (E);
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
544 |
\draw [<-,line width=0.6mm] (D) to node [pos=0.54,sloped,above=-0.5mm] {\fnt{}holding} (E1);
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
545 |
\draw [->,line width=0.6mm] (F) to node [pos=0.45,sloped,below=-0.5mm] {\fnt{}waiting} (E);
|
|
12
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
546 |
|
|
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
547 |
\draw [->,line width=0.6mm] (U1) to node [pos=0.45,sloped,below=-0.5mm] {\fnt{}waiting} (Y);
|
|
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
548 |
\draw [->,line width=0.6mm] (U2) to node [pos=0.45,sloped,below=-0.5mm] {\fnt{}waiting} (Z);
|
|
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
549 |
\draw [<-,line width=0.6mm] (X) to node [pos=0.54,sloped,below=-0.5mm] {\fnt{}holding} (Z);
|
|
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
550 |
\draw [<-,line width=0.6mm] (X) to node [pos=0.54,sloped,above=-0.5mm] {\fnt{}holding} (Y);
|
|
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
551 |
\draw [<-,line width=0.6mm] (U2) to node [pos=0.54,sloped,above=-0.5mm] {\fnt{}holding} (R);
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
552 |
\end{tikzpicture}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
553 |
\end{center}
|
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
554 |
\caption{An instance of a Resource Allocation Graph (RAG).\label{RAGgraph}}
|
|
12
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
555 |
\end{figure}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
556 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
557 |
\noindent |
|
22
9f0b78fcc894
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
20
diff
changeset
|
558 |
If there is no cycle, then every RAG can be pictured as a forrest of trees, as |
|
9f0b78fcc894
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
20
diff
changeset
|
559 |
for example in Figure~\ref{RAGgraph}.
|
|
9f0b78fcc894
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
20
diff
changeset
|
560 |
|
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
561 |
Because of the RAGs, we will need to formalise some results about |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
562 |
graphs. While there are few formalisations for graphs already |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
563 |
implemented in Isabelle, we choose to introduce our own library of |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
564 |
graphs for PIP. The justification for this is that we wanted to be able to |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
565 |
reason about potentially infinite graphs (in the sense of infinitely |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
566 |
branching and infinite size): the property that our RAGs are |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
567 |
actually forrests of finitely branching trees having only an finite |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
568 |
depth should be something we can \emph{prove} for our model of
|
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
569 |
PIP---it should not be an assumption we build already into our |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
570 |
model. It seemed for our purposes the most convenient |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
571 |
represeantation of graphs are binary relations given by sets of |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
572 |
pairs shown in \eqref{pairs}. The pairs stand for the edges in
|
|
136
fb3f52fe99d1
updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
135
diff
changeset
|
573 |
graphs. This relation-based representation is convenient since we |
|
fb3f52fe99d1
updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
135
diff
changeset
|
574 |
can use the notions of transitive closure operations @{term "trancl
|
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
575 |
DUMMY"} and @{term "rtrancl DUMMY"}, as well as relation
|
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
576 |
composition. A \emph{forrest} is defined as the relation @{text
|
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
577 |
rel} that is \emph{single valued} and \emph{acyclic}:
|
|
135
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
578 |
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
579 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
580 |
\begin{tabular}{@ {}l}
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
581 |
@{thm single_valued_def[where ?r="rel", THEN eq_reflection]}\\
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
582 |
@{thm acyclic_def[where ?r="rel", THEN eq_reflection]}
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
583 |
\end{tabular}
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
584 |
\end{isabelle}
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
585 |
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
586 |
\noindent |
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
587 |
The \emph{children}, \emph{subtree} and \emph{ancestors} of a node in a graph
|
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
588 |
can be easily defined relationally as |
|
135
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
589 |
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
590 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
591 |
\begin{tabular}{@ {}l}
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
592 |
@{thm children_def[where ?r="rel" and ?x="node", THEN eq_reflection]}\\
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
593 |
@{thm subtree_def[where ?r="rel" and ?x="node", THEN eq_reflection]}\\
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
594 |
@{thm ancestors_def[where ?r="rel" and ?x="node", THEN eq_reflection]}\\
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
595 |
\end{tabular}
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
596 |
\end{isabelle}
|
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
597 |
|
|
143
c5a65d98191a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
598 |
\noindent Note that forrests can have trees with infinte depth and |
|
135
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
599 |
containing nodes with infinitely many children. A \emph{finite
|
|
143
c5a65d98191a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
600 |
forrest} is a forrest which is well-founded and every node has |
|
135
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
601 |
finitely many children (is only finitely branching). |
|
9b5da0327d43
updated paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
134
diff
changeset
|
602 |
|
| 161 | 603 |
%\endnote{
|
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
604 |
%\begin{isabelle}\ \ \ \ \ %%%
|
| 161 | 605 |
%@ {thm rtrancl_path.intros}
|
606 |
%\end{isabelle}
|
|
607 |
% |
|
608 |
%\begin{isabelle}\ \ \ \ \ %%%
|
|
609 |
%@ {thm rpath_def}
|
|
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
610 |
%\end{isabelle}
|
| 161 | 611 |
%} |
|
136
fb3f52fe99d1
updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
135
diff
changeset
|
612 |
|
|
fb3f52fe99d1
updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
135
diff
changeset
|
613 |
|
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
614 |
%\endnote{{\bf Lemma about overlapping paths}}
|
|
136
fb3f52fe99d1
updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
135
diff
changeset
|
615 |
|
|
143
c5a65d98191a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
616 |
The locking mechanism of PIP ensures that for each thread node, |
|
c5a65d98191a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
617 |
there can be many incoming holding edges in the RAG, but at most one |
|
c5a65d98191a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
618 |
out going waiting edge. The reason is that when a thread asks for |
|
c5a65d98191a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
619 |
resource that is locked already, then the thread is blocked and |
|
c5a65d98191a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
620 |
cannot ask for another resource. Clearly, also every resource can |
|
c5a65d98191a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
621 |
only have at most one outgoing holding edge---indicating that the |
|
c5a65d98191a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
622 |
resource is locked. So if the @{text "RAG"} is well-founded and
|
|
c5a65d98191a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
623 |
finite, we can always start at a thread waiting for a resource and |
|
c5a65d98191a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
624 |
``chase'' outgoing arrows leading to a single root of a tree. |
|
17
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
625 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
626 |
The use of relations for representing RAGs allows us to conveniently define |
| 161 | 627 |
the notion of the \emph{dependants} of a thread
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
628 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
629 |
\begin{isabelle}\ \ \ \ \ %%%
|
| 178 | 630 |
@{thm dependants_raw_def}\hfill\numbered{dependants}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
631 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
632 |
|
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
633 |
\noindent This definition needs to account for all threads that wait |
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
634 |
for a thread to release a resource. This means we need to include |
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
635 |
threads that transitively wait for a resource to be released (in the |
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
636 |
picture above this means the dependants of @{text "th\<^sub>0"} are
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
637 |
@{text "th\<^sub>1"} and @{text "th\<^sub>2"}, which wait for
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
638 |
resource @{text "cs\<^sub>1"}, but also @{text "th\<^sub>3"}, which
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
639 |
cannot make any progress unless @{text "th\<^sub>2"} makes progress,
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
640 |
which in turn needs to wait for @{text "th\<^sub>0"} to finish). If
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
641 |
there is a circle of dependencies in a RAG, then clearly we have a |
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
642 |
deadlock. Therefore when a thread requests a resource, we must |
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
643 |
ensure that the resulting RAG is not circular. In practice, the |
|
144
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
644 |
programmer has to ensure this. Our model will enforce that critical |
|
143
c5a65d98191a
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
142
diff
changeset
|
645 |
resources can only be requested provided no circularity can arise. |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
646 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
647 |
Next we introduce the notion of the \emph{current precedence} of a thread @{text th} in a
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
648 |
state @{text s}. It is defined as
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
649 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
650 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
144
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
651 |
@{thm cpreced_def3}\hfill\numbered{cpreced}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
652 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
653 |
|
| 178 | 654 |
\noindent where the dependants of @{text th} are given by the
|
|
144
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
655 |
waiting queue function. While the precedence @{term prec} of any
|
|
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
656 |
thread is determined statically (for example when the thread is |
|
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
657 |
created), the point of the current precedence is to dynamically |
|
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
658 |
increase this precedence, if needed according to PIP. Therefore the |
|
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
659 |
current precedence of @{text th} is given as the maximum of the
|
|
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
660 |
precedence of @{text th} \emph{and} all threads that are dependants
|
|
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
661 |
of @{text th} in the state @{text s}. Since the notion @{term
|
|
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
662 |
"dependants"} is defined as the transitive closure of all dependent |
|
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
663 |
threads, we deal correctly with the problem in the informal |
|
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
664 |
algorithm by Sha et al.~\cite{Sha90} where a priority of a thread is
|
|
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
665 |
lowered prematurely (see Introduction). We again introduce an abbreviation for current |
| 166 | 666 |
precedeces of a set of threads, written @{text "cprecs wq s ths"}.
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
667 |
|
|
130
0f124691c191
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
127
diff
changeset
|
668 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
0f124691c191
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
127
diff
changeset
|
669 |
@{thm cpreceds_def}
|
|
0f124691c191
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
127
diff
changeset
|
670 |
\end{isabelle}
|
|
0f124691c191
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
127
diff
changeset
|
671 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
672 |
The next function, called @{term schs}, defines the behaviour of the scheduler. It will be defined
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
673 |
by recursion on the state (a list of events); this function returns a \emph{schedule state}, which
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
674 |
we represent as a record consisting of two |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
675 |
functions: |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
676 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
677 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
144
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
678 |
@{text "\<lparr>wq_fun, cprec_fun\<rparr>"}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
679 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
680 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
681 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
682 |
The first function is a waiting queue function (that is, it takes a |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
683 |
resource @{text "cs"} and returns the corresponding list of threads
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
684 |
that lock, respectively wait for, it); the second is a function that |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
685 |
takes a thread and returns its current precedence (see |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
686 |
the definition in \eqref{cpreced}). We assume the usual getter and setter methods for
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
687 |
such records. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
688 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
689 |
In the initial state, the scheduler starts with all resources unlocked (the corresponding |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
690 |
function is defined in \eqref{allunlocked}) and the
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
691 |
current precedence of every thread is initialised with @{term "Prc 0 0"}; that means
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
692 |
\mbox{@{abbrev initial_cprec}}. Therefore
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
693 |
we have for the initial shedule state |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
694 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
695 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
696 |
\begin{tabular}{@ {}l}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
697 |
@{thm (lhs) schs.simps(1)} @{text "\<equiv>"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
698 |
\hspace{5mm}@{term "(|wq_fun = all_unlocked, cprec_fun = (\<lambda>_::thread. Prc 0 0)|)"}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
699 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
700 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
701 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
702 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
703 |
The cases for @{term Create}, @{term Exit} and @{term Set} are also straightforward:
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
704 |
we calculate the waiting queue function of the (previous) state @{text s};
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
705 |
this waiting queue function @{text wq} is unchanged in the next schedule state---because
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
706 |
none of these events lock or release any resource; |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
707 |
for calculating the next @{term "cprec_fun"}, we use @{text wq} and
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
708 |
@{term cpreced}. This gives the following three clauses for @{term schs}:
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
709 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
710 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
711 |
\begin{tabular}{@ {}l}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
712 |
@{thm (lhs) schs.simps(2)} @{text "\<equiv>"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
713 |
\hspace{5mm}@{text "let"} @{text "wq = wq_fun (schs s)"} @{text "in"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
714 |
\hspace{8mm}@{term "(|wq_fun = wq\<iota>, cprec_fun = cpreced wq\<iota> (Create th prio # s)|)"}\smallskip\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
715 |
@{thm (lhs) schs.simps(3)} @{text "\<equiv>"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
716 |
\hspace{5mm}@{text "let"} @{text "wq = wq_fun (schs s)"} @{text "in"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
717 |
\hspace{8mm}@{term "(|wq_fun = wq\<iota>, cprec_fun = cpreced wq\<iota> (Exit th # s)|)"}\smallskip\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
718 |
@{thm (lhs) schs.simps(4)} @{text "\<equiv>"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
719 |
\hspace{5mm}@{text "let"} @{text "wq = wq_fun (schs s)"} @{text "in"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
720 |
\hspace{8mm}@{term "(|wq_fun = wq\<iota>, cprec_fun = cpreced wq\<iota> (Set th prio # s)|)"}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
721 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
722 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
723 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
724 |
\noindent |
|
144
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
725 |
More interesting are the cases where a resource, say @{text cs}, is requested or released. In these cases
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
726 |
we need to calculate a new waiting queue function. For the event @{term "P th cs"}, we have to update
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
727 |
the function so that the new thread list for @{text cs} is the old thread list plus the thread @{text th}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
728 |
appended to the end of that list (remember the head of this list is assigned to be in the possession of this |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
729 |
resource). This gives the clause |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
730 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
731 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
732 |
\begin{tabular}{@ {}l}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
733 |
@{thm (lhs) schs.simps(5)} @{text "\<equiv>"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
734 |
\hspace{5mm}@{text "let"} @{text "wq = wq_fun (schs s)"} @{text "in"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
735 |
\hspace{5mm}@{text "let"} @{text "new_wq = wq(cs := (wq cs @ [th]))"} @{text "in"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
736 |
\hspace{8mm}@{term "(|wq_fun = new_wq, cprec_fun = cpreced new_wq (P th cs # s)|)"}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
737 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
738 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
739 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
740 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
741 |
The clause for event @{term "V th cs"} is similar, except that we need to update the waiting queue function
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
742 |
so that the thread that possessed the lock is deleted from the corresponding thread list. For this |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
743 |
list transformation, we use |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
744 |
the auxiliary function @{term release}. A simple version of @{term release} would
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
745 |
just delete this thread and return the remaining threads, namely |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
746 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
747 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
748 |
\begin{tabular}{@ {}lcl}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
749 |
@{term "release []"} & @{text "\<equiv>"} & @{term "[]"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
750 |
@{term "release (DUMMY # qs)"} & @{text "\<equiv>"} & @{term "qs"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
751 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
752 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
753 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
754 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
755 |
In practice, however, often the thread with the highest precedence in the list will get the |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
756 |
lock next. We have implemented this choice, but later found out that the choice |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
757 |
of which thread is chosen next is actually irrelevant for the correctness of PIP. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
758 |
Therefore we prove the stronger result where @{term release} is defined as
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
759 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
760 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
761 |
\begin{tabular}{@ {}lcl}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
762 |
@{term "release []"} & @{text "\<equiv>"} & @{term "[]"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
763 |
@{term "release (DUMMY # qs)"} & @{text "\<equiv>"} & @{term "SOME qs'. distinct qs' \<and> set qs' = set qs"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
764 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
765 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
766 |
|
|
144
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
767 |
\noindent where @{text "SOME"} stands for Hilbert's epsilon and
|
|
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
768 |
implements an arbitrary choice for the next waiting list. It just |
|
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
769 |
has to be a list of distinctive threads and contains the same |
|
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
770 |
elements as @{text "qs"} (essentially @{text "qs'"} can be any
|
|
e4d151d761c4
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
143
diff
changeset
|
771 |
reordering of the list @{text "qs"}). This gives for @{term V} the clause:
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
772 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
773 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
774 |
\begin{tabular}{@ {}l}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
775 |
@{thm (lhs) schs.simps(6)} @{text "\<equiv>"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
776 |
\hspace{5mm}@{text "let"} @{text "wq = wq_fun (schs s)"} @{text "in"}\\
|
|
17
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
777 |
\hspace{5mm}@{text "let"} @{text "new_wq = wq(cs := release (wq cs))"} @{text "in"}\\
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
778 |
\hspace{8mm}@{term "(|wq_fun = new_wq, cprec_fun = cpreced new_wq (V th cs # s)|)"}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
779 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
780 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
781 |
|
|
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:
33
diff
changeset
|
782 |
Having the scheduler function @{term schs} at our disposal, we can
|
|
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:
33
diff
changeset
|
783 |
``lift'', or overload, the notions @{term waiting}, @{term holding},
|
|
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:
33
diff
changeset
|
784 |
@{term RAG}, @{term dependants} and @{term cp} to operate on states
|
|
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:
33
diff
changeset
|
785 |
only. |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
786 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
787 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
788 |
\begin{tabular}{@ {}rcl}
|
| 178 | 789 |
@{thm (lhs) s_holding_abv} & @{text "\<equiv>"} & @{thm (rhs) s_holding_abv[simplified wq_def]}\\
|
790 |
@{thm (lhs) s_waiting_abv} & @{text "\<equiv>"} & @{thm (rhs) s_waiting_abv[simplified wq_def]}\\
|
|
791 |
@{thm (lhs) s_RAG_abv} & @{text "\<equiv>"} & @{thm (rhs) s_RAG_abv[simplified wq_def]}\\
|
|
792 |
@{thm (lhs) s_dependants_abv}& @{text "\<equiv>"} & @{thm (rhs) s_dependants_abv[simplified wq_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:
33
diff
changeset
|
793 |
@{thm (lhs) cp_def} & @{text "\<equiv>"} & @{thm (rhs) cp_def}\\
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
794 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
795 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
796 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
797 |
\noindent |
| 178 | 798 |
With these abbreviations in place we can derive for every valid trace @{text s}
|
799 |
the following two facts about @{term dependants} and @{term cprec} (see \eqref{dependants}
|
|
800 |
and \eqref{cpreced}): Given @{thm (prem 1) valid_trace.cp_alt_def}, then
|
|
801 |
||
802 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
803 |
\begin{tabular}{@ {}rcl}
|
|
804 |
@{thm (concl) valid_trace.cp_alt_def3}\\
|
|
805 |
\end{tabular}\hfill\numbered{overloaded}
|
|
806 |
\end{isabelle}
|
|
807 |
||
808 |
||
809 |
can introduce |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
810 |
the notion of a thread being @{term ready} in a state (i.e.~threads
|
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
811 |
that do not wait for any resource, which are the roots of the trees |
|
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
812 |
in the RAG, see Figure~\ref{RAGgraph}). The @{term running} thread
|
|
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
813 |
is then the thread with the highest current precedence of all ready threads. |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
814 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
815 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
816 |
\begin{tabular}{@ {}l}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
817 |
@{thm readys_def}\\
|
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
818 |
@{thm running_def}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
819 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
820 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
821 |
|
| 164 | 822 |
\noindent In the second definition @{term "DUMMY ` DUMMY"} stands
|
823 |
for the image of a set under a function. Note that in the initial |
|
824 |
state, that is where the list of events is empty, the set @{term
|
|
825 |
threads} is empty and therefore there is neither a thread ready nor |
|
826 |
running. If there is one or more threads ready, then there can only |
|
827 |
be \emph{one} thread running, namely the one whose current
|
|
828 |
precedence is equal to the maximum of all ready threads. We use sets |
|
829 |
to capture both possibilities. We can now also conveniently define |
|
830 |
the set of resources that are locked by a thread in a given state |
|
831 |
and also when a thread is detached in a state (meaning the thread |
|
832 |
neither holds nor waits for a resource---in the RAG this would |
|
833 |
correspond to an isolated node without any incoming and outgoing |
|
834 |
edges, see Figure~\ref{RAGgraph}):
|
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
835 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
836 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
837 |
\begin{tabular}{@ {}l}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
838 |
@{thm holdents_def}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
839 |
@{thm detached_def}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
840 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
841 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
842 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
843 |
%\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
844 |
%The second definition states that @{text th} in @{text s}.
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
845 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
846 |
Finally we can define what a \emph{valid state} is in our model of PIP. For
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
847 |
example we cannot expect to be able to exit a thread, if it was not |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
848 |
created yet. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
849 |
These validity constraints on states are characterised by the |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
850 |
inductive predicate @{term "step"} and @{term vt}. We first give five inference rules
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
851 |
for @{term step} relating a state and an event that can happen next.
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
852 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
853 |
\begin{center}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
854 |
\begin{tabular}{c}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
855 |
@{thm[mode=Rule] thread_create[where thread=th]}\hspace{1cm}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
856 |
@{thm[mode=Rule] thread_exit[where thread=th]}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
857 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
858 |
\end{center}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
859 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
860 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
861 |
The first rule states that a thread can only be created, if it is not alive yet. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
862 |
Similarly, the second rule states that a thread can only be terminated if it was |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
863 |
running and does not lock any resources anymore (this simplifies slightly our model; |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
864 |
in practice we would expect the operating system releases all locks held by a |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
865 |
thread that is about to exit). The event @{text Set} can happen
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
866 |
if the corresponding thread is running. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
867 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
868 |
\begin{center}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
869 |
@{thm[mode=Rule] thread_set[where thread=th]}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
870 |
\end{center}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
871 |
|
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
872 |
\noindent If a thread wants to lock a resource, then the thread |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
873 |
needs to be running and also we have to make sure that the resource |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
874 |
lock does not lead to a cycle in the RAG (the prurpose of the second |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
875 |
premise in the rule below). In practice, ensuring the latter is the |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
876 |
responsibility of the programmer. In our formal model we brush |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
877 |
aside these problematic cases in order to be able to make some |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
878 |
meaningful statements about PIP.\footnote{This situation is similar
|
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
879 |
to the infamous \emph{occurs check} in Prolog: In order to say
|
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
880 |
anything meaningful about unification, one needs to perform an |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
881 |
occurs check. But in practice the occurs check is omitted and the |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
882 |
responsibility for avoiding problems rests with the programmer.} |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
883 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
884 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
885 |
\begin{center}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
886 |
@{thm[mode=Rule] thread_P[where thread=th]}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
887 |
\end{center}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
888 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
889 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
890 |
Similarly, if a thread wants to release a lock on a resource, then |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
891 |
it must be running and in the possession of that lock. This is |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
892 |
formally given by the last inference rule of @{term step}.
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
893 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
894 |
\begin{center}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
895 |
@{thm[mode=Rule] thread_V[where thread=th]}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
896 |
\end{center}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
897 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
898 |
\noindent |
|
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
899 |
Note, however, that apart from the circularity condition, we do not make any |
|
20
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
900 |
assumption on how different resources can be locked and released relative to each |
|
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
901 |
other. In our model it is possible that critical sections overlap. This is in |
|
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
902 |
contrast to Sha et al \cite{Sha90} who require that critical sections are
|
|
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
903 |
properly nested (recall Fig.~\ref{overlap}).
|
|
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
904 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
905 |
A valid state of PIP can then be conveniently be defined as follows: |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
906 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
907 |
\begin{center}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
908 |
\begin{tabular}{c}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
909 |
@{thm[mode=Axiom] vt_nil}\hspace{1cm}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
910 |
@{thm[mode=Rule] vt_cons}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
911 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
912 |
\end{center}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
913 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
914 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
915 |
This completes our formal model of PIP. In the next section we present |
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
916 |
a series of desirable properties derived from our model of PIP. This can |
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
917 |
be regarded as a validation of the correctness of our model. |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
918 |
*} |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
919 |
|
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
920 |
(* |
|
137
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
921 |
section {* Preliminaries *}
|
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
922 |
*) |
|
137
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
923 |
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
924 |
(*<*) |
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
925 |
context valid_trace |
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
926 |
begin |
| 162 | 927 |
(*>*) |
928 |
(*<*) |
|
|
137
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
929 |
text {*
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
930 |
|
|
142
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
931 |
\endnote{In this section we prove facts that immediately follow from
|
|
137
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
932 |
our definitions of valid traces. |
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
933 |
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
934 |
\begin{lemma}??\label{precedunique}
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
935 |
@{thm [mode=IfThen] preced_unique[where ?th1.0=th\<^sub>1 and ?th2.0=th\<^sub>2]}
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
936 |
\end{lemma}
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
937 |
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
938 |
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
939 |
We can verify that in any valid state, there can only be at most |
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
940 |
one running thread---if there are more than one running thread, |
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
941 |
say @{text "th\<^sub>1"} and @{text "th\<^sub>2"}, they must be
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
942 |
equal. |
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
943 |
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
944 |
\begin{lemma}
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
945 |
@{thm [mode=IfThen] running_unique[where ?th1.0=th\<^sub>1 and ?th2.0=th\<^sub>2]}
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
946 |
\end{lemma}
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
947 |
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
948 |
\begin{proof}
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
949 |
Since @{text "th\<^sub>1"} and @{text "th\<^sub>2"} are running, they must be
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
950 |
roots in the RAG. |
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
951 |
According to XXX, there exists a chain in the RAG-subtree of @{text "th\<^sub>1"},
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
952 |
say starting from @{text "th'\<^sub>1"}, such that @{text "th'\<^sub>1"} has the
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
953 |
highest precedence in this subtree (@{text "th\<^sub>1"} inherited
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
954 |
the precedence of @{text "th'\<^sub>1"}). We have a similar chain starting from, say
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
955 |
@{text "th'\<^sub>2"}, in the RAG-subtree of @{text "th\<^sub>2"}. Since @{text "th\<^sub>1"}
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
956 |
and @{text "th\<^sub>2"} are running we know their cp-value must be the same, that is
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
957 |
\begin{center}
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
958 |
@{term "cp s th\<^sub>1 = cp s th\<^sub>2"}
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
959 |
\end{center}
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
960 |
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
961 |
\noindent |
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
962 |
That means the precedences of @{text "th'\<^sub>1"} and @{text "th'\<^sub>2"}
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
963 |
must be the same (they are the maxima in the respective RAG-subtrees). From this we can |
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
964 |
infer by Lemma~\ref{precedunique} that @{text "th'\<^sub>1"}
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
965 |
and @{text "th'\<^sub>2"} are the same threads. However, this also means the
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
966 |
roots @{text "th\<^sub>1"} and @{text "th\<^sub>2"} must be the same.\qed
|
|
142
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
967 |
\end{proof}}
|
|
137
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
968 |
|
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
969 |
*} |
| 162 | 970 |
(*>*) |
|
137
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
971 |
(*<*)end(*>*) |
|
785c0f6b8184
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
136
diff
changeset
|
972 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
973 |
section {* The Correctness Proof *}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
974 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
975 |
(*<*) |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
976 |
context extend_highest_gen |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
977 |
begin |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
978 |
(*>*) |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
979 |
text {*
|
|
82
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
980 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
981 |
Sha et al.~state their first correctness criterion for PIP in terms |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
982 |
of the number of low-priority threads \cite[Theorem 3]{Sha90}: if
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
983 |
there are @{text n} low-priority threads, then a blocked job with
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
984 |
high priority can only be blocked a maximum of @{text n} times.
|
|
82
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
985 |
Their second correctness criterion is given in terms of the number |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
986 |
of critical resources \cite[Theorem 6]{Sha90}: if there are @{text
|
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
987 |
m} critical resources, then a blocked job with high priority can |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
988 |
only be blocked a maximum of @{text m} times. Both results on their
|
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
989 |
own, strictly speaking, do \emph{not} prevent indefinite, or
|
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
990 |
unbounded, Priority Inversion, because if a low-priority thread does |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
991 |
not give up its critical resource (the one the high-priority thread |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
992 |
is waiting for), then the high-priority thread can never run. The |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
993 |
argument of Sha et al.~is that \emph{if} threads release locked
|
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
994 |
resources in a finite amount of time, then indefinite Priority |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
995 |
Inversion cannot occur---the high-priority thread is guaranteed to |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
996 |
run eventually. The assumption is that programmers must ensure that |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
997 |
threads are programmed in this way. However, even taking this |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
998 |
assumption into account, the correctness properties of Sha et |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
999 |
al.~are \emph{not} true for their version of PIP---despite being
|
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1000 |
``proved''. As Yodaiken \cite{Yodaiken02} and Moylan et
|
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1001 |
al.~\cite{deinheritance} pointed out: If a low-priority thread
|
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1002 |
possesses locks to two resources for which two high-priority threads |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1003 |
are waiting for, then lowering the priority prematurely after giving |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1004 |
up only one lock, can cause indefinite Priority Inversion for one of |
| 164 | 1005 |
the high-priority threads, invalidating their two bounds (recall the |
1006 |
counter example described in the Introduction). |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1007 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1008 |
Even when fixed, their proof idea does not seem to go through for |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1009 |
us, because of the way we have set up our formal model of PIP. One |
|
82
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1010 |
reason is that we allow critical sections, which start with a @{text
|
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1011 |
P}-event and finish with a corresponding @{text V}-event, to
|
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1012 |
arbitrarily overlap (something Sha et al.~explicitly exclude). |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1013 |
Therefore we have designed a different correctness criterion for |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1014 |
PIP. The idea behind our criterion is as follows: for all states |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1015 |
@{text s}, we know the corresponding thread @{text th} with the
|
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1016 |
highest precedence; we show that in every future state (denoted by |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1017 |
@{text "s' @ s"}) in which @{text th} is still alive, either @{text
|
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1018 |
th} is running or it is blocked by a thread that was alive in the |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1019 |
state @{text s} and was waiting for or in the possession of a lock
|
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1020 |
in @{text s}. Since in @{text s}, as in every state, the set of
|
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1021 |
alive threads is finite, @{text th} can only be blocked a finite
|
| 165 | 1022 |
number of threads. We will actually prove a |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1023 |
stronger statement where we also provide the current precedence of |
|
82
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1024 |
the blocking thread. |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1025 |
|
| 170 | 1026 |
However, the theorem we are going to prove hinges upon a number of |
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1027 |
natural assumptions about the states @{text s} and @{text "s' @ s"}, the
|
|
82
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1028 |
thread @{text th} and the events happening in @{text s'}. We list
|
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1029 |
them next: |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1030 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1031 |
\begin{quote}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1032 |
{\bf Assumptions on the states {\boldmath@{text s}} and
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1033 |
{\boldmath@{text "s' @ s"}:}} We need to require that @{text "s"} and
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1034 |
@{text "s' @ s"} are valid states:
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1035 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1036 |
\begin{tabular}{l}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1037 |
@{term "vt s"}, @{term "vt (s' @ s)"}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1038 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1039 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1040 |
\end{quote}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1041 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1042 |
\begin{quote}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1043 |
{\bf Assumptions on the thread {\boldmath@{text "th"}:}}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1044 |
The thread @{text th} must be alive in @{text s} and
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1045 |
has the highest precedence of all alive threads in @{text s}. Furthermore the
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1046 |
priority of @{text th} is @{text prio} (we need this in the next assumptions).
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1047 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1048 |
\begin{tabular}{l}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1049 |
@{term "th \<in> threads s"}\\
|
| 178 | 1050 |
@{term "prec th s = Max (precs s (threads s))"}\\
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1051 |
@{term "prec th s = (prio, DUMMY)"}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1052 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1053 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1054 |
\end{quote}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1055 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1056 |
\begin{quote}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1057 |
{\bf Assumptions on the events in {\boldmath@{text "s'"}:}} We want to prove that @{text th} cannot
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1058 |
be blocked indefinitely. Of course this can happen if threads with higher priority |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1059 |
than @{text th} are continuously created in @{text s'}. Therefore we have to assume that
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1060 |
events in @{text s'} can only create (respectively set) threads with equal or lower
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1061 |
priority than @{text prio} of @{text th}. We also need to assume that the
|
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1062 |
priority of @{text "th"} does not get reset and all other reset priorities are either
|
|
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1063 |
less or equal. Moreover, we assume that @{text th} does
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1064 |
not get ``exited'' in @{text "s'"}. This can be ensured by assuming the following three implications.
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1065 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1066 |
\begin{tabular}{l}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1067 |
{If}~~@{text "Create th' prio' \<in> set s'"}~~{then}~~@{text "prio' \<le> prio"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1068 |
{If}~~@{text "Set th' prio' \<in> set s'"}~~{then}~~@{text "th' \<noteq> th"}~~{and}~~@{text "prio' \<le> prio"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1069 |
{If}~~@{text "Exit th' \<in> set s'"}~~{then}~~@{text "th' \<noteq> th"}\\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1070 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1071 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1072 |
\end{quote}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1073 |
|
|
82
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1074 |
\noindent The locale mechanism of Isabelle helps us to manage |
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1075 |
conveniently such assumptions~\cite{Haftmann08}. Under these
|
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1076 |
assumptions we shall prove the following correctness property: |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1077 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1078 |
\begin{theorem}\label{mainthm}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1079 |
Given the assumptions about states @{text "s"} and @{text "s' @ s"},
|
|
95
8d2cc27f45f3
changes to my repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
82
diff
changeset
|
1080 |
the thread @{text th} and the events in @{text "s'"}, then either
|
|
8d2cc27f45f3
changes to my repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
82
diff
changeset
|
1081 |
\begin{itemize}
|
|
142
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1082 |
\item[$\bullet$] @{term "th \<in> running (s' @ s)"} or\medskip
|
|
95
8d2cc27f45f3
changes to my repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
82
diff
changeset
|
1083 |
|
|
142
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1084 |
\item[$\bullet$] there exists a thread @{term "th'"} with @{term "th' \<noteq> th"}
|
|
95
8d2cc27f45f3
changes to my repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
82
diff
changeset
|
1085 |
and @{term "th' \<in> running (s' @ s)"} such that @{text "th' \<in> threads
|
|
8d2cc27f45f3
changes to my repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
82
diff
changeset
|
1086 |
s"}, @{text "\<not> detached s th'"} and @{term "cp (s' @ s) th' = prec
|
|
8d2cc27f45f3
changes to my repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
82
diff
changeset
|
1087 |
th s"}. |
|
8d2cc27f45f3
changes to my repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
82
diff
changeset
|
1088 |
\end{itemize}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1089 |
\end{theorem}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1090 |
|
|
82
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1091 |
\noindent This theorem ensures that the thread @{text th}, which has
|
|
95
8d2cc27f45f3
changes to my repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
82
diff
changeset
|
1092 |
the highest precedence in the state @{text s}, is either running in
|
|
8d2cc27f45f3
changes to my repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
82
diff
changeset
|
1093 |
state @{term "s' @ s"}, or can only be blocked in the state @{text
|
|
8d2cc27f45f3
changes to my repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
82
diff
changeset
|
1094 |
"s' @ s"} by a thread @{text th'} that already existed in @{text s}
|
| 166 | 1095 |
and is waiting for a resource or had a lock on at least one resource---that means |
|
95
8d2cc27f45f3
changes to my repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
82
diff
changeset
|
1096 |
the thread was not \emph{detached} in @{text s}. As we shall see
|
|
8d2cc27f45f3
changes to my repository
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
82
diff
changeset
|
1097 |
shortly, that means there are only finitely many threads that can |
| 166 | 1098 |
block @{text th} in this way.
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1099 |
|
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1100 |
|
| 168 | 1101 |
%% HERE |
|
139
289e362f7a76
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
138
diff
changeset
|
1102 |
|
| 168 | 1103 |
%Given our assumptions (on @{text th}), the first property we can
|
1104 |
%show is that any running thread, say @{text "th'"}, has the same
|
|
1105 |
%precedence as @{text th}:
|
|
|
139
289e362f7a76
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
138
diff
changeset
|
1106 |
|
| 168 | 1107 |
%\begin{lemma}\label{runningpreced}
|
1108 |
%@{thm [mode=IfThen] running_preced_inversion}
|
|
1109 |
%\end{lemma}
|
|
|
139
289e362f7a76
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
138
diff
changeset
|
1110 |
|
| 168 | 1111 |
%\begin{proof}
|
1112 |
%By definition, the running thread has as current precedence the maximum of |
|
1113 |
%all ready threads, that is |
|
|
139
289e362f7a76
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
138
diff
changeset
|
1114 |
|
| 168 | 1115 |
%\begin{center}
|
1116 |
%@{term "cp (t @ s) th' = Max (cp (t @ s) ` readys (t @ s))"}
|
|
1117 |
%\end{center}
|
|
|
140
389ef8b1959c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
139
diff
changeset
|
1118 |
|
| 168 | 1119 |
%\noindent |
1120 |
%We also know that this is equal to the maximum of current precedences of all threads, |
|
1121 |
%that is |
|
|
140
389ef8b1959c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
139
diff
changeset
|
1122 |
|
| 168 | 1123 |
%\begin{center}
|
1124 |
%@{term "cp (t @ s) th' = Max (cp (t @ s) ` threads (t @ s))"}
|
|
1125 |
%\end{center}
|
|
|
140
389ef8b1959c
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
139
diff
changeset
|
1126 |
|
| 168 | 1127 |
%\noindent |
1128 |
%This is because each ready thread, say @{text "th\<^sub>r"}, has the maximum
|
|
1129 |
%current precedence of the subtree located at @{text "th\<^sub>r"}. All these
|
|
1130 |
%subtrees together form the set of threads. |
|
1131 |
%But the maximum of all threads is the @{term "cp"} of @{text "th"},
|
|
1132 |
%which is equal to the @{term preced} of @{text th}.\qed
|
|
1133 |
%\end{proof}
|
|
|
139
289e362f7a76
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
138
diff
changeset
|
1134 |
|
| 163 | 1135 |
%\endnote{
|
1136 |
%@{thm "th_blockedE_pretty"} -- thm-blockedE??
|
|
1137 |
% |
|
1138 |
% @{text "th_kept"} shows that th is a thread in s'-s
|
|
1139 |
% } |
|
|
139
289e362f7a76
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
138
diff
changeset
|
1140 |
|
| 174 | 1141 |
Given our assumptions (on @{text th}), the first property we show
|
1142 |
that a running thread @{text "th'"} must either wait for or hold a
|
|
1143 |
resource in state @{text s}.
|
|
|
142
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1144 |
|
|
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1145 |
\begin{lemma}\label{notdetached}
|
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1146 |
If @{term "th' \<in> running (s' @ s)"} and @{term "th \<noteq> th'"} then @{term "\<not>detached s th'"}.
|
|
142
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1147 |
\end{lemma}
|
|
138
20c1d3a14143
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
137
diff
changeset
|
1148 |
|
|
142
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1149 |
\begin{proof} Let us assume @{text "th'"} is detached in state
|
|
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1150 |
@{text "s"}, then, according to the definition of detached, @{text
|
|
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1151 |
"th’"} does not hold or wait for any resource. Hence the @{text
|
|
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1152 |
cp}-value of @{text "th'"} in @{text s} is not boosted, that is
|
| 166 | 1153 |
@{term "cp s th' = prec th' s"}, and is therefore lower than the
|
|
142
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1154 |
precedence (as well as the @{text "cp"}-value) of @{term "th"}. This
|
| 166 | 1155 |
means @{text "th'"} will not run as long as @{text "th"} is a live
|
1156 |
thread. In turn this means @{text "th'"} cannot take any action in
|
|
1157 |
state @{text "s' @ s"} to change its current status; therefore
|
|
1158 |
@{text "th'"} is still detached in state @{text "s' @ s"}.
|
|
1159 |
Consequently @{text "th'"} is also not boosted in state @{text "s' @
|
|
1160 |
s"} and would not run. This contradicts our assumption.\qed |
|
|
142
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1161 |
\end{proof}
|
|
139
289e362f7a76
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
138
diff
changeset
|
1162 |
|
|
289e362f7a76
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
138
diff
changeset
|
1163 |
|
|
142
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1164 |
\begin{proof}[of Theorem 1] If @{term "th \<in> running (s' @ s)"},
|
|
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1165 |
then there is nothing to show. So let us assume otherwise. Since the |
|
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1166 |
@{text "RAG"} is well-founded, we know there exists an ancestor of
|
|
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1167 |
@{text "th"} that is the root of the corrsponding subtree and
|
| 167 | 1168 |
therefore is ready (it does not request any resources). Let us call |
1169 |
this thread @{text "th'"}. Since in PIP the @{term "cpreced"}-value
|
|
1170 |
of any thread equals the maximum precedence of all threads in its |
|
1171 |
@{term "RAG"}-subtree, and @{text "th"} is in the subtree of @{text
|
|
1172 |
"th'"}, the @{term "cpreced"}-value of @{text "th'"} cannot be lower
|
|
1173 |
than the precedence of @{text "th"}. But, it can also not be higher,
|
|
1174 |
because the precedence of @{text "th"} is the maximum among all
|
|
1175 |
threads. Therefore we know that the @{term "cpreced"}-value of
|
|
1176 |
@{text "th'"} is the same as the precedence of @{text "th"}. The
|
|
1177 |
result is that @{text "th'"} must be running. This is because @{term
|
|
1178 |
"cpreced"}-value of @{text "th'"} is the highest of all ready
|
|
1179 |
threads. This follows from the fact that the @{term "cpreced"}-value
|
|
1180 |
of any ready thread is the maximum of the precedences of all threads |
|
1181 |
in its subtrees (with @{text "th"} having the highest of all threads
|
|
1182 |
and being in the subtree of @{text "th'"}). We also have that @{term
|
|
1183 |
"th \<noteq> th'"} since we assumed @{text th} is not running.
|
|
1184 |
By |
|
|
142
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1185 |
Lem.~\ref{notdetached} we have that @{term "\<not>detached s th'"}.
|
|
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1186 |
If @{text "th'"} is not detached in @{text s}, that is either
|
|
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1187 |
holding or waiting for a resource, it must be that @{term "th' \<in>
|
| 168 | 1188 |
threads s"}.\medskip |
1189 |
||
|
139
289e362f7a76
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
138
diff
changeset
|
1190 |
\noindent |
|
142
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
1191 |
This concludes the proof of Theorem 1.\qed |
|
138
20c1d3a14143
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
137
diff
changeset
|
1192 |
\end{proof}
|
|
20c1d3a14143
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
137
diff
changeset
|
1193 |
|
|
139
289e362f7a76
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
138
diff
changeset
|
1194 |
|
| 161 | 1195 |
%\endnote{
|
| 166 | 1196 |
|
| 169 | 1197 |
%In what follows we will describe properties of PIP that allow us to |
1198 |
% prove Theorem~\ref{mainthm} and, when instructive, briefly describe
|
|
1199 |
% our argument. Recall we want to prove that in state @ {term "s' @ s"}
|
|
1200 |
%either @{term th} is either running or blocked by a thread @
|
|
1201 |
% {term "th'"} (@{term "th \<noteq> th'"}) which was alive in state
|
|
1202 |
% @{term s}. We can show that
|
|
|
82
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1203 |
|
|
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1204 |
|
| 169 | 1205 |
% \begin{lemma}
|
1206 |
% If @{thm (prem 2) eq_pv_blocked}
|
|
1207 |
% then @{thm (concl) eq_pv_blocked}
|
|
1208 |
% \end{lemma}
|
|
|
82
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1209 |
|
| 169 | 1210 |
% \begin{lemma}
|
1211 |
% If @{thm (prem 2) eq_pv_persist}
|
|
1212 |
% then @{thm (concl) eq_pv_persist}
|
|
1213 |
% \end{lemma}
|
|
| 166 | 1214 |
%%%} |
|
82
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1215 |
|
| 162 | 1216 |
% \endnote{{\bf OUTLINE}
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1217 |
|
| 169 | 1218 |
% Since @{term "th"} is the most urgent thread, if it is somehow
|
1219 |
% blocked, people want to know why and wether this blocking is |
|
1220 |
% reasonable. |
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1221 |
|
| 169 | 1222 |
% @{thm [source] th_blockedE} @{thm th_blockedE}
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1223 |
|
| 169 | 1224 |
% if @{term "th"} is blocked, then there is a path leading from
|
1225 |
% @{term "th"} to @{term "th'"}, which means:
|
|
1226 |
% there is a chain of demand leading from @{term th} to @{term th'}.
|
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1227 |
|
| 169 | 1228 |
% in other words |
1229 |
% th -> cs1 -> th1 -> cs2 -> th2 -> ... -> csn -> thn -> cs -> th'. |
|
| 166 | 1230 |
|
| 169 | 1231 |
% We says that th is blocked by @{text "th'"}.
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1232 |
|
| 169 | 1233 |
% THEN |
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1234 |
|
| 169 | 1235 |
% @{thm [source] vat_t.th_chain_to_ready} @{thm vat_t.th_chain_to_ready}
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1236 |
|
| 169 | 1237 |
% It is basic propery with non-trival proof. |
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1238 |
|
| 169 | 1239 |
% THEN |
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1240 |
|
| 169 | 1241 |
% @{thm [source] max_preced} @{thm max_preced}
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1242 |
|
| 169 | 1243 |
% which says @{term "th"} holds the max precedence.
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1244 |
|
| 169 | 1245 |
% THEN |
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1246 |
|
| 169 | 1247 |
% @ {thm [source] th_cp_max th_cp_preced th_kept}
|
1248 |
% @ {thm th_cp_max th_cp_preced th_kept}
|
|
| 162 | 1249 |
|
| 169 | 1250 |
% THEN |
| 162 | 1251 |
|
| 169 | 1252 |
% ??? %%@ {thm [source] running_inversion_4} @ {thm running_inversion_4}
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1253 |
|
| 169 | 1254 |
% which explains what the @{term "th'"} looks like. Now, we have found the
|
1255 |
% @{term "th'"} which blocks @{term th}, we need to know more about it.
|
|
1256 |
% To see what kind of thread can block @{term th}.
|
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1257 |
|
| 169 | 1258 |
% From these two lemmas we can see the correctness of PIP, which is |
1259 |
% that: the blockage of th is reasonable and under control. |
|
| 162 | 1260 |
|
| 169 | 1261 |
% Lemmas we want to describe: |
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1262 |
|
| 169 | 1263 |
% \begin{lemma}
|
1264 |
% @{thm running_cntP_cntV_inv}
|
|
1265 |
% \end{lemma}
|
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1266 |
|
| 169 | 1267 |
% \noindent |
1268 |
% Remember we do not have the well-nestedness restriction in our |
|
1269 |
% proof, which means the difference between the counters @{const cntV}
|
|
1270 |
% and @{const cntP} can be larger than @{term 1}.
|
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1271 |
|
| 169 | 1272 |
% \begin{lemma}\label{runninginversion}
|
1273 |
% @ {thm running_inversion}
|
|
1274 |
% \end{lemma}
|
|
| 162 | 1275 |
|
| 169 | 1276 |
% explain tRAG |
| 162 | 1277 |
%} |
|
76
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
1278 |
|
|
82
c0a4e840aefe
some small changes to Correctness and Paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
76
diff
changeset
|
1279 |
|
| 162 | 1280 |
% Suppose the thread @ {term th} is \emph{not} running in state @ {term
|
1281 |
% "t @ s"}, meaning that it should be blocked by some other thread. |
|
1282 |
% It is first shown that there is a path in the RAG leading from node |
|
1283 |
% @ {term th} to another thread @ {text "th'"}, which is also in the
|
|
1284 |
% @ {term readys}-set. Since @ {term readys}-set is non-empty, there
|
|
1285 |
% must be one in it which holds the highest @ {term cp}-value, which,
|
|
1286 |
% by definition, is currently the @ {term running}-thread. However, we
|
|
1287 |
% are going to show in the next lemma slightly more: this running |
|
1288 |
% thread is exactly @ {term "th'"}.
|
|
|
76
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
1289 |
|
| 162 | 1290 |
% \begin{lemma}
|
1291 |
% There exists a thread @{text "th'"}
|
|
1292 |
% such that @{thm (no_quants) th_blockedE_pretty}.
|
|
1293 |
% \end{lemma}
|
|
|
125
95e7933968f8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
124
diff
changeset
|
1294 |
|
| 162 | 1295 |
% \begin{proof}
|
1296 |
% We know that @{term th} cannot be in @{term readys}, because it has
|
|
1297 |
% the highest precedence and therefore must be running. This violates our |
|
1298 |
% assumption. So by ?? we have that there must be a @{term "th'"} such that
|
|
1299 |
% @{term "th' \<in> readys (t @ s)"} and @{term "(Th th, Th th') \<in> (RAG (t @ s))\<^sup>+"}.
|
|
1300 |
% We are going to first show that this @{term "th'"} must be running. For this we
|
|
1301 |
% need to show that @{term th'} holds the highest @{term cp}-value.
|
|
1302 |
% By ?? we know that the @{term "cp"}-value of @{term "th'"} must
|
|
1303 |
% be the highest all precedences of all thread nodes in its @{term tRAG}-subtree.
|
|
1304 |
% That is |
|
|
76
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
1305 |
|
| 162 | 1306 |
% \begin{center}
|
1307 |
% @ {term "cp (t @ s) th' = Max (the_preced (t @ s) `
|
|
1308 |
% (the_thread ` subtree (tRAG (t @ s)) (Th th')))"} |
|
1309 |
% \end{center}
|
|
|
126
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
1310 |
|
| 162 | 1311 |
% But since @{term th} is in this subtree the right-hand side is equal
|
1312 |
% to @{term "preced th (t @ s)"}.
|
|
|
126
a88af0e4731f
minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
125
diff
changeset
|
1313 |
|
|
125
95e7933968f8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
124
diff
changeset
|
1314 |
%Let me distinguish between cp (current precedence) and assigned precedence (the precedence the |
|
95e7933968f8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
124
diff
changeset
|
1315 |
%thread ``normally'' has). |
|
95e7933968f8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
124
diff
changeset
|
1316 |
%So we want to show what the cp of th' is in state t @ s. |
|
95e7933968f8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
124
diff
changeset
|
1317 |
%We look at all the assingned precedences in the subgraph starting from th' |
|
95e7933968f8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
124
diff
changeset
|
1318 |
%We are looking for the maximum of these assigned precedences. |
|
95e7933968f8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
124
diff
changeset
|
1319 |
%This subgraph must contain the thread th, which actually has the highest precednence |
|
95e7933968f8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
124
diff
changeset
|
1320 |
%so cp of th' in t @ s has this (assigned) precedence of th |
|
95e7933968f8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
124
diff
changeset
|
1321 |
%We know that cp (t @ s) th' |
|
95e7933968f8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
124
diff
changeset
|
1322 |
%is the Maximum of the threads in the subgraph starting from th' |
|
95e7933968f8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
124
diff
changeset
|
1323 |
%this happens to be the precedence of th |
|
95e7933968f8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
124
diff
changeset
|
1324 |
%th has the highest precedence of all threads |
| 162 | 1325 |
% \end{proof}
|
|
125
95e7933968f8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
124
diff
changeset
|
1326 |
|
| 162 | 1327 |
% \begin{corollary}
|
1328 |
% Using the lemma \ref{runninginversion} we can say more about the thread th'
|
|
1329 |
% \end{corollary}
|
|
|
125
95e7933968f8
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
124
diff
changeset
|
1330 |
|
| 162 | 1331 |
% \endnote{\subsection*{END OUTLINE}}
|
|
76
b6ea51cd2e88
some small changes to the paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
75
diff
changeset
|
1332 |
|
| 162 | 1333 |
% In what follows we will describe properties of PIP that allow us to prove |
1334 |
% Theorem~\ref{mainthm} and, when instructive, briefly describe our argument.
|
|
1335 |
% It is relatively easy to see that: |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1336 |
|
| 162 | 1337 |
% \begin{isabelle}\ \ \ \ \ %%%
|
1338 |
% \begin{tabular}{@ {}l}
|
|
1339 |
% @ {text "running s \<subseteq> ready s \<subseteq> threads s"}\\
|
|
1340 |
% @ {thm[mode=IfThen] finite_threads}
|
|
1341 |
% \end{tabular}
|
|
1342 |
% \end{isabelle}
|
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1343 |
|
| 162 | 1344 |
% \noindent |
1345 |
% The second property is by induction on @{term vt}. The next three
|
|
1346 |
% properties are: |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1347 |
|
| 162 | 1348 |
% \begin{isabelle}\ \ \ \ \ %%%
|
1349 |
% \begin{tabular}{@ {}l}
|
|
1350 |
% HERE?? |
|
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1351 |
%@ {thm[mode=IfThen] waiting_unique[of _ _ "cs1" "cs2"]}\\
|
|
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1352 |
%@ {thm[mode=IfThen] held_unique[of _ "th1" _ "th2"]}\\
|
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
1353 |
%@ {thm[mode=IfThen] running_unique[of _ "th1" "th2"]}
|
| 162 | 1354 |
% \end{tabular}
|
1355 |
% \end{isabelle}
|
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1356 |
|
| 162 | 1357 |
% \noindent |
1358 |
% The first property states that every waiting thread can only wait for a single |
|
1359 |
% resource (because it gets suspended after requesting that resource); the second |
|
1360 |
% that every resource can only be held by a single thread; |
|
1361 |
% the third property establishes that in every given valid state, there is |
|
1362 |
% at most one running thread. We can also show the following properties |
|
1363 |
% about the @{term RAG} in @{text "s"}.
|
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1364 |
|
| 162 | 1365 |
% \begin{isabelle}\ \ \ \ \ %%%
|
1366 |
% \begin{tabular}{@ {}l}
|
|
1367 |
% HERE?? %@{text If}~@ {thm (prem 1) acyclic_RAG}~@{text "then"}:\\
|
|
1368 |
% \hspace{5mm}@{thm (concl) acyclic_RAG},
|
|
1369 |
% @{thm (concl) finite_RAG} and
|
|
1370 |
% %@ {thm (concl) wf_dep_converse},\\
|
|
1371 |
% %\hspace{5mm}@{text "if"}~@ {thm (prem 2) dm_RAG_threads}~@{text "then"}~@{thm (concl) dm_RAG_threads}
|
|
1372 |
% and\\ |
|
1373 |
% %\hspace{5mm}@{text "if"}~@ {thm (prem 2) range_in}~@{text "then"}~% @ {thm (concl) range_in}.
|
|
1374 |
% \end{tabular}
|
|
1375 |
% \end{isabelle}
|
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1376 |
|
| 162 | 1377 |
% \noindent |
1378 |
% The acyclicity property follows from how we restricted the events in |
|
1379 |
% @{text step}; similarly the finiteness and well-foundedness property.
|
|
1380 |
% The last two properties establish that every thread in a @{text "RAG"}
|
|
1381 |
% (either holding or waiting for a resource) is a live thread. |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1382 |
|
| 162 | 1383 |
% The key lemma in our proof of Theorem~\ref{mainthm} is as follows:
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1384 |
|
| 162 | 1385 |
% \begin{lemma}\label{mainlem}
|
1386 |
% Given the assumptions about states @{text "s"} and @{text "s' @ s"},
|
|
1387 |
% the thread @{text th} and the events in @{text "s'"},
|
|
1388 |
% if @{term "th' \<in> threads (s' @ s)"}, @{text "th' \<noteq> th"} and @{text "detached (s' @ s) th'"}\\
|
|
1389 |
% then @{text "th' \<notin> running (s' @ s)"}.
|
|
1390 |
% \end{lemma}
|
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1391 |
|
| 162 | 1392 |
% \noindent |
1393 |
% The point of this lemma is that a thread different from @{text th} (which has the highest
|
|
1394 |
% precedence in @{text s}) and not holding any resource, cannot be running
|
|
1395 |
% in the state @{text "s' @ s"}.
|
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1396 |
|
| 162 | 1397 |
% \begin{proof}
|
1398 |
% Since thread @{text "th'"} does not hold any resource, no thread can depend on it.
|
|
1399 |
% Therefore its current precedence @{term "cp (s' @ s) th'"} equals its own precedence
|
|
1400 |
% @{term "prec th' (s' @ s)"}. Since @{text "th"} has the highest precedence in the
|
|
1401 |
% state @{text "(s' @ s)"} and precedences are distinct among threads, we have
|
|
1402 |
% @{term "prec th' (s' @s ) < prec th (s' @ s)"}. From this
|
|
1403 |
% we have @{term "cp (s' @ s) th' < prec th (s' @ s)"}.
|
|
1404 |
% Since @{text "prec th (s' @ s)"} is already the highest
|
|
1405 |
% @{term "cp (s' @ s) th"} can not be higher than this and can not be lower either (by
|
|
1406 |
% definition of @{term "cp"}). Consequently, we have @{term "prec th (s' @ s) = cp (s' @ s) th"}.
|
|
1407 |
% Finally we have @{term "cp (s' @ s) th' < cp (s' @ s) th"}.
|
|
1408 |
% By defintion of @{text "running"}, @{text "th'"} can not be running in state
|
|
1409 |
% @{text "s' @ s"}, as we had to show.\qed
|
|
1410 |
% \end{proof}
|
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1411 |
|
| 162 | 1412 |
% \noindent |
1413 |
% Since @{text "th'"} is not able to run in state @{text "s' @ s"}, it is not able to
|
|
1414 |
% issue a @{text "P"} or @{text "V"} event. Therefore if @{text "s' @ s"} is extended
|
|
1415 |
% one step further, @{text "th'"} still cannot hold any resource. The situation will
|
|
1416 |
% not change in further extensions as long as @{text "th"} holds the highest precedence.
|
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1417 |
|
| 162 | 1418 |
% From this lemma we can deduce Theorem~\ref{mainthm}: that @{text th} can only be
|
1419 |
% blocked by a thread @{text th'} that
|
|
1420 |
% held some resource in state @{text s} (that is not @{text "detached"}). And furthermore
|
|
1421 |
% that the current precedence of @{text th'} in state @{text "(s' @ s)"} must be equal to the
|
|
1422 |
% precedence of @{text th} in @{text "s"}.
|
|
1423 |
% We show this theorem by induction on @{text "s'"} using Lemma~\ref{mainlem}.
|
|
1424 |
% This theorem gives a stricter bound on the threads that can block @{text th} than the
|
|
1425 |
% one obtained by Sha et al.~\cite{Sha90}:
|
|
1426 |
% only threads that were alive in state @{text s} and moreover held a resource.
|
|
1427 |
% This means our bound is in terms of both---alive threads in state @{text s}
|
|
1428 |
% and number of critical resources. Finally, the theorem establishes that the blocking threads have the |
|
1429 |
% current precedence raised to the precedence of @{text th}.
|
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1430 |
|
| 162 | 1431 |
% We can furthermore prove that under our assumptions no deadlock exists in the state @{text "s' @ s"}
|
1432 |
% by showing that @{text "running (s' @ s)"} is not empty.
|
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1433 |
|
| 162 | 1434 |
% \begin{lemma}
|
1435 |
% Given the assumptions about states @{text "s"} and @{text "s' @ s"},
|
|
1436 |
% the thread @{text th} and the events in @{text "s'"},
|
|
1437 |
% @{term "running (s' @ s) \<noteq> {}"}.
|
|
1438 |
% \end{lemma}
|
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1439 |
|
| 162 | 1440 |
% \begin{proof}
|
1441 |
% If @{text th} is blocked, then by following its dependants graph, we can always
|
|
1442 |
% reach a ready thread @{text th'}, and that thread must have inherited the
|
|
1443 |
% precedence of @{text th}.\qed
|
|
1444 |
% \end{proof}
|
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1445 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1446 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1447 |
%The following lemmas show how every node in RAG can be chased to ready threads: |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1448 |
%\begin{enumerate}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1449 |
%\item Every node in RAG can be chased to a ready thread (@{text "chain_building"}):
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1450 |
% @ {thm [display] chain_building[rule_format]}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1451 |
%\item The ready thread chased to is unique (@{text "dchain_unique"}):
|
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1452 |
% @ {thm [display] dchain_unique[of _ _ "th1" "th2"]}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1453 |
%\end{enumerate}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1454 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1455 |
%Some deeper results about the system: |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1456 |
%\begin{enumerate}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1457 |
%\item The maximum of @{term "cp"} and @{term "preced"} are equal (@{text "max_cp_eq"}):
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1458 |
%@ {thm [display] max_cp_eq}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1459 |
%\item There must be one ready thread having the max @{term "cp"}-value
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1460 |
%(@{text "max_cp_readys_threads"}):
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1461 |
%@ {thm [display] max_cp_readys_threads}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1462 |
%\end{enumerate}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1463 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1464 |
%The relationship between the count of @{text "P"} and @{text "V"} and the number of
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1465 |
%critical resources held by a thread is given as follows: |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1466 |
%\begin{enumerate}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1467 |
%\item The @{term "V"}-operation decreases the number of critical resources
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1468 |
% one thread holds (@{text "cntCS_v_dec"})
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1469 |
% @ {thm [display] cntCS_v_dec}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1470 |
%\item The number of @{text "V"} never exceeds the number of @{text "P"}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1471 |
% (@ {text "cnp_cnv_cncs"}):
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1472 |
% @ {thm [display] cnp_cnv_cncs}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1473 |
%\item The number of @{text "V"} equals the number of @{text "P"} when
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1474 |
% the relevant thread is not living: |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1475 |
% (@{text "cnp_cnv_eq"}):
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1476 |
% @ {thm [display] cnp_cnv_eq}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1477 |
%\item When a thread is not living, it does not hold any critical resource |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1478 |
% (@{text "not_thread_holdents"}):
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1479 |
% @ {thm [display] not_thread_holdents}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1480 |
%\item When the number of @{text "P"} equals the number of @{text "V"}, the relevant
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1481 |
% thread does not hold any critical resource, therefore no thread can depend on it |
|
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
1482 |
% (@{text "count_eq_dependants"}):
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
1483 |
% @ {thm [display] count_eq_dependants}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1484 |
%\end{enumerate}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1485 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1486 |
%The reason that only threads which already held some resoures |
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
1487 |
%can be running and block @{text "th"} is that if , otherwise, one thread
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1488 |
%does not hold any resource, it may never have its prioirty raised |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1489 |
%and will not get a chance to run. This fact is supported by |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1490 |
%lemma @{text "moment_blocked"}:
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1491 |
%@ {thm [display] moment_blocked}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1492 |
%When instantiating @{text "i"} to @{text "0"}, the lemma means threads which did not hold any
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1493 |
%resource in state @{text "s"} will not have a change to run latter. Rephrased, it means
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1494 |
%any thread which is running after @{text "th"} became the highest must have already held
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1495 |
%some resource at state @{text "s"}.
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1496 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1497 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1498 |
%When instantiating @{text "i"} to a number larger than @{text "0"}, the lemma means
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1499 |
%if a thread releases all its resources at some moment in @{text "t"}, after that,
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1500 |
%it may never get a change to run. If every thread releases its resource in finite duration, |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1501 |
%then after a while, only thread @{text "th"} is left running. This shows how indefinite
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1502 |
%priority inversion can be avoided. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1503 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1504 |
%All these assumptions are put into a predicate @{term "extend_highest_gen"}.
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1505 |
%It can be proved that @{term "extend_highest_gen"} holds
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1506 |
%for any moment @{text "i"} in it @{term "t"} (@{text "red_moment"}):
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1507 |
%@ {thm [display] red_moment}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1508 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1509 |
%From this, an induction principle can be derived for @{text "t"}, so that
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1510 |
%properties already derived for @{term "t"} can be applied to any prefix
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1511 |
%of @{text "t"} in the proof of new properties
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1512 |
%about @{term "t"} (@{text "ind"}):
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1513 |
%\begin{center}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1514 |
%@ {thm[display] ind}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1515 |
%\end{center}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1516 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1517 |
%The following properties can be proved about @{term "th"} in @{term "t"}:
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1518 |
%\begin{enumerate}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1519 |
%\item In @{term "t"}, thread @{term "th"} is kept live and its
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1520 |
% precedence is preserved as well |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1521 |
% (@{text "th_kept"}):
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1522 |
% @ {thm [display] th_kept}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1523 |
%\item In @{term "t"}, thread @{term "th"}'s precedence is always the maximum among
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1524 |
% all living threads |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1525 |
% (@{text "max_preced"}):
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1526 |
% @ {thm [display] max_preced}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1527 |
%\item In @{term "t"}, thread @{term "th"}'s current precedence is always the maximum precedence
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1528 |
% among all living threads |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1529 |
% (@{text "th_cp_max_preced"}):
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1530 |
% @ {thm [display] th_cp_max_preced}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1531 |
%\item In @{term "t"}, thread @{term "th"}'s current precedence is always the maximum current
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1532 |
% precedence among all living threads |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1533 |
% (@{text "th_cp_max"}):
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1534 |
% @ {thm [display] th_cp_max}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1535 |
%\item In @{term "t"}, thread @{term "th"}'s current precedence equals its precedence at moment
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1536 |
% @{term "s"}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1537 |
% (@{text "th_cp_preced"}):
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1538 |
% @ {thm [display] th_cp_preced}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1539 |
%\end{enumerate}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1540 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1541 |
%The main theorem of this part is to characterizing the running thread during @{term "t"}
|
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
1542 |
%(@{text "running_inversion_2"}):
|
|
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
1543 |
%@ {thm [display] running_inversion_2}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1544 |
%According to this, if a thread is running, it is either @{term "th"} or was
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1545 |
%already live and held some resource |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1546 |
%at moment @{text "s"} (expressed by: @{text "cntV s th' < cntP s th'"}).
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1547 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1548 |
%Since there are only finite many threads live and holding some resource at any moment, |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1549 |
%if every such thread can release all its resources in finite duration, then after finite |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1550 |
%duration, none of them may block @{term "th"} anymore. So, no priority inversion may happen
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1551 |
%then. |
|
70
92ca2410b3d9
further simplificaton of Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
64
diff
changeset
|
1552 |
|
| 162 | 1553 |
% NOTE: about bounds in sha et al and ours: they prove a bound on the length of individual |
1554 |
% blocages. We prove a bound for the overall-blockage. |
|
|
70
92ca2410b3d9
further simplificaton of Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
64
diff
changeset
|
1555 |
|
| 162 | 1556 |
% There are low priority threads, |
1557 |
% which do not hold any resources, |
|
1558 |
% such thread will not block th. |
|
1559 |
% Their Theorem 3 does not exclude such threads. |
|
|
70
92ca2410b3d9
further simplificaton of Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
64
diff
changeset
|
1560 |
|
| 162 | 1561 |
% There are resources, which are not held by any low prioirty threads, |
1562 |
% such resources can not cause blockage of th neither. And similiary, |
|
1563 |
% theorem 6 does not exlude them. |
|
|
70
92ca2410b3d9
further simplificaton of Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
64
diff
changeset
|
1564 |
|
| 162 | 1565 |
% Our one bound excudle them by using a different formaulation. " |
|
70
92ca2410b3d9
further simplificaton of Moment.thy
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
64
diff
changeset
|
1566 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1567 |
*} |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1568 |
(*<*) |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1569 |
end |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1570 |
(*>*) |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1571 |
|
| 162 | 1572 |
(*text {*
|
|
155
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1573 |
explan why Thm1 roughly corresponds to Sha's Thm 3 |
| 162 | 1574 |
*}*) |
|
155
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1575 |
|
|
152
15f4481bc0c9
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1576 |
section {* A Finite Bound on Priority Inversion *}
|
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1577 |
|
|
154
9756a51f2223
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
152
diff
changeset
|
1578 |
(*<*) |
|
9756a51f2223
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
152
diff
changeset
|
1579 |
context extend_highest_gen |
|
9756a51f2223
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
152
diff
changeset
|
1580 |
begin |
|
9756a51f2223
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
152
diff
changeset
|
1581 |
(*>*) |
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1582 |
text {*
|
|
152
15f4481bc0c9
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1583 |
|
| 158 | 1584 |
Like in the work by Sha et al.~our result in Thm 1 does not yet |
1585 |
guarantee the absence of indefinite Priority Inversion. For this we |
|
1586 |
further need the property that every thread gives up its resources |
|
1587 |
after a finite amount of time. We found that this property is not so |
|
|
154
9756a51f2223
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
152
diff
changeset
|
1588 |
straightforward to formalise in our model. There are mainly two |
|
152
15f4481bc0c9
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1589 |
reasons for this: First, we do not specify what ``running'' the code |
|
15f4481bc0c9
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1590 |
of a thread means, for example by giving an operational semantics |
|
15f4481bc0c9
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1591 |
for machine instructions. Therefore we cannot characterise what are |
| 170 | 1592 |
``good'' programs that contain for every locking request for a |
1593 |
resource also a corresponding unlocking request. Second, we need to |
|
|
154
9756a51f2223
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
152
diff
changeset
|
1594 |
distinghish between a thread that ``just'' locks a resource for a |
|
9756a51f2223
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
152
diff
changeset
|
1595 |
finite amount of time (even if it is very long) and one that locks |
| 176 | 1596 |
it forever (there might be an unbounded loop in between the locking and |
| 158 | 1597 |
unlocking requests). |
|
155
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1598 |
|
|
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1599 |
Because of these problems, we decided in our earlier paper |
|
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1600 |
\cite{ZhangUrbanWu12} to leave out this property and let the
|
| 170 | 1601 |
programmer take on the responsibility to program threads in such a |
|
155
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1602 |
benign manner (in addition to causing no circularity in the |
| 158 | 1603 |
RAG). This leave-it-to-the-programmer was also the approach taken by |
1604 |
Sha et al.~in their paper. However, in this paper we can make an |
|
| 170 | 1605 |
improvement by establishing a finite bound on the duration of |
1606 |
Priority Inversion measured by the number of events. The events can |
|
1607 |
be seen as a \textit{rough(!)} abstraction of the ``runtime
|
|
1608 |
behaviour'' of threads and also as an abstract notion of |
|
| 176 | 1609 |
``time''---when a new event happens, some time must have passed. |
| 171 | 1610 |
|
|
152
15f4481bc0c9
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1611 |
|
|
155
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1612 |
What we will establish in this section is that there can only be a |
| 170 | 1613 |
finite number of states after state @{term s} in which the thread
|
| 176 | 1614 |
@{term th} is blocked (recall for this that a state is a list of
|
1615 |
events). For this finiteness bound to exist, Sha et al.~informally |
|
1616 |
make two assumtions: first, there is a finite pool of threads |
|
1617 |
(active or hibernating) and second, each of them giving up its |
|
1618 |
resources after a finite amount of time. However, we do not have |
|
1619 |
this concept of active or hibernating threads in our model. In fact |
|
1620 |
we can dispence with the first assumption altogether and allow that |
|
1621 |
in our model we can create new threads or exit existing threads |
|
| 170 | 1622 |
arbitrarily. Consequently, the avoidance of indefinite priority |
1623 |
inversion we are trying to establish is in our model not true, |
|
| 176 | 1624 |
unless we stipulate an upper bound on the number of threads that |
1625 |
have been created during the time leading to any future state |
|
1626 |
after @{term s}. Otherwise our PIP scheduler could be ``swamped''
|
|
1627 |
with @{text "Create"}-requests. So our first assumption states:
|
|
|
155
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1628 |
|
| 172 | 1629 |
\begin{quote} {\bf Assumption on the number of threads created
|
1630 |
after the state {\boldmath@{text s}}:} Given the
|
|
| 174 | 1631 |
state @{text s}, in every ``future'' valid state @{text "es @ s"}, we
|
|
155
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1632 |
require that the number of created threads is less than |
| 157 | 1633 |
a bound @{text "BC"}, that is
|
1634 |
||
| 174 | 1635 |
\[@{text "len (filter isCreate es) < BC"}\;\]
|
1636 |
||
1637 |
wherby @{text es} is a list of events.
|
|
| 157 | 1638 |
\end{quote}
|
|
155
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1639 |
|
| 157 | 1640 |
\noindent Note that it is not enough to just to state that there are |
| 171 | 1641 |
only finite number of threads created up until a single state @{text
|
1642 |
"s' @ s"} after @{text s}. Instead, we need to put this bound on
|
|
1643 |
the @{text "Create"} events for all valid states after @{text s}.
|
|
| 174 | 1644 |
This ensures that no matter which ``future'' state is reached, the |
1645 |
number of @{text "Create"}-events is finite. We use @{text "es @ s"}
|
|
1646 |
to stand for \emph{future states} after @{text s}---it is @{text s}
|
|
| 176 | 1647 |
extended with some list @{text es} of events.
|
|
155
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1648 |
|
|
156
550ab0f68960
updasted
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
155
diff
changeset
|
1649 |
For our second assumption about giving up resources after a finite |
| 157 | 1650 |
amount of ``time'', let us introduce the following definition about |
1651 |
threads that can potentially block @{text th}:
|
|
|
152
15f4481bc0c9
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
145
diff
changeset
|
1652 |
|
|
154
9756a51f2223
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
152
diff
changeset
|
1653 |
\begin{isabelle}\ \ \ \ \ %%%
|
| 170 | 1654 |
@{thm blockers_def[THEN eq_reflection]}
|
|
154
9756a51f2223
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
152
diff
changeset
|
1655 |
\end{isabelle}
|
|
9756a51f2223
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
152
diff
changeset
|
1656 |
|
| 157 | 1657 |
\noindent This set contains all treads that are not detached in |
| 172 | 1658 |
state @{text s}. According to our definiton of @{text "detached"},
|
1659 |
this means a thread in @{text "blockers"} either holds or waits for
|
|
| 176 | 1660 |
some resource in state @{text s} . Our Them~1 implies that any of
|
1661 |
those threads can all potentially block @{text th} after state
|
|
1662 |
@{text s}. We need to make the following assumption about the
|
|
1663 |
threads in the @{text "blockers"}-set:
|
|
|
155
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1664 |
|
|
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1665 |
\begin{quote}
|
|
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1666 |
{\bf Assumptions on the threads {\boldmath{@{term "th' \<in> blockers"}}}:}
|
| 158 | 1667 |
For each such @{text "th'"} there exists a finite bound @{text "BND(th')"}
|
|
156
550ab0f68960
updasted
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
155
diff
changeset
|
1668 |
such that for all future |
| 174 | 1669 |
valid states @{text "es @ s"},
|
1670 |
we have that if \mbox{@{term "\<not>(detached (es @ s) th')"}}, then
|
|
1671 |
\[@{text "len (actions_of {th'} es) < BND(th')"}\]
|
|
|
156
550ab0f68960
updasted
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
155
diff
changeset
|
1672 |
\end{quote}
|
|
155
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1673 |
|
| 157 | 1674 |
\noindent By this assumption we enforce that any thread potentially |
| 158 | 1675 |
blocking @{term th} must become detached (that is lock no resource
|
| 174 | 1676 |
anymore) after a finite number of events in @{text "es @ s"}. Again
|
| 158 | 1677 |
we have to state this bound to hold in all valid states after @{text
|
1678 |
s}. The bound reflects how each thread @{text "th'"} is programmed:
|
|
1679 |
Though we cannot express what instructions a thread is executing, |
|
1680 |
the events in our model correspond to the system calls made by |
|
| 176 | 1681 |
a thread. Our @{text "BND(th')"} binds the number of these ``calls''.
|
|
155
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1682 |
|
| 170 | 1683 |
The main reason for these two assumptions is that we can prove the |
1684 |
following: The number of states after @{text s} in which the thread
|
|
1685 |
@{text th} is not running (that is where Priority Inversion occurs)
|
|
1686 |
can be bounded by the number of actions the threads in @{text
|
|
| 176 | 1687 |
blockers} perform (i.e.~events) and how many threads are newly |
1688 |
created. To state our bound formally, we need to make a definition |
|
1689 |
of what we mean by intermediate states between a state @{text s} and
|
|
1690 |
a future state after @{text s}; they will be the list of states
|
|
1691 |
starting from @{text s} upto the state \mbox{@{text "es @ s"}}. For
|
|
1692 |
example, suppose $\textit{es} = [\textit{e}_n, \textit{e}_{n-1},
|
|
1693 |
\ldots, \textit{e}_2, \textit{e}_1]$, then the intermediate states
|
|
1694 |
from @{text s} upto @{text "es @ s"} are
|
|
| 157 | 1695 |
|
| 158 | 1696 |
\begin{center}
|
| 170 | 1697 |
\begin{tabular}{l}
|
1698 |
$\textit{s}$\\
|
|
1699 |
$\textit{e}_1 :: \textit{s}$\\
|
|
1700 |
$\textit{e}_2 :: \textit{e}_1 :: \textit{s}$\\
|
|
1701 |
\ldots\\ |
|
1702 |
$\textit{e}_{n - 1} :: \ldots :: \textit{e}_2 :: \textit{e}_1 :: \textit{s}$\\
|
|
1703 |
\end{tabular}
|
|
| 158 | 1704 |
\end{center}
|
| 157 | 1705 |
|
| 170 | 1706 |
|
| 172 | 1707 |
\noindent This list of \emph{intermediate states} can be defined by
|
1708 |
the following recursive function |
|
| 158 | 1709 |
|
1710 |
\begin{center}
|
|
| 170 | 1711 |
\begin{tabular}{lcl}
|
1712 |
@{text "s upto []"} & $\dn$ & $[]$\\
|
|
| 174 | 1713 |
@{text "s upto (_::es)"} & $\dn$ & @{text "(es @ s) :: s upto es"}
|
| 158 | 1714 |
\end{tabular}
|
1715 |
\end{center}
|
|
|
156
550ab0f68960
updasted
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
155
diff
changeset
|
1716 |
|
| 174 | 1717 |
\noindent |
1718 |
Our theorem can then be stated as follows: |
|
| 158 | 1719 |
|
| 173 | 1720 |
\begin{theorem}
|
1721 |
Given our assumptions about bounds, we have that |
|
1722 |
\[ |
|
1723 |
@{text "len"}\,[@{text "s'"}
|
|
| 174 | 1724 |
\leftarrow @{text "s upto es"}.\;\; @{text "th"} \not\in @{text "running s'"}] \;\;\leq\;\;
|
1725 |
@{text "BC"} + \sum @{text "th'"} \in @{text "blockers"}.\;\; @{text "BND(th')"}\;.
|
|
| 173 | 1726 |
\] |
| 176 | 1727 |
\end{theorem}
|
| 173 | 1728 |
|
| 176 | 1729 |
\noindent This theorem uses Isabelle's list-comprehension notation, |
1730 |
which lists all intermediate states between @{text s} and @{text "es
|
|
1731 |
@ s"}, and then filters this list according to states in which |
|
1732 |
@{text th} is not running. By calculating the number of elements in
|
|
1733 |
the filtered list using the function @{text len}, we have the number
|
|
1734 |
of intermediate states in which @{text th} is not running and which
|
|
1735 |
by the theorem is bounded by the term on the right-hand side. |
|
| 173 | 1736 |
|
| 174 | 1737 |
\begin{proof} There are two characterisations for the number of
|
| 176 | 1738 |
events in @{text es}: First, in each state in
|
1739 |
@{text "s upto es"}, clearly either @{text th} is running or
|
|
1740 |
not running. Together with @{text "len es = len (s upto es)"}, that
|
|
1741 |
implies % |
|
1742 |
||
1743 |
\begin{equation}
|
|
1744 |
\label{firsteq}
|
|
1745 |
\begin{array}{lcl}
|
|
1746 |
@{text "len es"} & \;=\; &
|
|
1747 |
@{text len}\, [@{text "s'"} \leftarrow @{text "s upto es"}.\;\; @{text "th"} \in @{text "running s'"}]\\
|
|
1748 |
& & +\; |
|
1749 |
@{text len}\, [@{text "s'"} \leftarrow @{text "s upto es"}.\;\; @{text "th"} \not\in @{text "running s'"}]
|
|
1750 |
\end{array}
|
|
| 174 | 1751 |
\end{equation}
|
| 158 | 1752 |
|
| 174 | 1753 |
\noindent Second by Thm~\ref{mainthm}, the events are either the
|
1754 |
actions of @{text th} or @{text "Create"}-events or actions of the
|
|
1755 |
threads in blockers. That is |
|
| 176 | 1756 |
% |
| 174 | 1757 |
\begin{equation}\label{secondeq}
|
| 176 | 1758 |
\begin{array}{lcl}
|
1759 |
@{text "len es"} & \;=\; & @{text "len (actions_of {th} es)"}\\
|
|
1760 |
& & +\; @{text "len (filter isCreate es)"}\\
|
|
1761 |
& & +\; @{text "len (actions_of blockers es)"}
|
|
1762 |
\end{array}
|
|
| 174 | 1763 |
\end{equation}
|
1764 |
||
| 176 | 1765 |
\noindent Furthermore we know that an action of @{text th} in the
|
1766 |
intermediate states @{text "s upto es"} can only be taken when
|
|
1767 |
@{text th} is running. Therefore
|
|
1768 |
% |
|
| 174 | 1769 |
\[ |
| 176 | 1770 |
@{text "len (actions_of {th} es)"} \;\leq\;
|
1771 |
@{text len}\,[@{text "s'"} \leftarrow @{text "s upto es"}.\;\; @{text "th"} \in @{text "running s'"}]
|
|
| 174 | 1772 |
\] |
| 158 | 1773 |
|
| 176 | 1774 |
\noindent holds. Substituting this into \eqref{firsteq} gives
|
1775 |
% |
|
| 174 | 1776 |
\[ |
| 176 | 1777 |
@{text len}\,[@{text "s'"} \leftarrow @{text "s upto es"}.\;\; @{text "th"} \not\in @{text "running s'"}]
|
1778 |
\;\leq\; @{text "len es"} - @{text "len (actions_of {th} es)"}
|
|
| 174 | 1779 |
\] |
|
155
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1780 |
|
| 176 | 1781 |
\noindent |
| 174 | 1782 |
into which we can substitute \eqref{secondeq} yielding
|
| 176 | 1783 |
% |
| 174 | 1784 |
\[ |
| 176 | 1785 |
\begin{array}{rcl}
|
1786 |
@{text len} [@{text "s'"} \leftarrow @{text "s upto es"}.\;\; @{text "th"} \not\in @{text "running s'"}] & \;\;\leq\;\; &
|
|
1787 |
@{text "len (filter isCreate es)"}\\
|
|
1788 |
& & \quad + @{text "len (actions_of blockers es)"}
|
|
1789 |
\end{array}
|
|
| 174 | 1790 |
\] |
|
155
eae86cba8b89
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
154
diff
changeset
|
1791 |
|
| 176 | 1792 |
\noindent By our first assumption we know that the number of @{text
|
| 174 | 1793 |
"Create"}-events are bounded by the bound @{text BC}. By our second
|
1794 |
assumption we can prove that the actions of all blockers is bounded |
|
1795 |
by the sum of bounds of the individual blocking threads, that is |
|
1796 |
||
1797 |
\[ |
|
1798 |
@{text "len (actions_of blockers es)"} \;\;\leq\;\;
|
|
1799 |
\sum @{text "th'"} \in @{text "blockers"}.\;\; @{text "BND(th')"}
|
|
1800 |
\] |
|
|
154
9756a51f2223
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
152
diff
changeset
|
1801 |
|
| 174 | 1802 |
\noindent With this in place we can conclude our theorem.\hfill\qed |
1803 |
\end{proof}
|
|
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1804 |
|
| 174 | 1805 |
\noindent This theorem is the main conclusion we obtain for the |
| 176 | 1806 |
Priority Inheritance Protocol. It is based on the fact that the set of |
1807 |
@{text blockers} is fixed at state @{text s} when @{text th} becomes
|
|
1808 |
the thread with highest priority. Then no additional blocker of |
|
1809 |
@{text th} can appear after the state @{text s}. And in this way we
|
|
1810 |
can bound the number of states where the thread @{text th} with the
|
|
1811 |
highest priority is prevented from running. |
|
1812 |
Our bound does not depend on the restriction of well-nested critical |
|
1813 |
sections in the Priority Inheritance Protocol as imposed by Sha et al. |
|
1814 |
*} (*<*) end (*>*) |
|
|
145
188fe0c81ac7
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
144
diff
changeset
|
1815 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1816 |
section {* Properties for an Implementation\label{implement} *}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1817 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1818 |
text {*
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1819 |
While our formalised proof gives us confidence about the correctness of our model of PIP, |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1820 |
we found that the formalisation can even help us with efficiently implementing it. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1821 |
For example Baker complained that calculating the current precedence |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1822 |
in PIP is quite ``heavy weight'' in Linux (see the Introduction). |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1823 |
In our model of PIP the current precedence of a thread in a state @{text s}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1824 |
depends on all its dependants---a ``global'' transitive notion, |
|
23
24e6884d9258
made some small chages
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
22
diff
changeset
|
1825 |
which is indeed heavy weight (see Definition shown in \eqref{cpreced}).
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1826 |
We can however improve upon this. For this let us define the notion |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1827 |
of @{term children} of a thread @{text th} in a state @{text s} as
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1828 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1829 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1830 |
\begin{tabular}{@ {}l}
|
| 176 | 1831 |
?? @{thm children_def}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1832 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1833 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1834 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1835 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1836 |
where a child is a thread that is only one ``hop'' away from the thread |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1837 |
@{text th} in the @{term RAG} (and waiting for @{text th} to release
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1838 |
a resource). We can prove the following lemma. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1839 |
|
| 178 | 1840 |
\begin{center}
|
1841 |
@{thm tG_alt_def}\\
|
|
1842 |
@{thm dependants_alt_tG}
|
|
1843 |
\end{center}
|
|
1844 |
||
1845 |
\begin{center}
|
|
1846 |
@{thm valid_trace.cp_alt_def3}
|
|
1847 |
\end{center}
|
|
1848 |
||
1849 |
||
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1850 |
\begin{lemma}\label{childrenlem}
|
| 176 | 1851 |
HERE %@ {text "If"} @ {thm (prem 1) cp_rec} @{text "then"}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1852 |
\begin{center}
|
| 178 | 1853 |
@{thm valid_trace.cp_rec_tG}.
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1854 |
%@ {thm (concl) cp_rec}.
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1855 |
\end{center}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1856 |
\end{lemma}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1857 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1858 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1859 |
That means the current precedence of a thread @{text th} can be
|
|
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:
33
diff
changeset
|
1860 |
computed locally by considering only the current precedences of the children of @{text th}. In
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1861 |
effect, it only needs to be recomputed for @{text th} when one of
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1862 |
its children changes its current precedence. Once the current |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1863 |
precedence is computed in this more efficient manner, the selection |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1864 |
of the thread with highest precedence from a set of ready threads is |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1865 |
a standard scheduling operation implemented in most operating |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1866 |
systems. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1867 |
|
|
45
fc83f79009bd
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
44
diff
changeset
|
1868 |
%\begin{proof}[of Lemma~\ref{childrenlem}]
|
|
fc83f79009bd
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
44
diff
changeset
|
1869 |
%Test |
|
fc83f79009bd
updated for Isabelle 2015
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
44
diff
changeset
|
1870 |
%\end{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:
33
diff
changeset
|
1871 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1872 |
Of course the main work for implementing PIP involves the |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1873 |
scheduler and coding how it should react to events. Below we |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1874 |
outline how our formalisation guides this implementation for each |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1875 |
kind of events.\smallskip |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1876 |
*} |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1877 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1878 |
text {*
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1879 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1880 |
\colorbox{mygrey}{@{term "Create th prio"}:} We assume that the current state @{text s'} and
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1881 |
the next state @{term "s \<equiv> Create th prio#s'"} are both valid (meaning the event
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1882 |
is allowed to occur). In this situation we can show that |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1883 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1884 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1885 |
\begin{tabular}{@ {}l}
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1886 |
HERE ?? %@ {thm eq_dep},\\
|
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1887 |
@{thm valid_trace_create.eq_cp_th}, and\\
|
|
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1888 |
@{thm[mode=IfThen] valid_trace_create.eq_cp}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1889 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1890 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1891 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1892 |
\noindent |
|
20
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
1893 |
This means in an implementation we do not have to recalculate the @{text RAG} and also none of the
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1894 |
current precedences of the other threads. The current precedence of the created |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1895 |
thread @{text th} is just its precedence, namely the pair @{term "(prio, length (s::event list))"}.
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1896 |
\smallskip |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1897 |
*} |
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1898 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1899 |
text {*
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1900 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1901 |
\colorbox{mygrey}{@{term "Exit th"}:} We again assume that the current state @{text s'} and
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1902 |
the next state @{term "s \<equiv> Exit th#s'"} are both valid. We can show that
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1903 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1904 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1905 |
\begin{tabular}{@ {}l}
|
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1906 |
HERE %@ {thm valid_trace_create.eq_dep}, and\\
|
|
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1907 |
@{thm[mode=IfThen] valid_trace_create.eq_cp}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1908 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1909 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1910 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1911 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1912 |
This means again we do not have to recalculate the @{text RAG} and
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1913 |
also not the current precedences for the other threads. Since @{term th} is not
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1914 |
alive anymore in state @{term "s"}, there is no need to calculate its
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1915 |
current precedence. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1916 |
\smallskip |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1917 |
*} |
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1918 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1919 |
text {*
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1920 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1921 |
\colorbox{mygrey}{@{term "Set th prio"}:} We assume that @{text s'} and
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1922 |
@{term "s \<equiv> Set th prio#s'"} are both valid. We can show that
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1923 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1924 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1925 |
\begin{tabular}{@ {}l}
|
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1926 |
%@ {thm[mode=IfThen] eq_dep}, and\\
|
|
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1927 |
%@ {thm[mode=IfThen] valid_trace_create.eq_cp_pre}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1928 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1929 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1930 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1931 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1932 |
The first property is again telling us we do not need to change the @{text RAG}.
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1933 |
The second shows that the @{term cp}-values of all threads other than @{text th}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1934 |
are unchanged. The reason is that @{text th} is running; therefore it is not in
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1935 |
the @{term dependants} relation of any other thread. This in turn means that the
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1936 |
change of its priority cannot affect other threads. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1937 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1938 |
%The second |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1939 |
%however states that only threads that are \emph{not} dependants of @{text th} have their
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1940 |
%current precedence unchanged. For the others we have to recalculate the current |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1941 |
%precedence. To do this we can start from @{term "th"}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1942 |
%and follow the @{term "depend"}-edges to recompute using Lemma~\ref{childrenlem}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1943 |
%the @{term "cp"} of every
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1944 |
%thread encountered on the way. Since the @{term "depend"}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1945 |
%is assumed to be loop free, this procedure will always stop. The following two lemmas show, however, |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1946 |
%that this procedure can actually stop often earlier without having to consider all |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1947 |
%dependants. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1948 |
% |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1949 |
%\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1950 |
%\begin{tabular}{@ {}l}
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1951 |
%@ {thm[mode=IfThen] eq_up_self}\\
|
|
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1952 |
%@{text "If"} @ {thm (prem 1) eq_up}, @ {thm (prem 2) eq_up} and @ {thm (prem 3) eq_up}\\
|
|
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
1953 |
%@{text "then"} @ {thm (concl) eq_up}.
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1954 |
%\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1955 |
%\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1956 |
% |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1957 |
%\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1958 |
%The first lemma states that if the current precedence of @{text th} is unchanged,
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1959 |
%then the procedure can stop immediately (all dependent threads have their @{term cp}-value unchanged).
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1960 |
%The second states that if an intermediate @{term cp}-value does not change, then
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1961 |
%the procedure can also stop, because none of its dependent threads will |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1962 |
%have their current precedence changed. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1963 |
\smallskip |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1964 |
*} |
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1965 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1966 |
text {*
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1967 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1968 |
\colorbox{mygrey}{@{term "V th cs"}:} We assume that @{text s'} and
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1969 |
@{term "s \<equiv> V th cs#s'"} are both valid. We have to consider two
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1970 |
subcases: one where there is a thread to ``take over'' the released |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1971 |
resource @{text cs}, and one where there is not. Let us consider them
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1972 |
in turn. Suppose in state @{text s}, the thread @{text th'} takes over
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1973 |
resource @{text cs} from thread @{text th}. We can prove
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1974 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1975 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1976 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1977 |
%@ {thm RAG_s}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1978 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1979 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1980 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1981 |
which shows how the @{text RAG} needs to be changed. The next lemma suggests
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1982 |
how the current precedences need to be recalculated. For threads that are |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1983 |
not @{text "th"} and @{text "th'"} nothing needs to be changed, since we
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1984 |
can show |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1985 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1986 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1987 |
%@ {thm[mode=IfThen] cp_kept}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1988 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1989 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1990 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1991 |
For @{text th} and @{text th'} we need to use Lemma~\ref{childrenlem} to
|
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1992 |
recalculate their current precedence since their children have changed. *} |
|
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1993 |
|
|
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
1994 |
text {*
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1995 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1996 |
In the other case where there is no thread that takes over @{text cs}, we can show how
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1997 |
to recalculate the @{text RAG} and also show that no current precedence needs
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1998 |
to be recalculated. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1999 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2000 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2001 |
\begin{tabular}{@ {}l}
|
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
2002 |
%@ {thm RAG_s}\\
|
|
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
2003 |
%@ {thm eq_cp}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2004 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2005 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2006 |
*} |
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
2007 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2008 |
text {*
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2009 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2010 |
\colorbox{mygrey}{@{term "P th cs"}:} We assume that @{text s'} and
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2011 |
@{term "s \<equiv> P th cs#s'"} are both valid. We again have to analyse two subcases, namely
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2012 |
the one where @{text cs} is not locked, and one where it is. We treat the former case
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2013 |
first by showing that |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2014 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2015 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2016 |
\begin{tabular}{@ {}l}
|
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
2017 |
%@ {thm RAG_s}\\
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
2018 |
HERE %@ {thm eq_cp}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2019 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2020 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2021 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2022 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2023 |
This means we need to add a holding edge to the @{text RAG} and no
|
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
2024 |
current precedence needs to be recalculated.*} |
|
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
2025 |
|
|
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
2026 |
text {*
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2027 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2028 |
In the second case we know that resource @{text cs} is locked. We can show that
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2029 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2030 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2031 |
\begin{tabular}{@ {}l}
|
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
2032 |
%@ {thm RAG_s}\\
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
2033 |
HERE %@ {thm[mode=IfThen] eq_cp}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2034 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2035 |
\end{isabelle}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2036 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2037 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2038 |
That means we have to add a waiting edge to the @{text RAG}. Furthermore
|
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
2039 |
the current precedence for all threads that are not dependants of @{text "th'"}
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2040 |
are unchanged. For the others we need to follow the edges |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2041 |
in the @{text RAG} and recompute the @{term "cp"}. To do this we can start from @{term "th"}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2042 |
and follow the @{term "depend"}-edges to recompute using Lemma~\ref{childrenlem}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2043 |
the @{term "cp"} of every
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2044 |
thread encountered on the way. Since the @{term "depend"}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2045 |
is loop free, this procedure will always stop. The following lemma shows, however, |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2046 |
that this procedure can actually stop often earlier without having to consider all |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2047 |
dependants. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2048 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2049 |
\begin{isabelle}\ \ \ \ \ %%%
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2050 |
\begin{tabular}{@ {}l}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2051 |
%%@ {t hm[mode=IfThen] eq_up_self}\\
|
|
64
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
2052 |
HERE |
|
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
2053 |
%@{text "If"} @ {thm (prem 1) eq_up}, @ {thm (prem 2) eq_up} and @ {thm (prem 3) eq_up}\\
|
|
b4bcd1edbb6d
renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
46
diff
changeset
|
2054 |
%@{text "then"} @ {thm (concl) eq_up}.
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2055 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2056 |
\end{isabelle}
|
|
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2057 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2058 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2059 |
This lemma states that if an intermediate @{term cp}-value does not change, then
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2060 |
the procedure can also stop, because none of its dependent threads will |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2061 |
have their current precedence changed. |
|
124
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
2062 |
*} |
|
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
2063 |
|
|
71a3300d497b
updated journal paper
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
95
diff
changeset
|
2064 |
text {*
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2065 |
As can be seen, a pleasing byproduct of our formalisation is that the properties in |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2066 |
this section closely inform an implementation of PIP, namely whether the |
|
12
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
2067 |
RAG needs to be reconfigured or current precedences need to |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2068 |
be recalculated for an event. This information is provided by the lemmas we proved. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2069 |
We confirmed that our observations translate into practice by implementing |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2070 |
our version of PIP on top of PINTOS, a small operating system written in C and used for teaching at |
|
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2071 |
Stanford University \cite{PINTOS}. An alternative would have been the small Xv6 operating
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2072 |
system used for teaching at MIT \cite{Xv6link,Xv6}. However this operating system implements
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2073 |
a simple round robin scheduler that lacks stubs for dealing with priorities. This |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2074 |
is inconvenient for our purposes. |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2075 |
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2076 |
|
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2077 |
To implement PIP in PINTOS, we only need to modify the kernel |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2078 |
functions corresponding to the events in our formal model. The events translate to the following |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2079 |
function interface in PINTOS: |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2080 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2081 |
\begin{center}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2082 |
\begin{tabular}{|l@ {\hspace{2mm}}|l@ {\hspace{2mm}}|}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2083 |
\hline |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2084 |
{\bf Event} & {\bf PINTOS function} \\
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2085 |
\hline |
|
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2086 |
@{text Create} & @{ML_text "thread_create"}\\
|
|
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2087 |
@{text Exit} & @{ML_text "thread_exit"}\\
|
|
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2088 |
@{text Set} & @{ML_text "thread_set_priority"}\\
|
|
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2089 |
@{text P} & @{ML_text "lock_acquire"}\\
|
|
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2090 |
@{text V} & @{ML_text "lock_release"}\\
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2091 |
\hline |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2092 |
\end{tabular}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2093 |
\end{center}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2094 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2095 |
\noindent |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2096 |
Our implicit assumption that every event is an atomic operation is ensured by the architecture of |
|
15
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2097 |
PINTOS (which allows disabling of interrupts when some operations are performed). The case where |
|
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2098 |
an unlocked resource is given next to the waiting thread with the |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2099 |
highest precedence is realised in our implementation by priority queues. We implemented |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2100 |
them as \emph{Braun trees} \cite{Paulson96}, which provide efficient @{text "O(log n)"}-operations
|
|
17
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
2101 |
for accessing and updating. In the code we shall describe below, we use the function |
|
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
2102 |
@{ML_text "queue_insert"}, for inserting a new element into a priority queue, and
|
|
22
9f0b78fcc894
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
20
diff
changeset
|
2103 |
the function @{ML_text "queue_update"}, for updating the position of an element that is already
|
|
17
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
2104 |
in a queue. Both functions take an extra argument that specifies the |
|
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
2105 |
comparison function used for organising the priority queue. |
|
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
2106 |
|
|
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
2107 |
Apart from having to implement relatively complex data\-structures in C |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2108 |
using pointers, our experience with the implementation has been very positive: our specification |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2109 |
and formalisation of PIP translates smoothly to an efficent implementation in PINTOS. |
|
15
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2110 |
Let us illustrate this with the C-code for the function @{ML_text "lock_acquire"},
|
|
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2111 |
shown in Figure~\ref{code}. This function implements the operation of requesting and, if free,
|
|
15
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2112 |
locking of a resource by the current running thread. The convention in the PINTOS |
|
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2113 |
code is to use the terminology \emph{locks} rather than resources.
|
|
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2114 |
A lock is represented as a pointer to the structure {\tt lock} (Line 1).
|
|
15
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2115 |
Lines 2 to 4 are taken from the original |
|
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2116 |
code of @{ML_text "lock_acquire"} in PINTOS. They contain diagnostic code: first,
|
|
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2117 |
there is a check that |
|
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2118 |
the lock is a ``valid'' lock |
|
15
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2119 |
by testing whether it is not {\tt NULL}; second, a check that the code is not called
|
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
2120 |
as part of an interrupt---acquiring a lock should only be initiated by a |
|
15
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2121 |
request from a (user) thread, not from an interrupt; third, it is ensured that the |
|
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2122 |
current thread does not ask twice for a lock. These assertions are supposed |
|
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2123 |
to be satisfied because of the assumptions in PINTOS about how this code is called. |
|
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2124 |
If not, then the assertions indicate a bug in PINTOS and the result will be |
|
17
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
2125 |
a ``kernel panic''. |
|
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2126 |
|
|
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2127 |
|
|
7
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2128 |
|
|
22
9f0b78fcc894
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
20
diff
changeset
|
2129 |
\begin{figure}[tph]
|
|
7
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2130 |
\begin{lstlisting}
|
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2131 |
void lock_acquire (struct lock *lock) |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2132 |
{ ASSERT (lock != NULL);
|
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2133 |
ASSERT (!intr_context()); |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2134 |
ASSERT (!lock_held_by_current_thread (lock)); |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2135 |
|
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2136 |
enum intr_level old_level; |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2137 |
old_level = intr_disable(); |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2138 |
if (lock->value == 0) {
|
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
2139 |
queue_insert(thread_cprec, &lock->wq, &thread_current()->helem); |
|
7
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2140 |
thread_current()->waiting = lock; |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2141 |
struct thread *pt; |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2142 |
pt = lock->holder; |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2143 |
while (pt) {
|
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
2144 |
queue_update(lock_cprec, &pt->held, &lock->helem); |
|
12
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
2145 |
if (!(update_cprec(pt))) |
|
7
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2146 |
break; |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2147 |
lock = pt->waiting; |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2148 |
if (!lock) {
|
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
2149 |
queue_update(higher_cprec, &ready_queue, &pt->helem); |
|
7
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2150 |
break; |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2151 |
}; |
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
2152 |
queue_update(thread_cprec, &lock->wq, &pt->helem); |
|
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2153 |
pt = lock->holder; |
|
7
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2154 |
}; |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2155 |
thread_block(); |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2156 |
} else {
|
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2157 |
lock->value--; |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2158 |
lock->holder = thread_current(); |
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
2159 |
queue_insert(lock_prec, &thread_current()->held, &lock->helem); |
|
7
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2160 |
}; |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2161 |
intr_set_level(old_level); |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2162 |
} |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2163 |
\end{lstlisting}
|
|
17
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
2164 |
\caption{Our version of the {\tt lock\_acquire} function for the small operating
|
|
15
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2165 |
system PINTOS. It implements the operation corresponding to a @{text P}-event.\label{code}}
|
|
7
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2166 |
\end{figure}
|
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2167 |
|
|
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2168 |
|
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
2169 |
Line 6 and 7 of {\tt lock\_acquire} make the operation of acquiring a lock atomic by disabling all
|
|
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2170 |
interrupts, but saving them for resumption at the end of the function (Line 31). |
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
2171 |
In Line 8, the interesting code with respect to scheduling starts: we |
|
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2172 |
first check whether the lock is already taken (its value is then 0 indicating ``already |
|
12
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
2173 |
taken'', or 1 for being ``free''). In case the lock is taken, we enter the |
|
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2174 |
if-branch inserting the current thread into the waiting queue of this lock (Line 9). |
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
2175 |
The waiting queue is referenced in the usual C-way as @{ML_text "&lock->wq"}.
|
|
15
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2176 |
Next, we record that the current thread is waiting for the lock (Line 10). |
|
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2177 |
Thus we established two pointers: one in the waiting queue of the lock pointing to the |
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2178 |
current thread, and the other from the currend thread pointing to the lock. |
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2179 |
According to our specification in Section~\ref{model} and the properties we were able
|
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2180 |
to prove for @{text P}, we need to ``chase'' all the dependants
|
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2181 |
in the RAG (Resource Allocation Graph) and update their |
|
15
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2182 |
current precedence; however we only have to do this as long as there is change in the |
|
17
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
2183 |
current precedence. |
|
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2184 |
|
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2185 |
The ``chase'' is implemented in the while-loop in Lines 13 to 24. |
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2186 |
To initialise the loop, we |
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
2187 |
assign in Lines 11 and 12 the variable @{ML_text pt} to the owner
|
|
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2188 |
of the lock. |
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
2189 |
Inside the loop, we first update the precedence of the lock held by @{ML_text pt} (Line 14).
|
|
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2190 |
Next, we check whether there is a change in the current precedence of @{ML_text pt}. If not,
|
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2191 |
then we leave the loop, since nothing else needs to be updated (Lines 15 and 16). |
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2192 |
If there is a change, then we have to continue our ``chase''. We check what lock the |
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2193 |
thread @{ML_text pt} is waiting for (Lines 17 and 18). If there is none, then
|
|
15
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2194 |
the thread @{ML_text pt} is ready (the ``chase'' is finished with finding a root in the RAG). In this
|
|
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2195 |
case we update the ready-queue accordingly (Lines 19 and 20). If there is a lock @{ML_text pt} is
|
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2196 |
waiting for, we update the waiting queue for this lock and we continue the loop with |
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2197 |
the holder of that lock |
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2198 |
(Lines 22 and 23). After all current precedences have been updated, we finally need |
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2199 |
to block the current thread, because the lock it asked for was taken (Line 25). |
|
7
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
2200 |
|
|
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2201 |
If the lock the current thread asked for is \emph{not} taken, we proceed with the else-branch
|
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2202 |
(Lines 26 to 30). We first decrease the value of the lock to 0, meaning |
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2203 |
it is taken now (Line 27). Second, we update the reference of the holder of |
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2204 |
the lock (Line 28), and finally update the queue of locks the current |
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2205 |
thread already possesses (Line 29). |
|
15
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2206 |
The very last step is to enable interrupts again thus leaving the protected section. |
|
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2207 |
|
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
2208 |
|
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
2209 |
Similar operations need to be implementated for the @{ML_text lock_release} function, which
|
|
15
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2210 |
we however do not show. The reader should note though that we did \emph{not} verify our C-code.
|
|
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2211 |
This is in contrast, for example, to the work on seL4, which actually verified in Isabelle/HOL |
|
17
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
2212 |
that their C-code satisfies its specification, thought this specification does not contain |
|
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
2213 |
anything about PIP \cite{sel4}.
|
|
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
2214 |
Our verification of PIP however provided us with the justification for designing |
|
15
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2215 |
the C-code. It gave us confidence that leaving the ``chase'' early, whenever |
|
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2216 |
there is no change in the calculated current precedence, does not break the |
|
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
2217 |
correctness of the algorithm. |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2218 |
*} |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2219 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2220 |
section {* Conclusion *}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2221 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2222 |
text {*
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2223 |
The Priority Inheritance Protocol (PIP) is a classic textbook |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2224 |
algorithm used in many real-time operating systems in order to avoid the problem of |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2225 |
Priority Inversion. Although classic and widely used, PIP does have |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2226 |
its faults: for example it does not prevent deadlocks in cases where threads |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2227 |
have circular lock dependencies. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2228 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2229 |
We had two goals in mind with our formalisation of PIP: One is to |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2230 |
make the notions in the correctness proof by Sha et al.~\cite{Sha90}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2231 |
precise so that they can be processed by a theorem prover. The reason is |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2232 |
that a mechanically checked proof avoids the flaws that crept into their |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2233 |
informal reasoning. We achieved this goal: The correctness of PIP now |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2234 |
only hinges on the assumptions behind our formal model. The reasoning, which is |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2235 |
sometimes quite intricate and tedious, has been checked by Isabelle/HOL. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2236 |
We can also confirm that Paulson's |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2237 |
inductive method for protocol verification~\cite{Paulson98} is quite
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2238 |
suitable for our formal model and proof. The traditional application |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2239 |
area of this method is security protocols. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2240 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2241 |
The second goal of our formalisation is to provide a specification for actually |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2242 |
implementing PIP. Textbooks, for example \cite[Section 5.6.5]{Vahalia96},
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2243 |
explain how to use various implementations of PIP and abstractly |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2244 |
discuss their properties, but surprisingly lack most details important for a |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2245 |
programmer who wants to implement PIP (similarly Sha et al.~\cite{Sha90}).
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2246 |
That this is an issue in practice is illustrated by the |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2247 |
email from Baker we cited in the Introduction. We achieved also this |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2248 |
goal: The formalisation allowed us to efficently implement our version |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2249 |
of PIP on top of PINTOS \cite{PINTOS}, a simple instructional operating system for the x86
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2250 |
architecture. It also gives the first author enough data to enable |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2251 |
his undergraduate students to implement PIP (as part of their OS course). |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2252 |
A byproduct of our formalisation effort is that nearly all |
|
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2253 |
design choices for the implementation of PIP scheduler are backed up with a proved |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2254 |
lemma. We were also able to establish the property that the choice of |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2255 |
the next thread which takes over a lock is irrelevant for the correctness |
|
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2256 |
of PIP. Moreover, we eliminated a crucial restriction present in |
|
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
2257 |
the proof of Sha et al.: they require that critical sections nest properly, |
|
32
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2258 |
whereas our scheduler allows critical sections to overlap. What we |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2259 |
are not able to do is to mechanically ``synthesise'' an actual implementation |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2260 |
from our formalisation. To do so for C-code seems quite hard and is beyond |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2261 |
current technology available for Isabelle. Also our proof-method based |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2262 |
on events is not ``computational'' in the sense of having a concrete |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2263 |
algorithm behind it: our formalisation is really more about the |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2264 |
specification of PIP and ensuring that it has the desired properties |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2265 |
(the informal specification by Sha et al.~did not). |
|
e861aff29655
made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
28
diff
changeset
|
2266 |
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2267 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2268 |
PIP is a scheduling algorithm for single-processor systems. We are |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2269 |
now living in a multi-processor world. Priority Inversion certainly |
|
20
b56616fd88dd
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
17
diff
changeset
|
2270 |
occurs also there, see for example \cite{Brandenburg11,Davis11}.
|
|
16
9764023f719e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
15
diff
changeset
|
2271 |
However, there is very little ``foundational'' |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2272 |
work about PIP-algorithms on multi-processor systems. We are not |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2273 |
aware of any correctness proofs, not even informal ones. There is an |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2274 |
implementation of a PIP-algorithm for multi-processors as part of the |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2275 |
``real-time'' effort in Linux, including an informal description of the implemented scheduling |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2276 |
algorithm given in \cite{LINUX}. We estimate that the formal
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2277 |
verification of this algorithm, involving more fine-grained events, |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2278 |
is a magnitude harder than the one we presented here, but still |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2279 |
within reach of current theorem proving technology. We leave this |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2280 |
for future work. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2281 |
|
|
22
9f0b78fcc894
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
20
diff
changeset
|
2282 |
To us, it seems sound reasoning about scheduling algorithms is fiendishly difficult |
|
27
6b1141c5e24c
cleaned up
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
26
diff
changeset
|
2283 |
if done informally by ``pencil-and-paper''. We infer this from the flawed proof |
|
6b1141c5e24c
cleaned up
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
26
diff
changeset
|
2284 |
in the paper by Sha et al.~\cite{Sha90} and also from \cite{Regehr} where Regehr
|
|
6b1141c5e24c
cleaned up
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
26
diff
changeset
|
2285 |
points out an error in a paper about Preemption |
|
22
9f0b78fcc894
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
20
diff
changeset
|
2286 |
Threshold Scheduling \cite{ThreadX}. The use of a theorem prover was
|
|
23
24e6884d9258
made some small chages
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
22
diff
changeset
|
2287 |
invaluable to us in order to be confident about the correctness of our reasoning |
|
41
66ed924aaa5c
added another book that makes the error, some more proofs
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
38
diff
changeset
|
2288 |
(for example no corner case can be overlooked). |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2289 |
The most closely related work to ours is the formal verification in |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2290 |
PVS of the Priority Ceiling Protocol done by Dutertre |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2291 |
\cite{dutertre99b}---another solution to the Priority Inversion
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2292 |
problem, which however needs static analysis of programs in order to |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2293 |
avoid it. There have been earlier formal investigations |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2294 |
into PIP \cite{Faria08,Jahier09,Wellings07}, but they employ model
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2295 |
checking techniques. The results obtained by them apply, |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2296 |
however, only to systems with a fixed size, such as a fixed number of |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2297 |
events and threads. In contrast, our result applies to systems of arbitrary |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2298 |
size. Moreover, our result is a good |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2299 |
witness for one of the major reasons to be interested in machine checked |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2300 |
reasoning: gaining deeper understanding of the subject matter. |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2301 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2302 |
Our formalisation |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2303 |
consists of around 210 lemmas and overall 6950 lines of readable Isabelle/Isar |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2304 |
code with a few apply-scripts interspersed. The formal model of PIP |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2305 |
is 385 lines long; the formal correctness proof 3800 lines. Some auxiliary |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2306 |
definitions and proofs span over 770 lines of code. The properties relevant |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2307 |
for an implementation require 2000 lines. |
|
22
9f0b78fcc894
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
20
diff
changeset
|
2308 |
The code of our formalisation |
|
9f0b78fcc894
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
20
diff
changeset
|
2309 |
can be downloaded from the Mercurial repository at |
|
27
6b1141c5e24c
cleaned up
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
26
diff
changeset
|
2310 |
\url{http://www.dcs.kcl.ac.uk/staff/urbanc/cgi-bin/repos.cgi/pip}.
|
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2311 |
|
|
16
9764023f719e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
15
diff
changeset
|
2312 |
%\medskip |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2313 |
|
|
16
9764023f719e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
15
diff
changeset
|
2314 |
%\noindent |
|
9764023f719e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
15
diff
changeset
|
2315 |
%{\bf Acknowledgements:}
|
|
9764023f719e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
15
diff
changeset
|
2316 |
%We are grateful for the comments we received from anonymous |
|
9764023f719e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
15
diff
changeset
|
2317 |
%referees. |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2318 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2319 |
\bibliographystyle{plain}
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2320 |
\bibliography{root}
|
|
142
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
2321 |
|
|
10c16b85a839
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
141
diff
changeset
|
2322 |
\theendnotes |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2323 |
*} |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2324 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2325 |
|
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2326 |
(*<*) |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
2327 |
end |
|
127
38c6acf03f68
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
126
diff
changeset
|
2328 |
(*>*) |