Journal/Paper.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Tue, 04 Mar 2014 16:38:38 +0000
changeset 28 7fa738a9615a
parent 27 6b1141c5e24c
child 32 e861aff29655
permissions -rw-r--r--
updated
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
     7
declare [[show_question_marks = false]]
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
     8
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     9
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    10
notation (latex output)
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    11
  Cons ("_::_" [78,77] 73) and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    12
  vt ("valid'_state") and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    13
  runing ("running") and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    14
  birthtime ("last'_set") and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    15
  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
    16
  Prc ("'(_, _')") and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    17
  holding ("holds") and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    18
  waiting ("waits") and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
  Th ("T") and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    20
  Cs ("C") and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
  readys ("ready") and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
  depend ("RAG") and 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    23
  preced ("prec") and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
  cpreced ("cprec") and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    25
  dependents ("dependants") and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    26
  cp ("cprec") and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
  holdents ("resources") and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    28
  original_priority ("priority") and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    29
  DUMMY  ("\<^raw:\mbox{$\_\!\_$}>")
28
7fa738a9615a updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 27
diff changeset
    30
 
25
a9c0eeb00cc3 added two more references
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
    31
  
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    32
(*>*)
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    33
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    34
section {* Introduction *}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    35
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    36
text {*
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    37
  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
    38
  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
    39
  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
    40
  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
    41
  finish their work within deadlines.  Unfortunately, both features
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    42
  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
    43
  \emph{Priority Inversion}. Suppose three threads having priorities
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    44
  $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
    45
  $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
    46
  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
    47
  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
    48
  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
    49
  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
    50
  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
    51
  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
    52
  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
    53
  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
    54
  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
    55
  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
    56
  Mesa programming language designed for concurrent programming.
22
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
    57
 
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    58
  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
    59
  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
    60
  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
    61
  controlled the Mars Pathfinder mission in 1997 \cite{Reeves98}.
25
a9c0eeb00cc3 added two more references
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
    62
  On Earth the software run mostly without any problem, but
a9c0eeb00cc3 added two more references
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
    63
  once the spacecraft landed on Mars, it shut down at irregular
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    64
  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
    65
  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
    66
  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
    67
  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
    68
  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
    69
  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
    70
  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
    71
  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
    72
  \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
    73
  Inheritance Protocol} \cite{Sha90} and others sometimes also call it
11
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
    74
  \emph{Priority Boosting}, \emph{Priority Donation} or \emph{Priority Lending}.} 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
    75
  in the scheduling software.
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    76
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    77
  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
    78
  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
    79
  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
    80
  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
    81
  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
    82
  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
    83
  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
    84
  in the Mars Pathfinder mission, in Boeing's 787 Dreamliner, Honda's
22
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
    85
  ASIMO robot, etc.) and ThreadX (another proprietary real-time OS
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
    86
  used in HP inkjet printers \cite{ThreadX}), but also
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
    87
  the POSIX 1003.1c Standard realised for
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    88
  example in libraries for FreeBSD, Solaris and Linux. 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    89
22
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
    90
  Two advantages of PIP are that it is deterministic and that increasing the priority of a thread
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
    91
  can be performed dynamically by the scheduler.
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    92
  This is in contrast to \emph{Priority Ceiling}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    93
  \cite{Sha90}, another solution to the Priority Inversion problem,
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    94
  which requires static analysis of the program in order to prevent
22
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
    95
  Priority Inversion, and also in contrast to the Windows NT scheduler, which avoids
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
    96
  this problem by randomly boosting the priority of ready low-priority threads
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
    97
  (see for instance~\cite{WINDOWSNT}).
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    98
  However, there has also been strong criticism against
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    99
  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
   100
  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
   101
  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
   102
  Though, most criticism against PIP centres around unreliable
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   103
  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
   104
  For example, Yodaiken writes in \cite{Yodaiken02}:
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   105
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   106
  \begin{quote}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   107
  \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
   108
  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
   109
  \end{quote}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   110
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   111
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   112
  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
   113
  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
   114
  not always be achievable in practice.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   115
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   116
  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
   117
  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
   118
  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
   119
  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
   120
  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
   121
  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
   122
  Kernel mailing list:
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   123
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   124
  \begin{quote}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   125
  \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
   126
  implementation is a nightmare: extremely heavy weight, involving
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   127
  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
   128
  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
   129
  wait operations.''
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   130
  \end{quote}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   131
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   132
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   133
  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
   134
  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
   135
  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
   136
  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
   137
  presentation of PIP~\cite{Sha90}, despite being informally
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   138
  ``proved'' correct, is actually \emph{flawed}. 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   139
25
a9c0eeb00cc3 added two more references
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   140
  Yodaiken \cite{Yodaiken02} and also Moylan et al.~\cite{deinheritance} 
a9c0eeb00cc3 added two more references
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   141
  point to a subtlety that had been
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   142
  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
   143
  \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
   144
  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
   145
  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
   146
  implementation of PIP is ``rather straightforward''~\cite{Sha90}.
25
a9c0eeb00cc3 added two more references
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   147
  Unfortunately, as Yodaiken and Moylan et al.~point out, this behaviour is too
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   148
  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
   149
  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
   150
  $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
   151
  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
   152
  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
   153
  behaviour for $L$ is to switch to the highest remaining priority of
25
a9c0eeb00cc3 added two more references
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   154
  the threads that it blocks. A similar error is made in the textbook
a9c0eeb00cc3 added two more references
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   155
  \cite[Section 2.3.1]{book} which specifies for a process that 
