PIPDefs.thy
author zhangx
Thu, 28 Jan 2016 15:36:48 +0800
changeset 87 33cb65e00ac0
parent 67 25fd656667a7
child 80 17305a85493d
child 94 44df6ac30bd2
permissions -rw-r--r--
Tracking ExtGG.thy etc., so that a update to 83 is possible.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
     1
chapter {* Definitions *}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     2
(*<*)
64
b4bcd1edbb6d renamed files
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 53
diff changeset
     3
theory PIPDefs
65
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
     4
imports Precedence_ord Moment RTree Max
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     5
begin
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     6
(*>*)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     7
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
     8
text {*
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
     9
  In this section, the formal model of  Priority Inheritance Protocol (PIP) is presented. 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    10
  The model is based on Paulson's inductive protocol verification method, where 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    11
  the state of the system is modelled as a list of events happened so far with the latest 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    12
  event put at the head. 
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    13
*}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    14
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    15
text {*
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    16
  To define events, the identifiers of {\em threads},
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    17
  {\em priority} and {\em critical resources } (abbreviated as @{text "cs"}) 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    18
  need to be represented. All three are represetned using standard 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    19
  Isabelle/HOL type @{typ "nat"}:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    20
*}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    21
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    22
type_synonym thread = nat -- {* Type for thread identifiers. *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    23
type_synonym priority = nat  -- {* Type for priorities. *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    24
type_synonym cs = nat -- {* Type for critical sections (or critical resources). *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    25
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    26
text {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    27
  \noindent
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    28
  The abstraction of Priority Inheritance Protocol (PIP) is set at the system call level.
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    29
  Every system call is represented as an event. The format of events is defined 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    30
  defined as follows:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    31
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    32
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    33
datatype event = 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    34
  Create thread priority | -- {* Thread @{text "thread"} is created with priority @{text "priority"}. *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    35
  Exit thread | -- {* Thread @{text "thread"} finishing its execution. *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    36
  P thread cs | -- {* Thread @{text "thread"} requesting critical resource @{text "cs"}. *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    37
  V thread cs | -- {* Thread @{text "thread"}  releasing critical resource @{text "cs"}. *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    38
  Set thread priority -- {* Thread @{text "thread"} resets its priority to @{text "priority"}. *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    39
67
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    40
fun actor :: "event \<Rightarrow> thread" where
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    41
  "actor (Create th pty) = th" |
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    42
  "actor (Exit th) = th" |
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    43
  "actor (P th cs) = th" |
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    44
  "actor (V th cs) = th" |
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    45
  "actor (Set th pty) = th"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    46
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    47
fun isCreate :: "event \<Rightarrow> bool" where
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    48
  "isCreate (Create th pty) = True" |
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    49
  "isCreate _ = False"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    50
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    51
fun isP :: "event \<Rightarrow> bool" where
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    52
  "isP (P th cs) = True" |
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    53
  "isP _ = False"
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    54
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    55
fun isV :: "event \<Rightarrow> bool" where
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    56
  "isV (V th cs) = True" |
25fd656667a7 Correctness simplified a great deal.
zhangx
parents: 65
diff changeset
    57
  "isV _ = False"
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    58
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    59
text {* 
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    60
  As mentioned earlier, in Paulson's inductive method, the states of system are represented as lists of events,
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    61
  which is defined by the following type @{text "state"}:
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    62
  *}
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    63
type_synonym state = "event list"
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    64
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    65
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    66
text {* 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    67
\noindent
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    68
  Resource Allocation Graph (RAG for short) is used extensively in our formal analysis. 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    69
  The following type @{text "node"} is used to represent nodes in RAG.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    70
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    71
datatype node = 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    72
   Th "thread" | -- {* Node for thread. *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    73
   Cs "cs" -- {* Node for critical resource. *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    74
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    75
text {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    76
  \noindent
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    77
  The following function
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    78
  @{text "threads"} is used to calculate the set of live threads (@{text "threads s"})
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    79
  in state @{text "s"}.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    80
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    81
fun threads :: "state \<Rightarrow> thread set"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    82
  where 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    83
  -- {* 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
    84
  "threads [] = {}" | 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    85
  -- {* New thread is added to the @{text "threads"}: *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    86
  "threads (Create thread prio#s) = {thread} \<union> threads s" | 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    87
  -- {* Finished thread is removed: *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    88
  "threads (Exit thread # s) = (threads s) - {thread}" | 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    89
  -- {* Other kind of events does not affect the value of @{text "threads"}: *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
    90
  "threads (e#s) = threads s" 
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
    91
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
    92
text {* 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
    93
  \noindent
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
    94
  The function @{text "threads"} defined above is one of 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
    95
  the so called {\em observation function}s which forms 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
    96
  the very basis of Paulson's inductive protocol verification method.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
    97
  Each observation function {\em observes} one particular aspect (or attribute)
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
    98
  of the system. For example, the attribute observed by  @{text "threads s"}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
    99
  is the set of threads living in state @{text "s"}. 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   100
  The protocol being modelled 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   101
  The decision made the protocol being modelled is based on the {\em observation}s
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   102
  returned by {\em observation function}s. Since {\observation function}s forms 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   103
  the very basis on which Paulson's inductive method is based, there will be 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   104
  a lot of such observation functions introduced in the following. In fact, any function 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   105
  which takes event list as argument is a {\em observation function}.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   106
  *}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   107
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   108
text {* \noindent
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   109
  Observation @{text "priority th s"} is
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   110
  the {\em original priority} of thread @{text "th"} in state @{text "s"}. 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   111
  The {\em original priority} is the priority 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   112
  assigned to a thread when it is created or when it is reset by system call 
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   113
  (represented by event @{text "Set thread priority"}).
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   114
*}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   115
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   116
fun priority :: "thread \<Rightarrow> state \<Rightarrow> priority"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   117
  where
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   118
  -- {* @{text "0"} is assigned to threads which have never been created: *}
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   119
  "priority thread [] = 0" |
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   120
  "priority thread (Create thread' prio#s) = 
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   121
     (if thread' = thread then prio else priority thread s)" |
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   122
  "priority thread (Set thread' prio#s) = 
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   123
     (if thread' = thread then prio else priority thread s)" |
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   124
  "priority thread (e#s) = priority thread s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   125
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   126
text {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   127
  \noindent
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   128
  Observation @{text "last_set th s"} is the last time when the priority of thread @{text "th"} is set, 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   129
  observed from state @{text "s"}.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   130
  The time in the system is measured by the number of events happened so far since the very beginning.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   131
*}
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   132
fun last_set :: "thread \<Rightarrow> state \<Rightarrow> nat"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   133
  where
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   134
  "last_set thread [] = 0" |
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   135
  "last_set thread ((Create thread' prio)#s) = 
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   136
       (if (thread = thread') then length s else last_set thread s)" |
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   137
  "last_set thread ((Set thread' prio)#s) = 
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   138
       (if (thread = thread') then length s else last_set thread s)" |
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   139
  "last_set thread (_#s) = last_set thread s"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   140
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   141
text {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   142
  \noindent 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   143
  The {\em precedence} is a notion derived from {\em priority}, where the {\em precedence} of 
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
   144
  a thread is the combination of its {\em original priority} and {\em time} the priority is set. 
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
   145
  The intention is to discriminate threads with the same priority by giving threads whose priority
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   146
  is assigned earlier higher precedences, becasue such threads are more urgent to finish. 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   147
  This explains the following definition:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   148
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   149
definition preced :: "thread \<Rightarrow> state \<Rightarrow> precedence"
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   150
  where "preced thread s \<equiv> Prc (priority thread s) (last_set thread s)"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   151
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   152
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   153
text {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   154
  \noindent
49
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   155
  A number of important notions in PIP are represented as the following functions, 
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   156
  defined in terms of the waiting queues of the system, where the waiting queues 
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
   157
  , as a whole, is represented by the @{text "wq"} argument of every notion function.
49
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   158
  The @{text "wq"} argument is itself a functions which maps every critical resource 
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   159
  @{text "cs"} to the list of threads which are holding or waiting for it. 
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
   160
  The thread at the head of this list is designated as the thread which is current 
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
   161
  holding the resrouce, which is slightly different from tradition where
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
   162
  all threads in 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
   163
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   164
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   165
consts 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   166
  holding :: "'b \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> bool" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   167
  waiting :: "'b \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> bool"
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
   168
  RAG :: "'b \<Rightarrow> (node \<times> node) set"
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   169
  dependants :: "'b \<Rightarrow> thread \<Rightarrow> thread set"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   170
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   171
defs (overloaded) 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   172
  -- {* 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   173
  \begin{minipage}{0.9\textwidth}
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
   174
  This meaning of @{text "wq"} is reflected in the following definition of @{text "holding wq th cs"},
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
   175
  where @{text "holding wq th cs"} means thread @{text "th"} is holding the critical 
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
   176
  resource @{text "cs"}. This decision is based on @{text "wq"}.
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   177
  \end{minipage}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   178
  *}
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
   179
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   180
  cs_holding_def: 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   181
  "holding wq thread cs \<equiv> (thread \<in> set (wq cs) \<and> thread = hd (wq cs))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   182
  -- {* 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   183
  \begin{minipage}{0.9\textwidth}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   184
  In accordance with the definition of @{text "holding wq th cs"}, 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   185
  a thread @{text "th"} is considered waiting for @{text "cs"} if 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   186
  it is in the {\em waiting queue} of critical resource @{text "cs"}, but not at the head.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   187
  This is reflected in the definition of @{text "waiting wq th cs"} as follows:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   188
  \end{minipage}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   189
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   190
  cs_waiting_def: 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   191
  "waiting wq thread cs \<equiv> (thread \<in> set (wq cs) \<and> thread \<noteq> hd (wq cs))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   192
  -- {* 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   193
  \begin{minipage}{0.9\textwidth}
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
   194
  @{text "RAG wq"} generates RAG (a binary relations on @{text "node"})
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
   195
  out of waiting queues of the system (represented by the @{text "wq"} argument):
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   196
  \end{minipage}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   197
  *}
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
   198
  cs_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
   199
  "RAG (wq::cs \<Rightarrow> thread list) \<equiv>
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   200
      {(Th th, Cs cs) | th cs. waiting wq th cs} \<union> {(Cs cs, Th th) | cs th. holding wq th cs}"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   201
  -- {* 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   202
  \begin{minipage}{0.9\textwidth}
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
   203
  The following @{text "dependants wq th"} represents the set of threads which are RAGing on
48
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
   204
  thread @{text "th"} in Resource Allocation Graph @{text "RAG wq"}. 
c0f14399c12f Some changes in the PrioGDef.thy.
xingyuan zhang <xingyuanzhang@126.com>
parents: 35
diff changeset
   205
  Here, "RAGing" means waiting directly or indirectly on the critical resource. 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   206
  \end{minipage}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   207
  *}
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   208
  cs_dependants_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
   209
  "dependants (wq::cs \<Rightarrow> thread list) th \<equiv> {th' . (Th th', Th th) \<in> (RAG wq)^+}"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   210
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   211
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   212
text {* \noindent 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   213
  The following
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   214
  @{text "cpreced s th"} gives the {\em current precedence} of thread @{text "th"} under
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   215
  state @{text "s"}. The definition of @{text "cpreced"} reflects the basic idea of 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   216
  Priority Inheritance that the {\em current precedence} of a thread is the precedence 
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   217
  inherited from the maximum of all its dependants, i.e. the threads which are waiting 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   218
  directly or indirectly waiting for some resources from it. If no such thread exits, 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   219
  @{text "th"}'s {\em current precedence} equals its original precedence, i.e. 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   220
  @{text "preced th s"}.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   221
  *}
33
9b9f2117561f simplified the cp_rec proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   222
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   223
definition cpreced :: "(cs \<Rightarrow> thread list) \<Rightarrow> state \<Rightarrow> thread \<Rightarrow> precedence"
33
9b9f2117561f simplified the cp_rec proof
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 32
diff changeset
   224
  where "cpreced wq s = (\<lambda>th. Max ((\<lambda>th'. preced th' s) ` ({th} \<union> dependants wq th)))"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   225
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   226
text {*
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   227
  Notice that the current precedence (@{text "cpreced"}) of one thread @{text "th"} can be boosted 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   228
  (becoming larger than its own precedence) by those threads in 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   229
  the @{text "dependants wq th"}-set. If one thread get boosted, we say 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   230
  it inherits the priority (or, more precisely, the precedence) of 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   231
  its dependants. This is how the word "Inheritance" in 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   232
  Priority Inheritance Protocol comes.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   233
*}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   234
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   235
(*<*)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   236
lemma 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   237
  cpreced_def2:
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   238
  "cpreced wq s th \<equiv> Max ({preced th s} \<union> {preced th' s | th'. th' \<in> dependants wq th})"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   239
  unfolding cpreced_def image_def
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   240
  apply(rule eq_reflection)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   241
  apply(rule_tac f="Max" in arg_cong)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   242
  by (auto)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   243
(*>*)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   244
49
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   245
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   246
text {* \noindent
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   247
  Assuming @{text "qs"} be the waiting queue of a critical resource, 
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   248
  the following abbreviation "release qs" is the waiting queue after the thread 
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   249
  holding the resource (which is thread at the head of @{text "qs"}) released
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   250
  the resource:
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   251
*}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   252
abbreviation
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   253
  "release qs \<equiv> case qs of
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   254
             [] => [] 
49
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   255
          |  (_#qs') => (SOME q. distinct q \<and> set q = set qs')"
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   256
text {* \noindent
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   257
  It can be seen from the definition that the thread at the head of @{text "qs"} is removed
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   258
  from the return value, and the value @{term "q"} is an reordering of @{text "qs'"}, the 
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   259
  tail of @{text "qs"}. Through this reordering, one of the waiting threads (those in @{text "qs'"} }
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   260
  is chosen nondeterministically to be the head of the new queue @{text "q"}. 
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   261
  Therefore, this thread is the one who takes over the resource. This is a little better different 
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   262
  from common sense that the thread who comes the earliest should take over.  
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   263
  The intention of this definition is to show that the choice of which thread to take over the 
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   264
  release resource does not affect the correctness of the PIP protocol. 
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   265
*}
8679d75b1d76 A little more change.
xingyuan zhang <xingyuanzhang@126.com>
parents: 48
diff changeset
   266
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   267
text {*
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   268
  The data structure used by the operating system for scheduling is referred to as 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   269
  {\em schedule state}. It is represented as a record consisting of 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   270
  a function assigning waiting queue to resources 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   271
  (to be used as the @{text "wq"} argument in @{text "holding"}, @{text "waiting"} 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   272
  and  @{text "RAG"}, etc) and a function assigning precedence to threads:
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
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   275
record schedule_state = 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   276
    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
   277
    cprec_fun :: "thread \<Rightarrow> precedence" -- {* The function assigning precedence. *}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   278
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   279
text {* \noindent
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   280
  The following two abbreviations (@{text "all_unlocked"} and @{text "initial_cprec"}) 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   281
  are used to set the initial values of the @{text "wq_fun"} @{text "cprec_fun"} fields 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   282
  respectively of the @{text "schedule_state"} record by the following function @{text "sch"},
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   283
  which is used to calculate the system's {\em schedule state}.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   284
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   285
  Since there is no thread at the very beginning to make request, all critical resources 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   286
  are free (or unlocked). This status is represented by the abbreviation
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   287
  @{text "all_unlocked"}. 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   288
  *}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   289
abbreviation
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   290
  "all_unlocked \<equiv> \<lambda>_::cs. ([]::thread list)"
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   291
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   292
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   293
text {* \noindent
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   294
  The initial current precedence for a thread can be anything, because there is no thread then. 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   295
  We simply assume every thread has precedence @{text "Prc 0 0"}.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   296
  *}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   297
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   298
abbreviation 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   299
  "initial_cprec \<equiv> \<lambda>_::thread. Prc 0 0"
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
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   302
text {* \noindent
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   303
  The following function @{text "schs"} is used to calculate the system's schedule state @{text "schs s"}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   304
  out of the current system state @{text "s"}. It is the central function to model Priority Inheritance:
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   305
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   306
fun schs :: "state \<Rightarrow> schedule_state"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   307
  where 
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   308
  -- {*
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   309
  \begin{minipage}{0.9\textwidth}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   310
    Setting the initial value of the @{text "schedule_state"} record (see the explanations above).
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   311
  \end{minipage}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   312
  *}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   313
  "schs [] = (| wq_fun = all_unlocked,  cprec_fun = initial_cprec |)" |
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   314
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   315
  -- {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   316
  \begin{minipage}{0.9\textwidth}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   317
  \begin{enumerate}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   318
  \item @{text "ps"} is the schedule state of last moment.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   319
  \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
   320
  \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
   321
  \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
   322
  \begin{enumerate}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   323
      \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
   324
            the end of @{text "cs"}'s waiting queue.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   325
      \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
   326
            @{text "th'"} must equal to @{text "thread"}, 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   327
            because @{text "thread"} is the one currently holding @{text "cs"}. 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   328
            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
   329
            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
   330
            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
   331
            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
   332
      \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
   333
  \end{enumerate}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   334
  \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
   335
        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
   336
        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
   337
  \end{enumerate}
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   338
  
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   339
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   340
  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
   341
  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
   342
  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
   343
  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
   344
  \end{minipage}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   345
     *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   346
   "schs (Create th prio # s) = 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   347
       (let wq = wq_fun (schs s) in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   348
          (|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
   349
|  "schs (Exit th # s) = 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   350
       (let wq = wq_fun (schs s) in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   351
          (|wq_fun = wq, cprec_fun = cpreced wq (Exit th # s)|))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   352
|  "schs (Set th prio # s) = 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   353
       (let wq = wq_fun (schs s) in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   354
          (|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
   355
   -- {*
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   356
   \begin{minipage}{0.9\textwidth}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   357
      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
   358
      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
   359
   \end{minipage}   
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   360
   *}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   361
|  "schs (P th cs # s) = 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   362
       (let wq = wq_fun (schs s) in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   363
        let new_wq = wq(cs := (wq cs @ [th])) in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   364
          (|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
   365
|  "schs (V th cs # s) = 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   366
       (let wq = wq_fun (schs s) in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   367
        let new_wq = wq(cs := release (wq cs)) in
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   368
          (|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
   369
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   370
lemma cpreced_initial: 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   371
  "cpreced (\<lambda> cs. []) [] = (\<lambda>_. (Prc 0 0))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   372
apply(simp add: cpreced_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
   373
apply(simp add: cs_dependants_def cs_RAG_def cs_waiting_def cs_holding_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   374
apply(simp add: preced_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   375
done
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 sch_old_def:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   378
  "schs (e#s) = (let ps = schs s in 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   379
                  let pwq = wq_fun ps in 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   380
                  let nwq = case e of
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   381
                             P th cs \<Rightarrow>  pwq(cs:=(pwq cs @ [th])) |
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   382
                             V th cs \<Rightarrow> let nq = case (pwq cs) of
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   383
                                                      [] \<Rightarrow> [] | 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   384
                                                      (_#qs) \<Rightarrow> (SOME q. distinct q \<and> set q = set qs)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   385
                                            in pwq(cs:=nq)                 |
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   386
                              _ \<Rightarrow> pwq
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   387
                  in let ncp = cpreced nwq (e#s) in 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   388
                     \<lparr>wq_fun = nwq, cprec_fun = ncp\<rparr>
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
apply(cases e)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   391
apply(simp_all)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   392
done
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   393
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   394
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   395
text {* 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   396
  \noindent
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   397
  The following @{text "wq"} is a shorthand for @{text "wq_fun"}. 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   398
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   399
definition wq :: "state \<Rightarrow> cs \<Rightarrow> thread list" 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   400
  where "wq s = wq_fun (schs s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   401
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   402
text {* \noindent 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   403
  The following @{text "cp"} is a shorthand for @{text "cprec_fun"}. 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   404
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   405
definition cp :: "state \<Rightarrow> thread \<Rightarrow> precedence"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   406
  where "cp s \<equiv> cprec_fun (schs s)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   407
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   408
text {* \noindent
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   409
  Functions @{text "holding"}, @{text "waiting"}, @{text "RAG"} and 
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   410
  @{text "dependants"} still have the 
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
   411
  same meaning, but redefined so that they no longer RAG on the 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   412
  fictitious {\em waiting queue function}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   413
  @{text "wq"}, but on system state @{text "s"}.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   414
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   415
defs (overloaded) 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   416
  s_holding_abv: 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   417
  "holding (s::state) \<equiv> holding (wq_fun (schs s))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   418
  s_waiting_abv: 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   419
  "waiting (s::state) \<equiv> waiting (wq_fun (schs s))"
35
92f61f6a0fe7 added a bit more text to the paper and separated a theory about Max
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 33
diff changeset
   420
  s_RAG_abv: 
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
   421
  "RAG (s::state) \<equiv> RAG (wq_fun (schs s))"
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   422
  s_dependants_abv: 
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   423
  "dependants (s::state) \<equiv> dependants (wq_fun (schs s))"
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   424
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   425
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   426
text {* 
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   427
  The following lemma can be proved easily, and the meaning is obvious.
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   428
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   429
lemma
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   430
  s_holding_def: 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   431
  "holding (s::state) th cs \<equiv> (th \<in> set (wq_fun (schs s) cs) \<and> th = hd (wq_fun (schs s) cs))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   432
  by (auto simp:s_holding_abv wq_def cs_holding_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   433
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   434
lemma s_waiting_def: 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   435
  "waiting (s::state) th cs \<equiv> (th \<in> set (wq_fun (schs s) cs) \<and> th \<noteq> hd (wq_fun (schs s) cs))"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   436
  by (auto simp:s_waiting_abv wq_def cs_waiting_def)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   437
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
   438
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
   439
  "RAG (s::state) =
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   440
    {(Th th, Cs cs) | th cs. waiting (wq s) th cs} \<union> {(Cs cs, Th th) | cs th. holding (wq s) th cs}"
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
   441
  by (auto simp:s_RAG_abv wq_def cs_RAG_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   442
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   443
lemma
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   444
  s_dependants_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
   445
  "dependants (s::state) th \<equiv> {th' . (Th th', Th th) \<in> (RAG (wq s))^+}"
32
e861aff29655 made some modifications.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 17
diff changeset
   446
  by (auto simp:s_dependants_abv wq_def cs_dependants_def)
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   447
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   448
text {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   449
  The following function @{text "readys"} calculates the set of ready threads. A thread is {\em ready} 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   450
  for running if it is a live thread and it is not waiting for any critical resource.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   451
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   452
definition readys :: "state \<Rightarrow> thread set"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   453
  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
   454
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   455
text {* \noindent
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   456
  The following function @{text "runing"} calculates the set of running thread, which is the ready 
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   457
  thread with the highest precedence.  
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   458
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   459
definition runing :: "state \<Rightarrow> thread set"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   460
  where "runing s \<equiv> {th . th \<in> readys s \<and> cp s th = Max ((cp s) ` (readys s))}"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   461
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   462
text {* \noindent
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   463
  Notice that the definition of @{text "running"} reflects the preemptive scheduling strategy,  
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   464
  because, if the @{text "running"}-thread (the one in @{text "runing"} set) 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   465
  lowered its precedence by resetting its own priority to a lower
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   466
  one, it will lose its status of being the max in @{text "ready"}-set and be superseded.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   467
*}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   468
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   469
text {* \noindent
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   470
  The following function @{text "holdents s th"} returns the set of resources held by thread 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   471
  @{text "th"} in state @{text "s"}.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   472
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   473
definition holdents :: "state \<Rightarrow> thread \<Rightarrow> cs set"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   474
  where "holdents s th \<equiv> {cs . holding s th cs}"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   475
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   476
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
   477
  "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
   478
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
   479
unfolding s_RAG_def
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   480
unfolding s_holding_abv
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   481
unfolding wq_def
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   482
by (simp)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   483
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   484
text {* \noindent
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   485
  Observation @{text "cntCS s th"} returns the number of resources held by thread @{text "th"} in
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   486
  state @{text "s"}:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   487
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   488
definition cntCS :: "state \<Rightarrow> thread \<Rightarrow> nat"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   489
  where "cntCS s th = card (holdents s th)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   490
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   491
text {* \noindent
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   492
  According to the convention of Paulson's inductive method,
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   493
  the decision made by a protocol that event @{text "e"} is eligible to happen next under state @{text "s"} 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   494
  is expressed as @{text "step s e"}. The predicate @{text "step"} is inductively defined as 
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   495
  follows (notice how the decision is based on the {\em observation function}s 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   496
  defined above, and also notice how a complicated protocol is modeled by a few simple 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   497
  observations, and how such a kind of simplicity gives rise to improved trust on
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   498
  faithfulness):
0
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
inductive step :: "state \<Rightarrow> event \<Rightarrow> bool"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   501
  where
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   502
  -- {* 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   503
  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
   504
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   505
  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
   506
  -- {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   507
  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
   508
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   509
  thread_exit: "\<lbrakk>thread \<in> runing s; holdents s thread = {}\<rbrakk> \<Longrightarrow> step s (Exit thread)" |
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   510
  -- {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   511
  \begin{minipage}{0.9\textwidth}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   512
  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
   513
  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
   514
  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
   515
  carefully programmed so that deadlock can not happen:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   516
  \end{minipage}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   517
  *}
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
   518
  thread_P: "\<lbrakk>thread \<in> runing 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
   519
                                                                step s (P thread cs)" |
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   520
  -- {*
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   521
  \begin{minipage}{0.9\textwidth}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   522
  A thread can release a critical resource @{text "cs"} 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   523
  if it is running and holding that resource:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   524
  \end{minipage}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   525
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   526
  thread_V: "\<lbrakk>thread \<in> runing s;  holding s thread cs\<rbrakk> \<Longrightarrow> step s (V thread cs)" |
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   527
  -- {*
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   528
  \begin{minipage}{0.9\textwidth}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   529
  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
   530
  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
   531
  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
   532
  function, 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   533
  \end{minipage}
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   534
  *}  
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   535
  thread_set: "\<lbrakk>thread \<in> runing s\<rbrakk> \<Longrightarrow> step s (Set thread prio)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   536
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   537
text {*
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   538
  In Paulson's inductive method, every protocol is defined by such a @{text "step"}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   539
  predicate. For instance, the predicate @{text "step"} given above 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   540
  defines the PIP protocol. So, it can also be called "PIP".
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   541
*}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   542
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   543
abbreviation
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   544
  "PIP \<equiv> step"
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   545
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   546
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   547
text {* \noindent
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   548
  For any protocol defined by a @{text "step"} predicate, 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   549
  the fact that @{text "s"} is a legal state in 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   550
  the protocol is expressed as: @{text "vt step s"}, where
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   551
  the predicate @{text "vt"} can be defined as the following:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   552
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   553
inductive vt :: "state \<Rightarrow> bool"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   554
  where
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   555
  -- {* Empty list @{text "[]"} is a legal state in any protocol:*}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   556
  vt_nil[intro]: "vt []" |
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   557
  -- {* 
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   558
  \begin{minipage}{0.9\textwidth}
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   559
  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
   560
  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
   561
  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
   562
  happening of @{text "e"}:
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   563
  \end{minipage}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   564
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   565
  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
   566
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   567
text {*  \noindent
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   568
  It is easy to see that the definition of @{text "vt"} is generic. It can be applied to 
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   569
  any specific protocol specified by a @{text "step"}-predicate to get the set of
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   570
  legal states of that particular protocol.
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   571
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   572
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   573
text {* 
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   574
  The following are two very basic properties of @{text "vt"}.
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   575
*}
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   576
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   577
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
   578
  by(ind_cases "vt (e#s)", simp)
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   579
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   580
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
   581
  by(ind_cases "vt (e#s)", simp)
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   582
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   583
text {* \noindent
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   584
  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
   585
  critical resource and thread respectively out of RAG nodes.
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
fun the_cs :: "node \<Rightarrow> cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   588
  where "the_cs (Cs cs) = cs"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   589
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   590
fun the_th :: "node \<Rightarrow> thread"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   591
  where "the_th (Th th) = th"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   592
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   593
text {* \noindent
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   594
  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
   595
  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
   596
  @{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
   597
  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
   598
  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
   599
  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
   600
  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
   601
  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
   602
  @{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
   603
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   604
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
   605
  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
   606
                                     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
   607
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   608
text {* \noindent
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   609
  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
   610
  in list @{text "l"}:
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   611
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   612
definition count :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> nat"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   613
  where "count Q l = length (filter Q l)"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   614
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   615
text {* \noindent
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   616
  The following observation @{text "cntP s"} returns the number of operation @{text "P"} happened 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   617
  before reaching state @{text "s"}.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   618
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   619
definition cntP :: "state \<Rightarrow> thread \<Rightarrow> nat"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   620
  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
   621
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   622
text {* \noindent
53
8142e80f5d58 Finished comments on PrioGDef.thy
xingyuan zhang <xingyuanzhang@126.com>
parents: 49
diff changeset
   623
  The following observation @{text "cntV s"} returns the number of operation @{text "V"} happened 
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   624
  before reaching state @{text "s"}.
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   625
  *}
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   626
definition cntV :: "state \<Rightarrow> thread \<Rightarrow> nat"
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   627
  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
   628
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   629
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
   630
       difference is the order of arguemts. *}
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   631
definition "the_preced s th = preced th s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   632
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   633
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
   634
fun the_thread :: "node \<Rightarrow> thread" where
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   635
   "the_thread (Th th) = th"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   636
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   637
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
   638
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
   639
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   640
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
   641
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
   642
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   643
text {* 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   644
  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
   645
  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
   646
  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
   647
  names @{term "wRAG"} and @{term "hRAG"}.
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   648
 *}
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   649
definition "tRAG s = wRAG s O hRAG s"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   650
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   651
text {* The following lemma splits @{term "RAG"} graph into the above two sub-graphs. *}
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   652
lemma RAG_split: "RAG s = (wRAG s \<union> hRAG s)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   653
  by (unfold s_RAG_abv wRAG_def hRAG_def s_waiting_abv 
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   654
             s_holding_abv cs_RAG_def, auto)
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   655
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   656
definition "cp_gen s x =
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   657
                  Max ((the_preced s \<circ> the_thread) ` subtree (tRAG s) x)"
633b1fc8631b Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
zhangx
parents: 64
diff changeset
   658
0
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   659
(*<*)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   660
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   661
end
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   662
(*>*)
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents:
diff changeset
   663