PIPDefs.thy
author Christian Urban <christian dot urban at kcl dot ac dot uk>
Fri, 09 Dec 2016 12:51:29 +0000
changeset 142 10c16b85a839
parent 138 20c1d3a14143
child 144 e4d151d761c4
permissions -rw-r--r--
updated
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     1
(*<*)
64
b4bcd1edbb6d renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 53
diff changeset
     2
theory PIPDefs
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
     3
imports Precedence_ord Max
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     4
begin
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
(*>*)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     6
94
44df6ac30bd2 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 67
diff changeset
     7
chapter {* Definitions *}
44df6ac30bd2 some small changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 67
diff changeset
     8
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     9
text {*
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    10
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
    11
  In this chapter, the formal model of the Priority Inheritance Protocol
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
    12
  (PIP) is presented. The model is based on Paulson's inductive protocol
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    13
  verification method, where the state of the system is modelled as a list
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    14
  of events (trace) happened so far with the latest event put at the head.
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    15
  *}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    16
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    17
text {*
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    18
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    19
  To define events, the identifiers of {\em threads}, {\em priority} and
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    20
  {\em critical resources } (abbreviated as @{text "cs"}) need to be
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
    21
  represented. All three are represented using standard Isabelle/HOL type
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    22
  @{typ "nat"}: *}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    23
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
    24
type_synonym thread = nat    -- {* Type for thread identifiers. *}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    25
type_synonym priority = nat  -- {* Type for priorities. *}
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
    26
type_synonym cs = nat        -- {* Type for critical sections (or critical resources). *}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    28
text {*
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    29
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    30
  \noindent The abstraction of Priority Inheritance Protocol (PIP) is set at
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    31
  the system call level. Every system call is represented as an event. The
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    32
  format of events is defined defined as follows: *}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    33
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    34
datatype event = 
126
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    35
  Create thread priority   -- {* Thread @{text "thread"} is created with priority @{text "priority"}. *}
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    36
| Exit thread              -- {* Thread @{text "thread"} finishing its execution. *}
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    37
| P thread cs              -- {* Thread @{text "thread"} requesting critical resource @{text "cs"}. *}
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    38
| V thread cs              -- {* Thread @{text "thread"} releasing critical resource @{text "cs"}. *}
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    39
| Set thread priority      -- {* Thread @{text "thread"} resets its priority to @{text "priority"}. *}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    40
125
95e7933968f8 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 118
diff changeset
    41
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    42
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    43
text {* 
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    44
  
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
    45
  As mentioned earlier, in Paulson's inductive method, the states of the
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
    46
  system are represented as lists of events, which is defined by the
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
    47
  following type @{text "state"}: *}
126
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    48
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    49
type_synonym state = "event list"
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    50
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    51
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    52
text {*
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    53
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    54
  The function @{text "threads"} is used to calculate the set of live
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    55
  threads (@{text "threads s"}) in state @{text "s"}. *}