a9c0eeb00cc3 added two more references
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   156
  inherited a higher priority and exits a critical section ``it resumes 
a9c0eeb00cc3 added two more references
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   157
  the priority it had at the point of entry into the critical section''.
27
6b1141c5e24c cleaned up
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
   158
   
25
a9c0eeb00cc3 added two more references
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   159
  While \cite{book} and
26
da7a6ccfa7a9 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 25
diff changeset
   160
  \cite{Sha90} are the only formal publications we have 
25
a9c0eeb00cc3 added two more references
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   161
  found that describe the incorrect behaviour, not all, but many
22
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   162
  informal\footnote{informal as in ``found on the Web''} 
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   163
  descriptions of PIP overlook the possibility that another
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   164
  high-priority might wait for a low-priority process to finish.
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   165
  The advantage of formalising the
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   166
  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
   167
  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
   168
  informal reasoning (since we have to analyse all possible behaviours
22
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   169
  of threads, i.e.~\emph{traces}, that could possibly happen).\medskip\smallskip
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   170
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   171
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   172
  {\bf Contributions:} There have been earlier formal investigations
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   173
  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
   174
  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
   175
  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
   176
  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
   177
  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
   178
  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
   179
  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
   180
  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
   181
  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
   182
  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
   183
  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
   184
  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
   185
  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
   186
  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
   187
  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
   188
  et al.~\cite{Sha90} to the practically relevant case where critical 
