author | Christian Urban <christian dot urban at kcl dot ac dot uk> |
Sat, 22 Dec 2012 14:50:29 +0000 | |
changeset 17 | 105715a0a807 |
parent 16 | 9764023f719e |
child 20 | b56616fd88dd |
permissions | -rwxr-xr-x |
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 |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
3 |
imports "../CpsG" "../ExtGG" "~~/src/HOL/Library/LaTeXsugar" |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
4 |
begin |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
5 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
6 |
(* |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
7 |
find_unused_assms CpsG |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
8 |
find_unused_assms ExtGG |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
9 |
find_unused_assms Moment |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
10 |
find_unused_assms Precedence_ord |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
11 |
find_unused_assms PrioG |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
12 |
find_unused_assms PrioGDef |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
13 |
*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
14 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
15 |
ML {* |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
16 |
open Printer; |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
17 |
show_question_marks_default := false; |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
18 |
*} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
19 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
20 |
notation (latex output) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
21 |
Cons ("_::_" [78,77] 73) and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
22 |
vt ("valid'_state") and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
23 |
runing ("running") and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
24 |
birthtime ("last'_set") and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
25 |
If ("(\<^raw:\textrm{>if\<^raw:}> (_)/ \<^raw:\textrm{>then\<^raw:}> (_)/ \<^raw:\textrm{>else\<^raw:}> (_))" 10) and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
26 |
Prc ("'(_, _')") and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
27 |
holding ("holds") and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
28 |
waiting ("waits") and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
29 |
Th ("T") and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
30 |
Cs ("C") and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
31 |
readys ("ready") and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
32 |
depend ("RAG") and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
33 |
preced ("prec") and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
34 |
cpreced ("cprec") and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
35 |
dependents ("dependants") and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
36 |
cp ("cprec") and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
37 |
holdents ("resources") and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
38 |
original_priority ("priority") and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
39 |
DUMMY ("\<^raw:\mbox{$\_\!\_$}>") |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
40 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
41 |
(*abbreviation |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
42 |
"detached s th \<equiv> cntP s th = cntV s th" |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
43 |
*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
44 |
(*>*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
45 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
46 |
section {* Introduction *} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
47 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
48 |
text {* |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
49 |
Many real-time systems need to support threads involving priorities and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
50 |
locking of resources. Locking of resources ensures mutual exclusion |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
51 |
when accessing shared data or devices that cannot be |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
52 |
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
|
53 |
finish their work within deadlines. Unfortunately, both features |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
54 |
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
|
55 |
\emph{Priority Inversion}. Suppose three threads having priorities |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
56 |
$H$(igh), $M$(edium) and $L$(ow). We would expect that the thread |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
57 |
$H$ blocks any other thread with lower priority and the thread itself cannot |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
58 |
be blocked indefinitely by threads with lower priority. Alas, in a naive |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
59 |
implementation of resource locking and priorities this property can |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
60 |
be violated. For this let $L$ be in the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
61 |
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
|
62 |
therefore wait for $L$ to exit the critical section and release this |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
63 |
lock. The problem is that $L$ might in turn be blocked by any |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
64 |
thread with priority $M$, and so $H$ sits there potentially waiting |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
65 |
indefinitely. Since $H$ is blocked by threads with lower |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
66 |
priorities, the problem is called Priority Inversion. It was first |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
67 |
described in \cite{Lampson80} in the context of the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
68 |
Mesa programming language designed for concurrent programming. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
69 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
70 |
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
|
71 |
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
|
72 |
The classic example where this happened is the software that |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
73 |
controlled the Mars Pathfinder mission in 1997 \cite{Reeves98}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
74 |
Once the spacecraft landed, the software shut down at irregular |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
75 |
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
|
76 |
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
|
77 |
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
|
78 |
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
|
79 |
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
|
80 |
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
|
81 |
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
|
82 |
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
|
83 |
\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
|
84 |
Inheritance Protocol} \cite{Sha90} and others sometimes also call it |
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
85 |
\emph{Priority Boosting}, \emph{Priority Donation} or \emph{Priority Lending}.} |
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
86 |
in the scheduling software. |
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
87 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
88 |
The idea behind PIP is to let the thread $L$ temporarily inherit |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
89 |
the high priority from $H$ until $L$ leaves the critical section |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
90 |
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
|
91 |
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
|
92 |
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
|
93 |
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
|
94 |
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
|
95 |
in the Mars Pathfinder mission, in Boeing's 787 Dreamliner, Honda's |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
96 |
ASIMO robot, etc.), but also the POSIX 1003.1c Standard realised for |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
97 |
example in libraries for FreeBSD, Solaris and Linux. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
98 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
99 |
One advantage of PIP is that increasing the priority of a thread |
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
100 |
can be performed dynamically by the scheduler. This is in contrast |
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
101 |
to, for example, \emph{Priority Ceiling} \cite{Sha90}, another |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
102 |
solution to the Priority Inversion problem, which requires static |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
103 |
analysis of the program in order to prevent Priority |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
104 |
Inversion. However, there has also been strong criticism against |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
105 |
PIP. For instance, PIP cannot prevent deadlocks when lock |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
106 |
dependencies are circular, and also blocking times can be |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
107 |
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
|
108 |
Though, most criticism against PIP centres around unreliable |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
109 |
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
|
110 |
For example, Yodaiken writes in \cite{Yodaiken02}: |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
111 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
112 |
\begin{quote} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
113 |
\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
|
114 |
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
|
115 |
\end{quote} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
116 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
117 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
118 |
He suggests avoiding PIP altogether by designing the system so that no |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
119 |
priority inversion may happen in the first place. However, such ideal designs may |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
120 |
not always be achievable in practice. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
121 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
122 |
In our opinion, there is clearly a need for investigating correct |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
123 |
algorithms for PIP. A few specifications for PIP exist (in English) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
124 |
and also a few high-level descriptions of implementations (e.g.~in |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
125 |
the textbook \cite[Section 5.6.5]{Vahalia96}), but they help little |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
126 |
with actual implementations. That this is a problem in practice is |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
127 |
proved by an email by Baker, who wrote on 13 July 2009 on the Linux |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
128 |
Kernel mailing list: |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
129 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
130 |
\begin{quote} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
131 |
\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
|
132 |
implementation is a nightmare: extremely heavy weight, involving |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
133 |
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
|
134 |
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
|
135 |
wait operations.'' |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
136 |
\end{quote} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
137 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
138 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
139 |
The criticism by Yodaiken, Baker and others suggests another look |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
140 |
at PIP from a more abstract level (but still concrete enough |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
141 |
to inform an implementation), and makes PIP a good candidate for a |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
142 |
formal verification. An additional reason is that the original |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
143 |
presentation of PIP~\cite{Sha90}, despite being informally |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
144 |
``proved'' correct, is actually \emph{flawed}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
145 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
146 |
Yodaiken \cite{Yodaiken02} points to a subtlety that had been |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
147 |
overlooked in the informal proof by Sha et al. They specify in |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
148 |
\cite{Sha90} that after the thread (whose priority has been raised) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
149 |
completes its critical section and releases the lock, it ``returns |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
150 |
to its original priority level.'' This leads them to believe that an |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
151 |
implementation of PIP is ``rather straightforward''~\cite{Sha90}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
152 |
Unfortunately, as Yodaiken points out, this behaviour is too |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
153 |
simplistic. Consider the case where the low priority thread $L$ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
154 |
locks \emph{two} resources, and two high-priority threads $H$ and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
155 |
$H'$ each wait for one of them. If $L$ releases one resource |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
156 |
so that $H$, say, can proceed, then we still have Priority Inversion |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
157 |
with $H'$ (which waits for the other resource). The correct |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
158 |
behaviour for $L$ is to switch to the highest remaining priority of |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
159 |
the threads that it blocks. The advantage of formalising the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
160 |
correctness of a high-level specification of PIP in a theorem prover |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
161 |
is that such issues clearly show up and cannot be overlooked as in |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
162 |
informal reasoning (since we have to analyse all possible behaviours |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
163 |
of threads, i.e.~\emph{traces}, that could possibly happen).\medskip |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
164 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
165 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
166 |
{\bf Contributions:} There have been earlier formal investigations |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
167 |
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
|
168 |
checking techniques. This paper presents a formalised and |
8
5ba3d79622da
added a paragraph about RAGS
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
7
diff
changeset
|
169 |
mechanically checked proof for the correctness of PIP. For this we |
12
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
170 |
needed to design a new correctness criterion for PIP. In contrast to model checking, our |
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
171 |
formalisation provides insight into why PIP is correct and allows us |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
172 |
to prove stronger properties that, as we will show, can help with an |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
173 |
efficient implementation of PIP in the educational PINTOS operating |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
174 |
system \cite{PINTOS}. For example, we found by ``playing'' with the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
175 |
formalisation that the choice of the next thread to take over a lock |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
176 |
when a resource is released is irrelevant for PIP being correct---a |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
177 |
fact that has not been mentioned in the literature and not been used |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
178 |
in the reference implementation of PIP in PINTOS. This fact, however, is important |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
179 |
for an efficient implementation of PIP, because we can give the lock |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
180 |
to the thread with the highest priority so that it terminates more |
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
181 |
quickly. We were also able to generalise the scheduler of Sha |
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
182 |
et al.~\cite{Sha90} to the practically relevant case where critical |
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
183 |
sections can overlap. |
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
184 |
*} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
185 |
|
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
186 |
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
|
187 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
188 |
text {* |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
189 |
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
|
190 |
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
|
191 |
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
|
192 |
Following good experience in earlier work \cite{Wang09}, |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
193 |
our model of PIP is based on Paulson's inductive approach to protocol |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
194 |
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
|
195 |
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
|
196 |
\emph{Events} of PIP fall |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
197 |
into five categories defined as the datatype: |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
198 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
199 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
200 |
\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
|
201 |
\isacommand{datatype} event |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
202 |
& @{text "="} & @{term "Create thread priority"}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
203 |
& @{text "|"} & @{term "Exit thread"} \\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
204 |
& @{text "|"} & @{term "Set thread priority"} & {\rm reset of the priority for} @{text thread}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
205 |
& @{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
|
206 |
& @{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
|
207 |
\end{tabular}} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
208 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
209 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
210 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
211 |
whereby threads, priorities and (critical) resources are represented |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
212 |
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
|
213 |
a thread obtains a new priority given by the programmer or |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
214 |
user (for example via the {\tt nice} utility under UNIX). As in Paulson's work, we |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
215 |
need to define functions that allow us to make some observations |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
216 |
about states. One, called @{term threads}, calculates the set of |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
217 |
``live'' threads that we have seen so far: |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
218 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
219 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
220 |
\mbox{\begin{tabular}{lcl} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
221 |
@{thm (lhs) threads.simps(1)} & @{text "\<equiv>"} & |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
222 |
@{thm (rhs) threads.simps(1)}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
223 |
@{thm (lhs) threads.simps(2)[where thread="th"]} & @{text "\<equiv>"} & |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
224 |
@{thm (rhs) threads.simps(2)[where thread="th"]}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
225 |
@{thm (lhs) threads.simps(3)[where thread="th"]} & @{text "\<equiv>"} & |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
226 |
@{thm (rhs) threads.simps(3)[where thread="th"]}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
227 |
@{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
|
228 |
\end{tabular}} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
229 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
230 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
231 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
232 |
In this definition @{term "DUMMY # DUMMY"} stands for list-cons. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
233 |
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
|
234 |
defined as |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
235 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
236 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
237 |
\mbox{\begin{tabular}{lcl} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
238 |
@{thm (lhs) original_priority.simps(1)[where thread="th"]} & @{text "\<equiv>"} & |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
239 |
@{thm (rhs) original_priority.simps(1)[where thread="th"]}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
240 |
@{thm (lhs) original_priority.simps(2)[where thread="th" and thread'="th'"]} & @{text "\<equiv>"} & |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
241 |
@{thm (rhs) original_priority.simps(2)[where thread="th" and thread'="th'"]}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
242 |
@{thm (lhs) original_priority.simps(3)[where thread="th" and thread'="th'"]} & @{text "\<equiv>"} & |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
243 |
@{thm (rhs) original_priority.simps(3)[where thread="th" and thread'="th'"]}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
244 |
@{term "original_priority th (DUMMY#s)"} & @{text "\<equiv>"} & @{term "original_priority th s"}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
245 |
\end{tabular}} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
246 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
247 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
248 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
249 |
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
|
250 |
threads that have not (yet) been created. The last function we need |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
251 |
calculates the ``time'', or index, at which time a process had its |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
252 |
priority last set. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
253 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
254 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
255 |
\mbox{\begin{tabular}{lcl} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
256 |
@{thm (lhs) birthtime.simps(1)[where thread="th"]} & @{text "\<equiv>"} & |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
257 |
@{thm (rhs) birthtime.simps(1)[where thread="th"]}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
258 |
@{thm (lhs) birthtime.simps(2)[where thread="th" and thread'="th'"]} & @{text "\<equiv>"} & |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
259 |
@{thm (rhs) birthtime.simps(2)[where thread="th" and thread'="th'"]}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
260 |
@{thm (lhs) birthtime.simps(3)[where thread="th" and thread'="th'"]} & @{text "\<equiv>"} & |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
261 |
@{thm (rhs) birthtime.simps(3)[where thread="th" and thread'="th'"]}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
262 |
@{term "birthtime th (DUMMY#s)"} & @{text "\<equiv>"} & @{term "birthtime th s"}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
263 |
\end{tabular}} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
264 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
265 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
266 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
267 |
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
|
268 |
of events @{text s}. Again the default value in this function is @{text 0} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
269 |
for threads that have not been created yet. A \emph{precedence} of a thread @{text th} in a |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
270 |
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
|
271 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
272 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
273 |
@{thm preced_def[where thread="th"]} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
274 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
275 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
276 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
277 |
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
|
278 |
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
|
279 |
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
|
280 |
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
|
281 |
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
|
282 |
set earlier, since for such threads it is more urgent to finish their work. In an implementation |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
283 |
this choice would translate to a quite natural FIFO-scheduling of processes with |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
284 |
the same priority. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
285 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
286 |
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
|
287 |
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
|
288 |
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
|
289 |
that is in possession of the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
290 |
``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
|
291 |
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
|
292 |
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
|
293 |
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
|
294 |
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
|
295 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
296 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
297 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
298 |
@{thm cs_holding_def[where thread="th"]}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
299 |
@{thm cs_waiting_def[where thread="th"]} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
300 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
301 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
302 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
303 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
304 |
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
|
305 |
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
|
306 |
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
|
307 |
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
|
308 |
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
|
309 |
empty list for every resource. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
310 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
311 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
312 |
@{abbrev all_unlocked}\hfill\numbered{allunlocked} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
313 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
314 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
315 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
316 |
Using @{term "holding"} and @{term waiting}, we can introduce \emph{Resource Allocation Graphs} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
317 |
(RAG), which represent the dependencies between threads and resources. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
318 |
We represent RAGs as relations using pairs of the form |
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 |
@{term "(Th th, Cs cs)"} \hspace{5mm}{\rm and}\hspace{5mm} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
322 |
@{term "(Cs cs, Th th)"} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
323 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
324 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
325 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
326 |
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
|
327 |
\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
|
328 |
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
|
329 |
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
|
330 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
331 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
332 |
@{thm cs_depend_def} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
333 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
334 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
335 |
\noindent |
12
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
336 |
If there is no cycle, then every RAG can be pictured as a forrest of trees, for example as follows: |
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
337 |
|
12
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
338 |
\begin{figure}[h] |
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
339 |
\begin{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
340 |
\newcommand{\fnt}{\fontsize{7}{8}\selectfont} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
341 |
\begin{tikzpicture}[scale=1] |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
342 |
%%\draw[step=2mm] (-3,2) grid (1,-1); |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
343 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
344 |
\node (A) at (0,0) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^isub>0"}}; |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
345 |
\node (B) at (2,0) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^isub>1"}}; |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
346 |
\node (C) at (4,0.7) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^isub>1"}}; |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
347 |
\node (D) at (4,-0.7) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^isub>2"}}; |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
348 |
\node (E) at (6,-0.7) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^isub>2"}}; |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
349 |
\node (E1) at (6, 0.2) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^isub>3"}}; |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
350 |
\node (F) at (8,-0.7) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^isub>3"}}; |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
351 |
|
12
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
352 |
\node (X) at (0,-2) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^isub>4"}}; |
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
353 |
\node (Y) at (2,-2) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^isub>4"}}; |
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
354 |
\node (Z) at (2,-2.9) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^isub>5"}}; |
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
355 |
\node (U1) at (4,-2) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^isub>5"}}; |
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
356 |
\node (U2) at (4,-2.9) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^isub>6"}}; |
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
357 |
\node (R) at (6,-2.9) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^isub>6"}}; |
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
358 |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
359 |
\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
|
360 |
\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
|
361 |
\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
|
362 |
\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
|
363 |
\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
|
364 |
\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
|
365 |
|
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
366 |
\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
|
367 |
\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
|
368 |
\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
|
369 |
\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
|
370 |
\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
|
371 |
\end{tikzpicture} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
372 |
\end{center} |
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
373 |
\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
|
374 |
\end{figure} |
6
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 |
8
5ba3d79622da
added a paragraph about RAGS
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
7
diff
changeset
|
377 |
We will design our scheduler |
5ba3d79622da
added a paragraph about RAGS
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
7
diff
changeset
|
378 |
so that every thread can be in the possession of several resources, that is |
12
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
379 |
it has potentially several incoming holding edges in the RAG, but has at most one outgoing |
8
5ba3d79622da
added a paragraph about RAGS
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
7
diff
changeset
|
380 |
waiting edge. The reason is that when a thread asks for resource that is locked |
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
381 |
already, then the process is blocked and cannot ask for another resource. |
8
5ba3d79622da
added a paragraph about RAGS
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
7
diff
changeset
|
382 |
Clearly, also every resource can only have at most one outgoing holding edge---indicating |
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
383 |
that the resource is locked. In this way we can always start at a thread waiting for a |
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
384 |
resource and ``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
|
385 |
|
8
5ba3d79622da
added a paragraph about RAGS
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
7
diff
changeset
|
386 |
|
5ba3d79622da
added a paragraph about RAGS
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
7
diff
changeset
|
387 |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
388 |
The use of relations for representing RAGs allows us to conveniently define |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
389 |
the notion of the \emph{dependants} of a thread using the transitive closure |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
390 |
operation for relations. This gives |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
391 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
392 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
393 |
@{thm cs_dependents_def} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
394 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
395 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
396 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
397 |
This definition needs to account for all threads that wait for a thread to |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
398 |
release a resource. This means we need to include threads that transitively |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
399 |
wait for a resource being released (in the picture above this means the dependants |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
400 |
of @{text "th\<^isub>0"} are @{text "th\<^isub>1"} and @{text "th\<^isub>2"}, which wait for resource @{text "cs\<^isub>1"}, |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
401 |
but also @{text "th\<^isub>3"}, |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
402 |
which cannot make any progress unless @{text "th\<^isub>2"} makes progress, which |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
403 |
in turn needs to wait for @{text "th\<^isub>0"} to finish). If there is a circle of dependencies |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
404 |
in a RAG, then clearly |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
405 |
we have a deadlock. Therefore when a thread requests a resource, |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
406 |
we must ensure that the resulting RAG is not circular. In practice, the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
407 |
programmer has to ensure this. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
408 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
409 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
410 |
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
|
411 |
state @{text s}. It is defined as |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
412 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
413 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
414 |
@{thm cpreced_def2}\hfill\numbered{cpreced} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
415 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
416 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
417 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
418 |
where the dependants of @{text th} are given by the waiting queue function. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
419 |
While the precedence @{term prec} of a thread is determined statically |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
420 |
(for example when the thread is |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
421 |
created), the point of the current precedence is to let the scheduler increase this |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
422 |
precedence, if needed according to PIP. Therefore the current precedence of @{text th} is |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
423 |
given as the maximum of the precedence @{text th} has in state @{text s} \emph{and} all |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
424 |
threads that are dependants of @{text th}. Since the notion @{term "dependants"} is |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
425 |
defined as the transitive closure of all dependent threads, we deal correctly with the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
426 |
problem in the informal algorithm by Sha et al.~\cite{Sha90} where a priority of a thread is |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
427 |
lowered prematurely. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
428 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
429 |
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
|
430 |
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
|
431 |
we represent as a record consisting of two |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
432 |
functions: |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
433 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
434 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
435 |
@{text "\<lparr>wq_fun, cprec_fun\<rparr>"} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
436 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
437 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
438 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
439 |
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
|
440 |
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
|
441 |
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
|
442 |
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
|
443 |
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
|
444 |
such records. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
445 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
446 |
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
|
447 |
function is defined in \eqref{allunlocked}) and the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
448 |
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
|
449 |
\mbox{@{abbrev initial_cprec}}. Therefore |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
450 |
we have for the initial shedule state |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
451 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
452 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
453 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
454 |
@{thm (lhs) schs.simps(1)} @{text "\<equiv>"}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
455 |
\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
|
456 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
457 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
458 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
459 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
460 |
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
|
461 |
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
|
462 |
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
|
463 |
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
|
464 |
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
|
465 |
@{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
|
466 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
467 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
468 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
469 |
@{thm (lhs) schs.simps(2)} @{text "\<equiv>"}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
470 |
\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
|
471 |
\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
|
472 |
@{thm (lhs) schs.simps(3)} @{text "\<equiv>"}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
473 |
\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
|
474 |
\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
|
475 |
@{thm (lhs) schs.simps(4)} @{text "\<equiv>"}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
476 |
\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
|
477 |
\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
|
478 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
479 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
480 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
481 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
482 |
More interesting are the cases where a resource, say @{text cs}, is locked or released. In these cases |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
483 |
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
|
484 |
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
|
485 |
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
|
486 |
resource). This gives the clause |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
487 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
488 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
489 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
490 |
@{thm (lhs) schs.simps(5)} @{text "\<equiv>"}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
491 |
\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
|
492 |
\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
|
493 |
\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
|
494 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
495 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
496 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
497 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
498 |
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
|
499 |
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
|
500 |
list transformation, we use |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
501 |
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
|
502 |
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
|
503 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
504 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
505 |
\begin{tabular}{@ {}lcl} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
506 |
@{term "release []"} & @{text "\<equiv>"} & @{term "[]"}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
507 |
@{term "release (DUMMY # qs)"} & @{text "\<equiv>"} & @{term "qs"}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
508 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
509 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
510 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
511 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
512 |
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
|
513 |
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
|
514 |
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
|
515 |
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
|
516 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
517 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
518 |
\begin{tabular}{@ {}lcl} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
519 |
@{term "release []"} & @{text "\<equiv>"} & @{term "[]"}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
520 |
@{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
|
521 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
522 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
523 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
524 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
525 |
where @{text "SOME"} stands for Hilbert's epsilon and implements an arbitrary |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
526 |
choice for the next waiting list. It just has to be a list of distinctive threads and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
527 |
contain the same elements as @{text "qs"}. This gives for @{term V} the clause: |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
528 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
529 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
530 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
531 |
@{thm (lhs) schs.simps(6)} @{text "\<equiv>"}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
532 |
\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
|
533 |
\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
|
534 |
\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
|
535 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
536 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
537 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
538 |
Having the scheduler function @{term schs} at our disposal, we can ``lift'', or |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
539 |
overload, the notions |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
540 |
@{term waiting}, @{term holding}, @{term depend} and @{term cp} to operate on states only. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
541 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
542 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
543 |
\begin{tabular}{@ {}rcl} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
544 |
@{thm (lhs) s_holding_abv} & @{text "\<equiv>"} & @{thm (rhs) s_holding_abv}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
545 |
@{thm (lhs) s_waiting_abv} & @{text "\<equiv>"} & @{thm (rhs) s_waiting_abv}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
546 |
@{thm (lhs) s_depend_abv} & @{text "\<equiv>"} & @{thm (rhs) s_depend_abv}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
547 |
@{thm (lhs) cp_def} & @{text "\<equiv>"} & @{thm (rhs) cp_def} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
548 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
549 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
550 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
551 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
552 |
With these abbreviations in place we can introduce |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
553 |
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
|
554 |
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
|
555 |
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
|
556 |
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
|
557 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
558 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
559 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
560 |
@{thm readys_def}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
561 |
@{thm runing_def} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
562 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
563 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
564 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
565 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
566 |
In the second definition @{term "DUMMY ` DUMMY"} stands for the image of a set under a function. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
567 |
Note that in the initial state, that is where the list of events is empty, the set |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
568 |
@{term threads} is empty and therefore there is neither a thread ready nor running. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
569 |
If there is one or more threads ready, then there can only be \emph{one} thread |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
570 |
running, namely the one whose current precedence is equal to the maximum of all ready |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
571 |
threads. We use sets to capture both possibilities. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
572 |
We can now also conveniently define the set of resources that are locked by a thread in a |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
573 |
given state and also when a thread is detached in a state (meaning the thread neither |
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
574 |
holds nor waits for a resource---in the RAG this would correspond to an |
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
575 |
isolated node without any incoming and outgoing edges, see Figure~\ref{RAGgraph}): |
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
576 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
577 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
578 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
579 |
@{thm holdents_def}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
580 |
@{thm detached_def} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
581 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
582 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
583 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
584 |
%\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
585 |
%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
|
586 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
587 |
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
|
588 |
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
|
589 |
created yet. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
590 |
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
|
591 |
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
|
592 |
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
|
593 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
594 |
\begin{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
595 |
\begin{tabular}{c} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
596 |
@{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
|
597 |
@{thm[mode=Rule] thread_exit[where thread=th]} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
598 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
599 |
\end{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
600 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
601 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
602 |
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
|
603 |
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
|
604 |
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
|
605 |
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
|
606 |
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
|
607 |
if the corresponding thread is running. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
608 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
609 |
\begin{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
610 |
@{thm[mode=Rule] thread_set[where thread=th]} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
611 |
\end{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
612 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
613 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
614 |
If a thread wants to lock a resource, then the thread needs to be |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
615 |
running and also we have to make sure that the resource lock does |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
616 |
not lead to a cycle in the RAG. In practice, ensuring the latter |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
617 |
is the responsibility of the programmer. In our formal |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
618 |
model we brush aside these problematic cases in order to be able to make |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
619 |
some meaningful statements about PIP.\footnote{This situation is |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
620 |
similar to the infamous \emph{occurs check} in Prolog: In order to say |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
621 |
anything meaningful about unification, one needs to perform an occurs |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
622 |
check. But in practice the occurs check is omitted and the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
623 |
responsibility for avoiding problems rests with the programmer.} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
624 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
625 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
626 |
\begin{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
627 |
@{thm[mode=Rule] thread_P[where thread=th]} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
628 |
\end{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
629 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
630 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
631 |
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
|
632 |
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
|
633 |
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
|
634 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
635 |
\begin{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
636 |
@{thm[mode=Rule] thread_V[where thread=th]} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
637 |
\end{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
638 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
639 |
\noindent |
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
640 |
Note, however, that apart from the circularity condition, we do not make any |
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
641 |
assumption on how different resources can locked and released relative to each |
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
642 |
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
|
643 |
contrast to Sha et al \cite{Sha90} who require that critical sections are |
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
644 |
properly nested. |
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
645 |
|
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
646 |
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
|
647 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
648 |
\begin{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
649 |
\begin{tabular}{c} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
650 |
@{thm[mode=Axiom] vt_nil}\hspace{1cm} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
651 |
@{thm[mode=Rule] vt_cons} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
652 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
653 |
\end{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
654 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
655 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
656 |
This completes our formal model of PIP. In the next section we present |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
657 |
properties that show our model of PIP is correct. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
658 |
*} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
659 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
660 |
section {* The Correctness Proof *} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
661 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
662 |
(*<*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
663 |
context extend_highest_gen |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
664 |
begin |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
665 |
(*>*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
666 |
text {* |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
667 |
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
|
668 |
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
|
669 |
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
|
670 |
high priority can only be blocked a maximum of @{text n} times. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
671 |
Their second correctness criterion is given |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
672 |
in terms of the number of critical resources \cite[Theorem 6]{Sha90}: if there are |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
673 |
@{text m} critical resources, then a blocked job with high priority |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
674 |
can only be blocked a maximum of @{text m} times. Both results on their own, strictly speaking, do |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
675 |
\emph{not} prevent indefinite, or unbounded, Priority Inversion, |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
676 |
because if a low-priority thread does not give up its critical |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
677 |
resource (the one the high-priority thread is waiting for), then the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
678 |
high-priority thread can never run. The argument of Sha et al.~is |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
679 |
that \emph{if} threads release locked resources in a finite amount |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
680 |
of time, then indefinite Priority Inversion cannot occur---the high-priority |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
681 |
thread is guaranteed to run eventually. The assumption is that |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
682 |
programmers must ensure that threads are programmed in this way. However, even |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
683 |
taking this assumption into account, the correctness properties of |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
684 |
Sha et al.~are |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
685 |
\emph{not} true for their version of PIP---despite being ``proved''. As Yodaiken |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
686 |
\cite{Yodaiken02} pointed out: If a low-priority thread possesses |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
687 |
locks to two resources for which two high-priority threads are |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
688 |
waiting for, then lowering the priority prematurely after giving up |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
689 |
only one lock, can cause indefinite Priority Inversion for one of the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
690 |
high-priority threads, invalidating their two bounds. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
691 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
692 |
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
|
693 |
us, because of the way we have set up our formal model of PIP. One |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
694 |
reason is that we allow critical sections, which start with a @{text P}-event |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
695 |
and finish with a corresponding @{text V}-event, to arbitrarily overlap |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
696 |
(something Sha et al.~explicitly exclude). Therefore we have |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
697 |
designed a different correctness criterion for PIP. The idea behind |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
698 |
our criterion is as follows: for all states @{text s}, we know the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
699 |
corresponding thread @{text th} with the highest precedence; we show |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
700 |
that in every future state (denoted by @{text "s' @ s"}) in which |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
701 |
@{text th} is still alive, either @{text th} is running or it is |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
702 |
blocked by a thread that was alive in the state @{text s} and was waiting |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
703 |
for or in the possession of a lock in @{text s}. Since in @{text s}, as in |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
704 |
every state, the set of alive threads is finite, @{text th} can only |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
705 |
be blocked a finite number of times. This is independent of how many |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
706 |
threads of lower priority are created in @{text "s'"}. We will actually prove a |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
707 |
stronger statement where we also provide the current precedence of |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
708 |
the blocking thread. However, this correctness criterion hinges upon |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
709 |
a number of assumptions about the states @{text s} and @{text "s' @ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
710 |
s"}, the thread @{text th} and the events happening in @{text |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
711 |
s'}. We list them next: |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
712 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
713 |
\begin{quote} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
714 |
{\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
|
715 |
{\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
|
716 |
@{text "s' @ s"} are valid states: |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
717 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
718 |
\begin{tabular}{l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
719 |
@{term "vt s"}, @{term "vt (s' @ s)"} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
720 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
721 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
722 |
\end{quote} |
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 |
\begin{quote} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
725 |
{\bf Assumptions on the thread {\boldmath@{text "th"}:}} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
726 |
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
|
727 |
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
|
728 |
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
|
729 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
730 |
\begin{tabular}{l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
731 |
@{term "th \<in> threads s"}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
732 |
@{term "prec th s = Max (cprec s ` threads s)"}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
733 |
@{term "prec th s = (prio, DUMMY)"} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
734 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
735 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
736 |
\end{quote} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
737 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
738 |
\begin{quote} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
739 |
{\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
|
740 |
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
|
741 |
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
|
742 |
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
|
743 |
priority than @{text prio} of @{text th}. We also need to assume that the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
744 |
priority of @{text "th"} does not get reset and also that @{text th} does |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
745 |
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
|
746 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
747 |
\begin{tabular}{l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
748 |
{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
|
749 |
{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
|
750 |
{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
|
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 |
\end{quote} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
754 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
755 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
756 |
The locale mechanism of Isabelle helps us to manage conveniently such assumptions~\cite{Haftmann08}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
757 |
Under these assumptions we shall prove the following correctness property: |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
758 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
759 |
\begin{theorem}\label{mainthm} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
760 |
Given the assumptions about states @{text "s"} and @{text "s' @ s"}, |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
761 |
the thread @{text th} and the events in @{text "s'"}, |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
762 |
if @{term "th' \<in> running (s' @ s)"} and @{text "th' \<noteq> th"} then |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
763 |
@{text "th' \<in> threads s"}, @{text "\<not> detached s th'"} and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
764 |
@{term "cp (s' @ s) th' = prec th s"}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
765 |
\end{theorem} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
766 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
767 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
768 |
This theorem ensures that the thread @{text th}, which has the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
769 |
highest precedence in the state @{text s}, can only be blocked in |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
770 |
the state @{text "s' @ s"} by a thread @{text th'} that already |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
771 |
existed in @{text s} and requested or had a lock on at least |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
772 |
one resource---that means the thread was not \emph{detached} in @{text s}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
773 |
As we shall see shortly, that means there are only finitely |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
774 |
many threads that can block @{text th} in this way and then they |
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
775 |
need to run with the same precedence as @{text th}. |
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
776 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
777 |
Like in the argument by Sha et al.~our |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
778 |
finite bound does not guarantee absence of indefinite Priority |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
779 |
Inversion. For this we further have to assume that every thread |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
780 |
gives up its resources after a finite amount of time. We found that |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
781 |
this assumption is awkward to formalise in our model. Therefore we |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
782 |
leave it out and let the programmer assume the responsibility to |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
783 |
program threads in such a benign manner (in addition to causing no |
12
85116bc854c0
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
11
diff
changeset
|
784 |
circularity in the RAG). In this detail, we do not |
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
785 |
make any progress in comparison with the work by Sha et al. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
786 |
However, we are able to combine their two separate bounds into a |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
787 |
single theorem improving their bound. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
788 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
789 |
In what follows we will describe properties of PIP that allow us to prove |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
790 |
Theorem~\ref{mainthm} and, when instructive, briefly describe our argument. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
791 |
It is relatively easy to see that |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
792 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
793 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
794 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
795 |
@{text "running s \<subseteq> ready s \<subseteq> threads s"}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
796 |
@{thm[mode=IfThen] finite_threads} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
797 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
798 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
799 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
800 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
801 |
The second property is by induction of @{term vt}. The next three |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
802 |
properties are |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
803 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
804 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
805 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
806 |
@{thm[mode=IfThen] waiting_unique[of _ _ "cs\<^isub>1" "cs\<^isub>2"]}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
807 |
@{thm[mode=IfThen] held_unique[of _ "th\<^isub>1" _ "th\<^isub>2"]}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
808 |
@{thm[mode=IfThen] runing_unique[of _ "th\<^isub>1" "th\<^isub>2"]} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
809 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
810 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
811 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
812 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
813 |
The first property states that every waiting thread can only wait for a single |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
814 |
resource (because it gets suspended after requesting that resource); the second |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
815 |
that every resource can only be held by a single thread; |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
816 |
the third property establishes that in every given valid state, there is |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
817 |
at most one running thread. We can also show the following properties |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
818 |
about the @{term RAG} in @{text "s"}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
819 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
820 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
821 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
822 |
@{text If}~@{thm (prem 1) acyclic_depend}~@{text "then"}:\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
823 |
\hspace{5mm}@{thm (concl) acyclic_depend}, |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
824 |
@{thm (concl) finite_depend} and |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
825 |
@{thm (concl) wf_dep_converse},\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
826 |
\hspace{5mm}@{text "if"}~@{thm (prem 2) dm_depend_threads}~@{text "then"}~@{thm (concl) dm_depend_threads} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
827 |
and\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
828 |
\hspace{5mm}@{text "if"}~@{thm (prem 2) range_in}~@{text "then"}~@{thm (concl) range_in}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
829 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
830 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
831 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
832 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
833 |
The acyclicity property follows from how we restricted the events in |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
834 |
@{text step}; similarly the finiteness and well-foundedness property. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
835 |
The last two properties establish that every thread in a @{text "RAG"} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
836 |
(either holding or waiting for a resource) is a live thread. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
837 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
838 |
The key lemma in our proof of Theorem~\ref{mainthm} is as follows: |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
839 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
840 |
\begin{lemma}\label{mainlem} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
841 |
Given the assumptions about states @{text "s"} and @{text "s' @ s"}, |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
842 |
the thread @{text th} and the events in @{text "s'"}, |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
843 |
if @{term "th' \<in> threads (s' @ s)"}, @{text "th' \<noteq> th"} and @{text "detached (s' @ s) th'"}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
844 |
then @{text "th' \<notin> running (s' @ s)"}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
845 |
\end{lemma} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
846 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
847 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
848 |
The point of this lemma is that a thread different from @{text th} (which has the highest |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
849 |
precedence in @{text s}) and not holding any resource, cannot be running |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
850 |
in the state @{text "s' @ s"}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
851 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
852 |
\begin{proof} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
853 |
Since thread @{text "th'"} does not hold any resource, no thread can depend on it. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
854 |
Therefore its current precedence @{term "cp (s' @ s) th'"} equals its own precedence |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
855 |
@{term "prec th' (s' @ s)"}. Since @{text "th"} has the highest precedence in the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
856 |
state @{text "(s' @ s)"} and precedences are distinct among threads, we have |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
857 |
@{term "prec th' (s' @s ) < prec th (s' @ s)"}. From this |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
858 |
we have @{term "cp (s' @ s) th' < prec th (s' @ s)"}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
859 |
Since @{text "prec th (s' @ s)"} is already the highest |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
860 |
@{term "cp (s' @ s) th"} can not be higher than this and can not be lower either (by |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
861 |
definition of @{term "cp"}). Consequently, we have @{term "prec th (s' @ s) = cp (s' @ s) th"}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
862 |
Finally we have @{term "cp (s' @ s) th' < cp (s' @ s) th"}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
863 |
By defintion of @{text "running"}, @{text "th'"} can not be running in state |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
864 |
@{text "s' @ s"}, as we had to show.\qed |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
865 |
\end{proof} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
866 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
867 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
868 |
Since @{text "th'"} is not able to run in state @{text "s' @ s"}, it is not able to |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
869 |
issue a @{text "P"} or @{text "V"} event. Therefore if @{text "s' @ s"} is extended |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
870 |
one step further, @{text "th'"} still cannot hold any resource. The situation will |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
871 |
not change in further extensions as long as @{text "th"} holds the highest precedence. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
872 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
873 |
From this lemma we can deduce Theorem~\ref{mainthm}: that @{text th} can only be |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
874 |
blocked by a thread @{text th'} that |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
875 |
held some resource in state @{text s} (that is not @{text "detached"}). And furthermore |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
876 |
that the current precedence of @{text th'} in state @{text "(s' @ s)"} must be equal to the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
877 |
precedence of @{text th} in @{text "s"}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
878 |
We show this theorem by induction on @{text "s'"} using Lemma~\ref{mainlem}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
879 |
This theorem gives a stricter bound on the threads that can block @{text th} than the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
880 |
one obtained by Sha et al.~\cite{Sha90}: |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
881 |
only threads that were alive in state @{text s} and moreover held a resource. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
882 |
This means our bound is in terms of both---alive threads in state @{text s} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
883 |
and number of critical resources. Finally, the theorem establishes that the blocking threads have the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
884 |
current precedence raised to the precedence of @{text th}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
885 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
886 |
We can furthermore prove that under our assumptions no deadlock exists in the state @{text "s' @ s"} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
887 |
by showing that @{text "running (s' @ s)"} is not empty. |
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 |
\begin{lemma} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
890 |
Given the assumptions about states @{text "s"} and @{text "s' @ s"}, |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
891 |
the thread @{text th} and the events in @{text "s'"}, |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
892 |
@{term "running (s' @ s) \<noteq> {}"}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
893 |
\end{lemma} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
894 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
895 |
\begin{proof} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
896 |
If @{text th} is blocked, then by following its dependants graph, we can always |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
897 |
reach a ready thread @{text th'}, and that thread must have inherited the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
898 |
precedence of @{text th}.\qed |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
899 |
\end{proof} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
900 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
901 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
902 |
%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
|
903 |
%\begin{enumerate} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
904 |
%\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
|
905 |
% @ {thm [display] chain_building[rule_format]} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
906 |
%\item The ready thread chased to is unique (@{text "dchain_unique"}): |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
907 |
% @ {thm [display] dchain_unique[of _ _ "th\<^isub>1" "th\<^isub>2"]} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
908 |
%\end{enumerate} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
909 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
910 |
%Some deeper results about the system: |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
911 |
%\begin{enumerate} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
912 |
%\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
|
913 |
%@ {thm [display] max_cp_eq} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
914 |
%\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
|
915 |
%(@{text "max_cp_readys_threads"}): |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
916 |
%@ {thm [display] max_cp_readys_threads} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
917 |
%\end{enumerate} |
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 |
%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
|
920 |
%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
|
921 |
%\begin{enumerate} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
922 |
%\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
|
923 |
% one thread holds (@{text "cntCS_v_dec"}) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
924 |
% @ {thm [display] cntCS_v_dec} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
925 |
%\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
|
926 |
% (@ {text "cnp_cnv_cncs"}): |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
927 |
% @ {thm [display] cnp_cnv_cncs} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
928 |
%\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
|
929 |
% the relevant thread is not living: |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
930 |
% (@{text "cnp_cnv_eq"}): |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
931 |
% @ {thm [display] cnp_cnv_eq} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
932 |
%\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
|
933 |
% (@{text "not_thread_holdents"}): |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
934 |
% @ {thm [display] not_thread_holdents} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
935 |
%\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
|
936 |
% thread does not hold any critical resource, therefore no thread can depend on it |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
937 |
% (@{text "count_eq_dependents"}): |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
938 |
% @ {thm [display] count_eq_dependents} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
939 |
%\end{enumerate} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
940 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
941 |
%The reason that only threads which already held some resoures |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
942 |
%can be runing and block @{text "th"} is that if , otherwise, one thread |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
943 |
%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
|
944 |
%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
|
945 |
%lemma @{text "moment_blocked"}: |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
946 |
%@ {thm [display] moment_blocked} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
947 |
%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
|
948 |
%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
|
949 |
%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
|
950 |
%some resource at state @{text "s"}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
951 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
952 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
953 |
%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
|
954 |
%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
|
955 |
%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
|
956 |
%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
|
957 |
%priority inversion can be avoided. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
958 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
959 |
%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
|
960 |
%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
|
961 |
%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
|
962 |
%@ {thm [display] red_moment} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
963 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
964 |
%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
|
965 |
%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
|
966 |
%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
|
967 |
%about @{term "t"} (@{text "ind"}): |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
968 |
%\begin{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
969 |
%@ {thm[display] ind} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
970 |
%\end{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
971 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
972 |
%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
|
973 |
%\begin{enumerate} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
974 |
%\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
|
975 |
% precedence is preserved as well |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
976 |
% (@{text "th_kept"}): |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
977 |
% @ {thm [display] th_kept} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
978 |
%\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
|
979 |
% all living threads |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
980 |
% (@{text "max_preced"}): |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
981 |
% @ {thm [display] max_preced} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
982 |
%\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
|
983 |
% among all living threads |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
984 |
% (@{text "th_cp_max_preced"}): |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
985 |
% @ {thm [display] th_cp_max_preced} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
986 |
%\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
|
987 |
% precedence among all living threads |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
988 |
% (@{text "th_cp_max"}): |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
989 |
% @ {thm [display] th_cp_max} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
990 |
%\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
|
991 |
% @{term "s"} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
992 |
% (@{text "th_cp_preced"}): |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
993 |
% @ {thm [display] th_cp_preced} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
994 |
%\end{enumerate} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
995 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
996 |
%The main theorem of this part is to characterizing the running thread during @{term "t"} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
997 |
%(@{text "runing_inversion_2"}): |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
998 |
%@ {thm [display] runing_inversion_2} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
999 |
%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
|
1000 |
%already live and held some resource |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1001 |
%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
|
1002 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1003 |
%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
|
1004 |
%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
|
1005 |
%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
|
1006 |
%then. |
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 |
(*<*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1009 |
end |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1010 |
(*>*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1011 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1012 |
section {* Properties for an Implementation\label{implement} *} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1013 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1014 |
text {* |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1015 |
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
|
1016 |
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
|
1017 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1018 |
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
|
1019 |
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
|
1020 |
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
|
1021 |
depends on all its dependants---a ``global'' transitive notion, |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1022 |
which is indeed heavy weight (see Def.~shown in \eqref{cpreced}). |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1023 |
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
|
1024 |
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
|
1025 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1026 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1027 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1028 |
@{thm children_def2} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1029 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1030 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1031 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1032 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1033 |
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
|
1034 |
@{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
|
1035 |
a resource). We can prove the following lemma. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1036 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1037 |
\begin{lemma}\label{childrenlem} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1038 |
@{text "If"} @{thm (prem 1) cp_rec} @{text "then"} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1039 |
\begin{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1040 |
@{thm (concl) cp_rec}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1041 |
\end{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1042 |
\end{lemma} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1043 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1044 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1045 |
That means the current precedence of a thread @{text th} can be |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1046 |
computed locally by considering only the children of @{text th}. In |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1047 |
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
|
1048 |
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
|
1049 |
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
|
1050 |
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
|
1051 |
a standard scheduling operation implemented in most operating |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1052 |
systems. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1053 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1054 |
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
|
1055 |
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
|
1056 |
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
|
1057 |
kind of events.\smallskip |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1058 |
*} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1059 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1060 |
(*<*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1061 |
context step_create_cps |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1062 |
begin |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1063 |
(*>*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1064 |
text {* |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1065 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1066 |
\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
|
1067 |
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
|
1068 |
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
|
1069 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1070 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1071 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1072 |
@{thm eq_dep},\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1073 |
@{thm eq_cp_th}, and\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1074 |
@{thm[mode=IfThen] eq_cp} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1075 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1076 |
\end{isabelle} |
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 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1079 |
This means in an implementation we do not have recalculate the @{text RAG} and also none of the |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1080 |
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
|
1081 |
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
|
1082 |
\smallskip |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1083 |
*} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1084 |
(*<*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1085 |
end |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1086 |
context step_exit_cps |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1087 |
begin |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1088 |
(*>*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1089 |
text {* |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1090 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1091 |
\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
|
1092 |
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
|
1093 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1094 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1095 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1096 |
@{thm eq_dep}, and\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1097 |
@{thm[mode=IfThen] eq_cp} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1098 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1099 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1100 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1101 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1102 |
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
|
1103 |
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
|
1104 |
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
|
1105 |
current precedence. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1106 |
\smallskip |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1107 |
*} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1108 |
(*<*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1109 |
end |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1110 |
context step_set_cps |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1111 |
begin |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1112 |
(*>*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1113 |
text {* |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1114 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1115 |
\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
|
1116 |
@{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
|
1117 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1118 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1119 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1120 |
@{thm[mode=IfThen] eq_dep}, and\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1121 |
@{thm[mode=IfThen] eq_cp_pre} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1122 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1123 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1124 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1125 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1126 |
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
|
1127 |
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
|
1128 |
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
|
1129 |
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
|
1130 |
change of its priority cannot affect other threads. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1131 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1132 |
%The second |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1133 |
%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
|
1134 |
%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
|
1135 |
%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
|
1136 |
%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
|
1137 |
%the @{term "cp"} of every |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1138 |
%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
|
1139 |
%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
|
1140 |
%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
|
1141 |
%dependants. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1142 |
% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1143 |
%\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1144 |
%\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1145 |
%@{thm[mode=IfThen] eq_up_self}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1146 |
%@{text "If"} @{thm (prem 1) eq_up}, @{thm (prem 2) eq_up} and @{thm (prem 3) eq_up}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1147 |
%@{text "then"} @{thm (concl) eq_up}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1148 |
%\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1149 |
%\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1150 |
% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1151 |
%\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1152 |
%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
|
1153 |
%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
|
1154 |
%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
|
1155 |
%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
|
1156 |
%have their current precedence changed. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1157 |
\smallskip |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1158 |
*} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1159 |
(*<*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1160 |
end |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1161 |
context step_v_cps_nt |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1162 |
begin |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1163 |
(*>*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1164 |
text {* |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1165 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1166 |
\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
|
1167 |
@{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
|
1168 |
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
|
1169 |
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
|
1170 |
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
|
1171 |
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
|
1172 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1173 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1174 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1175 |
@{thm depend_s} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1176 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1177 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1178 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1179 |
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
|
1180 |
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
|
1181 |
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
|
1182 |
can show |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1183 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1184 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1185 |
@{thm[mode=IfThen] cp_kept} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1186 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1187 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1188 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1189 |
For @{text th} and @{text th'} we need to use Lemma~\ref{childrenlem} to |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1190 |
recalculate their current precedence since their children have changed. *}(*<*)end context step_v_cps_nnt begin (*>*)text {* |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1191 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1192 |
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
|
1193 |
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
|
1194 |
to be recalculated. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1195 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1196 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1197 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1198 |
@{thm depend_s}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1199 |
@{thm eq_cp} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1200 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1201 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1202 |
*} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1203 |
(*<*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1204 |
end |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1205 |
context step_P_cps_e |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1206 |
begin |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1207 |
(*>*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1208 |
text {* |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1209 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1210 |
\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
|
1211 |
@{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
|
1212 |
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
|
1213 |
first by showing that |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1214 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1215 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1216 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1217 |
@{thm depend_s}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1218 |
@{thm eq_cp} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1219 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1220 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1221 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1222 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1223 |
This means we need to add a holding edge to the @{text RAG} and no |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1224 |
current precedence needs to be recalculated.*}(*<*)end context step_P_cps_ne begin(*>*) text {* |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1225 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1226 |
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
|
1227 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1228 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1229 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1230 |
@{thm depend_s}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1231 |
@{thm[mode=IfThen] eq_cp} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1232 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1233 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1234 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1235 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1236 |
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
|
1237 |
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
|
1238 |
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
|
1239 |
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
|
1240 |
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
|
1241 |
the @{term "cp"} of every |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1242 |
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
|
1243 |
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
|
1244 |
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
|
1245 |
dependants. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1246 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1247 |
\begin{isabelle}\ \ \ \ \ %%% |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1248 |
\begin{tabular}{@ {}l} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1249 |
%%@ {t hm[mode=IfThen] eq_up_self}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1250 |
@{text "If"} @{thm (prem 1) eq_up}, @{thm (prem 2) eq_up} and @{thm (prem 3) eq_up}\\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1251 |
@{text "then"} @{thm (concl) eq_up}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1252 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1253 |
\end{isabelle} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1254 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1255 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1256 |
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
|
1257 |
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
|
1258 |
have their current precedence changed. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1259 |
*} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1260 |
(*<*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1261 |
end |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1262 |
(*>*) |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1263 |
text {* |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1264 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1265 |
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
|
1266 |
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
|
1267 |
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
|
1268 |
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
|
1269 |
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
|
1270 |
our version of PIP on top of PINTOS, a small operating system written in C and used for teaching at |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1271 |
Stanford University \cite{PINTOS}. To implement PIP, we only need to modify the kernel |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1272 |
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
|
1273 |
function interface in PINTOS: |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1274 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1275 |
\begin{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1276 |
\begin{tabular}{|l@ {\hspace{2mm}}|l@ {\hspace{2mm}}|} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1277 |
\hline |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1278 |
{\bf Event} & {\bf PINTOS function} \\ |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1279 |
\hline |
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
1280 |
@{text Create} & @{ML_text "thread_create"}\\ |
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
1281 |
@{text Exit} & @{ML_text "thread_exit"}\\ |
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
1282 |
@{text Set} & @{ML_text "thread_set_priority"}\\ |
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
1283 |
@{text P} & @{ML_text "lock_acquire"}\\ |
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
1284 |
@{text V} & @{ML_text "lock_release"}\\ |
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1285 |
\hline |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1286 |
\end{tabular} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1287 |
\end{center} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1288 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1289 |
\noindent |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1290 |
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
|
1291 |
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
|
1292 |
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
|
1293 |
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
|
1294 |
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
|
1295 |
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
|
1296 |
@{ML_text "queue_insert"}, for inserting a new element into a priority queue, and |
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
1297 |
@{ML_text "queue_update"}, for updating the position of an element that is already |
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
1298 |
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
|
1299 |
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
|
1300 |
|
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
1301 |
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
|
1302 |
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
|
1303 |
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
|
1304 |
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
|
1305 |
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
|
1306 |
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
|
1307 |
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
|
1308 |
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
|
1309 |
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
|
1310 |
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
|
1311 |
there is a check that |
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
1312 |
the lock is a ``valid'' lock |
15
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
1313 |
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
|
1314 |
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
|
1315 |
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
|
1316 |
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
|
1317 |
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
|
1318 |
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
|
1319 |
a ``kernel panic''. |
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
1320 |
|
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
1321 |
|
7
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1322 |
|
15
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
1323 |
\begin{figure}[t] |
7
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1324 |
\begin{lstlisting} |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1325 |
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
|
1326 |
{ ASSERT (lock != NULL); |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1327 |
ASSERT (!intr_context()); |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1328 |
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
|
1329 |
|
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1330 |
enum intr_level old_level; |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1331 |
old_level = intr_disable(); |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1332 |
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
|
1333 |
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
|
1334 |
thread_current()->waiting = lock; |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1335 |
struct thread *pt; |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1336 |
pt = lock->holder; |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1337 |
while (pt) { |
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
1338 |
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
|
1339 |
if (!(update_cprec(pt))) |
7
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1340 |
break; |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1341 |
lock = pt->waiting; |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1342 |
if (!lock) { |
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
1343 |
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
|
1344 |
break; |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1345 |
}; |
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
1346 |
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
|
1347 |
pt = lock->holder; |
7
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1348 |
}; |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1349 |
thread_block(); |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1350 |
} else { |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1351 |
lock->value--; |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1352 |
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
|
1353 |
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
|
1354 |
}; |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1355 |
intr_set_level(old_level); |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1356 |
} |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1357 |
\end{lstlisting} |
17
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
1358 |
\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
|
1359 |
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
|
1360 |
\end{figure} |
0514be2ad83e
started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
6
diff
changeset
|
1361 |
|
11
8e02fb168350
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
8
diff
changeset
|
1362 |
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
1363 |
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
|
1364 |
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
|
1365 |
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
|
1366 |
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
|
1367 |
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
|
1368 |
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
|
1369 |
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
|
1370 |
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
|
1371 |
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
|
1372 |
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
|
1373 |
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
|
1374 |
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
|
1375 |
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
|
1376 |
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
|
1377 |
current precedence. |
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
1378 |
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
1379 |
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
|
1380 |
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
|
1381 |
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
|
1382 |
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
|
1383 |
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
|
1384 |
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
|
1385 |
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
|
1386 |
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
|
1387 |
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
|
1388 |
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
|
1389 |
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
|
1390 |
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
|
1391 |
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
|
1392 |
(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
|
1393 |
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
|
1394 |
|
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
1395 |
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
|
1396 |
(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
|
1397 |
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
|
1398 |
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
|
1399 |
thread already possesses (Line 29). |
15
9e664c268e25
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
14
diff
changeset
|
1400 |
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
|
1401 |
|
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
1402 |
|
13
735e36c64a71
added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
12
diff
changeset
|
1403 |
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
|
1404 |
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
|
1405 |
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
|
1406 |
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
|
1407 |
anything about PIP \cite{sel4}. |
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
1408 |
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
|
1409 |
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
|
1410 |
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
|
1411 |
correctness of the algorithm. |
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1412 |
*} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1413 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1414 |
section {* Conclusion *} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1415 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1416 |
text {* |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1417 |
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
|
1418 |
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
|
1419 |
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
|
1420 |
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
|
1421 |
have circular lock dependencies. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1422 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1423 |
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
|
1424 |
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
|
1425 |
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
|
1426 |
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
|
1427 |
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
|
1428 |
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
|
1429 |
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
|
1430 |
We can also confirm that Paulson's |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1431 |
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
|
1432 |
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
|
1433 |
area of this method is security protocols. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1434 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1435 |
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
|
1436 |
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
|
1437 |
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
|
1438 |
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
|
1439 |
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
|
1440 |
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
|
1441 |
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
|
1442 |
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
|
1443 |
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
|
1444 |
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
|
1445 |
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
|
1446 |
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
|
1447 |
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
|
1448 |
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
|
1449 |
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
|
1450 |
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
|
1451 |
the proof of Sha et al.: they require that critical sections nest properly, |
14
1bf194825a4e
more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
13
diff
changeset
|
1452 |
whereas our scheduler allows critical sections to overlap. |
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1453 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1454 |
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
|
1455 |
now living in a multi-processor world. Priority Inversion certainly |
17
105715a0a807
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
16
diff
changeset
|
1456 |
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
|
1457 |
However, there is very little ``foundational'' |
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1458 |
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
|
1459 |
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
|
1460 |
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
|
1461 |
``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
|
1462 |
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
|
1463 |
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
|
1464 |
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
|
1465 |
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
|
1466 |
for future work. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1467 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1468 |
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
|
1469 |
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
|
1470 |
\cite{dutertre99b}---another solution to the Priority Inversion |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1471 |
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
|
1472 |
avoid it. There have been earlier formal investigations |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1473 |
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
|
1474 |
checking techniques. The results obtained by them apply, |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1475 |
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
|
1476 |
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
|
1477 |
size. Moreover, our result is a good |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1478 |
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
|
1479 |
reasoning: gaining deeper understanding of the subject matter. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1480 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1481 |
Our formalisation |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1482 |
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
|
1483 |
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
|
1484 |
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
|
1485 |
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
|
1486 |
for an implementation require 2000 lines. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1487 |
%The code of our formalisation |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1488 |
%can be downloaded from |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1489 |
%\url{http://www.inf.kcl.ac.uk/staff/urbanc/pip.html}. |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1490 |
|
16
9764023f719e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
15
diff
changeset
|
1491 |
%\medskip |
6
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1492 |
|
16
9764023f719e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
15
diff
changeset
|
1493 |
%\noindent |
9764023f719e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
15
diff
changeset
|
1494 |
%{\bf Acknowledgements:} |
9764023f719e
added
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
15
diff
changeset
|
1495 |
%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
|
1496 |
%referees. |
6
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 |
\bibliographystyle{plain} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1499 |
\bibliography{root} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1500 |
*} |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1501 |
|
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1502 |
|
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 |
end |
7f2493296c39
updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff
changeset
|
1505 |
(*>*) |