author | urbanc |
Sat, 11 Feb 2012 09:14:14 +0000 | |
changeset 293 | cab43c4a96d2 |
parent 292 | 1f16ff7fea94 |
child 294 | bc5bf9e9ada2 |
permissions | -rwxr-xr-x |
262 | 1 |
(*<*) |
2 |
theory Paper |
|
292 | 3 |
imports CpsG ExtGG "~~/src/HOL/Library/LaTeXsugar" |
262 | 4 |
begin |
266 | 5 |
ML {* |
273 | 6 |
open Printer; |
272 | 7 |
show_question_marks_default := false; |
266 | 8 |
*} |
284 | 9 |
|
10 |
notation (latex output) |
|
11 |
Cons ("_::_" [78,77] 73) and |
|
12 |
vt ("valid'_state") and |
|
13 |
runing ("running") and |
|
286 | 14 |
birthtime ("last'_set") and |
284 | 15 |
If ("(\<^raw:\textrm{>if\<^raw:}> (_)/ \<^raw:\textrm{>then\<^raw:}> (_)/ \<^raw:\textrm{>else\<^raw:}> (_))" 10) and |
286 | 16 |
Prc ("'(_, _')") and |
287 | 17 |
holding ("holds") and |
18 |
waiting ("waits") and |
|
290 | 19 |
Th ("T") and |
20 |
Cs ("C") and |
|
287 | 21 |
readys ("ready") and |
290 | 22 |
depend ("RAG") and |
23 |
preced ("prec") and |
|
24 |
cpreced ("cprec") and |
|
291
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
25 |
dependents ("dependants") and |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
26 |
waiting_queue ("wq_fun") and |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
27 |
cur_preced ("cprec_fun") and |
284 | 28 |
DUMMY ("\<^raw:\mbox{$\_\!\_$}>") |
262 | 29 |
(*>*) |
30 |
||
31 |
section {* Introduction *} |
|
32 |
||
33 |
text {* |
|
284 | 34 |
Many real-time systems need to support threads involving priorities and |
267 | 35 |
locking of resources. Locking of resources ensures mutual exclusion |
275 | 36 |
when accessing shared data or devices that cannot be |
284 | 37 |
preempted. Priorities allow scheduling of threads that need to |
275 | 38 |
finish their work within deadlines. Unfortunately, both features |
39 |
can interact in subtle ways leading to a problem, called |
|
284 | 40 |
\emph{Priority Inversion}. Suppose three threads having priorities |
41 |
$H$(igh), $M$(edium) and $L$(ow). We would expect that the thread |
|
42 |
$H$ blocks any other thread with lower priority and itself cannot |
|
43 |
be blocked by any thread with lower priority. Alas, in a naive |
|
275 | 44 |
implementation of resource looking and priorities this property can |
45 |
be violated. Even worse, $H$ can be delayed indefinitely by |
|
284 | 46 |
threads with lower priorities. For this let $L$ be in the |
275 | 47 |
possession of a lock for a resource that also $H$ needs. $H$ must |
48 |
therefore wait for $L$ to exit the critical section and release this |
|
49 |
lock. The problem is that $L$ might in turn be blocked by any |
|
284 | 50 |
thread with priority $M$, and so $H$ sits there potentially waiting |
51 |
indefinitely. Since $H$ is blocked by threads with lower |
|
275 | 52 |
priorities, the problem is called Priority Inversion. It was first |
277 | 53 |
described in \cite{Lampson80} in the context of the |
275 | 54 |
Mesa programming language designed for concurrent programming. |
265 | 55 |
|
273 | 56 |
If the problem of Priority Inversion is ignored, real-time systems |
267 | 57 |
can become unpredictable and resulting bugs can be hard to diagnose. |
58 |
The classic example where this happened is the software that |
|
284 | 59 |
controlled the Mars Pathfinder mission in 1997 \cite{Reeves98}. |
60 |
Once the spacecraft landed, the software shut down at irregular |
|
61 |
intervals leading to loss of project time as normal operation of the |
|
62 |
craft could only resume the next day (the mission and data already |
|
63 |
collected were fortunately not lost, because of a clever system |
|
64 |
design). The reason for the shutdowns was that the scheduling |
|
65 |
software fell victim of Priority Inversion: a low priority thread |
|
66 |
locking a resource prevented a high priority thread from running in |
|
67 |
time leading to a system reset. Once the problem was found, it was |
|
68 |
rectified by enabling the \emph{Priority Inheritance Protocol} (PIP) |
|
69 |
\cite{Sha90}\footnote{Sha et al.~call it the \emph{Basic Priority |
|
286 | 70 |
Inheritance Protocol} \cite{Sha90} and others sometimes also call it |
71 |
\emph{Priority Boosting}.} in the scheduling software. |
|
262 | 72 |
|
284 | 73 |
The idea behind PIP is to let the thread $L$ temporarily inherit |
286 | 74 |
the high priority from $H$ until $L$ leaves the critical section |
284 | 75 |
unlocking the resource. This solves the problem of $H$ having to |
76 |
wait indefinitely, because $L$ cannot be blocked by threads having |
|
77 |
priority $M$. While a few other solutions exist for the Priority |
|
78 |
Inversion problem, PIP is one that is widely deployed and |
|
79 |
implemented. This includes VxWorks (a proprietary real-time OS used |
|
80 |
in the Mars Pathfinder mission, in Boeing's 787 Dreamliner, Honda's |
|
81 |
ASIMO robot, etc.), but also the POSIX 1003.1c Standard realised for |
|
82 |
example in libraries for FreeBSD, Solaris and Linux. |
|
274 | 83 |
|
284 | 84 |
One advantage of PIP is that increasing the priority of a thread |
275 | 85 |
can be dynamically calculated by the scheduler. This is in contrast |
277 | 86 |
to, for example, \emph{Priority Ceiling} \cite{Sha90}, another |
87 |
solution to the Priority Inversion problem, which requires static |
|
284 | 88 |
analysis of the program in order to prevent Priority |
89 |
Inversion. However, there has also been strong criticism against |
|
90 |
PIP. For instance, PIP cannot prevent deadlocks when lock |
|
91 |
dependencies are circular, and also blocking times can be |
|
92 |
substantial (more than just the duration of a critical section). |
|
93 |
Though, most criticism against PIP centres around unreliable |
|
94 |
implementations and PIP being too complicated and too inefficient. |
|
95 |
For example, Yodaiken writes in \cite{Yodaiken02}: |
|
274 | 96 |
|
97 |
\begin{quote} |
|
98 |
\it{}``Priority inheritance is neither efficient nor reliable. Implementations |
|
99 |
are either incomplete (and unreliable) or surprisingly complex and intrusive.'' |
|
100 |
\end{quote} |
|
273 | 101 |
|
274 | 102 |
\noindent |
275 | 103 |
He suggests to avoid PIP altogether by not allowing critical |
286 | 104 |
sections to be preempted. Unfortunately, this solution does not |
105 |
help in real-time systems with low latency \emph{requirements}. |
|
278 | 106 |
|
286 | 107 |
In our opinion, there is clearly a need for investigating correct |
278 | 108 |
algorithms for PIP. A few specifications for PIP exist (in English) |
109 |
and also a few high-level descriptions of implementations (e.g.~in |
|
110 |
the textbook \cite[Section 5.6.5]{Vahalia96}), but they help little |
|
111 |
with actual implementations. That this is a problem in practise is |
|
283 | 112 |
proved by an email from Baker, who wrote on 13 July 2009 on the Linux |
278 | 113 |
Kernel mailing list: |
274 | 114 |
|
115 |
\begin{quote} |
|
275 | 116 |
\it{}``I observed in the kernel code (to my disgust), the Linux PIP |
117 |
implementation is a nightmare: extremely heavy weight, involving |
|
118 |
maintenance of a full wait-for graph, and requiring updates for a |
|
119 |
range of events, including priority changes and interruptions of |
|
120 |
wait operations.'' |
|
274 | 121 |
\end{quote} |
122 |
||
123 |
\noindent |
|
277 | 124 |
The criticism by Yodaiken, Baker and others suggests to us to look |
125 |
again at PIP from a more abstract level (but still concrete enough |
|
286 | 126 |
to inform an implementation), and makes PIP an ideal candidate for a |
277 | 127 |
formal verification. One reason, of course, is that the original |
284 | 128 |
presentation of PIP~\cite{Sha90}, despite being informally |
283 | 129 |
``proved'' correct, is actually \emph{flawed}. |
130 |
||
131 |
Yodaiken \cite{Yodaiken02} points to a subtlety that had been |
|
132 |
overlooked in the informal proof by Sha et al. They specify in |
|
284 | 133 |
\cite{Sha90} that after the thread (whose priority has been raised) |
283 | 134 |
completes its critical section and releases the lock, it ``returns |
135 |
to its original priority level.'' This leads them to believe that an |
|
284 | 136 |
implementation of PIP is ``rather straightforward''~\cite{Sha90}. |
137 |
Unfortunately, as Yodaiken points out, this behaviour is too |
|
138 |
simplistic. Consider the case where the low priority thread $L$ |
|
139 |
locks \emph{two} resources, and two high-priority threads $H$ and |
|
283 | 140 |
$H'$ each wait for one of them. If $L$ then releases one resource |
141 |
so that $H$, say, can proceed, then we still have Priority Inversion |
|
142 |
with $H'$ (which waits for the other resource). The correct |
|
143 |
behaviour for $L$ is to revert to the highest remaining priority of |
|
284 | 144 |
the threads that it blocks. The advantage of formalising the |
145 |
correctness of a high-level specification of PIP in a theorem prover |
|
146 |
is that such issues clearly show up and cannot be overlooked as in |
|
147 |
informal reasoning (since we have to analyse all possible behaviours |
|
148 |
of threads, i.e.~\emph{traces}, that could possibly happen). |
|
274 | 149 |
|
279 | 150 |
There have been earlier formal investigations into PIP, but ...\cite{Faria08} |
284 | 151 |
|
152 |
vt (valid trace) was introduced earlier, cite |
|
153 |
||
154 |
distributed PIP |
|
286 | 155 |
|
156 |
Paulson's method has not been used outside security field, except |
|
157 |
work by Zhang et al. |
|
158 |
||
159 |
no clue about multi-processor case according to \cite{Steinberg10} |
|
280 | 160 |
*} |
278 | 161 |
|
283 | 162 |
section {* Formal Model of the Priority Inheritance Protocol *} |
267 | 163 |
|
280 | 164 |
text {* |
286 | 165 |
The Priority Inheritance Protocol, short PIP, is a scheduling |
166 |
algorithm for a single-processor system.\footnote{We shall come back |
|
167 |
later to the case of PIP on multi-processor systems.} Our model of |
|
168 |
PIP is based on Paulson's inductive approach to protocol |
|
169 |
verification \cite{Paulson98}, where the \emph{state} of a system is |
|
287 | 170 |
given by a list of events that happened so far. \emph{Events} in PIP fall |
290 | 171 |
into five categories defined as the datatype: |
283 | 172 |
|
173 |
\begin{isabelle}\ \ \ \ \ %%% |
|
284 | 174 |
\mbox{\begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{7mm}}l} |
175 |
\isacommand{datatype} event |
|
176 |
& @{text "="} & @{term "Create thread priority"}\\ |
|
177 |
& @{text "|"} & @{term "Exit thread"} \\ |
|
286 | 178 |
& @{text "|"} & @{term "Set thread priority"} & {\rm reset of the priority for} @{text thread}\\ |
284 | 179 |
& @{text "|"} & @{term "P thread cs"} & {\rm request of resource} @{text "cs"} {\rm by} @{text "thread"}\\ |
180 |
& @{text "|"} & @{term "V thread cs"} & {\rm release of resource} @{text "cs"} {\rm by} @{text "thread"} |
|
181 |
\end{tabular}} |
|
182 |
\end{isabelle} |
|
183 |
||
184 |
\noindent |
|
286 | 185 |
whereby threads, priorities and (critical) resources are represented |
186 |
as natural numbers. The event @{term Set} models the situation that |
|
187 |
a thread obtains a new priority given by the programmer or |
|
188 |
user (for example via the {\tt nice} utility under UNIX). As in Paulson's work, we |
|
189 |
need to define functions that allow one to make some observations |
|
293 | 190 |
about states. One, called @{term threads}, calculates, given a state @{text s}, the set of |
191 |
``live'' threads that we have seen so far: |
|
284 | 192 |
|
193 |
\begin{isabelle}\ \ \ \ \ %%% |
|
194 |
\mbox{\begin{tabular}{lcl} |
|
195 |
@{thm (lhs) threads.simps(1)} & @{text "\<equiv>"} & |
|
196 |
@{thm (rhs) threads.simps(1)}\\ |
|
197 |
@{thm (lhs) threads.simps(2)[where thread="th"]} & @{text "\<equiv>"} & |
|
198 |
@{thm (rhs) threads.simps(2)[where thread="th"]}\\ |
|
199 |
@{thm (lhs) threads.simps(3)[where thread="th"]} & @{text "\<equiv>"} & |
|
200 |
@{thm (rhs) threads.simps(3)[where thread="th"]}\\ |
|
201 |
@{term "threads (DUMMY#s)"} & @{text "\<equiv>"} & @{term "threads s"}\\ |
|
202 |
\end{tabular}} |
|
283 | 203 |
\end{isabelle} |
204 |
||
205 |
\noindent |
|
290 | 206 |
Another function calculates the priority for a thread @{text "th"}, which is |
207 |
defined as |
|
284 | 208 |
|
209 |
\begin{isabelle}\ \ \ \ \ %%% |
|
210 |
\mbox{\begin{tabular}{lcl} |
|
211 |
@{thm (lhs) original_priority.simps(1)[where thread="th"]} & @{text "\<equiv>"} & |
|
212 |
@{thm (rhs) original_priority.simps(1)[where thread="th"]}\\ |
|
213 |
@{thm (lhs) original_priority.simps(2)[where thread="th" and thread'="th'"]} & @{text "\<equiv>"} & |
|
214 |
@{thm (rhs) original_priority.simps(2)[where thread="th" and thread'="th'"]}\\ |
|
215 |
@{thm (lhs) original_priority.simps(3)[where thread="th" and thread'="th'"]} & @{text "\<equiv>"} & |
|
216 |
@{thm (rhs) original_priority.simps(3)[where thread="th" and thread'="th'"]}\\ |
|
217 |
@{term "original_priority th (DUMMY#s)"} & @{text "\<equiv>"} & @{term "original_priority th s"}\\ |
|
218 |
\end{tabular}} |
|
219 |
\end{isabelle} |
|
220 |
||
221 |
\noindent |
|
222 |
In this definition we set @{text 0} as the default priority for |
|
223 |
threads that have not (yet) been created. The last function we need |
|
285 | 224 |
calculates the ``time'', or index, at which time a process had its |
290 | 225 |
priority last set. |
284 | 226 |
|
227 |
\begin{isabelle}\ \ \ \ \ %%% |
|
228 |
\mbox{\begin{tabular}{lcl} |
|
229 |
@{thm (lhs) birthtime.simps(1)[where thread="th"]} & @{text "\<equiv>"} & |
|
230 |
@{thm (rhs) birthtime.simps(1)[where thread="th"]}\\ |
|
231 |
@{thm (lhs) birthtime.simps(2)[where thread="th" and thread'="th'"]} & @{text "\<equiv>"} & |
|
232 |
@{thm (rhs) birthtime.simps(2)[where thread="th" and thread'="th'"]}\\ |
|
233 |
@{thm (lhs) birthtime.simps(3)[where thread="th" and thread'="th'"]} & @{text "\<equiv>"} & |
|
234 |
@{thm (rhs) birthtime.simps(3)[where thread="th" and thread'="th'"]}\\ |
|
235 |
@{term "birthtime th (DUMMY#s)"} & @{text "\<equiv>"} & @{term "birthtime th s"}\\ |
|
236 |
\end{tabular}} |
|
237 |
\end{isabelle} |
|
286 | 238 |
|
239 |
\noindent |
|
287 | 240 |
In this definition @{term "length s"} stands for the length of the list |
241 |
of events @{text s}. Again the default value in this function is @{text 0} |
|
242 |
for threads that have not been created yet. A \emph{precedence} of a thread @{text th} in a |
|
290 | 243 |
state @{text s} is the pair of natural numbers defined as |
284 | 244 |
|
286 | 245 |
\begin{isabelle}\ \ \ \ \ %%% |
290 | 246 |
@{thm preced_def[where thread="th"]} |
286 | 247 |
\end{isabelle} |
248 |
||
249 |
\noindent |
|
287 | 250 |
The point of precedences is to schedule threads not according to priorities (because what should |
286 | 251 |
we do in case two threads have the same priority), but according to precedences. |
290 | 252 |
Precedences allow us to always discriminate between two threads with equal priority by |
286 | 253 |
tacking into account the time when the priority was last set. We order precedences so |
254 |
that threads with the same priority get a higher precedence if their priority has been |
|
293 | 255 |
set earlier, since for such threads it is more urgent to finish their work. In an implementation |
256 |
this choice would translate to a quite natural FIFO-scheduling of processes with |
|
286 | 257 |
the same priority. |
258 |
||
259 |
Next, we introduce the concept of \emph{waiting queues}. They are |
|
260 |
lists of threads associated with every resource. The first thread in |
|
291
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
261 |
this list (the head, or short @{term hd}) is chosen to be the one |
290 | 262 |
that is in possession of the |
286 | 263 |
``lock'' of the corresponding resource. We model waiting queues as |
293 | 264 |
functions, below abbreviated as @{text wq}. They take a resource as |
265 |
argument and return a list of threads. This allows us to define |
|
290 | 266 |
when a thread \emph{holds}, respectively \emph{waits} for, a |
293 | 267 |
resource @{text cs} given a waiting queue function @{text wq}. |
287 | 268 |
|
269 |
\begin{isabelle}\ \ \ \ \ %%% |
|
270 |
\begin{tabular}{@ {}l} |
|
290 | 271 |
@{thm cs_holding_def[where thread="th"]}\\ |
272 |
@{thm cs_waiting_def[where thread="th"]} |
|
287 | 273 |
\end{tabular} |
274 |
\end{isabelle} |
|
275 |
||
276 |
\noindent |
|
277 |
In this definition we assume @{text "set"} converts a list into a set. |
|
291
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
278 |
At the beginning, that is in the state where no process is created yet, |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
279 |
the waiting queue function will be just the function that returns the |
293 | 280 |
empty list for every resource. |
291
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
281 |
|
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
282 |
\begin{isabelle}\ \ \ \ \ %%% |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
283 |
@{abbrev all_unlocked} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
284 |
\end{isabelle} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
285 |
|
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
286 |
\noindent |
290 | 287 |
Using @{term "holding"} and @{term waiting}, we can introduce \emph{Resource Allocation Graphs} |
288 |
(RAG), which represent the dependencies between threads and resources. |
|
289 |
We represent RAGs as relations using pairs of the form |
|
290 |
||
291 |
\begin{isabelle}\ \ \ \ \ %%% |
|
292 |
@{term "(Th th, Cs cs)"} \hspace{5mm}{\rm and}\hspace{5mm} |
|
293 |
@{term "(Cs cs, Th th)"} |
|
294 |
\end{isabelle} |
|
295 |
||
296 |
\noindent |
|
297 |
where the first stands for a \emph{waiting edge} and the second for a |
|
298 |
\emph{holding edge} (@{term Cs} and @{term Th} are constructors of a |
|
299 |
datatype for vertices). Given a waiting queue function, a RAG is defined |
|
300 |
as |
|
301 |
||
302 |
\begin{isabelle}\ \ \ \ \ %%% |
|
303 |
@{thm cs_depend_def} |
|
304 |
\end{isabelle} |
|
305 |
||
306 |
\noindent |
|
307 |
An instance of a RAG is as follows: |
|
308 |
||
309 |
\begin{center} |
|
291
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
310 |
\begin{tikzpicture}[scale=1] |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
311 |
\draw[step=2mm] (0,0) grid (3,2); |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
312 |
|
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
313 |
\draw[rounded corners=1mm, very thick] (0,0) rectangle (0.6,0.6); |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
314 |
\draw[rounded corners=1mm, very thick] (3,0) rectangle (3.6,0.6); |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
315 |
|
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
316 |
\draw (0.3,0.3) node {\small$th_0$}; |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
317 |
\draw (3.3,0.3) node {\small$th_1$}; |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
318 |
\end{tikzpicture} |
290 | 319 |
\end{center} |
320 |
||
321 |
\noindent |
|
322 |
The use or relations for representing RAGs allows us to conveninetly define |
|
323 |
the notion of the \emph{dependants} of a thread. This is defined as |
|
324 |
||
325 |
\begin{isabelle}\ \ \ \ \ %%% |
|
326 |
@{thm cs_dependents_def} |
|
327 |
\end{isabelle} |
|
328 |
||
329 |
\noindent |
|
293 | 330 |
This definition needs to account for all threads that wait for a tread to |
290 | 331 |
release a resource. This means we need to include threads that transitively |
293 | 332 |
wait for a resource being realeased (in the picture above this means also @{text "th\<^isub>3"}, |
290 | 333 |
which cannot make any progress unless @{text "th\<^isub>2"} makes progress which |
293 | 334 |
in turn needs to wait for @{text "th\<^isub>1"}). If there is a cirle in a RAG, then clearly |
291
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
335 |
we have a deadlock. Therefore when a thread requests a resource, |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
336 |
we must ensure that the resulting RAG is not not circular. |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
337 |
|
293 | 338 |
Next we introduce the notion of the \emph{current precedence} for a thread @{text th} in a |
291
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
339 |
state @{text s}, which is defined as |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
340 |
|
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
341 |
\begin{isabelle}\ \ \ \ \ %%% |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
342 |
@{thm cpreced_def2} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
343 |
\end{isabelle} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
344 |
|
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
345 |
\noindent |
293 | 346 |
While the precedence @{term prec} of a thread is determined by the programmer |
347 |
(for example when the thread is |
|
291
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
348 |
created), the point of the current precedence is to let scheduler increase this |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
349 |
priority, if needed according to PIP. Therefore the current precedence of @{text th} is |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
350 |
given as the maximum of the precedence @{text th} has in state @{text s} \emph{and} all |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
351 |
processes that are dependants of @{text th}. Since the notion @{term "dependents"} is |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
352 |
defined as the transitive closure of all dependent threads, we deal correctly with the |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
353 |
problem in the algorithm by Sha et al.~\cite{Sha90} where a priority of a thread is |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
354 |
lowered prematurely. |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
355 |
|
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
356 |
The next function, called @{term schs}, defines the behaviour of the scheduler. It is defined |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
357 |
by recursion on the state (a list of events); @{term "schs"} takes a state as argument |
293 | 358 |
and returns a \emph{schedule state}, which we define as a record consisting of two |
359 |
functions |
|
360 |
||
361 |
\begin{isabelle}\ \ \ \ \ %%% |
|
362 |
@{text "\<lparr>wq_fun, cprec_fun\<rparr>"} |
|
363 |
\end{isabelle} |
|
291
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
364 |
|
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
365 |
In the initial state, the scheduler starts with all resources unlocked and the |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
366 |
precedence of every thread is initialised with @{term "Prc 0 0"}. |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
367 |
|
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
368 |
\begin{isabelle}\ \ \ \ \ %%% |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
369 |
\begin{tabular}{@ {}l} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
370 |
@{thm (lhs) schs.simps(1)} @{text "\<equiv>"}\\ |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
371 |
\hspace{5mm}@{term "(|waiting_queue = all_unlocked, cur_preced = (\<lambda>_::thread. Prc 0 0)|)"} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
372 |
\end{tabular} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
373 |
\end{isabelle} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
374 |
|
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
375 |
\noindent |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
376 |
The cases for @{term Create}, @{term Exit} and @{term Set} are straightforward: |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
377 |
we calculuate the waiting queue function of the (previous) state @{text s}; |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
378 |
this waiting queue function @{text wq} is unchanged in the next schedule state; for calculuating |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
379 |
the next @{term "cprec_fun"} we use @{text wq} and the function @{term cpreced}. |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
380 |
This gives the following clauses: |
290 | 381 |
|
382 |
\begin{isabelle}\ \ \ \ \ %%% |
|
291
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
383 |
\begin{tabular}{@ {}l} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
384 |
@{thm (lhs) schs.simps(2)} @{text "\<equiv>"}\\ |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
385 |
\hspace{5mm}@{text "let"} @{text "wq = waiting_queue (schs s)"} @{text "in"}\\ |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
386 |
\hspace{8mm}@{term "(|waiting_queue = wq\<iota>, cur_preced = cpreced wq\<iota> (Create th prio # s)|)"}\\ |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
387 |
@{thm (lhs) schs.simps(3)} @{text "\<equiv>"}\\ |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
388 |
\hspace{5mm}@{text "let"} @{text "wq = waiting_queue (schs s)"} @{text "in"}\\ |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
389 |
\hspace{8mm}@{term "(|waiting_queue = wq\<iota>, cur_preced = cpreced wq\<iota> (Exit th # s)|)"}\smallskip\\ |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
390 |
@{thm (lhs) schs.simps(4)} @{text "\<equiv>"}\\ |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
391 |
\hspace{5mm}@{text "let"} @{text "wq = waiting_queue (schs s)"} @{text "in"}\\ |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
392 |
\hspace{8mm}@{term "(|waiting_queue = wq\<iota>, cur_preced = cpreced wq\<iota> (Set th prio # s)|)"} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
393 |
\end{tabular} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
394 |
\end{isabelle} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
395 |
|
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
396 |
\noindent |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
397 |
More interesting are the cases when a resource, say @{text cs}, is locked or relaesed. In this case |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
398 |
we need to calculate a new waiting queue function. In case of @{term P} we update |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
399 |
the function so that the new thread list for @{text cs} is old thread list with the waiting |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
400 |
thread appended to the end (remember the head of this list is in the possession of the |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
401 |
resource). |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
402 |
|
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
403 |
\begin{isabelle}\ \ \ \ \ %%% |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
404 |
\begin{tabular}{@ {}l} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
405 |
@{thm (lhs) schs.simps(5)} @{text "\<equiv>"}\\ |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
406 |
\hspace{5mm}@{text "let"} @{text "wq = waiting_queue (schs s)"} @{text "in"}\\ |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
407 |
\hspace{5mm}@{text "let"} @{text "new_wq = wq(cs := (wq cs @ [th]))"} @{text "in"}\\ |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
408 |
\hspace{8mm}@{term "(|waiting_queue = new_wq, cur_preced = cpreced new_wq (P th cs # s)|)"} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
409 |
\end{tabular} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
410 |
\end{isabelle} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
411 |
|
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
412 |
\noindent |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
413 |
The case for @{term V} is similarly, except that we update the waiting queue function |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
414 |
using the auxiliary definition |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
415 |
|
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
416 |
\begin{isabelle}\ \ \ \ \ %%% |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
417 |
@{abbrev release} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
418 |
\end{isabelle} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
419 |
|
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
420 |
\noindent |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
421 |
This gives for @{term schs} the clause: |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
422 |
|
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
423 |
\begin{isabelle}\ \ \ \ \ %%% |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
424 |
\begin{tabular}{@ {}l} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
425 |
@{thm (lhs) schs.simps(6)} @{text "\<equiv>"}\\ |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
426 |
\hspace{5mm}@{text "let"} @{text "wq = waiting_queue (schs s)"} @{text "in"}\\ |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
427 |
\hspace{5mm}@{text "let"} @{text "new_wq = release (wq cs)"} @{text "in"}\\ |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
428 |
\hspace{8mm}@{term "(|waiting_queue = new_wq, cur_preced = cpreced new_wq (V th cs # s)|)"} |
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
429 |
\end{tabular} |
290 | 430 |
\end{isabelle} |
431 |
||
432 |
||
291
5ef9f6ebe827
more on paper; modified schs functions; it is still compatible with the old definition
urbanc
parents:
290
diff
changeset
|
433 |
TODO |
286 | 434 |
|
435 |
\begin{isabelle}\ \ \ \ \ %%% |
|
436 |
\begin{tabular}{@ {}l} |
|
287 | 437 |
@{thm s_depend_def}\\ |
438 |
\end{tabular} |
|
439 |
\end{isabelle} |
|
440 |
||
441 |
\begin{isabelle}\ \ \ \ \ %%% |
|
442 |
\begin{tabular}{@ {}l} |
|
443 |
@{thm readys_def}\\ |
|
444 |
\end{tabular} |
|
445 |
\end{isabelle} |
|
446 |
||
447 |
\begin{isabelle}\ \ \ \ \ %%% |
|
448 |
\begin{tabular}{@ {}l} |
|
449 |
@{thm runing_def}\\ |
|
286 | 450 |
\end{tabular} |
451 |
\end{isabelle} |
|
284 | 452 |
|
453 |
||
287 | 454 |
resources |
284 | 455 |
|
287 | 456 |
done: threads not done: running |
284 | 457 |
|
458 |
step relation: |
|
459 |
||
460 |
\begin{center} |
|
461 |
\begin{tabular}{c} |
|
462 |
@{thm[mode=Rule] thread_create[where thread=th]}\hspace{1cm} |
|
463 |
@{thm[mode=Rule] thread_exit[where thread=th]}\medskip\\ |
|
464 |
||
287 | 465 |
@{thm[mode=Rule] thread_set[where thread=th]}\medskip\\ |
284 | 466 |
@{thm[mode=Rule] thread_P[where thread=th]}\medskip\\ |
467 |
@{thm[mode=Rule] thread_V[where thread=th]}\\ |
|
468 |
\end{tabular} |
|
469 |
\end{center} |
|
470 |
||
471 |
valid state: |
|
472 |
||
473 |
\begin{center} |
|
474 |
\begin{tabular}{c} |
|
475 |
@{thm[mode=Axiom] vt_nil[where cs=step]}\hspace{1cm} |
|
476 |
@{thm[mode=Rule] vt_cons[where cs=step]} |
|
477 |
\end{tabular} |
|
478 |
\end{center} |
|
479 |
||
274 | 480 |
|
280 | 481 |
To define events, the identifiers of {\em threads}, |
482 |
{\em priority} and {\em critical resources } (abbreviated as @{text "cs"}) |
|
483 |
need to be represented. All three are represetned using standard |
|
484 |
Isabelle/HOL type @{typ "nat"}: |
|
485 |
*} |
|
273 | 486 |
|
280 | 487 |
text {* |
488 |
\bigskip |
|
284 | 489 |
The priority inversion phenomenon was first published in |
490 |
\cite{Lampson80}. The two protocols widely used to eliminate |
|
491 |
priority inversion, namely PI (Priority Inheritance) and PCE |
|
492 |
(Priority Ceiling Emulation), were proposed in \cite{Sha90}. PCE is |
|
493 |
less convenient to use because it requires static analysis of |
|
494 |
programs. Therefore, PI is more commonly used in |
|
495 |
practice\cite{locke-july02}. However, as pointed out in the |
|
496 |
literature, the analysis of priority inheritance protocol is quite |
|
497 |
subtle\cite{yodaiken-july02}. A formal analysis will certainly be |
|
498 |
helpful for us to understand and correctly implement PI. All |
|
499 |
existing formal analysis of PI |
|
500 |
\cite{conf/fase/JahierHR09,WellingsBSB07,Faria08} are based on the |
|
501 |
model checking technology. Because of the state explosion problem, |
|
502 |
model check is much like an exhaustive testing of finite models with |
|
503 |
limited size. The results obtained can not be safely generalized to |
|
504 |
models with arbitrarily large size. Worse still, since model |
|
505 |
checking is fully automatic, it give little insight on why the |
|
506 |
formal model is correct. It is therefore definitely desirable to |
|
507 |
analyze PI using theorem proving, which gives more general results |
|
508 |
as well as deeper insight. And this is the purpose of this paper |
|
509 |
which gives a formal analysis of PI in the interactive theorem |
|
510 |
prover Isabelle using Higher Order Logic (HOL). The formalization |
|
262 | 511 |
focuses on on two issues: |
512 |
||
513 |
\begin{enumerate} |
|
514 |
\item The correctness of the protocol model itself. A series of desirable properties is |
|
515 |
derived until we are fully convinced that the formal model of PI does |
|
516 |
eliminate priority inversion. And a better understanding of PI is so obtained |
|
517 |
in due course. For example, we find through formalization that the choice of |
|
518 |
next thread to take hold when a |
|
519 |
resource is released is irrelevant for the very basic property of PI to hold. |
|
520 |
A point never mentioned in literature. |
|
521 |
\item The correctness of the implementation. A series of properties is derived the meaning |
|
522 |
of which can be used as guidelines on how PI can be implemented efficiently and correctly. |
|
523 |
\end{enumerate} |
|
524 |
||
525 |
The rest of the paper is organized as follows: Section \ref{overview} gives an overview |
|
526 |
of PI. Section \ref{model} introduces the formal model of PI. Section \ref{general} |
|
527 |
discusses a series of basic properties of PI. Section \ref{extension} shows formally |
|
528 |
how priority inversion is controlled by PI. Section \ref{implement} gives properties |
|
529 |
which can be used for guidelines of implementation. Section \ref{related} discusses |
|
530 |
related works. Section \ref{conclusion} concludes the whole paper. |
|
265 | 531 |
|
273 | 532 |
The basic priority inheritance protocol has two problems: |
533 |
||
534 |
It does not prevent a deadlock from happening in a program with circular lock dependencies. |
|
535 |
||
536 |
A chain of blocking may be formed; blocking duration can be substantial, though bounded. |
|
537 |
||
265 | 538 |
|
539 |
Contributions |
|
540 |
||
541 |
Despite the wide use of Priority Inheritance Protocol in real time operating |
|
542 |
system, it's correctness has never been formally proved and mechanically checked. |
|
543 |
All existing verification are based on model checking technology. Full automatic |
|
544 |
verification gives little help to understand why the protocol is correct. |
|
545 |
And results such obtained only apply to models of limited size. |
|
546 |
This paper presents a formal verification based on theorem proving. |
|
547 |
Machine checked formal proof does help to get deeper understanding. We found |
|
548 |
the fact which is not mentioned in the literature, that the choice of next |
|
549 |
thread to take over when an critical resource is release does not affect the correctness |
|
550 |
of the protocol. The paper also shows how formal proof can help to construct |
|
551 |
correct and efficient implementation.\bigskip |
|
552 |
||
262 | 553 |
*} |
554 |
||
555 |
section {* An overview of priority inversion and priority inheritance \label{overview} *} |
|
556 |
||
557 |
text {* |
|
558 |
||
559 |
Priority inversion refers to the phenomenon when a thread with high priority is blocked |
|
560 |
by a thread with low priority. Priority happens when the high priority thread requests |
|
561 |
for some critical resource already taken by the low priority thread. Since the high |
|
562 |
priority thread has to wait for the low priority thread to complete, it is said to be |
|
563 |
blocked by the low priority thread. Priority inversion might prevent high priority |
|
564 |
thread from fulfill its task in time if the duration of priority inversion is indefinite |
|
565 |
and unpredictable. Indefinite priority inversion happens when indefinite number |
|
566 |
of threads with medium priorities is activated during the period when the high |
|
567 |
priority thread is blocked by the low priority thread. Although these medium |
|
568 |
priority threads can not preempt the high priority thread directly, they are able |
|
569 |
to preempt the low priority threads and cause it to stay in critical section for |
|
570 |
an indefinite long duration. In this way, the high priority thread may be blocked indefinitely. |
|
571 |
||
572 |
Priority inheritance is one protocol proposed to avoid indefinite priority inversion. |
|
573 |
The basic idea is to let the high priority thread donate its priority to the low priority |
|
574 |
thread holding the critical resource, so that it will not be preempted by medium priority |
|
575 |
threads. The thread with highest priority will not be blocked unless it is requesting |
|
576 |
some critical resource already taken by other threads. Viewed from a different angle, |
|
577 |
any thread which is able to block the highest priority threads must already hold some |
|
578 |
critical resource. Further more, it must have hold some critical resource at the |
|
579 |
moment the highest priority is created, otherwise, it may never get change to run and |
|
580 |
get hold. Since the number of such resource holding lower priority threads is finite, |
|
581 |
if every one of them finishes with its own critical section in a definite duration, |
|
582 |
the duration the highest priority thread is blocked is definite as well. The key to |
|
583 |
guarantee lower priority threads to finish in definite is to donate them the highest |
|
584 |
priority. In such cases, the lower priority threads is said to have inherited the |
|
585 |
highest priority. And this explains the name of the protocol: |
|
586 |
{\em Priority Inheritance} and how Priority Inheritance prevents indefinite delay. |
|
587 |
||
588 |
The objectives of this paper are: |
|
589 |
\begin{enumerate} |
|
590 |
\item Build the above mentioned idea into formal model and prove a series of properties |
|
591 |
until we are convinced that the formal model does fulfill the original idea. |
|
592 |
\item Show how formally derived properties can be used as guidelines for correct |
|
593 |
and efficient implementation. |
|
594 |
\end{enumerate} |
|
595 |
The proof is totally formal in the sense that every detail is reduced to the |
|
596 |
very first principles of Higher Order Logic. The nature of interactive theorem |
|
597 |
proving is for the human user to persuade computer program to accept its arguments. |
|
598 |
A clear and simple understanding of the problem at hand is both a prerequisite and a |
|
599 |
byproduct of such an effort, because everything has finally be reduced to the very |
|
600 |
first principle to be checked mechanically. The former intuitive explanation of |
|
601 |
Priority Inheritance is just such a byproduct. |
|
602 |
*} |
|
603 |
||
604 |
section {* Formal model of Priority Inheritance \label{model} *} |
|
605 |
text {* |
|
606 |
\input{../../generated/PrioGDef} |
|
607 |
*} |
|
608 |
||
609 |
section {* General properties of Priority Inheritance \label{general} *} |
|
264 | 610 |
|
611 |
text {* |
|
612 |
The following are several very basic prioprites: |
|
613 |
\begin{enumerate} |
|
614 |
\item All runing threads must be ready (@{text "runing_ready"}): |
|
615 |
@{thm[display] "runing_ready"} |
|
616 |
\item All ready threads must be living (@{text "readys_threads"}): |
|
617 |
@{thm[display] "readys_threads"} |
|
618 |
\item There are finite many living threads at any moment (@{text "finite_threads"}): |
|
619 |
@{thm[display] "finite_threads"} |
|
620 |
\item Every waiting queue does not contain duplcated elements (@{text "wq_distinct"}): |
|
621 |
@{thm[display] "wq_distinct"} |
|
622 |
\item All threads in waiting queues are living threads (@{text "wq_threads"}): |
|
623 |
@{thm[display] "wq_threads"} |
|
624 |
\item The event which can get a thread into waiting queue must be @{term "P"}-events |
|
625 |
(@{text "block_pre"}): |
|
626 |
@{thm[display] "block_pre"} |
|
627 |
\item A thread may never wait for two different critical resources |
|
628 |
(@{text "waiting_unique"}): |
|
629 |
@{thm[display] waiting_unique[of _ _ "cs\<^isub>1" "cs\<^isub>2"]} |
|
630 |
\item Every resource can only be held by one thread |
|
631 |
(@{text "held_unique"}): |
|
632 |
@{thm[display] held_unique[of _ "th\<^isub>1" _ "th\<^isub>2"]} |
|
633 |
\item Every living thread has an unique precedence |
|
634 |
(@{text "preced_unique"}): |
|
635 |
@{thm[display] preced_unique[of "th\<^isub>1" _ "th\<^isub>2"]} |
|
636 |
\end{enumerate} |
|
637 |
*} |
|
638 |
||
639 |
text {* \noindent |
|
640 |
The following lemmas show how RAG is changed with the execution of events: |
|
641 |
\begin{enumerate} |
|
642 |
\item Execution of @{term "Set"} does not change RAG (@{text "depend_set_unchanged"}): |
|
643 |
@{thm[display] depend_set_unchanged} |
|
644 |
\item Execution of @{term "Create"} does not change RAG (@{text "depend_create_unchanged"}): |
|
645 |
@{thm[display] depend_create_unchanged} |
|
646 |
\item Execution of @{term "Exit"} does not change RAG (@{text "depend_exit_unchanged"}): |
|
647 |
@{thm[display] depend_exit_unchanged} |
|
648 |
\item Execution of @{term "P"} (@{text "step_depend_p"}): |
|
649 |
@{thm[display] step_depend_p} |
|
650 |
\item Execution of @{term "V"} (@{text "step_depend_v"}): |
|
651 |
@{thm[display] step_depend_v} |
|
652 |
\end{enumerate} |
|
653 |
*} |
|
654 |
||
655 |
text {* \noindent |
|
656 |
These properties are used to derive the following important results about RAG: |
|
657 |
\begin{enumerate} |
|
658 |
\item RAG is loop free (@{text "acyclic_depend"}): |
|
659 |
@{thm [display] acyclic_depend} |
|
660 |
\item RAGs are finite (@{text "finite_depend"}): |
|
661 |
@{thm [display] finite_depend} |
|
662 |
\item Reverse paths in RAG are well founded (@{text "wf_dep_converse"}): |
|
663 |
@{thm [display] wf_dep_converse} |
|
664 |
\item The dependence relation represented by RAG has a tree structure (@{text "unique_depend"}): |
|
665 |
@{thm [display] unique_depend[of _ _ "n\<^isub>1" "n\<^isub>2"]} |
|
666 |
\item All threads in RAG are living threads |
|
667 |
(@{text "dm_depend_threads"} and @{text "range_in"}): |
|
668 |
@{thm [display] dm_depend_threads range_in} |
|
669 |
\end{enumerate} |
|
670 |
*} |
|
671 |
||
672 |
text {* \noindent |
|
673 |
The following lemmas show how every node in RAG can be chased to ready threads: |
|
674 |
\begin{enumerate} |
|
675 |
\item Every node in RAG can be chased to a ready thread (@{text "chain_building"}): |
|
676 |
@{thm [display] chain_building[rule_format]} |
|
677 |
\item The ready thread chased to is unique (@{text "dchain_unique"}): |
|
678 |
@{thm [display] dchain_unique[of _ _ "th\<^isub>1" "th\<^isub>2"]} |
|
679 |
\end{enumerate} |
|
680 |
*} |
|
681 |
||
682 |
text {* \noindent |
|
683 |
Properties about @{term "next_th"}: |
|
684 |
\begin{enumerate} |
|
685 |
\item The thread taking over is different from the thread which is releasing |
|
686 |
(@{text "next_th_neq"}): |
|
687 |
@{thm [display] next_th_neq} |
|
688 |
\item The thread taking over is unique |
|
689 |
(@{text "next_th_unique"}): |
|
690 |
@{thm [display] next_th_unique[of _ _ _ "th\<^isub>1" "th\<^isub>2"]} |
|
691 |
\end{enumerate} |
|
692 |
*} |
|
693 |
||
694 |
text {* \noindent |
|
695 |
Some deeper results about the system: |
|
696 |
\begin{enumerate} |
|
697 |
\item There can only be one running thread (@{text "runing_unique"}): |
|
698 |
@{thm [display] runing_unique[of _ "th\<^isub>1" "th\<^isub>2"]} |
|
699 |
\item The maximum of @{term "cp"} and @{term "preced"} are equal (@{text "max_cp_eq"}): |
|
700 |
@{thm [display] max_cp_eq} |
|
701 |
\item There must be one ready thread having the max @{term "cp"}-value |
|
702 |
(@{text "max_cp_readys_threads"}): |
|
703 |
@{thm [display] max_cp_readys_threads} |
|
704 |
\end{enumerate} |
|
705 |
*} |
|
706 |
||
707 |
text {* \noindent |
|
708 |
The relationship between the count of @{text "P"} and @{text "V"} and the number of |
|
709 |
critical resources held by a thread is given as follows: |
|
710 |
\begin{enumerate} |
|
711 |
\item The @{term "V"}-operation decreases the number of critical resources |
|
712 |
one thread holds (@{text "cntCS_v_dec"}) |
|
713 |
@{thm [display] cntCS_v_dec} |
|
714 |
\item The number of @{text "V"} never exceeds the number of @{text "P"} |
|
715 |
(@{text "cnp_cnv_cncs"}): |
|
716 |
@{thm [display] cnp_cnv_cncs} |
|
717 |
\item The number of @{text "V"} equals the number of @{text "P"} when |
|
718 |
the relevant thread is not living: |
|
719 |
(@{text "cnp_cnv_eq"}): |
|
720 |
@{thm [display] cnp_cnv_eq} |
|
721 |
\item When a thread is not living, it does not hold any critical resource |
|
722 |
(@{text "not_thread_holdents"}): |
|
723 |
@{thm [display] not_thread_holdents} |
|
724 |
\item When the number of @{text "P"} equals the number of @{text "V"}, the relevant |
|
725 |
thread does not hold any critical resource, therefore no thread can depend on it |
|
726 |
(@{text "count_eq_dependents"}): |
|
727 |
@{thm [display] count_eq_dependents} |
|
728 |
\end{enumerate} |
|
729 |
*} |
|
262 | 730 |
|
731 |
section {* Key properties \label{extension} *} |
|
732 |
||
264 | 733 |
(*<*) |
734 |
context extend_highest_gen |
|
735 |
begin |
|
736 |
(*>*) |
|
737 |
||
738 |
text {* |
|
739 |
The essential of {\em Priority Inheritance} is to avoid indefinite priority inversion. For this |
|
740 |
purpose, we need to investigate what happens after one thread takes the highest precedence. |
|
741 |
A locale is used to describe such a situation, which assumes: |
|
742 |
\begin{enumerate} |
|
743 |
\item @{term "s"} is a valid state (@{text "vt_s"}): |
|
744 |
@{thm vt_s}. |
|
745 |
\item @{term "th"} is a living thread in @{term "s"} (@{text "threads_s"}): |
|
746 |
@{thm threads_s}. |
|
747 |
\item @{term "th"} has the highest precedence in @{term "s"} (@{text "highest"}): |
|
748 |
@{thm highest}. |
|
749 |
\item The precedence of @{term "th"} is @{term "Prc prio tm"} (@{text "preced_th"}): |
|
750 |
@{thm preced_th}. |
|
751 |
\end{enumerate} |
|
752 |
*} |
|
753 |
||
754 |
text {* \noindent |
|
755 |
Under these assumptions, some basic priority can be derived for @{term "th"}: |
|
756 |
\begin{enumerate} |
|
757 |
\item The current precedence of @{term "th"} equals its own precedence (@{text "eq_cp_s_th"}): |
|
758 |
@{thm [display] eq_cp_s_th} |
|
759 |
\item The current precedence of @{term "th"} is the highest precedence in |
|
760 |
the system (@{text "highest_cp_preced"}): |
|
761 |
@{thm [display] highest_cp_preced} |
|
762 |
\item The precedence of @{term "th"} is the highest precedence |
|
763 |
in the system (@{text "highest_preced_thread"}): |
|
764 |
@{thm [display] highest_preced_thread} |
|
765 |
\item The current precedence of @{term "th"} is the highest current precedence |
|
766 |
in the system (@{text "highest'"}): |
|
767 |
@{thm [display] highest'} |
|
768 |
\end{enumerate} |
|
769 |
*} |
|
770 |
||
771 |
text {* \noindent |
|
772 |
To analysis what happens after state @{term "s"} a sub-locale is defined, which |
|
773 |
assumes: |
|
774 |
\begin{enumerate} |
|
775 |
\item @{term "t"} is a valid extension of @{term "s"} (@{text "vt_t"}): @{thm vt_t}. |
|
776 |
\item Any thread created in @{term "t"} has priority no higher than @{term "prio"}, therefore |
|
777 |
its precedence can not be higher than @{term "th"}, therefore |
|
778 |
@{term "th"} remain to be the one with the highest precedence |
|
779 |
(@{text "create_low"}): |
|
780 |
@{thm [display] create_low} |
|
781 |
\item Any adjustment of priority in |
|
782 |
@{term "t"} does not happen to @{term "th"} and |
|
783 |
the priority set is no higher than @{term "prio"}, therefore |
|
784 |
@{term "th"} remain to be the one with the highest precedence (@{text "set_diff_low"}): |
|
785 |
@{thm [display] set_diff_low} |
|
786 |
\item Since we are investigating what happens to @{term "th"}, it is assumed |
|
787 |
@{term "th"} does not exit during @{term "t"} (@{text "exit_diff"}): |
|
788 |
@{thm [display] exit_diff} |
|
789 |
\end{enumerate} |
|
790 |
*} |
|
791 |
||
792 |
text {* \noindent |
|
793 |
All these assumptions are put into a predicate @{term "extend_highest_gen"}. |
|
794 |
It can be proved that @{term "extend_highest_gen"} holds |
|
795 |
for any moment @{text "i"} in it @{term "t"} (@{text "red_moment"}): |
|
796 |
@{thm [display] red_moment} |
|
797 |
||
798 |
From this, an induction principle can be derived for @{text "t"}, so that |
|
799 |
properties already derived for @{term "t"} can be applied to any prefix |
|
800 |
of @{text "t"} in the proof of new properties |
|
801 |
about @{term "t"} (@{text "ind"}): |
|
802 |
\begin{center} |
|
803 |
@{thm[display] ind} |
|
804 |
\end{center} |
|
805 |
||
806 |
The following properties can be proved about @{term "th"} in @{term "t"}: |
|
807 |
\begin{enumerate} |
|
808 |
\item In @{term "t"}, thread @{term "th"} is kept live and its |
|
809 |
precedence is preserved as well |
|
810 |
(@{text "th_kept"}): |
|
811 |
@{thm [display] th_kept} |
|
812 |
\item In @{term "t"}, thread @{term "th"}'s precedence is always the maximum among |
|
813 |
all living threads |
|
814 |
(@{text "max_preced"}): |
|
815 |
@{thm [display] max_preced} |
|
816 |
\item In @{term "t"}, thread @{term "th"}'s current precedence is always the maximum precedence |
|
817 |
among all living threads |
|
818 |
(@{text "th_cp_max_preced"}): |
|
819 |
@{thm [display] th_cp_max_preced} |
|
820 |
\item In @{term "t"}, thread @{term "th"}'s current precedence is always the maximum current |
|
821 |
precedence among all living threads |
|
822 |
(@{text "th_cp_max"}): |
|
823 |
@{thm [display] th_cp_max} |
|
824 |
\item In @{term "t"}, thread @{term "th"}'s current precedence equals its precedence at moment |
|
825 |
@{term "s"} |
|
826 |
(@{text "th_cp_preced"}): |
|
827 |
@{thm [display] th_cp_preced} |
|
828 |
\end{enumerate} |
|
829 |
*} |
|
830 |
||
831 |
text {* \noindent |
|
266 | 832 |
The main theorem of this part is to characterizing the running thread during @{term "t"} |
264 | 833 |
(@{text "runing_inversion_2"}): |
834 |
@{thm [display] runing_inversion_2} |
|
835 |
According to this, if a thread is running, it is either @{term "th"} or was |
|
836 |
already live and held some resource |
|
837 |
at moment @{text "s"} (expressed by: @{text "cntV s th' < cntP s th'"}). |
|
838 |
||
839 |
Since there are only finite many threads live and holding some resource at any moment, |
|
840 |
if every such thread can release all its resources in finite duration, then after finite |
|
841 |
duration, none of them may block @{term "th"} anymore. So, no priority inversion may happen |
|
842 |
then. |
|
843 |
*} |
|
844 |
||
845 |
(*<*) |
|
846 |
end |
|
847 |
(*>*) |
|
848 |
||
262 | 849 |
section {* Properties to guide implementation \label{implement} *} |
850 |
||
264 | 851 |
text {* |
266 | 852 |
The properties (especially @{text "runing_inversion_2"}) convinced us that the model defined |
853 |
in Section \ref{model} does prevent indefinite priority inversion and therefore fulfills |
|
264 | 854 |
the fundamental requirement of Priority Inheritance protocol. Another purpose of this paper |
266 | 855 |
is to show how this model can be used to guide a concrete implementation. As discussed in |
276 | 856 |
Section 5.6.5 of \cite{Vahalia96}, the implementation of Priority Inheritance in Solaris |
266 | 857 |
uses sophisticated linking data structure. Except discussing two scenarios to show how |
858 |
the data structure should be manipulated, a lot of details of the implementation are missing. |
|
859 |
In \cite{Faria08,conf/fase/JahierHR09,WellingsBSB07} the protocol is described formally |
|
860 |
using different notations, but little information is given on how this protocol can be |
|
861 |
implemented efficiently, especially there is no information on how these data structure |
|
862 |
should be manipulated. |
|
863 |
||
864 |
Because the scheduling of threads is based on current precedence, |
|
865 |
the central issue in implementation of Priority Inheritance is how to compute the precedence |
|
866 |
correctly and efficiently. As long as the precedence is correct, it is very easy to |
|
867 |
modify the scheduling algorithm to select the correct thread to execute. |
|
868 |
||
869 |
First, it can be proved that the computation of current precedence @{term "cp"} of a threads |
|
870 |
only involves its children (@{text "cp_rec"}): |
|
871 |
@{thm [display] cp_rec} |
|
872 |
where @{term "children s th"} represents the set of children of @{term "th"} in the current |
|
873 |
RAG: |
|
874 |
\[ |
|
875 |
@{thm (lhs) children_def} @{text "\<equiv>"} @{thm (rhs) children_def} |
|
876 |
\] |
|
877 |
where the definition of @{term "child"} is: |
|
878 |
\[ @{thm (lhs) child_def} @{text "\<equiv>"} @{thm (rhs) child_def} |
|
879 |
\] |
|
880 |
||
881 |
The aim of this section is to fill the missing details of how current precedence should |
|
882 |
be changed with the happening of events, with each event type treated by one subsection, |
|
883 |
where the computation of @{term "cp"} uses lemma @{text "cp_rec"}. |
|
884 |
*} |
|
885 |
||
886 |
subsection {* Event @{text "Set th prio"} *} |
|
887 |
||
888 |
(*<*) |
|
889 |
context step_set_cps |
|
890 |
begin |
|
891 |
(*>*) |
|
892 |
||
893 |
text {* |
|
894 |
The context under which event @{text "Set th prio"} happens is formalized as follows: |
|
895 |
\begin{enumerate} |
|
896 |
\item The formation of @{term "s"} (@{text "s_def"}): @{thm s_def}. |
|
897 |
\item State @{term "s"} is a valid state (@{text "vt_s"}): @{thm vt_s}. This implies |
|
898 |
event @{text "Set th prio"} is eligible to happen under state @{term "s'"} and |
|
899 |
state @{term "s'"} is a valid state. |
|
900 |
\end{enumerate} |
|
264 | 901 |
*} |
902 |
||
266 | 903 |
text {* \noindent |
904 |
Under such a context, we investigated how the current precedence @{term "cp"} of |
|
905 |
threads change from state @{term "s'"} to @{term "s"} and obtained the following |
|
906 |
conclusions: |
|
907 |
\begin{enumerate} |
|
908 |
%% \item The RAG does not change (@{text "eq_dep"}): @{thm "eq_dep"}. |
|
909 |
\item All threads with no dependence relation with thread @{term "th"} have their |
|
910 |
@{term "cp"}-value unchanged (@{text "eq_cp"}): |
|
911 |
@{thm [display] eq_cp} |
|
912 |
This lemma implies the @{term "cp"}-value of @{term "th"} |
|
913 |
and those threads which have a dependence relation with @{term "th"} might need |
|
914 |
to be recomputed. The way to do this is to start from @{term "th"} |
|
915 |
and follow the @{term "depend"}-chain to recompute the @{term "cp"}-value of every |
|
916 |
encountered thread using lemma @{text "cp_rec"}. |
|
917 |
Since the @{term "depend"}-relation is loop free, this procedure |
|
918 |
can always stop. The the following lemma shows this procedure actually could stop earlier. |
|
919 |
\item The following two lemma shows, if a thread the re-computation of which |
|
920 |
gives an unchanged @{term "cp"}-value, the procedure described above can stop. |
|
921 |
\begin{enumerate} |
|
922 |
\item Lemma @{text "eq_up_self"} shows if the re-computation of |
|
923 |
@{term "th"}'s @{term "cp"} gives the same result, the procedure can stop: |
|
924 |
@{thm [display] eq_up_self} |
|
925 |
\item Lemma @{text "eq_up"}) shows if the re-computation at intermediate threads |
|
926 |
gives unchanged result, the procedure can stop: |
|
927 |
@{thm [display] eq_up} |
|
928 |
\end{enumerate} |
|
929 |
\end{enumerate} |
|
930 |
*} |
|
931 |
||
932 |
(*<*) |
|
933 |
end |
|
934 |
(*>*) |
|
264 | 935 |
|
272 | 936 |
subsection {* Event @{text "V th cs"} *} |
937 |
||
938 |
(*<*) |
|
939 |
context step_v_cps_nt |
|
940 |
begin |
|
941 |
(*>*) |
|
942 |
||
943 |
text {* |
|
944 |
The context under which event @{text "V th cs"} happens is formalized as follows: |
|
945 |
\begin{enumerate} |
|
946 |
\item The formation of @{term "s"} (@{text "s_def"}): @{thm s_def}. |
|
947 |
\item State @{term "s"} is a valid state (@{text "vt_s"}): @{thm vt_s}. This implies |
|
948 |
event @{text "V th cs"} is eligible to happen under state @{term "s'"} and |
|
949 |
state @{term "s'"} is a valid state. |
|
950 |
\end{enumerate} |
|
951 |
*} |
|
952 |
||
953 |
text {* \noindent |
|
954 |
Under such a context, we investigated how the current precedence @{term "cp"} of |
|
955 |
threads change from state @{term "s'"} to @{term "s"}. |
|
956 |
||
957 |
||
958 |
Two subcases are considerted, |
|
959 |
where the first is that there exits @{term "th'"} |
|
960 |
such that |
|
961 |
@{thm [display] nt} |
|
962 |
holds, which means there exists a thread @{term "th'"} to take over |
|
963 |
the resource release by thread @{term "th"}. |
|
964 |
In this sub-case, the following results are obtained: |
|
965 |
\begin{enumerate} |
|
966 |
\item The change of RAG is given by lemma @{text "depend_s"}: |
|
967 |
@{thm [display] "depend_s"} |
|
968 |
which shows two edges are removed while one is added. These changes imply how |
|
969 |
the current precedences should be re-computed. |
|
970 |
\item First all threads different from @{term "th"} and @{term "th'"} have their |
|
971 |
@{term "cp"}-value kept, therefore do not need a re-computation |
|
972 |
(@{text "cp_kept"}): @{thm [display] cp_kept} |
|
973 |
This lemma also implies, only the @{term "cp"}-values of @{term "th"} and @{term "th'"} |
|
974 |
need to be recomputed. |
|
975 |
\end{enumerate} |
|
976 |
*} |
|
977 |
||
978 |
(*<*) |
|
979 |
end |
|
980 |
||
981 |
context step_v_cps_nnt |
|
982 |
begin |
|
983 |
(*>*) |
|
984 |
||
985 |
text {* |
|
986 |
The other sub-case is when for all @{text "th'"} |
|
987 |
@{thm [display] nnt} |
|
988 |
holds, no such thread exists. The following results can be obtained for this |
|
989 |
sub-case: |
|
990 |
\begin{enumerate} |
|
991 |
\item The change of RAG is given by lemma @{text "depend_s"}: |
|
992 |
@{thm [display] depend_s} |
|
993 |
which means only one edge is removed. |
|
994 |
\item In this case, no re-computation is needed (@{text "eq_cp"}): |
|
995 |
@{thm [display] eq_cp} |
|
996 |
\end{enumerate} |
|
997 |
*} |
|
998 |
||
999 |
(*<*) |
|
1000 |
end |
|
1001 |
(*>*) |
|
1002 |
||
1003 |
||
1004 |
subsection {* Event @{text "P th cs"} *} |
|
1005 |
||
1006 |
(*<*) |
|
1007 |
context step_P_cps_e |
|
1008 |
begin |
|
1009 |
(*>*) |
|
1010 |
||
1011 |
text {* |
|
1012 |
The context under which event @{text "P th cs"} happens is formalized as follows: |
|
1013 |
\begin{enumerate} |
|
1014 |
\item The formation of @{term "s"} (@{text "s_def"}): @{thm s_def}. |
|
1015 |
\item State @{term "s"} is a valid state (@{text "vt_s"}): @{thm vt_s}. This implies |
|
1016 |
event @{text "P th cs"} is eligible to happen under state @{term "s'"} and |
|
1017 |
state @{term "s'"} is a valid state. |
|
1018 |
\end{enumerate} |
|
1019 |
||
1020 |
This case is further divided into two sub-cases. The first is when @{thm ee} holds. |
|
1021 |
The following results can be obtained: |
|
1022 |
\begin{enumerate} |
|
1023 |
\item One edge is added to the RAG (@{text "depend_s"}): |
|
1024 |
@{thm [display] depend_s} |
|
1025 |
\item No re-computation is needed (@{text "eq_cp"}): |
|
1026 |
@{thm [display] eq_cp} |
|
1027 |
\end{enumerate} |
|
1028 |
*} |
|
1029 |
||
1030 |
(*<*) |
|
1031 |
end |
|
1032 |
||
1033 |
context step_P_cps_ne |
|
1034 |
begin |
|
1035 |
(*>*) |
|
1036 |
||
1037 |
text {* |
|
1038 |
The second is when @{thm ne} holds. |
|
1039 |
The following results can be obtained: |
|
1040 |
\begin{enumerate} |
|
1041 |
\item One edge is added to the RAG (@{text "depend_s"}): |
|
1042 |
@{thm [display] depend_s} |
|
1043 |
\item Threads with no dependence relation with @{term "th"} do not need a re-computation |
|
1044 |
of their @{term "cp"}-values (@{text "eq_cp"}): |
|
1045 |
@{thm [display] eq_cp} |
|
1046 |
This lemma implies all threads with a dependence relation with @{term "th"} may need |
|
1047 |
re-computation. |
|
1048 |
\item Similar to the case of @{term "Set"}, the computation procedure could stop earlier |
|
1049 |
(@{text "eq_up"}): |
|
1050 |
@{thm [display] eq_up} |
|
1051 |
\end{enumerate} |
|
1052 |
||
1053 |
*} |
|
1054 |
||
1055 |
(*<*) |
|
1056 |
end |
|
1057 |
(*>*) |
|
1058 |
||
1059 |
subsection {* Event @{text "Create th prio"} *} |
|
1060 |
||
1061 |
(*<*) |
|
1062 |
context step_create_cps |
|
1063 |
begin |
|
1064 |
(*>*) |
|
1065 |
||
1066 |
text {* |
|
1067 |
The context under which event @{text "Create th prio"} happens is formalized as follows: |
|
1068 |
\begin{enumerate} |
|
1069 |
\item The formation of @{term "s"} (@{text "s_def"}): @{thm s_def}. |
|
1070 |
\item State @{term "s"} is a valid state (@{text "vt_s"}): @{thm vt_s}. This implies |
|
1071 |
event @{text "Create th prio"} is eligible to happen under state @{term "s'"} and |
|
1072 |
state @{term "s'"} is a valid state. |
|
1073 |
\end{enumerate} |
|
1074 |
The following results can be obtained under this context: |
|
1075 |
\begin{enumerate} |
|
1076 |
\item The RAG does not change (@{text "eq_dep"}): |
|
1077 |
@{thm [display] eq_dep} |
|
1078 |
\item All threads other than @{term "th"} do not need re-computation (@{text "eq_cp"}): |
|
1079 |
@{thm [display] eq_cp} |
|
1080 |
\item The @{term "cp"}-value of @{term "th"} equals its precedence |
|
1081 |
(@{text "eq_cp_th"}): |
|
1082 |
@{thm [display] eq_cp_th} |
|
1083 |
\end{enumerate} |
|
1084 |
||
1085 |
*} |
|
1086 |
||
1087 |
||
1088 |
(*<*) |
|
1089 |
end |
|
1090 |
(*>*) |
|
1091 |
||
1092 |
subsection {* Event @{text "Exit th"} *} |
|
1093 |
||
1094 |
(*<*) |
|
1095 |
context step_exit_cps |
|
1096 |
begin |
|
1097 |
(*>*) |
|
1098 |
||
1099 |
text {* |
|
1100 |
The context under which event @{text "Exit th"} happens is formalized as follows: |
|
1101 |
\begin{enumerate} |
|
1102 |
\item The formation of @{term "s"} (@{text "s_def"}): @{thm s_def}. |
|
1103 |
\item State @{term "s"} is a valid state (@{text "vt_s"}): @{thm vt_s}. This implies |
|
1104 |
event @{text "Exit th"} is eligible to happen under state @{term "s'"} and |
|
1105 |
state @{term "s'"} is a valid state. |
|
1106 |
\end{enumerate} |
|
1107 |
The following results can be obtained under this context: |
|
1108 |
\begin{enumerate} |
|
1109 |
\item The RAG does not change (@{text "eq_dep"}): |
|
1110 |
@{thm [display] eq_dep} |
|
1111 |
\item All threads other than @{term "th"} do not need re-computation (@{text "eq_cp"}): |
|
1112 |
@{thm [display] eq_cp} |
|
1113 |
\end{enumerate} |
|
1114 |
Since @{term th} does not live in state @{term "s"}, there is no need to compute |
|
1115 |
its @{term cp}-value. |
|
1116 |
*} |
|
1117 |
||
1118 |
(*<*) |
|
1119 |
end |
|
1120 |
(*>*) |
|
1121 |
||
1122 |
||
262 | 1123 |
section {* Related works \label{related} *} |
1124 |
||
1125 |
text {* |
|
1126 |
\begin{enumerate} |
|
1127 |
\item {\em Integrating Priority Inheritance Algorithms in the Real-Time Specification for Java} |
|
1128 |
\cite{WellingsBSB07} models and verifies the combination of Priority Inheritance (PI) and |
|
1129 |
Priority Ceiling Emulation (PCE) protocols in the setting of Java virtual machine |
|
1130 |
using extended Timed Automata(TA) formalism of the UPPAAL tool. Although a detailed |
|
1131 |
formal model of combined PI and PCE is given, the number of properties is quite |
|
1132 |
small and the focus is put on the harmonious working of PI and PCE. Most key features of PI |
|
1133 |
(as well as PCE) are not shown. Because of the limitation of the model checking technique |
|
1134 |
used there, properties are shown only for a small number of scenarios. Therefore, |
|
1135 |
the verification does not show the correctness of the formal model itself in a |
|
1136 |
convincing way. |
|
1137 |
\item {\em Formal Development of Solutions for Real-Time Operating Systems with TLA+/TLC} |
|
1138 |
\cite{Faria08}. A formal model of PI is given in TLA+. Only 3 properties are shown |
|
1139 |
for PI using model checking. The limitation of model checking is intrinsic to the work. |
|
1140 |
\item {\em Synchronous modeling and validation of priority inheritance schedulers} |
|
1141 |
\cite{conf/fase/JahierHR09}. Gives a formal model |
|
1142 |
of PI and PCE in AADL (Architecture Analysis \& Design Language) and checked |
|
1143 |
several properties using model checking. The number of properties shown there is |
|
1144 |
less than here and the scale is also limited by the model checking technique. |
|
1145 |
\item {\em The Priority Ceiling Protocol: Formalization and Analysis Using PVS} |
|
1146 |
\cite{dutertre99b}. Formalized another protocol for Priority Inversion in the |
|
1147 |
interactive theorem proving system PVS. |
|
1148 |
\end{enumerate} |
|
1149 |
||
1150 |
||
1151 |
There are several works on inversion avoidance: |
|
1152 |
\begin{enumerate} |
|
1153 |
\item {\em Solving the group priority inversion problem in a timed asynchronous system} |
|
1154 |
\cite{Wang:2002:SGP}. The notion of Group Priority Inversion is introduced. The main |
|
1155 |
strategy is still inversion avoidance. The method is by reordering requests |
|
1156 |
in the setting of Client-Server. |
|
1157 |
\item {\em A Formalization of Priority Inversion} \cite{journals/rts/BabaogluMS93}. |
|
1158 |
Formalized the notion of Priority |
|
1159 |
Inversion and proposes methods to avoid it. |
|
1160 |
\end{enumerate} |
|
1161 |
||
1162 |
{\em Examples of inaccurate specification of the protocol ???}. |
|
1163 |
||
1164 |
*} |
|
1165 |
||
1166 |
section {* Conclusions \label{conclusion} *} |
|
1167 |
||
286 | 1168 |
text {* |
1169 |
The work in this paper only deals with single CPU configurations. The |
|
1170 |
"one CPU" assumption is essential for our formalisation, because the |
|
1171 |
main lemma fails in multi-CPU configuration. The lemma says that any |
|
1172 |
runing thead must be the one with the highest prioirty or already held |
|
1173 |
some resource when the highest priority thread was initiated. When |
|
1174 |
there are multiple CPUs, it may well be the case that a threads did |
|
1175 |
not hold any resource when the highest priority thread was initiated, |
|
1176 |
but that thread still runs after that moment on a separate CPU. In |
|
1177 |
this way, the main lemma does not hold anymore. |
|
1178 |
||
1179 |
||
1180 |
There are some works deals with priority inversion in multi-CPU |
|
1181 |
configurations[???], but none of them have given a formal correctness |
|
1182 |
proof. The extension of our formal proof to deal with multi-CPU |
|
1183 |
configurations is not obvious. One possibility, as suggested in paper |
|
1184 |
[???], is change our formal model (the defiintion of "schs") to give |
|
1185 |
the released resource to the thread with the highest prioirty. In this |
|
1186 |
way, indefinite prioirty inversion can be avoided, but for a quite |
|
1187 |
different reason from the one formalized in this paper (because the |
|
1188 |
"mail lemma" will be different). This means a formal correctness proof |
|
1189 |
for milt-CPU configuration would be quite different from the one given |
|
1190 |
in this paper. The solution of prioirty inversion problem in mult-CPU |
|
1191 |
configurations is a different problem which needs different solutions |
|
1192 |
which is outside the scope of this paper. |
|
1193 |
||
1194 |
*} |
|
1195 |
||
262 | 1196 |
(*<*) |
1197 |
end |
|
1198 |
(*>*) |