126
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    56
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    57
fun threads :: "state \<Rightarrow> thread set"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    58
  where 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    59
  -- {* At the start of the system, the set of threads is empty: *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    60
  "threads [] = {}" | 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    61
  -- {* New thread is added to the @{text "threads"}: *}
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    62
  "threads (Create th prio#s) = {th} \<union> threads s" | 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    63
  -- {* Finished thread is removed: *}
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    64
  "threads (Exit th # s) = (threads s) - {th}" | 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    65
  -- {* Other kind of events does not affect the value of @{text "threads"}: *}
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    66
  "threads (e # s) = threads s" 
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    67
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
    68
text {* 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
    69
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    70
  \noindent The function @{text "threads"} is one of the {\em observation
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    71
  function}s which forms the very basis of Paulson's inductive protocol
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    72
  verification method. Each observation function {\em observes} one
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    73
  particular aspect (or attribute) of the system. For example, the attribute
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
    74
  observed by @{text "threads s"} is the set of threads being live in state
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    75
  @{text "s"}. The protocol being modelled The decision made the protocol
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    76
  being modelled is based on the {\em observation}s returned by {\em
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    77
  observation function}s. Since {\observation function}s forms the very
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    78
  basis on which Paulson's inductive method is based, there will be a lot of
126
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    79
  such observation functions introduced in the following. In fact, any
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    80
  function which takes event list as argument is a {\em observation
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    81
  function}. *}
126
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    82
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    83
text {* 
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
    84
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
    85
  Observation @{text "priority th s"} is the {\em original priority} of
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
    86
  thread @{text "th"} in state @{text "s"}. The {\em original priority} is
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
    87
  the priority assigned to a thread when it is created or when it is reset
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
    88
  by system call (represented by event @{text "Set thread priority"}). *}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    89
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
    90
fun priority :: "thread \<Rightarrow> state \<Rightarrow> priority"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    91
  where
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    92
  -- {* @{text "0"} is assigned to threads which have never been created: *}
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    93
  "priority th [] = 0" |
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    94
  "priority th (Create th' prio # s) = 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    95
     (if th' = th then prio else priority th s)" |
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    96
  "priority th (Set th' prio # s) = 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    97
     (if th' = th then prio else priority th s)" |
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
    98
  "priority th (e # s) = priority th s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    99
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   100
text {*
126
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   101
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   102
  \noindent Observation @{text "last_set th s"} is the last time when the
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   103
  priority of thread @{text "th"} is set, observed from state @{text "s"}.
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   104
  The time in the system is measured by the number of events happened so far
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   105
  since the very beginning. *}
126
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   106
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   107
fun last_set :: "thread \<Rightarrow> state \<Rightarrow> nat"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   108
  where
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   109
  "last_set th [] = 0" |
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   110
  "last_set th ((Create th' prio) # s) = 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   111
       (if (th = th') then length s else last_set th s)" |
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   112
  "last_set th ((Set th' prio) # s) = 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   113
       (if (th = th') then length s else last_set th s)" |
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   114
  "last_set th (_ # s) = last_set th s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   115
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   116
text {*
126
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   117
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   118
  \noindent The {\em precedence} is a notion derived from {\em priority},
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   119
  where the {\em precedence} of a thread is the combination of its {\em
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   120
  original priority} and the {\em time} the priority was set. The intention
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   121
  is to discriminate threads with the same priority by giving threads whose
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   122
  priority is assigned earlier higher precedences, because such threads are
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   123
  assumed more urgent to finish. *}
126
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   124
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   125
definition preced :: "thread \<Rightarrow> state \<Rightarrow> precedence"
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   126
  where "preced th s \<equiv> Prc (priority th s) (last_set th s)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   127
138
20c1d3a14143 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
   128
abbreviation preceds :: "thread set \<Rightarrow> state \<Rightarrow> precedence set"
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   129
  where "preceds ths s \<equiv> {preced th s | th. th \<in> ths}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   130
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   131
text {*
126
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   132
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   133
  \noindent A number of important notions in PIP are represented as the
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   134
  following functions, defined in terms of the waiting queues of the system,
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   135
  where the waiting queues, as a whole, is represented by the @{text "wq"}
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   136
  argument of every notion function. The @{text "wq"} argument is itself a
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   137
  functions which maps every critical resource @{text "cs"} to the list of
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   138
  threads which are holding or waiting for it. The thread at the head of
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   139
  this list is designated as the thread which is current holding the
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   140
  resource, which is slightly different from tradition where all threads in
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   141
  the waiting queue are considered as waiting for the resource. *}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   142
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   143
text {* 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   144
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   145
  \begin{minipage}{0.9\textwidth} This meaning of @{text "wq"} is reflected
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   146
  in the following definition of @{text "holding wq th cs"}, where @{text
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   147
  "holding wq th cs"} means thread @{text "th"} is holding the critical
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   148
  resource @{text "cs"}. This decision is based on @{text "wq"}.
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   149
  \end{minipage} *}
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   150
126
a88af0e4731f minor update
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 125
diff changeset
   151
definition
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   152
  "holding_raw wq thread cs \<equiv> (thread \<in> set (wq cs) \<and> thread = hd (wq cs))"
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   153
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   154
text {* 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   155
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   156
  \begin{minipage}{0.9\textwidth} In accordance with the definition of
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   157
  @{text "holding wq th cs"}, a thread @{text "th"} is considered waiting
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   158
  for @{text "cs"} if it is in the {\em waiting queue} of critical resource
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   159
  @{text "cs"}, but not at the head. This is reflected in the definition of
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   160
  @{text "waiting wq th cs"} as follows: \end{minipage} *}
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   161
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   162
definition
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   163
  "waiting_raw wq thread cs \<equiv> (thread \<in> set (wq cs) \<and> thread \<noteq> hd (wq cs))"
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   164
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   165
text {* 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   166
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   167
  Resource Allocation Graphs (RAG for short) are used extensively in our
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   168
  formal analysis. The following type @{text "node"} is used to represent
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   169
  the two types of nodes in RAGs. *}
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
   170
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   171
datatype node = 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   172
  Th "thread" -- {* Node for a thread. *}
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   173
| Cs "cs"     -- {* Node for a critical resource. *}
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   174
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   175
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   176
text {*
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   177
  
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   178
  \begin{minipage}{0.9\textwidth} @{text "RAG wq"} generates RAG (a binary
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   179
  relations on @{text "node"}) out of waiting queues of the system
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   180
  (represented by the @{text "wq"} argument): \end{minipage} *}
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   181
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   182
definition
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   183
  RAG_raw :: "(cs \<Rightarrow> thread list) => (node * node) set"
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   184
where
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   185
  "RAG_raw wq \<equiv>
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   186
      {(Th th, Cs cs) | th cs. waiting_raw wq th cs} \<union> 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   187
      {(Cs cs, Th th) | cs th. holding_raw wq th cs}"
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   188
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   189
text {*
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   190
 
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   191
  \noindent The following @{text "dependants wq th"} represents the set of
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   192
  threads which are waiting on thread @{text "th"} in Resource Allocation
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   193
  Graph @{text "RAG wq"}. Here, "waiting" means waiting directly or
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   194
  indirectly on the critical resource. *}
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   195
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   196
definition
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   197
  dependants_raw :: "(cs \<Rightarrow> thread list) \<Rightarrow> thread \<Rightarrow> thread set"
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   198
where
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   199
  "dependants_raw wq th \<equiv> {th' . (Th th', Th th) \<in> (RAG_raw wq)^+}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   200
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   201
text {* 
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   202
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   203
  \noindent The following @{text "cpreced s th"} gives the {\em current
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   204
  precedence} of thread @{text "th"} under state @{text "s"}. The definition
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   205
  of @{text "cpreced"} reflects the basic idea of Priority Inheritance that
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   206
  the {\em current precedence} of a thread is the precedence inherited from
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   207
  the maximum of all its dependants, i.e. the threads which are waiting
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   208
  directly or indirectly waiting for some resources from it. If no such
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   209
  thread exits, @{text "th"}'s {\em current precedence} equals its original
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   210
  precedence, i.e. @{text "preced th s"}. *}
33
9b9f2117561f simplified the cp_rec proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   211
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   212
definition 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   213
  cpreced :: "(cs \<Rightarrow> thread list) \<Rightarrow> state \<Rightarrow> thread \<Rightarrow> precedence"
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   214
  where 
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   215
  "cpreced wq s th \<equiv> Max ({preced th s} \<union> preceds (dependants_raw wq th) s)"
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   216
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   217
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   218
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   219
text {*
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   220
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   221
  Notice that the current precedence (@{text "cpreced"}) of one thread
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   222
  @{text "th"} can be boosted (increased) by those threads in the @{text
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   223
  "dependants wq th"}-set. If one thread gets boosted, we say it inherits
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   224
  the priority (or, more precisely, the precedence) of its dependants. This
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   225
  is whrere the word "Inheritance" in Priority Inheritance Protocol comes
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   226
  from. *}
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   227
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   228
lemma cpreced_def2:
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   229
  "cpreced wq s th \<equiv> Max ((\<lambda>th'. preced th' s) ` ({th} \<union> dependants_raw wq th))"
138
20c1d3a14143 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
   230
  unfolding cpreced_def image_def
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   231
  apply(rule eq_reflection)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   232
  apply(rule_tac f="Max" in arg_cong)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   233
  by (auto)
49
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   234
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   235
definition 
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   236
  cpreceds :: "(cs \<Rightarrow> thread list) \<Rightarrow> state \<Rightarrow> thread set \<Rightarrow> precedence set"
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   237
  where 
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   238
  "cpreceds wq s ths \<equiv> {cpreced wq s th | th. th \<in> ths}"
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   239
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   240
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   241
text {* 
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   242
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   243
  \noindent Assuming @{text "qs"} be the waiting queue of a critical
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   244
  resource, the following abbreviation "release qs" is the waiting queue
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   245
  after the thread holding the resource (which is thread at the head of
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   246
  @{text "qs"}) released the resource: *}
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   247
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   248
abbreviation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   249
  "release qs \<equiv> case qs of
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   250
             [] => [] 
49
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   251
          |  (_#qs') => (SOME q. distinct q \<and> set q = set qs')"
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   252
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   253
text {* 
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   254
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   255
  \noindent It can be seen from the definition that the thread at the head
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   256
  of @{text "qs"} is removed from the return value, and the value @{term
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   257
  "q"} is an reordering of @{text "qs'"}, the tail of @{text "qs"}. Through
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   258
  this reordering, one of the waiting threads (those in @{text "qs'"} } is
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   259
  chosen nondeterministically to be the head of the new queue @{text "q"}.
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   260
  Therefore, this thread is the one who takes over the resource. This is a
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   261
  little better different from common sense that the thread who comes the
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   262
  earliest should take over. The intention of this definition is to show
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   263
  that the choice of which thread to take over the release resource does not
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   264
  affect the correctness of the PIP protocol. *}
49
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   265
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   266
text {*
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   267
  
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   268
  The data structure used by the operating system for scheduling is referred
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   269
  to as {\em schedule state}. It is represented as a record consisting of a
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   270
  function assigning waiting queue to resources (to be used as the @{text
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   271
  "wq"} argument in @{text "holding"}, @{text "waiting"} and @{text "RAG"},
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   272
  etc) and a function assigning precedence to threads: *}
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   273
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   274
record schedule_state = 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   275
    wq_fun :: "cs \<Rightarrow> thread list" -- {* The function assigning waiting queue. *}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   276
    cprec_fun :: "thread \<Rightarrow> precedence" -- {* The function assigning precedence. *}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   277
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   278
text {* 
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   279
  
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   280
  \noindent The following two abbreviations (@{text "all_unlocked"} and
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   281
  @{text "initial_cprec"}) are used to set the initial values of the @{text
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   282
  "wq_fun"} @{text "cprec_fun"} fields respectively of the @{text
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   283
  "schedule_state"} record by the following function @{text "sch"}, which is
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   284
  used to calculate the system's {\em schedule state}.
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   285
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   286
  Since there is no thread at the very beginning to make request, all
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   287
  critical resources are free (or unlocked). This status is represented by
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   288
  the abbreviation @{text "all_unlocked"}.
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   289
*}
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   290
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   291
abbreviation
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   292
  "all_unlocked \<equiv> \<lambda>_::cs. ([]::thread list)"
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   293
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   294
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   295
text {* 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   296
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   297
  \noindent The initial current precedence for a thread can be anything,
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   298
  because there is no thread then. We simply assume every thread has
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   299
  precedence @{text "Prc 0 0"}. *}
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   300
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   301
abbreviation 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   302
  "initial_cprec \<equiv> \<lambda>_::thread. Prc 0 0"
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   303
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   304
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   305
text {* 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   306
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   307
  \noindent The following function @{text "schs"} is used to calculate the
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   308
  system's schedule state @{text "schs s"} out of the current system state
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   309
  @{text "s"}. It is the central function to model Priority Inheritance.
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   310
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   311
  \begin{minipage}{0.9\textwidth}
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   312
    Setting the initial value of the @{text "schedule_state"} record 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   313
    (see the explanations above).
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   314
  \end{minipage}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   315
  *}
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   316
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   317
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   318
  -- {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   319
  \begin{minipage}{0.9\textwidth}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   320
  \begin{enumerate}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   321
  \item @{text "ps"} is the schedule state of last moment.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   322
  \item @{text "pwq"} is the waiting queue function of last moment.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   323
  \item @{text "pcp"} is the precedence function of last moment (NOT USED). 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   324
  \item @{text "nwq"} is the new waiting queue function. It is calculated using a @{text "case"} statement:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   325
  \begin{enumerate}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   326
      \item If the happening event is @{text "P thread cs"}, @{text "thread"} is added to 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   327
            the end of @{text "cs"}'s waiting queue.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   328
      \item If the happening event is @{text "V thread cs"} and @{text "s"} is a legal state,
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   329
            @{text "th'"} must equal to @{text "thread"}, 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   330
            because @{text "thread"} is the one currently holding @{text "cs"}. 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   331
            The case @{text "[] \<Longrightarrow> []"} may never be executed in a legal state.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   332
            the @{text "(SOME q. distinct q \<and> set q = set qs)"} is used to choose arbitrarily one 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   333
            thread in waiting to take over the released resource @{text "cs"}. In our representation,
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   334
            this amounts to rearrange elements in waiting queue, so that one of them is put at the head.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   335
      \item For other happening event, the schedule state just does not change.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   336
  \end{enumerate}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   337
  \item @{text "ncp"} is new precedence function, it is calculated from the newly updated waiting queue 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   338
        function. The RAGency of precedence function on waiting queue function is the reason to 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   339
        put them in the same record so that they can evolve together.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   340
  \end{enumerate}
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   341
  
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   342
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   343
  The calculation of @{text "cprec_fun"} depends on the value of @{text "wq_fun"}. 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   344
  Therefore, in the following cases, @{text "wq_fun"} is always calculated first, in 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   345
  the name of @{text "wq"} (if  @{text "wq_fun"} is not changed
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   346
  by the happening event) or @{text "new_wq"} (if the value of @{text "wq_fun"} is changed).
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   347
  \end{minipage}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   348
     *}
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   349
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   350
fun schs :: "state \<Rightarrow> schedule_state"
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   351
  where 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   352
  "schs [] = (| wq_fun = all_unlocked,  cprec_fun = initial_cprec |)" 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   353
| "schs (Create th prio # s) = 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   354
       (let wq = wq_fun (schs s) in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   355
          (|wq_fun = wq, cprec_fun = cpreced wq (Create th prio # s)|))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   356
|  "schs (Exit th # s) = 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   357
       (let wq = wq_fun (schs s) in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   358
          (|wq_fun = wq, cprec_fun = cpreced wq (Exit th # s)|))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   359
|  "schs (Set th prio # s) = 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   360
       (let wq = wq_fun (schs s) in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   361
          (|wq_fun = wq, cprec_fun = cpreced wq (Set th prio # s)|))"
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   362
   -- {*
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   363
   \begin{minipage}{0.9\textwidth}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   364
      Different from the forth coming cases, the @{text "wq_fun"} field of the schedule state 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   365
      is changed. So, the new value is calculated first, in the name of @{text "new_wq"}.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   366
   \end{minipage}   
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   367
   *}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   368
|  "schs (P th cs # s) = 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   369
       (let wq = wq_fun (schs s) in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   370
        let new_wq = wq(cs := (wq cs @ [th])) in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   371
          (|wq_fun = new_wq, cprec_fun = cpreced new_wq (P th cs # s)|))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   372
|  "schs (V th cs # s) = 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   373
       (let wq = wq_fun (schs s) in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   374
        let new_wq = wq(cs := release (wq cs)) in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   375
          (|wq_fun = new_wq, cprec_fun = cpreced new_wq (V th cs # s)|))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   376
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   377
lemma cpreced_initial: 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   378
  "cpreced (\<lambda> cs. []) [] = (\<lambda>_. (Prc 0 0))"
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   379
apply(rule ext)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   380
apply(simp add: cpreced_def)
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   381
apply(simp add: dependants_raw_def RAG_raw_def waiting_raw_def holding_raw_def)
138
20c1d3a14143 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 136
diff changeset
   382
apply(simp add: preced_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   383
done
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   384
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   385
text {* 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   386
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   387
  \noindent The following @{text "wq"} is a shorthand for @{text "wq_fun"}.
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   388
  *}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   389
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   390
definition wq :: "state \<Rightarrow> cs \<Rightarrow> thread list" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   391
  where "wq s = wq_fun (schs s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   392
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   393
text {* 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   394
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   395
  \noindent The following @{text "cp"} is a shorthand for @{text
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   396
  "cprec_fun"}. *}
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   397
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   398
definition cp :: "state \<Rightarrow> thread \<Rightarrow> precedence"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   399
  where "cp s \<equiv> cprec_fun (schs s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   400
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   401
text {* 
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   402
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   403
  \noindent Functions @{text "holding"}, @{text "waiting"}, @{text "RAG"}
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   404
  and @{text "dependants"} still have the same meaning, but redefined so
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   405
  that they no longer RAG on the fictitious {\em waiting queue function}
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   406
  @{text "wq"}, but on system state @{text "s"}. *}
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   407
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   408
definition 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   409
  s_holding_abv: 
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   410
  "holding (s::state) \<equiv> holding_raw (wq s)"
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   411
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   412
definition
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   413
  s_waiting_abv: 
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   414
  "waiting (s::state) \<equiv> waiting_raw (wq s)"
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   415
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   416
definition
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   417
  s_RAG_abv: 
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   418
  "RAG (s::state) \<equiv> RAG_raw (wq s)"
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   419
  
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   420
definition
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   421
  s_dependants_abv: 
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   422
  "dependants (s::state) \<equiv> dependants_raw (wq s)"
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   423
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   424
lemma cp_eq: 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   425
  shows "cp s th = cpreced (wq s) s th"
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   426
unfolding cp_def wq_def
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   427
apply(induct s rule: schs.induct)
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   428
apply(simp add: Let_def cpreced_initial)
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   429
apply(simp add: Let_def)
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   430
apply(simp add: Let_def)
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   431
apply(simp add: Let_def)
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   432
apply(subst (2) schs.simps)
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   433
apply(simp add: Let_def)
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   434
apply(subst (2) schs.simps)
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   435
apply(simp add: Let_def)
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   436
done
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   437
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   438
text {* 
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   439
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   440
  The following lemma can be proved easily, and the meaning is obvious. *}
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   441
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   442
lemma
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   443
  s_holding_def: 
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   444
  "holding (s::state) th cs \<equiv> (th \<in> set (wq s cs) \<and> th = hd (wq s cs))"
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   445
  by(simp add: s_holding_abv holding_raw_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   446
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   447
lemma s_waiting_def: 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   448
  "waiting (s::state) th cs \<equiv> (th \<in> set (wq_fun (schs s) cs) \<and> th \<noteq> hd (wq_fun (schs s) cs))"
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   449
  by (auto simp:s_waiting_abv wq_def waiting_raw_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   450
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   451
lemma s_RAG_def: 
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   452
  "RAG (s::state) =
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   453
    {(Th th, Cs cs) | th cs. waiting_raw (wq s) th cs} \<union> 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   454
    {(Cs cs, Th th) | cs th. holding_raw (wq s) th cs}"
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   455
  by (auto simp:s_RAG_abv wq_def RAG_raw_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   456
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   457
lemma eq_RAG: 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   458
  "RAG_raw (wq s) = RAG s"
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   459
  by (unfold RAG_raw_def s_RAG_def, auto)
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   460
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   461
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   462
lemma s_dependants_def: 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   463
  shows "dependants (s::state) th \<equiv> {th' . (Th th', Th th) \<in> (RAG s)^+}"
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   464
using dependants_raw_def eq_RAG s_dependants_abv wq_def by auto
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   465
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   466
text {*
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   467
  
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   468
  The following function @{text "readys"} calculates the set of ready
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   469
  threads. A thread is {\em ready} for running if it is a live thread and it
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   470
  is not waiting for any critical resource. *}
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   471
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   472
definition readys :: "state \<Rightarrow> thread set"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   473
  where "readys s \<equiv> {th . th \<in> threads s \<and> (\<forall> cs. \<not> waiting s th cs)}"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   474
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   475
text {* 
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   476
  
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   477
  \noindent The following function @{text "running"} calculates the set of
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   478
  running thread, which is the ready thread with the highest precedence. *}
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   479
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   480
definition running :: "state \<Rightarrow> thread set"
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   481
  where "running s \<equiv> {th . th \<in> readys s \<and> cp s th = Max ((cp s) ` (readys s))}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   482
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   483
text {* 
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   484
  
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   485
  \noindent Notice that the definition of @{text "running"} reflects the
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   486
  preemptive scheduling strategy, because, if the @{text "running"}-thread
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   487
  (the one in @{text "running"} set) lowered its precedence by resetting its
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   488
  own priority to a lower one, it will lose its status of being the max in
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   489
  @{text "ready"}-set and be superseded. *}
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   490
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   491
text {* 
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   492
  
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   493
  \noindent The following function @{text "holdents s th"} returns the set
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   494
  of resources held by thread @{text "th"} in state @{text "s"}. *}
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   495
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   496
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   497
definition holdents :: "state \<Rightarrow> thread \<Rightarrow> cs set" 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   498
  where "holdents s th \<equiv> {cs . holding s th cs}"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   499
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   500
lemma holdents_test: 
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   501
  "holdents s th = {cs . (Cs cs, Th th) \<in> RAG s}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   502
unfolding holdents_def
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   503
unfolding s_RAG_def
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   504
unfolding s_holding_abv
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   505
unfolding wq_def
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   506
by (simp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   507
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   508
text {* 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   509
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   510
  \noindent According to the convention of Paulson's inductive method, the
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   511
  decision made by a protocol that event @{text "e"} is eligible to happen
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   512
  next under state @{text "s"} is expressed as @{text "step s e"}. The
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   513
  predicate @{text "step"} is inductively defined as follows (notice how the
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   514
  decision is based on the {\em observation function}s defined above, and
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   515
  also notice how a complicated protocol is modeled by a few simple
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   516
  observations, and how such a kind of simplicity gives rise to improved
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   517
  trust on faithfulness): *}
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   518
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   519
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   520
inductive step :: "state \<Rightarrow> event \<Rightarrow> bool"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   521
  where
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   522
  -- {* 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   523
  A thread can be created if it is not a live thread:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   524
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   525
  thread_create: "\<lbrakk>thread \<notin> threads s\<rbrakk> \<Longrightarrow> step s (Create thread prio)" |
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   526
  -- {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   527
  A thread can exit if it no longer hold any resource:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   528
  *}
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   529
  thread_exit: "\<lbrakk>thread \<in> running s; holdents s thread = {}\<rbrakk> \<Longrightarrow> step s (Exit thread)" |
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   530
  -- {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   531
  \begin{minipage}{0.9\textwidth}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   532
  A thread can request for an critical resource @{text "cs"}, if it is running and 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   533
  the request does not form a loop in the current RAG. The latter condition 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   534
  is set up to avoid deadlock. The condition also reflects our assumption all threads are 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   535
  carefully programmed so that deadlock can not happen:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   536
  \end{minipage}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   537
  *}
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   538
  thread_P: "\<lbrakk>thread \<in> running s;  (Cs cs, Th thread)  \<notin> (RAG s)^+\<rbrakk> \<Longrightarrow> 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   539
                                                                step s (P thread cs)" |
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   540
  -- {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   541
  \begin{minipage}{0.9\textwidth}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   542
  A thread can release a critical resource @{text "cs"} 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   543
  if it is running and holding that resource:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   544
  \end{minipage}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   545
  *}
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   546
  thread_V: "\<lbrakk>thread \<in> running s;  holding s thread cs\<rbrakk> \<Longrightarrow> step s (V thread cs)" |
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   547
  -- {*
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   548
  \begin{minipage}{0.9\textwidth}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   549
  A thread can adjust its own priority as long as it is current running. 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   550
  With the resetting of one thread's priority, its precedence may change. 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   551
  If this change lowered the precedence, according to the definition of @{text "running"}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   552
  function, 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   553
  \end{minipage}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   554
  *}  
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   555
  thread_set: "\<lbrakk>thread \<in> running s\<rbrakk> \<Longrightarrow> step s (Set thread prio)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   556
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   557
text {*
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   558
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   559
  In Paulson's inductive method, every protocol is defined by such a @{text
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   560
  "step"} predicate. For instance, the predicate @{text "step"} given above
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   561
  defines the PIP protocol. So, it can also be called "PIP". *}
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   562
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   563
abbreviation
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   564
  "PIP \<equiv> step"
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   565
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   566
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   567
text {* 
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   568
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   569
  \noindent For any protocol defined by a @{text "step"} predicate, the fact
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   570
  that @{text "s"} is a legal state in the protocol is expressed as: @{text
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   571
  "vt step s"}, where the predicate @{text "vt"} can be defined as the
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   572
  following: *}
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   573
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   574
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   575
inductive vt :: "state \<Rightarrow> bool"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   576
  where
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   577
  -- {* Empty list @{text "[]"} is a legal state in any protocol:*}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   578
  vt_nil[intro]: "vt []" |
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   579
  -- {* 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   580
  \begin{minipage}{0.9\textwidth}
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   581
  If @{text "s"} a legal state of the protocol defined by predicate @{text "step"}, 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   582
  and event @{text "e"} is allowed to happen under state @{text "s"} by the protocol 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   583
  predicate @{text "step"}, then @{text "e#s"} is a new legal state rendered by the 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   584
  happening of @{text "e"}:
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   585
  \end{minipage}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   586
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   587
  vt_cons[intro]: "\<lbrakk>vt s; step s e\<rbrakk> \<Longrightarrow> vt (e#s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   588
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   589
text {*  
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   590
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   591
  \noindent It is easy to see that the definition of @{text "vt"} is
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   592
  generic. It can be applied to any specific protocol specified by a @{text
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   593
  "step"}-predicate to get the set of legal states of that particular
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   594
  protocol. *}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   595
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   596
text {* 
127
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   597
  
38c6acf03f68 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 126
diff changeset
   598
  The following are two very basic properties of @{text "vt"}. *}
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   599
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   600
lemma step_back_vt: "vt (e#s) \<Longrightarrow> vt s"
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   601
  by(ind_cases "vt (e#s)", simp)
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   602
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   603
lemma step_back_step: "vt (e#s) \<Longrightarrow> step s e"
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   604
  by(ind_cases "vt (e#s)", simp)
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   605
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   606
text {* \noindent
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   607
  The following two auxiliary functions @{text "the_cs"} and @{text "the_th"} are used to extract
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   608
  critical resource and thread respectively out of RAG nodes.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   609
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   610
fun the_cs :: "node \<Rightarrow> cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   611
  where "the_cs (Cs cs) = cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   612
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   613
fun the_th :: "node \<Rightarrow> thread"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   614
  where "the_th (Th th) = th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   615
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   616
text {* \noindent
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   617
  The following predicate @{text "next_th"} describe the next thread to 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   618
  take over when a critical resource is released. In @{text "next_th s th cs t"}, 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   619
  @{text "th"} is the thread to release, @{text "t"} is the one to take over.
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   620
  Notice how this definition is backed up by the @{text "release"} function and its use 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   621
  in the @{text "V"}-branch of @{text "schs"} function. This @{text "next_th"} function
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   622
  is not needed for the execution of PIP. It is introduced as an auxiliary function 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   623
  to state lemmas. The correctness of this definition will be confirmed by 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   624
  lemmas @{text "step_v_hold_inv"}, @{text " step_v_wait_inv"}, 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   625
  @{text "step_v_get_hold"} and @{text "step_v_not_wait"}.
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   626
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   627
definition next_th:: "state \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> thread \<Rightarrow> bool"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   628
  where "next_th s th cs t = (\<exists> rest. wq s cs = th#rest \<and> rest \<noteq> [] \<and> 
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   629
                                     t = hd (SOME q. distinct q \<and> set q = set rest))"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   630
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   631
text {* \noindent
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   632
  The aux function @{text "count Q l"} is used to count the occurrence of situation @{text "Q"}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   633
  in list @{text "l"}:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   634
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   635
definition count :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> nat"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   636
  where "count Q l = length (filter Q l)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   637
115
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   638
text {*
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   639
  The operation @{term P} is used by a thread to request for resources, while
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   640
  @{term V} is used to release. Therefore, the number of resources
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   641
  held by a thread roughly equals to the number of @{term P} it made minus 
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   642
  the number of @{term V}. The equality is rough because the @{term P}-operation
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   643
  might be blocked and fail to give back the holding of the requested resource.
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   644
  In the following, @{text "cntP"} and @{text "cntV"} are the number of @{term P}
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   645
  and @{term "V"} respectively, while @{term cntCS} is the number 
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   646
  resources held by a thread and @{text "pvD"} is introduced to account for 
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   647
  the above mentioned situation when @{term P} is blocked, so that 
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   648
  a equation between @{text cntP}, @{text "cntV"}, @{text "cntCS"} can be established
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   649
  (in the lemma named @{text "valid_trace.cnp_cnv_cncs"}).
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   650
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   651
  Such a equation, once established, is very handy, because the number of resources 
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   652
  held by a thread can be calculated by counting the number of @{term P} and @{text V},
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   653
  which is relatively easy.
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   654
*}
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   655
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   656
definition cntP :: "state \<Rightarrow> thread \<Rightarrow> nat"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   657
  where "cntP s th = count (\<lambda> e. \<exists> cs. e = P th cs) s"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   658
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   659
definition cntV :: "state \<Rightarrow> thread \<Rightarrow> nat"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   660
  where "cntV s th = count (\<lambda> e. \<exists> cs. e = V th cs) s"
65
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   661
115
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   662
definition cntCS :: "state \<Rightarrow> thread \<Rightarrow> nat"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   663
  where "cntCS s th = card (holdents s th)"
74fc1eae4605 Commenting of PIPBasics.thy almost completed. The last section needs to be distributed to Correctness.thy and Implementation.thy
zhangx
parents: 104
diff changeset
   664
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 67
diff changeset
   665
definition "pvD s th = (if (th \<in> readys s \<or> th \<notin> threads s) then 0 else (1::nat))"
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 67
diff changeset
   666
65
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   667
text {* @{text "the_preced"} is also the same as @{text "preced"}, the only
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   668
       difference is the order of arguemts. *}
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   669
definition "the_preced s th = preced th s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   670
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   671
text {* @{term "the_thread"} extracts thread out of RAG node. *}
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   672
fun the_thread :: "node \<Rightarrow> thread" where
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   673
   "the_thread (Th th) = th"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   674
80
17305a85493d CpsG.thy retrofiting almost completed. An important mile stone.
zhangx
parents: 67
diff changeset
   675
65
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   676
text {* The following @{text "wRAG"} is the waiting sub-graph of @{text "RAG"}. *}
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   677
definition "wRAG (s::state) = {(Th th, Cs cs) | th cs. waiting s th cs}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   678
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   679
text {* The following @{text "hRAG"} is the holding sub-graph of @{text "RAG"}. *}
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   680
definition "hRAG (s::state) =  {(Cs cs, Th th) | th cs. holding s th cs}"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   681
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   682
text {* 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   683
  The following @{text "tRAG"} is the thread-graph derived from @{term "RAG"}.
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   684
  It characterizes the dependency between threads when calculating current
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   685
  precedences. It is defined as the composition of the above two sub-graphs, 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   686
  names @{term "wRAG"} and @{term "hRAG"}.
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   687
 *}
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   688
definition "tRAG s = wRAG s O hRAG s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   689
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   690
text {* The following lemma splits @{term "RAG"} graph into the above two sub-graphs. *}
118
a4bb5525b7b6 updated to Isabelle 2016
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 115
diff changeset
   691
65
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   692
lemma RAG_split: "RAG s = (wRAG s \<union> hRAG s)"
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   693
using hRAG_def s_RAG_def s_holding_abv s_waiting_abv wRAG_def wq_def by auto
65
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   694
100
3d2b59f15f26 Reorganizing PIPBasics.thy
zhangx
parents: 80
diff changeset
   695
lemma tRAG_alt_def: 
3d2b59f15f26 Reorganizing PIPBasics.thy
zhangx
parents: 80
diff changeset
   696
  "tRAG s = {(Th th1, Th th2) | th1 th2. 
3d2b59f15f26 Reorganizing PIPBasics.thy
zhangx
parents: 80
diff changeset
   697
                  \<exists> cs. (Th th1, Cs cs) \<in> RAG s \<and> (Cs cs, Th th2) \<in> RAG s}"
3d2b59f15f26 Reorganizing PIPBasics.thy
zhangx
parents: 80
diff changeset
   698
 by (auto simp:tRAG_def RAG_split wRAG_def hRAG_def)
3d2b59f15f26 Reorganizing PIPBasics.thy
zhangx
parents: 80
diff changeset
   699
136
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   700
text {*
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   701
  A notion of {\em thread graph} (@{text tG}) is introduced to hide derivations using 
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   702
  tRAG, so that it is easier to understand.
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   703
*}
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   704
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   705
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   706
definition "tG s = (map_prod the_thread the_thread) `(tRAG s)"
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   707
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   708
lemma tG_alt_def: 
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   709
  "tG s = {(th1, th2) | th1 th2. 
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   710
                  \<exists> cs. (Th th1, Cs cs) \<in> RAG s \<and> (Cs cs, Th th2) \<in> RAG s}" (is "?L = ?R")
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   711
proof -
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   712
  { fix th1 th2
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   713
    assume "(th1, th2) \<in> ?L" 
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   714
    hence "(th1, th2) \<in> ?R" by (auto simp:tG_def tRAG_alt_def)
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   715
  } moreover {
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   716
    fix th1 th2
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   717
    assume "(th1, th2) \<in> ?R"
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   718
    then obtain cs where "(Th th1, Cs cs) \<in> RAG s" "(Cs cs, Th th2) \<in> RAG s" by auto
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   719
    hence "(Th th1, Th th2) \<in> (wRAG s O hRAG s)" by (auto simp:RAG_split wRAG_def hRAG_def)
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   720
    hence "(Th th1, Th th2) \<in> tRAG s" by (unfold tRAG_def, simp)
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   721
    hence "(th1, th2) \<in> ?L" by (simp add: tG_def rev_image_eqI)
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   722
  } ultimately show ?thesis by (meson pred_equals_eq2)
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   723
qed
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   724
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   725
lemma tRAG_def_tG: "tRAG s = (map_prod Th Th) ` (tG s)" (is "?L = ?R")
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   726
proof -
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   727
  have "?L = (\<lambda> x. x) ` ?L" by simp
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   728
  also have "... = ?R"
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   729
  proof(unfold tG_def image_comp, induct rule:image_cong)
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   730
    case (2 e)
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   731
    thus ?case by (unfold tRAG_alt_def, auto)
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   732
  qed auto
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   733
  finally show ?thesis .
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   734
qed
fb3f52fe99d1 updated tG definition
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 130
diff changeset
   735
130
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   736
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   737
fun actor  where
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   738
  "actor (Exit th) = th" |
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   739
  "actor (P th cs) = th" |
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   740
  "actor (V th cs) = th" |
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   741
  "actor (Set th pty) = th" |
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   742
  "actor (Create th prio) = th"
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   743
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   744
-- {* The actions of a set of threads *}
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   745
definition "actions_of ths s = filter (\<lambda> e. actor e \<in> ths) s"
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   746
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   747
fun isCreate :: "event \<Rightarrow> bool" where
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   748
  "isCreate (Create th pty) = True" |
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   749
  "isCreate _ = False"
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   750
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   751
fun isP :: "event \<Rightarrow> bool" where
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   752
  "isP (P th cs) = True" |
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   753
  "isP _ = False"
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   754
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   755
fun isV :: "event \<Rightarrow> bool" where
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   756
  "isV (V th cs) = True" |
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   757
  "isV _ = False"
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   758
0f124691c191 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 129
diff changeset
   759
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   760
(*<*)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   761
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   762
end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   763
(*>*)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   764