11
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
   189
  sections can overlap.
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   190
*}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   191
14
1bf194825a4e more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 13
diff changeset
   192
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
   193
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   194
text {*
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   195
  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
   196
  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
   197
  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
   198
  Following good experience in earlier work \cite{Wang09},  
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   199
  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
   200
  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
   201
  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
   202
  \emph{Events} of PIP fall
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   203
  into five categories defined as the datatype:
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   204
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   205
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   206
  \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
   207
  \isacommand{datatype} event 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   208
  & @{text "="} & @{term "Create thread priority"}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   209
  & @{text "|"} & @{term "Exit thread"} \\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   210
  & @{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
   211
  & @{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
   212
  & @{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
   213
  \end{tabular}}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   214
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   215
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   216
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   217
  whereby threads, priorities and (critical) resources are represented
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   218
  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
   219
  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
   220
  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
   221
  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
   222
  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
   223
  ``live'' threads that we have seen so far:
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   224
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   225
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   226
  \mbox{\begin{tabular}{lcl}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   227
  @{thm (lhs) threads.simps(1)} & @{text "\<equiv>"} & 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   228
    @{thm (rhs) threads.simps(1)}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   229
  @{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
   230
    @{thm (rhs) threads.simps(2)[where thread="th"]}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   231
  @{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
   232
    @{thm (rhs) threads.simps(3)[where thread="th"]}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   233
  @{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
   234
  \end{tabular}}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   235
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   236
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   237
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   238
  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
   239
  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
   240
  defined as
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   241
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   242
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   243
  \mbox{\begin{tabular}{lcl}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   244
  @{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
   245
    @{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
   246
  @{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
   247
    @{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
   248
  @{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
   249
    @{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
   250
  @{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
   251
  \end{tabular}}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   252
  \end{isabelle}
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
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   255
  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
   256
  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
   257
  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
   258
  priority last set.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   259
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   260
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   261
  \mbox{\begin{tabular}{lcl}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   262
  @{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
   263
    @{thm (rhs) birthtime.simps(1)[where thread="th"]}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   264
  @{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
   265
    @{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
   266
  @{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
   267
    @{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
   268
  @{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
   269
  \end{tabular}}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   270
  \end{isabelle}
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
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   273
  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
   274
  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
   275
  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
   276
  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
   277
  
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   278
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   279
  @{thm preced_def[where thread="th"]}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   280
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   281
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   282
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   283
  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
   284
  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
   285
  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
   286
  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
   287
  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
   288
  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
   289
  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
   290
  the same priority.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   291
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   292
  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
   293
  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
   294
  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
   295
  that is in possession of the
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   296
  ``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
   297
  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
   298
  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
   299
  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
   300
  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
   301
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   302
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   303
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   304
  @{thm cs_holding_def[where thread="th"]}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   305
  @{thm cs_waiting_def[where thread="th"]}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   306
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   307
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   308
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   309
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   310
  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
   311
  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
   312
  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
   313
  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
   314
  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
   315
  empty list for every resource.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   316
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   317
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   318
  @{abbrev all_unlocked}\hfill\numbered{allunlocked}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   319
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   320
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   321
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   322
  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
   323
  (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
   324
  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
   325
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   326
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   327
  @{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
   328
  @{term "(Cs cs, Th th)"}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   329
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   330
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   331
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   332
  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
   333
  \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
   334
  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
   335
  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
   336
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   337
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   338
  @{thm cs_depend_def}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   339
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   340
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   341
22
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   342
  \begin{figure}[t]
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   343
  \begin{center}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   344
  \newcommand{\fnt}{\fontsize{7}{8}\selectfont}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   345
  \begin{tikzpicture}[scale=1]
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   346
  %%\draw[step=2mm] (-3,2) grid (1,-1);
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   347
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   348
  \node (A) at (0,0) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^sub>0"}};
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   349
  \node (B) at (2,0) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^sub>1"}};
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   350
  \node (C) at (4,0.7) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^sub>1"}};
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   351
  \node (D) at (4,-0.7) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^sub>2"}};
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   352
  \node (E) at (6,-0.7) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^sub>2"}};
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   353
  \node (E1) at (6, 0.2) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^sub>3"}};
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   354
  \node (F) at (8,-0.7) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^sub>3"}};
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   355
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   356
  \node (X) at (0,-2) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^sub>4"}};
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   357
  \node (Y) at (2,-2) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^sub>4"}};
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   358
  \node (Z) at (2,-2.9) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^sub>5"}};
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   359
  \node (U1) at (4,-2) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^sub>5"}};
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   360
  \node (U2) at (4,-2.9) [draw, rounded corners=1mm, rectangle, very thick] {@{text "th\<^sub>6"}};
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   361
   \node (R) at (6,-2.9) [draw, circle, very thick, inner sep=0.4mm] {@{text "cs\<^sub>6"}};
12
85116bc854c0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 11
diff changeset
   362
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   363
  \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
   364
  \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
   365
  \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
   366
  \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
   367
  \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
   368
  \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
   369
85116bc854c0 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 11
diff changeset
   370
  \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
   371
  \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
   372
  \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
   373
  \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
   374
  \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
   375
  \end{tikzpicture}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   376
  \end{center}
13
735e36c64a71 added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
   377
  \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
   378
  \end{figure}
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   379
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   380
  \noindent
22
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   381
  If there is no cycle, then every RAG can be pictured as a forrest of trees, as
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   382
  for example in Figure~\ref{RAGgraph}.
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
   383
8
5ba3d79622da added a paragraph about RAGS
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   384
  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
   385
  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
   386
  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
   387
  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
   388
  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
   389
  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
   390
  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
   391
  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
   392
  
8
5ba3d79622da added a paragraph about RAGS
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   393
5ba3d79622da added a paragraph about RAGS
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 7
diff changeset
   394
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   395
  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
   396
  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
   397
  operation for relations. This gives
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   398
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   399
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   400
  @{thm cs_dependents_def}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   401
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   402
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   403
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   404
  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
   405
  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
   406
  wait for a resource being released (in the picture above this means the dependants
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   407
  of @{text "th\<^sub>0"} are @{text "th\<^sub>1"} and @{text "th\<^sub>2"}, which wait for resource @{text "cs\<^sub>1"}, 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   408
  but also @{text "th\<^sub>3"}, 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   409
  which cannot make any progress unless @{text "th\<^sub>2"} makes progress, which
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   410
  in turn needs to wait for @{text "th\<^sub>0"} to finish). If there is a circle of dependencies 
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   411
  in a RAG, then clearly
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   412
  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
   413
  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
   414
  programmer has to ensure this.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   415
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
  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
   418
  state @{text s}. It is defined as
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   419
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   420
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   421
  @{thm cpreced_def2}\hfill\numbered{cpreced}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   422
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   423
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   424
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   425
  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
   426
  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
   427
  (for example when the thread is
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   428
  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
   429
  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
   430
  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
   431
  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
   432
  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
   433
  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
   434
  lowered prematurely.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   435
  
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   436
  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
   437
  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
   438
  we represent as a record consisting of two
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   439
  functions:
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   440
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   441
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   442
  @{text "\<lparr>wq_fun, cprec_fun\<rparr>"}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   443
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   444
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   445
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   446
  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
   447
  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
   448
  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
   449
  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
   450
  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
   451
  such records.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   452
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   453
  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
   454
  function is defined in \eqref{allunlocked}) and the
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   455
  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
   456
  \mbox{@{abbrev initial_cprec}}. Therefore
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   457
  we have for the initial shedule state
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
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   460
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   461
  @{thm (lhs) schs.simps(1)} @{text "\<equiv>"}\\ 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   462
  \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
   463
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   464
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   465
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   466
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   467
  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
   468
  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
   469
  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
   470
  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
   471
  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
   472
  @{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
   473
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   474
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   475
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   476
  @{thm (lhs) schs.simps(2)} @{text "\<equiv>"}\\ 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   477
  \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
   478
  \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
   479
  @{thm (lhs) schs.simps(3)} @{text "\<equiv>"}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   480
  \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
   481
  \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
   482
  @{thm (lhs) schs.simps(4)} @{text "\<equiv>"}\\ 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   483
  \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
   484
  \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
   485
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   486
  \end{isabelle}
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
  \noindent 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   489
  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
   490
  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
   491
  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
   492
  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
   493
  resource). This gives the clause
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   494
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   495
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   496
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   497
  @{thm (lhs) schs.simps(5)} @{text "\<equiv>"}\\ 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   498
  \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
   499
  \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
   500
  \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
   501
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   502
  \end{isabelle}
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
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   505
  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
   506
  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
   507
  list transformation, we use
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   508
  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
   509
  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
   510
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   511
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   512
  \begin{tabular}{@ {}lcl}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   513
  @{term "release []"} & @{text "\<equiv>"} & @{term "[]"}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   514
  @{term "release (DUMMY # qs)"} & @{text "\<equiv>"} & @{term "qs"}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   515
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   516
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   517
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   518
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   519
  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
   520
  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
   521
  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
   522
  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
   523
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   524
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   525
  \begin{tabular}{@ {}lcl}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   526
  @{term "release []"} & @{text "\<equiv>"} & @{term "[]"}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   527
  @{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
   528
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   529
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   530
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   531
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   532
  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
   533
  choice for the next waiting list. It just has to be a list of distinctive threads and
23
24e6884d9258 made some small chages
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 22
diff changeset
   534
  contains the same elements as @{text "qs"}. This gives for @{term V} the clause:
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   535
 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   536
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   537
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   538
  @{thm (lhs) schs.simps(6)} @{text "\<equiv>"}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   539
  \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
   540
  \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
   541
  \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
   542
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   543
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   544
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   545
  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
   546
  overload, the notions
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   547
  @{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
   548
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   549
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   550
  \begin{tabular}{@ {}rcl}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   551
  @{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
   552
  @{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
   553
  @{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
   554
  @{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
   555
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   556
  \end{isabelle}
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
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   559
  With these abbreviations in place we can introduce 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   560
  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
   561
  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
   562
  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
   563
  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
   564
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   565
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   566
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   567
  @{thm readys_def}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   568
  @{thm runing_def}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   569
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   570
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   571
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   572
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   573
  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
   574
  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
   575
  @{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
   576
  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
   577
  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
   578
  threads. We use sets to capture both possibilities.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   579
  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
   580
  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
   581
  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
   582
  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
   583
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   584
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   585
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   586
  @{thm holdents_def}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   587
  @{thm detached_def}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   588
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   589
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   590
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   591
  %\noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   592
  %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
   593
  
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   594
  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
   595
  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
   596
  created yet. 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   597
  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
   598
  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
   599
  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
   600
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   601
  \begin{center}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   602
  \begin{tabular}{c}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   603
  @{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
   604
  @{thm[mode=Rule] thread_exit[where thread=th]}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   605
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   606
  \end{center}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   607
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   608
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   609
  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
   610
  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
   611
  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
   612
  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
   613
  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
   614
  if the corresponding thread is running. 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   615
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   616
  \begin{center}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   617
  @{thm[mode=Rule] thread_set[where thread=th]}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   618
  \end{center}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   619
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   620
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   621
  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
   622
  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
   623
  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
   624
  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
   625
  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
   626
  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
   627
  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
   628
  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
   629
  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
   630
  responsibility for avoiding problems rests with the programmer.}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   631
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   632
 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   633
  \begin{center}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   634
  @{thm[mode=Rule] thread_P[where thread=th]}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   635
  \end{center}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   636
 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   637
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   638
  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
   639
  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
   640
  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
   641
 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   642
  \begin{center}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   643
  @{thm[mode=Rule] thread_V[where thread=th]}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   644
  \end{center}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   645
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   646
  \noindent
11
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
   647
  Note, however, that apart from the circularity condition, we do not make any 
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   648
  assumption on how different resources can be locked and released relative to each 
11
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
   649
  other. In our model it is possible that critical sections overlap. This is in 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
   650
  contrast to Sha et al \cite{Sha90} who require that critical sections are 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
   651
  properly nested.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
   652
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   653
  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
   654
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   655
  \begin{center}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   656
  \begin{tabular}{c}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   657
  @{thm[mode=Axiom] vt_nil}\hspace{1cm}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   658
  @{thm[mode=Rule] vt_cons}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   659
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   660
  \end{center}
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
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   663
  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
   664
  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
   665
*}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   666
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   667
section {* The Correctness Proof *}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   668
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   669
(*<*)
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   670
context extend_highest_gen
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   671
begin
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   672
(*>*)
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   673
text {* 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   674
  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
   675
  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
   676
  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
   677
  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
   678
  Their second correctness criterion is given
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   679
  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
   680
  @{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
   681
  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
   682
  \emph{not} prevent indefinite, or unbounded, Priority Inversion,
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   683
  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
   684
  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
   685
  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
   686
  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
   687
  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
   688
  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
   689
  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
   690
  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
   691
  Sha et al.~are
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   692
  \emph{not} true for their version of PIP---despite being ``proved''. As Yodaiken
25
a9c0eeb00cc3 added two more references
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 24
diff changeset
   693
  \cite{Yodaiken02} and Moylan et al.~\cite{deinheritance} pointed out: If a low-priority thread possesses
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   694
  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
   695
  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
   696
  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
   697
  high-priority threads, invalidating their two bounds.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   698
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   699
  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
   700
  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
   701
  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
   702
  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
   703
  (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
   704
  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
   705
  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
   706
  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
   707
  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
   708
  @{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
   709
  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
   710
  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
   711
  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
   712
  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
   713
  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
   714
  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
   715
  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
   716
  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
   717
  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
   718
  s'}. We list them next:
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   719
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   720
  \begin{quote}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   721
  {\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
   722
  {\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
   723
  @{text "s' @ s"} are valid states:
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   724
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   725
  \begin{tabular}{l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   726
  @{term "vt s"}, @{term "vt (s' @ s)"} 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   727
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   728
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   729
  \end{quote}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   730
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   731
  \begin{quote}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   732
  {\bf Assumptions on the thread {\boldmath@{text "th"}:}} 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   733
  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
   734
  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
   735
  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
   736
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   737
  \begin{tabular}{l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   738
  @{term "th \<in> threads s"}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   739
  @{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
   740
  @{term "prec th s = (prio, DUMMY)"}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   741
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   742
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   743
  \end{quote}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   744
  
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   745
  \begin{quote}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   746
  {\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
   747
  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
   748
  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
   749
  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
   750
  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
   751
  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
   752
  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
   753
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   754
  \begin{tabular}{l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   755
  {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
   756
  {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
   757
  {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
   758
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   759
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   760
  \end{quote}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   761
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   762
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   763
  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
   764
  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
   765
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   766
  \begin{theorem}\label{mainthm}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   767
  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
   768
  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
   769
  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
   770
  @{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
   771
  @{term "cp (s' @ s) th' = prec th s"}.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   772
  \end{theorem}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   773
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   774
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   775
  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
   776
  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
   777
  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
   778
  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
   779
  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
   780
  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
   781
  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
   782
  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
   783
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   784
  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
   785
  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
   786
  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
   787
  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
   788
  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
   789
  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
   790
  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
   791
  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
   792
  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
   793
  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
   794
  single theorem improving their bound.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   795
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   796
  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
   797
  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
   798
  It is relatively easy to see that
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
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   801
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   802
  @{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
   803
  @{thm[mode=IfThen]  finite_threads}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   804
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   805
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   806
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   807
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   808
  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
   809
  properties are 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   810
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   811
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   812
  \begin{tabular}{@ {}l}
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   813
  @{thm[mode=IfThen] waiting_unique[of _ _ "cs\<^sub>1" "cs\<^sub>2"]}\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   814
  @{thm[mode=IfThen] held_unique[of _ "th\<^sub>1" _ "th\<^sub>2"]}\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   815
  @{thm[mode=IfThen] runing_unique[of _ "th\<^sub>1" "th\<^sub>2"]}
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   816
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   817
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   818
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   819
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   820
  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
   821
  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
   822
  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
   823
  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
   824
  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
   825
  about the @{term RAG} in @{text "s"}.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   826
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   827
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   828
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   829
  @{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
   830
  \hspace{5mm}@{thm (concl) acyclic_depend},
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   831
  @{thm (concl) finite_depend} and
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   832
  @{thm (concl) wf_dep_converse},\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   833
  \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
   834
  and\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   835
  \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
   836
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   837
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   838
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   839
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   840
  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
   841
  @{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
   842
  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
   843
  (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
   844
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   845
  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
   846
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   847
  \begin{lemma}\label{mainlem}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   848
  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
   849
  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
   850
  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
   851
  then @{text "th' \<notin> running (s' @ s)"}.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   852
  \end{lemma}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   853
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   854
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   855
  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
   856
  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
   857
  in the state @{text "s' @ s"}.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   858
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   859
  \begin{proof}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   860
  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
   861
  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
   862
  @{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
   863
  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
   864
  @{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
   865
  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
   866
  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
   867
  @{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
   868
  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
   869
  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
   870
  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
   871
  @{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
   872
  \end{proof}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   873
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   874
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   875
  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
   876
  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
   877
  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
   878
  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
   879
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   880
  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
   881
  blocked by a thread @{text th'} that
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   882
  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
   883
  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
   884
  precedence of @{text th} in @{text "s"}.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   885
  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
   886
  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
   887
  one obtained by Sha et al.~\cite{Sha90}:
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   888
  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
   889
  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
   890
  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
   891
  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
   892
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   893
  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
   894
  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
   895
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   896
  \begin{lemma}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   897
  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
   898
  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
   899
  @{term "running (s' @ s) \<noteq> {}"}.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   900
  \end{lemma}
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
  \begin{proof}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   903
  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
   904
  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
   905
  precedence of @{text th}.\qed
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   906
  \end{proof}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   907
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   908
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   909
  %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
   910
  %\begin{enumerate}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   911
  %\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
   912
  %  @   {thm [display] chain_building[rule_format]}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   913
  %\item The ready thread chased to is unique (@{text "dchain_unique"}):
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   914
  %  @   {thm [display] dchain_unique[of _ _ "th\<^sub>1" "th\<^sub>2"]}
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   915
  %\end{enumerate}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   916
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   917
  %Some deeper results about the system:
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   918
  %\begin{enumerate}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   919
  %\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
   920
  %@  {thm [display] max_cp_eq}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   921
  %\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
   922
  %(@{text "max_cp_readys_threads"}):
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   923
  %@  {thm [display] max_cp_readys_threads}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   924
  %\end{enumerate}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   925
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   926
  %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
   927
  %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
   928
  %\begin{enumerate}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   929
  %\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
   930
  %  one thread holds (@{text "cntCS_v_dec"})
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   931
  %   @  {thm [display]  cntCS_v_dec}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   932
  %\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
   933
  %  (@  {text "cnp_cnv_cncs"}):
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   934
  %  @  {thm [display]  cnp_cnv_cncs}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   935
  %\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
   936
  %  the relevant thread is not living:
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   937
  %  (@{text "cnp_cnv_eq"}):
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   938
  %  @  {thm [display]  cnp_cnv_eq}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   939
  %\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
   940
  %  (@{text "not_thread_holdents"}):
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   941
  %  @  {thm [display] not_thread_holdents}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   942
  %\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
   943
  %  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
   944
  %  (@{text "count_eq_dependents"}):
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   945
  %  @  {thm [display] count_eq_dependents}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   946
  %\end{enumerate}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   947
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   948
  %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
   949
  %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
   950
  %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
   951
  %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
   952
  %lemma @{text "moment_blocked"}:
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   953
  %@   {thm [display] moment_blocked}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   954
  %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
   955
  %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
   956
  %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
   957
  %some resource at state @{text "s"}.
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
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   960
  %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
   961
  %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
   962
  %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
   963
  %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
   964
  %priority inversion can be avoided. 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   965
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   966
  %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
   967
  %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
   968
  %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
   969
  %@   {thm [display] red_moment}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   970
  
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   971
  %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
   972
  %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
   973
  %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
   974
  %about @{term "t"} (@{text "ind"}):
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   975
  %\begin{center}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   976
  %@   {thm[display] ind}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   977
  %\end{center}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   978
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   979
  %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
   980
  %\begin{enumerate}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   981
  %\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
   982
  %  precedence is preserved as well
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   983
  %  (@{text "th_kept"}): 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   984
  %  @   {thm [display] th_kept}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   985
  %\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
   986
  %  all living threads
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   987
  %  (@{text "max_preced"}): 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   988
  %  @   {thm [display] max_preced}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   989
  %\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
   990
  %  among all living threads
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   991
  %  (@{text "th_cp_max_preced"}): 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   992
  %  @   {thm [display] th_cp_max_preced}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   993
  %\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
   994
  %  precedence among all living threads
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   995
  %  (@{text "th_cp_max"}): 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   996
  %  @   {thm [display] th_cp_max}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   997
  %\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
   998
  %  @{term "s"}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   999
  %  (@{text "th_cp_preced"}): 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1000
  %  @   {thm [display] th_cp_preced}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1001
  %\end{enumerate}
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
  %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
  1004
  %(@{text "runing_inversion_2"}):
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1005
  %@   {thm [display] runing_inversion_2}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1006
  %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
  1007
  %already live and held some resource 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1008
  %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
  1009
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1010
  %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
  1011
  %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
  1012
  %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
  1013
  %then.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1014
  *}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1015
(*<*)
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1016
end
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
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1019
section {* Properties for an Implementation\label{implement} *}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1020
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1021
text {*
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1022
  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
  1023
  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
  1024
  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
  1025
  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
  1026
  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
  1027
  depends on all its dependants---a ``global'' transitive notion,
23
24e6884d9258 made some small chages
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 22
diff changeset
  1028
  which is indeed heavy weight (see Definition shown in \eqref{cpreced}).
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1029
  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
  1030
  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
  1031
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1032
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1033
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1034
  @{thm children_def2}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1035
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1036
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1037
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1038
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1039
  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
  1040
  @{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
  1041
  a resource). We can prove the following lemma.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1042
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1043
  \begin{lemma}\label{childrenlem}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1044
  @{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
  1045
  \begin{center}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1046
  @{thm (concl) cp_rec}.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1047
  \end{center}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1048
  \end{lemma}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1049
  
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1050
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1051
  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
  1052
  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
  1053
  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
  1054
  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
  1055
  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
  1056
  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
  1057
  a standard scheduling operation implemented in most operating
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1058
  systems.
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
  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
  1061
  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
  1062
  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
  1063
  kind of events.\smallskip
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1064
*}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1065
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1066
(*<*)
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1067
context step_create_cps
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1068
begin
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
text {*
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1071
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1072
  \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
  1073
  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
  1074
  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
  1075
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1076
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1077
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1078
  @{thm eq_dep},\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1079
  @{thm eq_cp_th}, and\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1080
  @{thm[mode=IfThen] eq_cp}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1081
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1082
  \end{isabelle}
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
  \noindent
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
  1085
  This means in an implementation we do not have to recalculate the @{text RAG} and also none of the
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1086
  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
  1087
  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
  1088
  \smallskip
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1089
  *}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1090
(*<*)
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1091
end
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1092
context step_exit_cps
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1093
begin
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1094
(*>*)
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1095
text {*
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1096
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1097
  \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
  1098
  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
  1099
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1100
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1101
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1102
  @{thm eq_dep}, and\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1103
  @{thm[mode=IfThen] eq_cp}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1104
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1105
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1106
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1107
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1108
  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
  1109
  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
  1110
  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
  1111
  current precedence.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1112
  \smallskip
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1113
*}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1114
(*<*)
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1115
end
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1116
context step_set_cps
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1117
begin
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1118
(*>*)
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1119
text {*
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1120
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1121
  \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
  1122
  @{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
  1123
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1124
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1125
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1126
  @{thm[mode=IfThen] eq_dep}, and\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1127
  @{thm[mode=IfThen] eq_cp_pre}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1128
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1129
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1130
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1131
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1132
  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
  1133
  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
  1134
  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
  1135
  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
  1136
  change of its priority cannot affect other threads.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1137
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1138
  %The second
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1139
  %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
  1140
  %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
  1141
  %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
  1142
  %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
  1143
  %the @{term "cp"} of every 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1144
  %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
  1145
  %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
  1146
  %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
  1147
  %dependants.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1148
  %
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1149
  %\begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1150
  %\begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1151
  %@{thm[mode=IfThen] eq_up_self}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1152
  %@{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
  1153
  %@{text "then"} @{thm (concl) eq_up}.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1154
  %\end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1155
  %\end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1156
  %
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1157
  %\noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1158
  %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
  1159
  %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
  1160
  %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
  1161
  %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
  1162
  %have their current precedence changed.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1163
  \smallskip
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1164
  *}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1165
(*<*)
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1166
end
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1167
context step_v_cps_nt
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1168
begin
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1169
(*>*)
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1170
text {*
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1171
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1172
  \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
  1173
  @{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
  1174
  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
  1175
  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
  1176
  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
  1177
  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
  1178
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1179
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1180
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1181
  @{thm depend_s}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1182
  \end{isabelle}
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
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1185
  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
  1186
  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
  1187
  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
  1188
  can show
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1189
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1190
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1191
  @{thm[mode=IfThen] cp_kept}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1192
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1193
  
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1194
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1195
  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
  1196
  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
  1197
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1198
  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
  1199
  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
  1200
  to be recalculated.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1201
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1202
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1203
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1204
  @{thm depend_s}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1205
  @{thm eq_cp}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1206
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1207
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1208
  *}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1209
(*<*)
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1210
end
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1211
context step_P_cps_e
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1212
begin
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1213
(*>*)
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1214
text {*
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1215
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1216
  \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
  1217
  @{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
  1218
  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
  1219
  first by showing that
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1220
  
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1221
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1222
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1223
  @{thm depend_s}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1224
  @{thm eq_cp}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1225
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1226
  \end{isabelle}
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
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1229
  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
  1230
  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
  1231
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1232
  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
  1233
  
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1234
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1235
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1236
  @{thm depend_s}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1237
  @{thm[mode=IfThen] eq_cp}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1238
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1239
  \end{isabelle}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1240
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1241
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1242
  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
  1243
  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
  1244
  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
  1245
  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
  1246
  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
  1247
  the @{term "cp"} of every 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1248
  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
  1249
  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
  1250
  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
  1251
  dependants.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1252
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1253
  \begin{isabelle}\ \ \ \ \ %%%
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1254
  \begin{tabular}{@ {}l}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1255
  %%@ {t hm[mode=IfThen] eq_up_self}\\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1256
  @{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
  1257
  @{text "then"} @{thm (concl) eq_up}.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1258
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1259
  \end{isabelle}
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
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1262
  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
  1263
  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
  1264
  have their current precedence changed.
22
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  1265
  *}(*<*)end(*>*)text {*
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1266
  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
  1267
  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
  1268
  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
  1269
  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
  1270
  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
  1271
  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
  1272
  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
  1273
  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
  1274
  function interface in PINTOS:
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1275
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1276
  \begin{center}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1277
  \begin{tabular}{|l@ {\hspace{2mm}}|l@ {\hspace{2mm}}|}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1278
  \hline
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1279
  {\bf Event} & {\bf PINTOS function} \\
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1280
  \hline
11
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1281
  @{text Create} & @{ML_text "thread_create"}\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1282
  @{text Exit}   & @{ML_text "thread_exit"}\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1283
  @{text Set}    & @{ML_text "thread_set_priority"}\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1284
  @{text P}      & @{ML_text "lock_acquire"}\\
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1285
  @{text V}      & @{ML_text "lock_release"}\\ 
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1286
  \hline
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1287
  \end{tabular}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1288
  \end{center}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1289
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1290
  \noindent
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1291
  Our implicit assumption that every event is an atomic operation is ensured by the architecture of 
15
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1292
  PINTOS (which allows disabling of interrupts when some operations are performed). The case where 
11
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1293
  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
  1294
  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
  1295
  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
  1296
  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
  1297
  @{ML_text "queue_insert"}, for inserting a new element into a priority queue, and 
22
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  1298
  the function @{ML_text "queue_update"}, for updating the position of an element that is already
17
105715a0a807 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
  1299
  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
  1300
  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
  1301
  
105715a0a807 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
  1302
  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
  1303
  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
  1304
  and formalisation of PIP translates smoothly to an efficent implementation in PINTOS. 
15
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1305
  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
  1306
  shown in Figure~\ref{code}.  This function implements the operation of requesting and, if free, 
15
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1307
  locking of a resource by the current running thread. The convention in the PINTOS
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1308
  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
  1309
  A lock is represented as a pointer to the structure {\tt lock} (Line 1). 
15
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1310
  Lines 2 to 4 are taken from the original 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1311
  code of @{ML_text "lock_acquire"} in PINTOS. They contain diagnostic code: first, 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1312
  there is a check that 
11
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1313
  the lock is a ``valid'' lock 
15
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1314
  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
  1315
  as part of an interrupt---acquiring a lock should only be initiated by a 
15
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1316
  request from a (user) thread, not from an interrupt; third, it is ensured that the 
11
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1317
  current thread does not ask twice for a lock. These assertions are supposed
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1318
  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
  1319
  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
  1320
  a ``kernel panic''. 
11
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1321
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1322
7
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1323
22
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  1324
  \begin{figure}[tph]
7
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1325
  \begin{lstlisting}
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1326
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
  1327
{ ASSERT (lock != NULL);
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1328
  ASSERT (!intr_context());
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1329
  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
  1330
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1331
  enum intr_level old_level;
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1332
  old_level = intr_disable();
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1333
  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
  1334
    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
  1335
    thread_current()->waiting = lock;
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1336
    struct thread *pt;
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1337
    pt = lock->holder;
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1338
    while (pt) {
13
735e36c64a71 added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1339
      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
  1340
      if (!(update_cprec(pt)))
7
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1341
        break;
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1342
      lock = pt->waiting;
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1343
      if (!lock) {
13
735e36c64a71 added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1344
        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
  1345
        break;
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1346
      };
13
735e36c64a71 added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1347
      queue_update(thread_cprec, &lock->wq, &pt->helem);
11
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1348
      pt = lock->holder;
7
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1349
    };
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1350
    thread_block();
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1351
  } else {
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1352
    lock->value--;
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1353
    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
  1354
    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
  1355
  };
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1356
  intr_set_level(old_level);
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1357
}
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1358
  \end{lstlisting}
17
105715a0a807 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
  1359
  \caption{Our version of the {\tt lock\_acquire} function for the small operating 
15
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1360
  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
  1361
  \end{figure}
0514be2ad83e started code explanation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 6
diff changeset
  1362
11
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1363
 
13
735e36c64a71 added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1364
  Line 6 and 7 of {\tt lock\_acquire} make the operation of acquiring a lock atomic by disabling all 
11
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1365
  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
  1366
  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
  1367
  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
  1368
  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
  1369
  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
  1370
  The waiting queue is referenced in the usual C-way as @{ML_text "&lock->wq"}. 
15
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1371
  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
  1372
  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
  1373
  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
  1374
  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
  1375
  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
  1376
  in the RAG (Resource Allocation Graph) and update their
15
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1377
  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
  1378
  current precedence.
14
1bf194825a4e more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 13
diff changeset
  1379
1bf194825a4e more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 13
diff changeset
  1380
  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
  1381
  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
  1382
  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
  1383
  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
  1384
  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
  1385
  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
  1386
  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
  1387
  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
  1388
  thread @{ML_text pt} is waiting for (Lines 17 and 18). If there is none, then 
15
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1389
  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
  1390
  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
  1391
  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
  1392
  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
  1393
  (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
  1394
  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
  1395
14
1bf194825a4e more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 13
diff changeset
  1396
  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
  1397
  (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
  1398
  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
  1399
  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
  1400
  thread already possesses (Line 29).
15
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1401
  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
  1402
  
1bf194825a4e more one the implementation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 13
diff changeset
  1403
13
735e36c64a71 added explanation of the code
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 12
diff changeset
  1404
  Similar operations need to be implementated for the @{ML_text lock_release} function, which
15
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1405
  we however do not show. The reader should note though that we did \emph{not} verify our C-code. 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1406
  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
  1407
  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
  1408
  anything about PIP \cite{sel4}.
105715a0a807 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 16
diff changeset
  1409
  Our verification of PIP however provided us with the justification for designing 
15
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1410
  the C-code. It gave us confidence that leaving the ``chase'' early, whenever
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1411
  there is no change in the calculated current precedence, does not break the
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 14
diff changeset
  1412
  correctness of the algorithm.
6
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
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1415
section {* Conclusion *}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1416
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1417
text {* 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1418
  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
  1419
  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
  1420
  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
  1421
  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
  1422
  have circular lock dependencies.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1423
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1424
  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
  1425
  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
  1426
  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
  1427
  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
  1428
  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
  1429
  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
  1430
  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
  1431
  We can also confirm that Paulson's
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1432
  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
  1433
  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
  1434
  area of this method is security protocols. 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1435
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1436
  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
  1437
  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
  1438
  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
  1439
  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
  1440
  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
  1441
  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
  1442
  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
  1443
  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
  1444
  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
  1445
  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
  1446
  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
  1447
  A byproduct of our formalisation effort is that nearly all
11
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1448
  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
  1449
  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
  1450
  the next thread which takes over a lock is irrelevant for the correctness
11
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1451
  of PIP. Moreover, we eliminated a crucial restriction present in 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 8
diff changeset
  1452
  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
  1453
  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
  1454
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1455
  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
  1456
  now living in a multi-processor world. Priority Inversion certainly
20
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
  1457
  occurs also there, see for example \cite{Brandenburg11,Davis11}.  
16
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
  1458
  However, there is very little ``foundational''
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1459
  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
  1460
  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
  1461
  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
  1462
  ``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
  1463
  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
  1464
  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
  1465
  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
  1466
  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
  1467
  for future work.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1468
22
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  1469
  To us, it seems sound reasoning about scheduling algorithms is fiendishly difficult
27
6b1141c5e24c cleaned up
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
  1470
  if done informally by ``pencil-and-paper''. We infer this from the flawed proof
6b1141c5e24c cleaned up
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
  1471
  in the paper by Sha et al.~\cite{Sha90} and also from \cite{Regehr} where Regehr 
6b1141c5e24c cleaned up
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
  1472
  points out an error in a paper about Preemption 
22
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  1473
  Threshold Scheduling \cite{ThreadX}. The use of a theorem prover was
23
24e6884d9258 made some small chages
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 22
diff changeset
  1474
  invaluable to us in order to be confident about the correctness of our reasoning 
24
6f50e6a8c6e0 some additions
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 23
diff changeset
  1475
  (no case can be overlooked).   
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1476
  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
  1477
  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
  1478
  \cite{dutertre99b}---another solution to the Priority Inversion
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1479
  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
  1480
  avoid it. There have been earlier formal investigations
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1481
  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
  1482
  checking techniques. The results obtained by them apply,
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1483
  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
  1484
  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
  1485
  size. Moreover, our result is a good 
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1486
  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
  1487
  reasoning: gaining deeper understanding of the subject matter.
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1488
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1489
  Our formalisation
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1490
  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
  1491
  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
  1492
  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
  1493
  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
  1494
  for an implementation require 2000 lines. 
22
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  1495
  The code of our formalisation 
9f0b78fcc894 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 20
diff changeset
  1496
  can be downloaded from the Mercurial repository at
27
6b1141c5e24c cleaned up
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 26
diff changeset
  1497
  \url{http://www.dcs.kcl.ac.uk/staff/urbanc/cgi-bin/repos.cgi/pip}.
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1498
16
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
  1499
  %\medskip
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1500
16
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
  1501
  %\noindent
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
  1502
  %{\bf Acknowledgements:}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
  1503
  %We are grateful for the comments we received from anonymous
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 15
diff changeset
  1504
  %referees.
6
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1505
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1506
  \bibliographystyle{plain}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1507
  \bibliography{root}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1508
*}
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1509
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1510
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1511
(*<*)
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1512
end
7f2493296c39 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
  1513
(*>*)