prio/Attic/Prio.thy
author urbanc
Tue, 28 Feb 2012 13:13:32 +0000
changeset 336 f9e0d3274c14
parent 282 a3b4eed091d2
permissions -rw-r--r--
fixed typo
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     1
theory Prio
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     2
imports Precedence_ord Moment Lsp Happen_within
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     3
begin
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     4
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     5
type_synonym thread = nat
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     6
type_synonym priority = nat
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     7
type_synonym cs = nat
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     8
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     9
datatype event = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    10
  Create thread priority |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    11
  Exit thread |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    12
  P thread cs |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    13
  V thread cs |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    14
  Set thread priority
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    15
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    16
datatype node = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    17
   Th "thread" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    18
   Cs "cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    19
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    20
type_synonym state = "event list"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    21
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    22
fun threads :: "state \<Rightarrow> thread set"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    23
where 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    24
  "threads [] = {}" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    25
  "threads (Create thread prio#s) = {thread} \<union> threads s" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    26
  "threads (Exit thread # s) = (threads s) - {thread}" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    27
  "threads (e#s) = threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    28
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    29
fun original_priority :: "thread \<Rightarrow> state \<Rightarrow> nat"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    30
where
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    31
  "original_priority thread [] = 0" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    32
  "original_priority thread (Create thread' prio#s) = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    33
     (if thread' = thread then prio else original_priority thread s)" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    34
  "original_priority thread (Set thread' prio#s) = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    35
     (if thread' = thread then prio else original_priority thread s)" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    36
  "original_priority thread (e#s) = original_priority thread s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    37
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    38
fun birthtime :: "thread \<Rightarrow> state \<Rightarrow> nat"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    39
where
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    40
  "birthtime thread [] = 0" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    41
  "birthtime thread ((Create thread' prio)#s) = (if (thread = thread') then length s 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    42
                                                  else birthtime thread s)" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    43
  "birthtime thread ((Set thread' prio)#s) = (if (thread = thread') then length s 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    44
                                                  else birthtime thread s)" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    45
  "birthtime thread (e#s) = birthtime thread s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    46
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    47
definition preced :: "thread \<Rightarrow> state \<Rightarrow> precedence"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    48
  where "preced thread s = Prc (original_priority thread s) (birthtime thread s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    49
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    50
consts holding :: "'b \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> bool"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    51
       waiting :: "'b \<Rightarrow> thread \<Rightarrow> cs \<Rightarrow> bool"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    52
       depend :: "'b \<Rightarrow> (node \<times> node) set"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    53
       dependents :: "'b \<Rightarrow> thread \<Rightarrow> thread set"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    54
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    55
defs (overloaded) cs_holding_def: "holding wq thread cs == (thread \<in> set (wq cs) \<and> thread = hd (wq cs))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    56
                  cs_waiting_def: "waiting wq thread cs == (thread \<in> set (wq cs) \<and> thread \<noteq> hd (wq cs))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    57
                  cs_depend_def: "depend (wq::cs \<Rightarrow> thread list) == {(Th t, Cs c) | t c. waiting wq t c} \<union> 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    58
                                               {(Cs c, Th t) | c t. holding wq t c}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    59
                  cs_dependents_def: "dependents (wq::cs \<Rightarrow> thread list) th == {th' . (Th th', Th th) \<in> (depend wq)^+}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    60
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    61
record schedule_state = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    62
    waiting_queue :: "cs \<Rightarrow> thread list"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    63
    cur_preced :: "thread \<Rightarrow> precedence"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    64
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    65
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    66
definition cpreced :: "state \<Rightarrow> (cs \<Rightarrow> thread list) \<Rightarrow> thread \<Rightarrow> precedence"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    67
where "cpreced s wq = (\<lambda> th. Max ((\<lambda> th. preced th s) ` ({th} \<union> dependents wq th)))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    68
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    69
fun schs :: "state \<Rightarrow> schedule_state"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    70
where
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    71
   "schs [] = \<lparr>waiting_queue = \<lambda> cs. [], 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    72
               cur_preced = cpreced [] (\<lambda> cs. [])\<rparr>" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    73
   "schs (e#s) = (let ps = schs s in
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    74
                  let pwq = waiting_queue ps in
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    75
                  let pcp = cur_preced ps in
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    76
                  let nwq = case e of
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    77
                             P thread cs \<Rightarrow>  pwq(cs:=(pwq cs @ [thread])) |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    78
                             V thread cs \<Rightarrow> let nq = case (pwq cs) of
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    79
                                                      [] \<Rightarrow> [] | 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    80
                                                      (th#pq) \<Rightarrow> case (lsp pcp pq) of
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    81
                                                                   (l, [], r) \<Rightarrow> []
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    82
                                                                 | (l, m#ms, r) \<Rightarrow> m#(l@ms@r)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    83
                                            in pwq(cs:=nq)                 |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    84
                              _ \<Rightarrow> pwq
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    85
                  in let ncp = cpreced (e#s) nwq in 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    86
                     \<lparr>waiting_queue = nwq, cur_preced = ncp\<rparr>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    87
                 )"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    88
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    89
definition wq :: "state \<Rightarrow> cs \<Rightarrow> thread list" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    90
where "wq s == waiting_queue (schs s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    91
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    92
definition cp :: "state \<Rightarrow> thread \<Rightarrow> precedence"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    93
where "cp s = cur_preced (schs s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    94
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    95
defs (overloaded) s_holding_def: "holding (s::state) thread cs == (thread \<in> set (wq s cs) \<and> thread = hd (wq s cs))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    96
                  s_waiting_def: "waiting (s::state) thread cs == (thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    97
                  s_depend_def: "depend (s::state) == {(Th t, Cs c) | t c. waiting (wq s) t c} \<union> 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    98
                                               {(Cs c, Th t) | c t. holding (wq s) t c}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    99
                  s_dependents_def: "dependents (s::state) th == {th' . (Th th', Th th) \<in> (depend (wq s))^+}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   100
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   101
definition readys :: "state \<Rightarrow> thread set"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   102
where
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   103
  "readys s = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   104
     {thread . thread \<in> threads s \<and> (\<forall> cs. \<not> waiting s thread cs)}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   105
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   106
definition runing :: "state \<Rightarrow> thread set"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   107
where "runing s = {th . th \<in> readys s \<and> cp s th = Max ((cp s) ` (readys s))}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   108
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   109
definition holdents :: "state \<Rightarrow> thread \<Rightarrow> cs set"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   110
  where "holdents s th = {cs . (Cs cs, Th th) \<in> depend s}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   111
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   112
inductive step :: "state \<Rightarrow> event \<Rightarrow> bool"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   113
where
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   114
  thread_create: "\<lbrakk>prio \<le> max_prio; thread \<notin> threads s\<rbrakk> \<Longrightarrow> step s (Create thread prio)" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   115
  thread_exit: "\<lbrakk>thread \<in> runing s; holdents s thread = {}\<rbrakk> \<Longrightarrow> step s (Exit thread)" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   116
  thread_P: "\<lbrakk>thread \<in> runing s;  (Cs cs, Th thread)  \<notin> (depend s)^+\<rbrakk> \<Longrightarrow> step s (P thread cs)" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   117
  thread_V: "\<lbrakk>thread \<in> runing s;  holding s thread cs\<rbrakk> \<Longrightarrow> step s (V thread cs)" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   118
  thread_set: "\<lbrakk>thread \<in> runing s\<rbrakk> \<Longrightarrow> step s (Set thread prio)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   119
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   120
inductive vt :: "(state \<Rightarrow> event \<Rightarrow> bool) \<Rightarrow> state \<Rightarrow> bool"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   121
 for cs
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   122
where
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   123
  vt_nil[intro]: "vt cs []" |
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   124
  vt_cons[intro]: "\<lbrakk>vt cs s; cs s e\<rbrakk> \<Longrightarrow> vt cs (e#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   125
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   126
lemma runing_ready: "runing s \<subseteq> readys s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   127
  by (auto simp only:runing_def readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   128
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   129
lemma wq_v_eq_nil: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   130
  fixes s cs thread rest
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   131
  assumes eq_wq: "wq s cs = thread # rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   132
  and eq_lsp: "lsp (cp s) rest = (l, [], r)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   133
  shows "wq (V thread cs#s) cs = []"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   134
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   135
  from prems show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   136
    by (auto simp:wq_def Let_def cp_def split:list.splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   137
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   138
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   139
lemma wq_v_eq: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   140
  fixes s cs thread rest
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   141
  assumes eq_wq: "wq s cs = thread # rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   142
  and eq_lsp: "lsp (cp s) rest = (l, [th'], r)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   143
  shows "wq (V thread cs#s) cs = th'#l@r"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   144
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   145
  from prems show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   146
    by (auto simp:wq_def Let_def cp_def split:list.splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   147
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   148
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   149
lemma wq_v_neq:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   150
   "cs \<noteq> cs' \<Longrightarrow> wq (V thread cs#s) cs' = wq s cs'"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   151
  by (auto simp:wq_def Let_def cp_def split:list.splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   152
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   153
lemma wq_distinct: "vt step s \<Longrightarrow> distinct (wq s cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   154
proof(erule_tac vt.induct, simp add:wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   155
  fix s e
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   156
  assume h1: "step s e"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   157
  and h2: "distinct (wq s cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   158
  thus "distinct (wq (e # s) cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   159
  proof(induct rule:step.induct, auto simp: wq_def Let_def split:list.splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   160
    fix thread s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   161
    assume h1: "(Cs cs, Th thread) \<notin> (depend s)\<^sup>+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   162
      and h2: "thread \<in> set (waiting_queue (schs s) cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   163
      and h3: "thread \<in> runing s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   164
    show "False" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   165
    proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   166
      from h3 have "\<And> cs. thread \<in>  set (waiting_queue (schs s) cs) \<Longrightarrow> 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   167
                             thread = hd ((waiting_queue (schs s) cs))" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   168
        by (simp add:runing_def readys_def s_waiting_def wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   169
      from this [OF h2] have "thread = hd (waiting_queue (schs s) cs)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   170
      with h2
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   171
      have "(Cs cs, Th thread) \<in> (depend s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   172
        by (simp add:s_depend_def s_holding_def wq_def cs_holding_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   173
      with h1 show False by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   174
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   175
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   176
    fix thread s a list
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   177
    assume h1: "thread \<in> runing s" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   178
      and h2: "holding s thread cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   179
      and h3: "waiting_queue (schs s) cs = a # list"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   180
      and h4: "a \<notin> set list"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   181
      and h5: "distinct list"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   182
    thus "distinct
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   183
           ((\<lambda>(l, a, r). case a of [] \<Rightarrow> [] | m # ms \<Rightarrow> m # l @ ms @ r)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   184
             (lsp (cur_preced (schs s)) list))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   185
    apply (cases "(lsp (cur_preced (schs s)) list)", simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   186
    apply (case_tac b, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   187
    by (drule_tac lsp_set_eq, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   188
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   189
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   190
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   191
lemma block_pre: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   192
  fixes thread cs s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   193
  assumes s_ni: "thread \<notin>  set (wq s cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   194
  and s_i: "thread \<in> set (wq (e#s) cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   195
  shows "e = P thread cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   196
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   197
  have ee: "\<And> x y. \<lbrakk>x = y\<rbrakk> \<Longrightarrow> set x = set y"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   198
    by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   199
  from s_ni s_i show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   200
  proof (cases e, auto split:if_splits simp add:Let_def wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   201
    fix uu uub uuc uud uue
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   202
    assume h: "(uuc, thread # uu, uub) = lsp (cur_preced (schs s)) uud"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   203
      and h1 [symmetric]: "uue # uud = waiting_queue (schs s) cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   204
      and h2: "thread \<notin> set (waiting_queue (schs s) cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   205
    from lsp_set [OF h] have "set (uuc @ (thread # uu) @ uub) = set uud" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   206
    hence "thread \<in> set uud" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   207
    with h1 have "thread \<in> set (waiting_queue (schs s) cs)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   208
    with h2 show False by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   209
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   210
    fix uu uua uub uuc uud uue
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   211
    assume h1: "thread \<notin> set (waiting_queue (schs s) cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   212
      and h2: "uue # uud = waiting_queue (schs s) cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   213
      and h3: "(uuc, uua # uu, uub) = lsp (cur_preced (schs s)) uud"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   214
      and h4: "thread \<in> set uuc"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   215
    from lsp_set [OF h3] have "set (uuc @ (uua # uu) @ uub) = set uud" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   216
    with h4 have "thread \<in> set uud" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   217
    with h2 have "thread \<in> set (waiting_queue (schs s) cs)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   218
      apply (drule_tac ee) by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   219
    with h1 show "False" by fast
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   220
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   221
    fix uu uua uub uuc uud uue
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   222
    assume h1: "thread \<notin> set (waiting_queue (schs s) cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   223
      and h2: "uue # uud = waiting_queue (schs s) cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   224
      and h3: "(uuc, uua # uu, uub) = lsp (cur_preced (schs s)) uud"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   225
      and h4: "thread \<in> set uu"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   226
    from lsp_set [OF h3] have "set (uuc @ (uua # uu) @ uub) = set uud" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   227
    with h4 have "thread \<in> set uud" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   228
    with h2 have "thread \<in> set (waiting_queue (schs s) cs)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   229
      apply (drule_tac ee) by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   230
    with h1 show "False" by fast
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   231
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   232
    fix uu uua uub uuc uud uue
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   233
    assume h1: "thread \<notin> set (waiting_queue (schs s) cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   234
      and h2: "uue # uud = waiting_queue (schs s) cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   235
      and h3: "(uuc, uua # uu, uub) = lsp (cur_preced (schs s)) uud"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   236
      and h4: "thread \<in> set uub"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   237
    from lsp_set [OF h3] have "set (uuc @ (uua # uu) @ uub) = set uud" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   238
    with h4 have "thread \<in> set uud" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   239
    with h2 have "thread \<in> set (waiting_queue (schs s) cs)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   240
      apply (drule_tac ee) by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   241
    with h1 show "False" by fast
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   242
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   243
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   244
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   245
lemma p_pre: "\<lbrakk>vt step ((P thread cs)#s)\<rbrakk> \<Longrightarrow> 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   246
  thread \<in> runing s \<and> (Cs cs, Th thread)  \<notin> (depend s)^+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   247
apply (ind_cases "vt step ((P thread cs)#s)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   248
apply (ind_cases "step s (P thread cs)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   249
by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   250
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   251
lemma abs1:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   252
  fixes e es
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   253
  assumes ein: "e \<in> set es"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   254
  and neq: "hd es \<noteq> hd (es @ [x])"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   255
  shows "False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   256
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   257
  from ein have "es \<noteq> []" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   258
  then obtain e ess where "es = e # ess" by (cases es, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   259
  with neq show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   260
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   261
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   262
lemma q_head: "Q (hd es) \<Longrightarrow> hd es = hd [th\<leftarrow>es . Q th]"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   263
  by (cases es, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   264
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   265
inductive_cases evt_cons: "vt cs (a#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   266
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   267
lemma abs2:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   268
  assumes vt: "vt step (e#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   269
  and inq: "thread \<in> set (wq s cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   270
  and nh: "thread = hd (wq s cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   271
  and qt: "thread \<noteq> hd (wq (e#s) cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   272
  and inq': "thread \<in> set (wq (e#s) cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   273
  shows "False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   274
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   275
  have ee: "\<And> uuc thread uu uub s list. (uuc, thread # uu, uub) = lsp (cur_preced (schs s)) list \<Longrightarrow> 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   276
                 lsp (cur_preced (schs s)) list = (uuc, thread # uu, uub) 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   277
               " by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   278
  from prems show "False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   279
    apply (cases e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   280
    apply ((simp split:if_splits add:Let_def wq_def)[1])+
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   281
    apply (insert abs1, fast)[1] 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   282
    apply ((simp split:if_splits add:Let_def)[1])+
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   283
    apply (simp split:if_splits list.splits add:Let_def wq_def) 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   284
    apply (auto dest!:ee)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   285
    apply (drule_tac lsp_set_eq, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   286
    apply (subgoal_tac "distinct (waiting_queue (schs s) cs)", simp, fold wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   287
    apply (rule_tac wq_distinct, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   288
    apply (erule_tac evt_cons, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   289
    apply (drule_tac lsp_set_eq, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   290
    apply (subgoal_tac "distinct (wq s cs)", simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   291
    apply (rule_tac wq_distinct, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   292
    apply (erule_tac evt_cons, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   293
    apply (drule_tac lsp_set_eq, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   294
    apply (subgoal_tac "distinct (wq s cs)", simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   295
    apply (rule_tac wq_distinct, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   296
    apply (erule_tac evt_cons, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   297
    apply (auto simp:wq_def Let_def split:if_splits prod.splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   298
    done
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   299
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   300
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   301
lemma vt_moment: "\<And> t. \<lbrakk>vt cs s; t \<le> length s\<rbrakk> \<Longrightarrow> vt cs (moment t s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   302
proof(induct s, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   303
  fix a s t
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   304
  assume h: "\<And>t.\<lbrakk>vt cs s; t \<le> length s\<rbrakk> \<Longrightarrow> vt cs (moment t s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   305
    and vt_a: "vt cs (a # s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   306
    and le_t: "t \<le> length (a # s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   307
  show "vt cs (moment t (a # s))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   308
  proof(cases "t = length (a#s)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   309
    case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   310
    from True have "moment t (a#s) = a#s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   311
    with vt_a show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   312
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   313
    case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   314
    with le_t have le_t1: "t \<le> length s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   315
    from vt_a have "vt cs s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   316
      by (erule_tac evt_cons, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   317
    from h [OF this le_t1] have "vt cs (moment t s)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   318
    moreover have "moment t (a#s) = moment t s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   319
    proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   320
      from moment_app [OF le_t1, of "[a]"] 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   321
      show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   322
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   323
    ultimately show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   324
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   325
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   326
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   327
(* Wrong:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   328
    lemma \<lbrakk>thread \<in> set (waiting_queue cs1 s); thread \<in> set (waiting_queue cs2 s)\<rbrakk> \<Longrightarrow> cs1 = cs2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   329
*)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   330
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   331
lemma waiting_unique_pre:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   332
  fixes cs1 cs2 s thread
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   333
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   334
  and h11: "thread \<in> set (wq s cs1)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   335
  and h12: "thread \<noteq> hd (wq s cs1)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   336
  assumes h21: "thread \<in> set (wq s cs2)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   337
  and h22: "thread \<noteq> hd (wq s cs2)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   338
  and neq12: "cs1 \<noteq> cs2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   339
  shows "False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   340
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   341
  let "?Q cs s" = "thread \<in> set (wq s cs) \<and> thread \<noteq> hd (wq s cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   342
  from h11 and h12 have q1: "?Q cs1 s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   343
  from h21 and h22 have q2: "?Q cs2 s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   344
  have nq1: "\<not> ?Q cs1 []" by (simp add:wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   345
  have nq2: "\<not> ?Q cs2 []" by (simp add:wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   346
  from p_split [of "?Q cs1", OF q1 nq1]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   347
  obtain t1 where lt1: "t1 < length s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   348
    and np1: "\<not>(thread \<in> set (wq (moment t1 s) cs1) \<and>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   349
        thread \<noteq> hd (wq (moment t1 s) cs1))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   350
    and nn1: "(\<forall>i'>t1. thread \<in> set (wq (moment i' s) cs1) \<and>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   351
             thread \<noteq> hd (wq (moment i' s) cs1))" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   352
  from p_split [of "?Q cs2", OF q2 nq2]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   353
  obtain t2 where lt2: "t2 < length s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   354
    and np2: "\<not>(thread \<in> set (wq (moment t2 s) cs2) \<and>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   355
        thread \<noteq> hd (wq (moment t2 s) cs2))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   356
    and nn2: "(\<forall>i'>t2. thread \<in> set (wq (moment i' s) cs2) \<and>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   357
             thread \<noteq> hd (wq (moment i' s) cs2))" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   358
  show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   359
  proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   360
    { 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   361
      assume lt12: "t1 < t2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   362
      let ?t3 = "Suc t2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   363
      from lt2 have le_t3: "?t3 \<le> length s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   364
      from moment_plus [OF this] 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   365
      obtain e where eq_m: "moment ?t3 s = e#moment t2 s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   366
      have "t2 < ?t3" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   367
      from nn2 [rule_format, OF this] and eq_m
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   368
      have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   369
        h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   370
      have vt_e: "vt step (e#moment t2 s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   371
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   372
        from vt_moment [OF vt le_t3]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   373
        have "vt step (moment ?t3 s)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   374
        with eq_m show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   375
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   376
      have ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   377
      proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   378
        case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   379
        from True and np2 have eq_th: "thread = hd (wq (moment t2 s) cs2)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   380
          by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   381
        from abs2 [OF vt_e True eq_th h2 h1]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   382
        show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   383
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   384
        case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   385
        from block_pre [OF False h1]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   386
        have "e = P thread cs2" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   387
        with vt_e have "vt step ((P thread cs2)# moment t2 s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   388
        from p_pre [OF this] have "thread \<in> runing (moment t2 s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   389
        with runing_ready have "thread \<in> readys (moment t2 s)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   390
        with nn1 [rule_format, OF lt12]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   391
        show ?thesis  by (simp add:readys_def s_waiting_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   392
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   393
    } moreover {
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   394
      assume lt12: "t2 < t1"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   395
      let ?t3 = "Suc t1"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   396
      from lt1 have le_t3: "?t3 \<le> length s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   397
      from moment_plus [OF this] 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   398
      obtain e where eq_m: "moment ?t3 s = e#moment t1 s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   399
      have lt_t3: "t1 < ?t3" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   400
      from nn1 [rule_format, OF this] and eq_m
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   401
      have h1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   402
        h2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   403
      have vt_e: "vt step (e#moment t1 s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   404
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   405
        from vt_moment [OF vt le_t3]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   406
        have "vt step (moment ?t3 s)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   407
        with eq_m show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   408
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   409
      have ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   410
      proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   411
        case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   412
        from True and np1 have eq_th: "thread = hd (wq (moment t1 s) cs1)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   413
          by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   414
        from abs2 [OF vt_e True eq_th h2 h1]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   415
        show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   416
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   417
        case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   418
        from block_pre [OF False h1]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   419
        have "e = P thread cs1" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   420
        with vt_e have "vt step ((P thread cs1)# moment t1 s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   421
        from p_pre [OF this] have "thread \<in> runing (moment t1 s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   422
        with runing_ready have "thread \<in> readys (moment t1 s)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   423
        with nn2 [rule_format, OF lt12]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   424
        show ?thesis  by (simp add:readys_def s_waiting_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   425
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   426
    } moreover {
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   427
      assume eqt12: "t1 = t2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   428
      let ?t3 = "Suc t1"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   429
      from lt1 have le_t3: "?t3 \<le> length s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   430
      from moment_plus [OF this] 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   431
      obtain e where eq_m: "moment ?t3 s = e#moment t1 s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   432
      have lt_t3: "t1 < ?t3" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   433
      from nn1 [rule_format, OF this] and eq_m
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   434
      have h1: "thread \<in> set (wq (e#moment t1 s) cs1)" and
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   435
        h2: "thread \<noteq> hd (wq (e#moment t1 s) cs1)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   436
      have vt_e: "vt step (e#moment t1 s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   437
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   438
        from vt_moment [OF vt le_t3]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   439
        have "vt step (moment ?t3 s)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   440
        with eq_m show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   441
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   442
      have ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   443
      proof(cases "thread \<in> set (wq (moment t1 s) cs1)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   444
        case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   445
        from True and np1 have eq_th: "thread = hd (wq (moment t1 s) cs1)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   446
          by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   447
        from abs2 [OF vt_e True eq_th h2 h1]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   448
        show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   449
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   450
        case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   451
        from block_pre [OF False h1]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   452
        have eq_e1: "e = P thread cs1" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   453
        have lt_t3: "t1 < ?t3" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   454
        with eqt12 have "t2 < ?t3" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   455
        from nn2 [rule_format, OF this] and eq_m and eqt12
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   456
        have h1: "thread \<in> set (wq (e#moment t2 s) cs2)" and
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   457
          h2: "thread \<noteq> hd (wq (e#moment t2 s) cs2)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   458
        show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   459
        proof(cases "thread \<in> set (wq (moment t2 s) cs2)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   460
          case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   461
          from True and np2 have eq_th: "thread = hd (wq (moment t2 s) cs2)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   462
            by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   463
          from vt_e and eqt12 have "vt step (e#moment t2 s)" by simp 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   464
          from abs2 [OF this True eq_th h2 h1]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   465
          show ?thesis .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   466
        next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   467
          case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   468
          from block_pre [OF False h1]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   469
          have "e = P thread cs2" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   470
          with eq_e1 neq12 show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   471
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   472
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   473
    } ultimately show ?thesis by arith
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   474
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   475
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   476
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   477
lemma waiting_unique:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   478
  assumes "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   479
  and "waiting s th cs1"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   480
  and "waiting s th cs2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   481
  shows "cs1 = cs2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   482
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   483
  from waiting_unique_pre and prems
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   484
  show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   485
    by (auto simp add:s_waiting_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   486
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   487
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   488
lemma holded_unique:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   489
  assumes "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   490
  and "holding s th1 cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   491
  and "holding s th2 cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   492
  shows "th1 = th2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   493
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   494
  from prems show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   495
    unfolding s_holding_def
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   496
    by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   497
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   498
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   499
lemma birthtime_lt: "th \<in> threads s \<Longrightarrow> birthtime th s < length s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   500
  apply (induct s, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   501
  by (case_tac a, auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   502
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   503
lemma birthtime_unique: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   504
  "\<lbrakk>birthtime th1 s = birthtime th2 s; th1 \<in> threads s; th2 \<in> threads s\<rbrakk>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   505
          \<Longrightarrow> th1 = th2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   506
  apply (induct s, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   507
  by (case_tac a, auto split:if_splits dest:birthtime_lt)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   508
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   509
lemma preced_unique : 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   510
  assumes pcd_eq: "preced th1 s = preced th2 s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   511
  and th_in1: "th1 \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   512
  and th_in2: " th2 \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   513
  shows "th1 = th2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   514
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   515
  from pcd_eq have "birthtime th1 s = birthtime th2 s" by (simp add:preced_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   516
  from birthtime_unique [OF this th_in1 th_in2]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   517
  show ?thesis .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   518
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   519
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   520
lemma preced_linorder: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   521
  assumes neq_12: "th1 \<noteq> th2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   522
  and th_in1: "th1 \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   523
  and th_in2: " th2 \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   524
  shows "preced th1 s < preced th2 s \<or> preced th1 s > preced th2 s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   525
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   526
  from preced_unique [OF _ th_in1 th_in2] and neq_12 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   527
  have "preced th1 s \<noteq> preced th2 s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   528
  thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   529
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   530
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   531
lemma unique_minus:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   532
  fixes x y z r
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   533
  assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   534
  and xy: "(x, y) \<in> r"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   535
  and xz: "(x, z) \<in> r^+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   536
  and neq: "y \<noteq> z"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   537
  shows "(y, z) \<in> r^+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   538
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   539
 from xz and neq show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   540
 proof(induct)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   541
   case (base ya)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   542
   have "(x, ya) \<in> r" by fact
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   543
   from unique [OF xy this] have "y = ya" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   544
   with base show ?case by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   545
 next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   546
   case (step ya z)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   547
   show ?case
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   548
   proof(cases "y = ya")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   549
     case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   550
     from step True show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   551
   next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   552
     case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   553
     from step False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   554
     show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   555
   qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   556
 qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   557
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   558
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   559
lemma unique_base:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   560
  fixes r x y z
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   561
  assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   562
  and xy: "(x, y) \<in> r"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   563
  and xz: "(x, z) \<in> r^+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   564
  and neq_yz: "y \<noteq> z"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   565
  shows "(y, z) \<in> r^+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   566
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   567
  from xz neq_yz show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   568
  proof(induct)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   569
    case (base ya)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   570
    from xy unique base show ?case by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   571
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   572
    case (step ya z)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   573
    show ?case
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   574
    proof(cases "y = ya")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   575
      case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   576
      from True step show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   577
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   578
      case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   579
      from False step 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   580
      have "(y, ya) \<in> r\<^sup>+" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   581
      with step show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   582
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   583
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   584
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   585
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   586
lemma unique_chain:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   587
  fixes r x y z
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   588
  assumes unique: "\<And> a b c. \<lbrakk>(a, b) \<in> r; (a, c) \<in> r\<rbrakk> \<Longrightarrow> b = c"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   589
  and xy: "(x, y) \<in> r^+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   590
  and xz: "(x, z) \<in> r^+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   591
  and neq_yz: "y \<noteq> z"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   592
  shows "(y, z) \<in> r^+ \<or> (z, y) \<in> r^+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   593
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   594
  from xy xz neq_yz show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   595
  proof(induct)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   596
    case (base y)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   597
    have h1: "(x, y) \<in> r" and h2: "(x, z) \<in> r\<^sup>+" and h3: "y \<noteq> z" using base by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   598
    from unique_base [OF _ h1 h2 h3] and unique show ?case by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   599
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   600
    case (step y za)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   601
    show ?case
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   602
    proof(cases "y = z")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   603
      case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   604
      from True step show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   605
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   606
      case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   607
      from False step have "(y, z) \<in> r\<^sup>+ \<or> (z, y) \<in> r\<^sup>+" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   608
      thus ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   609
      proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   610
        assume "(z, y) \<in> r\<^sup>+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   611
        with step have "(z, za) \<in> r\<^sup>+" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   612
        thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   613
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   614
        assume h: "(y, z) \<in> r\<^sup>+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   615
        from step have yza: "(y, za) \<in> r" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   616
        from step have "za \<noteq> z" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   617
        from unique_minus [OF _ yza h this] and unique
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   618
        have "(za, z) \<in> r\<^sup>+" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   619
        thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   620
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   621
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   622
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   623
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   624
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   625
lemma depend_set_unchanged: "(depend (Set th prio # s)) = depend s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   626
apply (unfold s_depend_def s_waiting_def wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   627
by (simp add:Let_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   628
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   629
lemma depend_create_unchanged: "(depend (Create th prio # s)) = depend s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   630
apply (unfold s_depend_def s_waiting_def wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   631
by (simp add:Let_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   632
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   633
lemma depend_exit_unchanged: "(depend (Exit th # s)) = depend s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   634
apply (unfold s_depend_def s_waiting_def wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   635
by (simp add:Let_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   636
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   637
definition head_of :: "('a \<Rightarrow> 'b::linorder) \<Rightarrow> 'a set \<Rightarrow> 'a set"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   638
  where "head_of f A = {a . a \<in> A \<and> f a = Max (f ` A)}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   639
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   640
definition wq_head :: "state \<Rightarrow> cs \<Rightarrow> thread set"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   641
  where "wq_head s cs = head_of (cp s) (set (wq s cs))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   642
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   643
lemma f_nil_simp: "\<lbrakk>f cs = []\<rbrakk> \<Longrightarrow> f(cs:=[]) = f"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   644
proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   645
  fix x
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   646
  assume h:"f cs = []"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   647
  show "(f(cs := [])) x = f x"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   648
  proof(cases "cs = x")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   649
    case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   650
    with h show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   651
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   652
    case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   653
    with h show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   654
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   655
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   656
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   657
lemma step_back_vt: "vt ccs (e#s) \<Longrightarrow> vt ccs s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   658
  by(ind_cases "vt ccs (e#s)", simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   659
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   660
lemma step_back_step: "vt ccs (e#s) \<Longrightarrow> ccs s e"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   661
  by(ind_cases "vt ccs (e#s)", simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   662
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   663
lemma holding_nil:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   664
    "\<lbrakk>wq s cs = []; holding (wq s) th cs\<rbrakk> \<Longrightarrow> False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   665
  by (unfold cs_holding_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   666
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   667
lemma waiting_kept_1: "
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   668
       \<lbrakk>vt step (V th cs#s); wq s cs = a # list; waiting ((wq s)(cs := ab # aa @ lista @ ca)) t c;
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   669
        lsp (cp s) list = (aa, ab # lista, ca)\<rbrakk>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   670
       \<Longrightarrow> waiting (wq s) t c"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   671
  apply (drule_tac step_back_vt, drule_tac wq_distinct[of _ cs])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   672
  apply(drule_tac lsp_set_eq)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   673
  by (unfold cs_waiting_def, auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   674
 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   675
lemma waiting_kept_2: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   676
  "\<And>a list t c aa ca.
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   677
       \<lbrakk>wq s cs = a # list; waiting ((wq s)(cs := [])) t c; lsp (cp s) list = (aa, [], ca)\<rbrakk>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   678
       \<Longrightarrow> waiting (wq s) t c"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   679
  apply(drule_tac lsp_set_eq)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   680
  by (unfold cs_waiting_def, auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   681
  
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   682
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   683
lemma holding_nil_simp: "\<lbrakk>holding ((wq s)(cs := [])) t c\<rbrakk> \<Longrightarrow> holding (wq s) t c"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   684
  by(unfold cs_holding_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   685
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   686
lemma step_wq_elim: "\<lbrakk>vt step (V th cs#s); wq s cs = a # list; a \<noteq> th\<rbrakk> \<Longrightarrow> False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   687
  apply(drule_tac step_back_step)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   688
  apply(ind_cases "step s (V th cs)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   689
  by(unfold s_holding_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   690
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   691
lemma holding_cs_neq_simp: "c \<noteq> cs \<Longrightarrow> holding ((wq s)(cs := u)) t c = holding (wq s) t c"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   692
  by (unfold cs_holding_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   693
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   694
lemma holding_th_neq_elim:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   695
  "\<And>a list c t aa ca ab lista.
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   696
       \<lbrakk>\<not> holding (wq s) t c; holding ((wq s)(cs := ab # aa @ lista @ ca)) t c;
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   697
         ab \<noteq> t\<rbrakk>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   698
       \<Longrightarrow> False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   699
  by (unfold cs_holding_def, auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   700
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   701
lemma holding_nil_abs:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   702
  "\<not> holding ((wq s)(cs := [])) th cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   703
  by (unfold cs_holding_def, auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   704
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   705
lemma holding_abs: "\<lbrakk>holding ((wq s)(cs := ab # aa @ lista @ c)) th cs; ab \<noteq> th\<rbrakk>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   706
       \<Longrightarrow> False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   707
    by (unfold cs_holding_def, auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   708
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   709
lemma waiting_abs: "\<not> waiting ((wq s)(cs := t # l @ r)) t cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   710
    by (unfold cs_waiting_def, auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   711
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   712
lemma waiting_abs_1: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   713
  "\<lbrakk>\<not> waiting ((wq s)(cs := [])) t c; waiting (wq s) t c; c \<noteq> cs\<rbrakk>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   714
       \<Longrightarrow> False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   715
    by (unfold cs_waiting_def, auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   716
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   717
lemma waiting_abs_2: "
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   718
       \<lbrakk>\<not> waiting ((wq s)(cs := ab # aa @ lista @ ca)) t c; waiting (wq s) t c;
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   719
        c \<noteq> cs\<rbrakk>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   720
       \<Longrightarrow> False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   721
  by (unfold cs_waiting_def, auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   722
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   723
lemma waiting_abs_3:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   724
     "\<lbrakk>wq s cs = a # list; \<not> waiting ((wq s)(cs := [])) t c;
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   725
        waiting (wq s) t c; lsp (cp s) list = (aa, [], ca)\<rbrakk>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   726
       \<Longrightarrow> False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   727
  apply(drule_tac lsp_mid_nil, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   728
  by(unfold cs_waiting_def, auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   729
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   730
lemma waiting_simp: "c \<noteq> cs \<Longrightarrow> waiting ((wq s)(cs:=z)) t c = waiting (wq s) t c"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   731
   by(unfold cs_waiting_def, auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   732
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   733
lemma holding_cs_eq:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   734
  "\<lbrakk>\<not> holding ((wq s)(cs := [])) t c; holding (wq s) t c\<rbrakk> \<Longrightarrow> c = cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   735
   by(unfold cs_holding_def, auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   736
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   737
lemma holding_cs_eq_1:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   738
  "\<lbrakk>\<not> holding ((wq s)(cs := ab # aa @ lista @ ca)) t c; holding (wq s) t c\<rbrakk>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   739
       \<Longrightarrow> c = cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   740
   by(unfold cs_holding_def, auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   741
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   742
lemma holding_th_eq: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   743
       "\<lbrakk>vt step (V th cs#s); wq s cs = a # list; \<not> holding ((wq s)(cs := [])) t c; holding (wq s) t c;
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   744
        lsp (cp s) list = (aa, [], ca)\<rbrakk>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   745
       \<Longrightarrow> t = th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   746
  apply(drule_tac lsp_mid_nil, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   747
  apply(unfold cs_holding_def, auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   748
  apply(drule_tac step_back_step)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   749
  apply(ind_cases "step s (V th cs)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   750
  by (unfold s_holding_def, auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   751
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   752
lemma holding_th_eq_1:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   753
  "\<lbrakk>vt step (V th cs#s); 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   754
     wq s cs = a # list; \<not> holding ((wq s)(cs := ab # aa @ lista @ ca)) t c; holding (wq s) t c;
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   755
        lsp (cp s) list = (aa, ab # lista, ca)\<rbrakk>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   756
       \<Longrightarrow> t = th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   757
  apply(drule_tac step_back_step)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   758
  apply(ind_cases "step s (V th cs)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   759
  apply(unfold s_holding_def cs_holding_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   760
  by (auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   761
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   762
lemma holding_th_eq_2: "\<lbrakk>holding ((wq s)(cs := ac # x)) th cs\<rbrakk>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   763
       \<Longrightarrow> ac = th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   764
  by (unfold cs_holding_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   765
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   766
lemma holding_th_eq_3: "
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   767
       \<lbrakk>\<not> holding (wq s) t c;
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   768
        holding ((wq s)(cs := ac # x)) t c\<rbrakk>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   769
       \<Longrightarrow> ac = t"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   770
  by (unfold cs_holding_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   771
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   772
lemma holding_wq_eq: "holding ((wq s)(cs := th' # l @ r)) th' cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   773
   by (unfold cs_holding_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   774
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   775
lemma waiting_th_eq: "
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   776
       \<lbrakk>waiting (wq s) t c; wq s cs = a # list;
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   777
        lsp (cp s) list = (aa, ac # lista, ba); \<not> waiting ((wq s)(cs := ac # aa @ lista @ ba)) t c\<rbrakk>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   778
       \<Longrightarrow> ac = t"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   779
  apply(drule_tac lsp_set_eq, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   780
  by (unfold cs_waiting_def, auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   781
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   782
lemma step_depend_v:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   783
  "vt step (V th cs#s) \<Longrightarrow>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   784
  depend (V th cs # s) =
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   785
  depend s - {(Cs cs, Th th)} -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   786
  {(Th th', Cs cs) |th'. \<exists>rest. wq s cs = th # rest \<and> (\<exists>l r. lsp (cp s) rest = (l, [th'], r))} \<union>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   787
  {(Cs cs, Th th') |th'. \<exists>rest. wq s cs = th # rest \<and> (\<exists>l r. lsp (cp s) rest = (l, [th'], r))}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   788
  apply (unfold s_depend_def wq_def, 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   789
         auto split:list.splits simp:Let_def f_nil_simp holding_wq_eq, fold wq_def cp_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   790
  apply (auto split:list.splits prod.splits  
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   791
               simp:Let_def f_nil_simp holding_nil_simp holding_cs_neq_simp holding_nil_abs
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   792
                    waiting_abs waiting_simp holding_wq_eq
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   793
               elim:holding_nil waiting_kept_1 waiting_kept_2 step_wq_elim holding_th_neq_elim 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   794
               holding_abs waiting_abs_1 waiting_abs_3 holding_cs_eq holding_cs_eq_1
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   795
               holding_th_eq holding_th_eq_1 holding_th_eq_2 holding_th_eq_3 waiting_th_eq
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   796
               dest:lsp_mid_length)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   797
  done
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   798
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   799
lemma step_depend_p:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   800
  "vt step (P th cs#s) \<Longrightarrow>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   801
  depend (P th cs # s) =  (if (wq s cs = []) then depend s \<union> {(Cs cs, Th th)}
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   802
                                             else depend s \<union> {(Th th, Cs cs)})"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   803
  apply(unfold s_depend_def wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   804
  apply (auto split:list.splits prod.splits simp:Let_def cs_waiting_def cs_holding_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   805
  apply(case_tac "c = cs", auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   806
  apply(fold wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   807
  apply(drule_tac step_back_step)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   808
  by (ind_cases " step s (P (hd (wq s cs)) cs)", 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   809
    auto simp:s_depend_def wq_def cs_holding_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   810
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   811
lemma simple_A:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   812
  fixes A
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   813
  assumes h: "\<And> x y. \<lbrakk>x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> x = y"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   814
  shows "A = {} \<or> (\<exists> a. A = {a})"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   815
proof(cases "A = {}")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   816
  case True thus ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   817
next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   818
  case False then obtain a where "a \<in> A" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   819
  with h have "A = {a}" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   820
  thus ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   821
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   822
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   823
lemma depend_target_th: "(Th th, x) \<in> depend (s::state) \<Longrightarrow> \<exists> cs. x = Cs cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   824
  by (unfold s_depend_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   825
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   826
lemma acyclic_depend: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   827
  fixes s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   828
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   829
  shows "acyclic (depend s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   830
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   831
  from vt show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   832
  proof(induct)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   833
    case (vt_cons s e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   834
    assume ih: "acyclic (depend s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   835
      and stp: "step s e"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   836
      and vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   837
    show ?case
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   838
    proof(cases e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   839
      case (Create th prio)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   840
      with ih
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   841
      show ?thesis by (simp add:depend_create_unchanged)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   842
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   843
      case (Exit th)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   844
      with ih show ?thesis by (simp add:depend_exit_unchanged)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   845
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   846
      case (V th cs)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   847
      from V vt stp have vtt: "vt step (V th cs#s)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   848
      from step_depend_v [OF this]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   849
      have eq_de: "depend (e # s) = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   850
        depend s - {(Cs cs, Th th)} -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   851
        {(Th th', Cs cs) |th'. \<exists>rest. wq s cs = th # rest \<and> (\<exists>l r. lsp (cp s) rest = (l, [th'], r))} \<union>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   852
        {(Cs cs, Th th') |th'. \<exists>rest. wq s cs = th # rest \<and> (\<exists>l r. lsp (cp s) rest = (l, [th'], r))}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   853
        (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   854
      from ih have ac: "acyclic (?A - ?B - ?C)" by (auto elim:acyclic_subset)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   855
      have "?D = {} \<or> (\<exists> a. ?D = {a})" by (rule simple_A, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   856
      thus ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   857
      proof(cases "wq s cs")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   858
        case Nil
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   859
        hence "?D = {}" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   860
        with ac and eq_de show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   861
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   862
        case (Cons tth rest)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   863
        from stp and V have "step s (V th cs)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   864
        hence eq_wq: "wq s cs = th # rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   865
        proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   866
          show "step s (V th cs) \<Longrightarrow> wq s cs = th # rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   867
            apply(ind_cases "step s (V th cs)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   868
            by(insert Cons, unfold s_holding_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   869
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   870
        show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   871
        proof(cases "lsp (cp s) rest")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   872
          fix l b r
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   873
          assume eq_lsp: "lsp (cp s) rest = (l, b, r) "
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   874
          show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   875
          proof(cases "b")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   876
            case Nil
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   877
            with eq_lsp and eq_wq have "?D = {}" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   878
            with ac and eq_de show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   879
          next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   880
            case (Cons th' m)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   881
            with eq_lsp 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   882
            have eq_lsp: "lsp (cp s) rest = (l, [th'], r)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   883
              apply simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   884
              by (drule_tac lsp_mid_length, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   885
            from eq_wq and eq_lsp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   886
            have eq_D: "?D = {(Cs cs, Th th')}" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   887
            from eq_wq and eq_lsp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   888
            have eq_C: "?C = {(Th th', Cs cs)}" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   889
            let ?E = "(?A - ?B - ?C)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   890
            have "(Th th', Cs cs) \<notin> ?E\<^sup>*"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   891
            proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   892
              assume "(Th th', Cs cs) \<in> ?E\<^sup>*"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   893
              hence " (Th th', Cs cs) \<in> ?E\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   894
              from tranclD [OF this]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   895
              obtain x where th'_e: "(Th th', x) \<in> ?E" by blast
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   896
              hence th_d: "(Th th', x) \<in> ?A" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   897
              from depend_target_th [OF this]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   898
              obtain cs' where eq_x: "x = Cs cs'" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   899
              with th_d have "(Th th', Cs cs') \<in> ?A" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   900
              hence wt_th': "waiting s th' cs'"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   901
                unfolding s_depend_def s_waiting_def cs_waiting_def by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   902
              hence "cs' = cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   903
              proof(rule waiting_unique [OF vt])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   904
                from eq_wq eq_lsp wq_distinct[OF vt, of cs]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   905
                show "waiting s th' cs" by(unfold s_waiting_def, auto dest:lsp_set_eq)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   906
              qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   907
              with th'_e eq_x have "(Th th', Cs cs) \<in> ?E" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   908
              with eq_C show "False" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   909
            qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   910
            with acyclic_insert[symmetric] and ac and eq_D
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   911
            and eq_de show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   912
          qed 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   913
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   914
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   915
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   916
      case (P th cs)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   917
      from P vt stp have vtt: "vt step (P th cs#s)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   918
      from step_depend_p [OF this] P
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   919
      have "depend (e # s) = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   920
              (if wq s cs = [] then depend s \<union> {(Cs cs, Th th)} else 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   921
                                    depend s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   922
        by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   923
      moreover have "acyclic ?R"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   924
      proof(cases "wq s cs = []")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   925
        case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   926
        hence eq_r: "?R =  depend s \<union> {(Cs cs, Th th)}" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   927
        have "(Th th, Cs cs) \<notin> (depend s)\<^sup>*"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   928
        proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   929
          assume "(Th th, Cs cs) \<in> (depend s)\<^sup>*"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   930
          hence "(Th th, Cs cs) \<in> (depend s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   931
          from tranclD2 [OF this]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   932
          obtain x where "(x, Cs cs) \<in> depend s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   933
          with True show False by (auto simp:s_depend_def cs_waiting_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   934
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   935
        with acyclic_insert ih eq_r show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   936
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   937
        case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   938
        hence eq_r: "?R =  depend s \<union> {(Th th, Cs cs)}" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   939
        have "(Cs cs, Th th) \<notin> (depend s)\<^sup>*"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   940
        proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   941
          assume "(Cs cs, Th th) \<in> (depend s)\<^sup>*"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   942
          hence "(Cs cs, Th th) \<in> (depend s)\<^sup>+" by (simp add: rtrancl_eq_or_trancl)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   943
          moreover from step_back_step [OF vtt] have "step s (P th cs)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   944
          ultimately show False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   945
          proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   946
            show " \<lbrakk>(Cs cs, Th th) \<in> (depend s)\<^sup>+; step s (P th cs)\<rbrakk> \<Longrightarrow> False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   947
              by (ind_cases "step s (P th cs)", simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   948
          qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   949
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   950
        with acyclic_insert ih eq_r show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   951
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   952
      ultimately show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   953
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   954
      case (Set thread prio)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   955
      with ih
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   956
      thm depend_set_unchanged
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   957
      show ?thesis by (simp add:depend_set_unchanged)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   958
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   959
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   960
    case vt_nil
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   961
    show "acyclic (depend ([]::state))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   962
      by (auto simp: s_depend_def cs_waiting_def 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   963
                      cs_holding_def wq_def acyclic_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   964
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   965
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   966
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   967
lemma finite_depend: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   968
  fixes s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   969
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   970
  shows "finite (depend s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   971
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   972
  from vt show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   973
  proof(induct)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   974
    case (vt_cons s e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   975
    assume ih: "finite (depend s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   976
      and stp: "step s e"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   977
      and vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   978
    show ?case
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   979
    proof(cases e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   980
      case (Create th prio)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   981
      with ih
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   982
      show ?thesis by (simp add:depend_create_unchanged)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   983
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   984
      case (Exit th)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   985
      with ih show ?thesis by (simp add:depend_exit_unchanged)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   986
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   987
      case (V th cs)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   988
      from V vt stp have vtt: "vt step (V th cs#s)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   989
      from step_depend_v [OF this]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   990
      have eq_de: "depend (e # s) = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   991
        depend s - {(Cs cs, Th th)} -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   992
        {(Th th', Cs cs) |th'. \<exists>rest. wq s cs = th # rest \<and> (\<exists>l r. lsp (cp s) rest = (l, [th'], r))} \<union>
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   993
        {(Cs cs, Th th') |th'. \<exists>rest. wq s cs = th # rest \<and> (\<exists>l r. lsp (cp s) rest = (l, [th'], r))}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   994
        (is "?L = (?A - ?B - ?C) \<union> ?D") by (simp add:V)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   995
      moreover from ih have ac: "finite (?A - ?B - ?C)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   996
      moreover have "finite ?D"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   997
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   998
        have "?D = {} \<or> (\<exists> a. ?D = {a})" by (rule simple_A, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
   999
        thus ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1000
        proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1001
          assume h: "?D = {}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1002
          show ?thesis by (unfold h, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1003
        next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1004
          assume "\<exists> a. ?D = {a}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1005
          thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1006
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1007
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1008
      ultimately show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1009
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1010
      case (P th cs)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1011
      from P vt stp have vtt: "vt step (P th cs#s)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1012
      from step_depend_p [OF this] P
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1013
      have "depend (e # s) = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1014
              (if wq s cs = [] then depend s \<union> {(Cs cs, Th th)} else 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1015
                                    depend s \<union> {(Th th, Cs cs)})" (is "?L = ?R")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1016
        by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1017
      moreover have "finite ?R"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1018
      proof(cases "wq s cs = []")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1019
        case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1020
        hence eq_r: "?R =  depend s \<union> {(Cs cs, Th th)}" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1021
        with True and ih show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1022
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1023
        case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1024
        hence "?R = depend s \<union> {(Th th, Cs cs)}" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1025
        with False and ih show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1026
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1027
      ultimately show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1028
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1029
      case (Set thread prio)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1030
      with ih
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1031
      show ?thesis by (simp add:depend_set_unchanged)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1032
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1033
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1034
    case vt_nil
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1035
    show "finite (depend ([]::state))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1036
      by (auto simp: s_depend_def cs_waiting_def 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1037
                   cs_holding_def wq_def acyclic_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1038
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1039
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1040
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1041
text {* Several useful lemmas *}
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1042
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1043
thm wf_trancl
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1044
thm finite_acyclic_wf
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1045
thm finite_acyclic_wf_converse
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1046
thm wf_induct
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1047
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1048
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1049
lemma wf_dep_converse: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1050
  fixes s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1051
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1052
  shows "wf ((depend s)^-1)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1053
proof(rule finite_acyclic_wf_converse)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1054
  from finite_depend [OF vt]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1055
  show "finite (depend s)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1056
next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1057
  from acyclic_depend[OF vt]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1058
  show "acyclic (depend s)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1059
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1060
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1061
lemma hd_np_in: "x \<in> set l \<Longrightarrow> hd l \<in> set l"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1062
by (induct l, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1063
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1064
lemma th_chasing: "(Th th, Cs cs) \<in> depend (s::state) \<Longrightarrow> \<exists> th'. (Cs cs, Th th') \<in> depend s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1065
  by (auto simp:s_depend_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1066
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1067
lemma wq_threads: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1068
  fixes s cs
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1069
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1070
  and h: "th \<in> set (wq s cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1071
  shows "th \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1072
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1073
 from vt and h show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1074
  proof(induct arbitrary: th cs)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1075
    case (vt_cons s e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1076
    assume ih: "\<And>th cs. th \<in> set (wq s cs) \<Longrightarrow> th \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1077
      and stp: "step s e"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1078
      and vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1079
      and h: "th \<in> set (wq (e # s) cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1080
    show ?case
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1081
    proof(cases e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1082
      case (Create th' prio)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1083
      with ih h show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1084
        by (auto simp:wq_def Let_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1085
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1086
      case (Exit th')
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1087
      with stp ih h show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1088
        apply (auto simp:wq_def Let_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1089
        apply (ind_cases "step s (Exit th')")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1090
        apply (auto simp:runing_def readys_def s_holding_def s_waiting_def holdents_def
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1091
               s_depend_def s_holding_def cs_holding_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1092
        by (fold wq_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1093
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1094
      case (V th' cs')
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1095
      show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1096
      proof(cases "cs' = cs")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1097
        case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1098
        with h
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1099
        show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1100
          apply(unfold wq_def V, auto simp:Let_def V split:prod.splits, fold wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1101
          by (drule_tac ih, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1102
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1103
        case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1104
        from h
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1105
        show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1106
        proof(unfold V wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1107
          assume th_in: "th \<in> set (waiting_queue (schs (V th' cs' # s)) cs)" (is "th \<in> set ?l")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1108
          show "th \<in> threads (V th' cs' # s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1109
          proof(cases "cs = cs'")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1110
            case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1111
            hence "?l = waiting_queue (schs s) cs" by (simp add:Let_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1112
            with th_in have " th \<in> set (wq s cs)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1113
              by (fold wq_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1114
            from ih [OF this] show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1115
          next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1116
            case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1117
            show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1118
            proof(cases "waiting_queue (schs s) cs'")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1119
              case Nil
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1120
              with h V show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1121
                apply (auto simp:wq_def Let_def split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1122
                by (fold wq_def, drule_tac ih, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1123
            next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1124
              case (Cons a rest)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1125
              assume eq_wq: "waiting_queue (schs s) cs' = a # rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1126
              with h V show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1127
              proof(cases "(lsp (cur_preced (schs s)) rest)", unfold V)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1128
                fix l m r
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1129
                assume eq_lsp: "lsp (cur_preced (schs s)) rest = (l, m, r)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1130
                  and eq_wq: "waiting_queue (schs s) cs' = a # rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1131
                  and th_in_set: "th \<in> set (wq (V th' cs' # s) cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1132
                show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1133
                proof(cases "m")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1134
                  case Nil
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1135
                  with eq_lsp have "rest = []" using lsp_mid_nil by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1136
                  with eq_wq have "waiting_queue (schs s) cs' = [a]" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1137
                  with h[unfolded V wq_def] True 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1138
                  show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1139
                    by (simp add:Let_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1140
                next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1141
                  case (Cons b rb)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1142
                  with lsp_mid_length[OF eq_lsp] have eq_m: "m = [b]" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1143
                  with eq_lsp have "lsp (cur_preced (schs s)) rest = (l, [b], r)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1144
                  with h[unfolded V wq_def] True lsp_set_eq [OF this] eq_wq
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1145
                  show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1146
                    apply (auto simp:Let_def, fold wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1147
                    by (rule_tac ih [of _ cs'], auto)+
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1148
                qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1149
              qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1150
            qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1151
          qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1152
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1153
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1154
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1155
      case (P th' cs')
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1156
      from h stp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1157
      show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1158
        apply (unfold P wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1159
        apply (auto simp:Let_def split:if_splits, fold wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1160
        apply (auto intro:ih)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1161
        apply(ind_cases "step s (P th' cs')")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1162
        by (unfold runing_def readys_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1163
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1164
      case (Set thread prio)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1165
      with ih h show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1166
        by (auto simp:wq_def Let_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1167
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1168
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1169
    case vt_nil
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1170
    thus ?case by (auto simp:wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1171
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1172
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1173
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1174
lemma range_in: "\<lbrakk>vt step s; (Th th) \<in> Range (depend (s::state))\<rbrakk> \<Longrightarrow> th \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1175
  apply(unfold s_depend_def cs_waiting_def cs_holding_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1176
  by (auto intro:wq_threads)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1177
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1178
lemma readys_v_eq:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1179
  fixes th thread cs rest
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1180
  assumes neq_th: "th \<noteq> thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1181
  and eq_wq: "wq s cs = thread#rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1182
  and not_in: "th \<notin>  set rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1183
  shows "(th \<in> readys (V thread cs#s)) = (th \<in> readys s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1184
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1185
  from prems show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1186
    apply (auto simp:readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1187
    apply (case_tac "cs = csa", simp add:s_waiting_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1188
    apply (erule_tac x = csa in allE)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1189
    apply (simp add:s_waiting_def wq_def Let_def split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1190
    apply (case_tac "csa = cs", simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1191
    apply (erule_tac x = cs in allE)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1192
    by (auto simp:s_waiting_def wq_def Let_def split:list.splits prod.splits 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1193
            dest:lsp_set_eq)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1194
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1195
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1196
lemma readys_v_eq_1:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1197
  fixes th thread cs rest
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1198
  assumes neq_th: "th \<noteq> thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1199
  and eq_wq: "wq s cs = thread#rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1200
  and eq_lsp: "lsp (cp s) rest = (l, [th'], r)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1201
  and neq_th': "th \<noteq> th'"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1202
  shows "(th \<in> readys (V thread cs#s)) = (th \<in> readys s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1203
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1204
  from prems show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1205
    apply (auto simp:readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1206
    apply (case_tac "cs = csa", simp add:s_waiting_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1207
    apply (erule_tac x = cs in allE)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1208
    apply (simp add:s_waiting_def wq_def Let_def split:prod.splits list.splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1209
    apply (drule_tac lsp_mid_nil,simp, clarify, fold cp_def, clarsimp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1210
    apply (frule_tac lsp_set_eq, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1211
    apply (erule_tac x = csa in allE)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1212
    apply (subst (asm) (2) s_waiting_def, unfold wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1213
    apply (auto simp:Let_def split:list.splits prod.splits if_splits 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1214
            dest:lsp_set_eq)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1215
    apply (unfold s_waiting_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1216
    apply (fold wq_def, clarsimp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1217
    apply (clarsimp)+
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1218
    apply (case_tac "csa = cs", simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1219
    apply (erule_tac x = cs in allE, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1220
    apply (unfold wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1221
    by (auto simp:Let_def split:list.splits prod.splits if_splits 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1222
            dest:lsp_set_eq)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1223
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1224
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1225
lemma readys_v_eq_2:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1226
  fixes th thread cs rest
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1227
  assumes neq_th: "th \<noteq> thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1228
  and eq_wq: "wq s cs = thread#rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1229
  and eq_lsp: "lsp (cp s) rest = (l, [th'], r)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1230
  and neq_th': "th = th'"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1231
  and vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1232
  shows "(th \<in> readys (V thread cs#s))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1233
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1234
  from prems show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1235
    apply (auto simp:readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1236
    apply (rule_tac wq_threads [of s _ cs], auto dest:lsp_set_eq)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1237
    apply (unfold s_waiting_def wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1238
    apply (auto simp:Let_def split:list.splits prod.splits if_splits 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1239
            dest:lsp_set_eq lsp_mid_nil lsp_mid_length)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1240
    apply (fold cp_def, simp+, clarsimp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1241
    apply (frule_tac lsp_set_eq, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1242
    apply (fold wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1243
    apply (subgoal_tac "csa = cs", simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1244
    apply (rule_tac waiting_unique [of s th'], simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1245
    by (auto simp:s_waiting_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1246
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1247
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1248
lemma chain_building:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1249
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1250
  shows "node \<in> Domain (depend s) \<longrightarrow> (\<exists> th'. th' \<in> readys s \<and> (node, Th th') \<in> (depend s)^+)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1251
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1252
  from wf_dep_converse [OF vt]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1253
  have h: "wf ((depend s)\<inverse>)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1254
  show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1255
  proof(induct rule:wf_induct [OF h])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1256
    fix x
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1257
    assume ih [rule_format]: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1258
      "\<forall>y. (y, x) \<in> (depend s)\<inverse> \<longrightarrow> 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1259
           y \<in> Domain (depend s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (y, Th th') \<in> (depend s)\<^sup>+)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1260
    show "x \<in> Domain (depend s) \<longrightarrow> (\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (depend s)\<^sup>+)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1261
    proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1262
      assume x_d: "x \<in> Domain (depend s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1263
      show "\<exists>th'. th' \<in> readys s \<and> (x, Th th') \<in> (depend s)\<^sup>+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1264
      proof(cases x)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1265
        case (Th th)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1266
        from x_d Th obtain cs where x_in: "(Th th, Cs cs) \<in> depend s" by (auto simp:s_depend_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1267
        with Th have x_in_r: "(Cs cs, x) \<in> (depend s)^-1" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1268
        from th_chasing [OF x_in] obtain th' where "(Cs cs, Th th') \<in> depend s" by blast
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1269
        hence "Cs cs \<in> Domain (depend s)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1270
        from ih [OF x_in_r this] obtain th'
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1271
          where th'_ready: " th' \<in> readys s" and cs_in: "(Cs cs, Th th') \<in> (depend s)\<^sup>+" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1272
        have "(x, Th th') \<in> (depend s)\<^sup>+" using Th x_in cs_in by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1273
        with th'_ready show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1274
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1275
        case (Cs cs)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1276
        from x_d Cs obtain th' where th'_d: "(Th th', x) \<in> (depend s)^-1" by (auto simp:s_depend_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1277
        show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1278
        proof(cases "th' \<in> readys s")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1279
          case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1280
          from True and th'_d show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1281
        next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1282
          case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1283
          from th'_d and range_in [OF vt] have "th' \<in> threads s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1284
          with False have "Th th' \<in> Domain (depend s)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1285
            by (auto simp:readys_def s_waiting_def s_depend_def cs_waiting_def Domain_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1286
          from ih [OF th'_d this]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1287
          obtain th'' where 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1288
            th''_r: "th'' \<in> readys s" and 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1289
            th''_in: "(Th th', Th th'') \<in> (depend s)\<^sup>+" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1290
          from th'_d and th''_in 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1291
          have "(x, Th th'') \<in> (depend s)\<^sup>+" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1292
          with th''_r show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1293
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1294
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1295
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1296
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1297
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1298
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1299
lemma th_chain_to_ready:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1300
  fixes s th
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1301
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1302
  and th_in: "th \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1303
  shows "th \<in> readys s \<or> (\<exists> th'. th' \<in> readys s \<and> (Th th, Th th') \<in> (depend s)^+)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1304
proof(cases "th \<in> readys s")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1305
  case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1306
  thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1307
next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1308
  case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1309
  from False and th_in have "Th th \<in> Domain (depend s)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1310
    by (auto simp:readys_def s_waiting_def s_depend_def cs_waiting_def Domain_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1311
  from chain_building [rule_format, OF vt this]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1312
  show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1313
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1314
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1315
lemma waiting_eq: "waiting s th cs = waiting (wq s) th cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1316
  by  (unfold s_waiting_def cs_waiting_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1317
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1318
lemma holding_eq: "holding (s::state) th cs = holding (wq s) th cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1319
  by (unfold s_holding_def cs_holding_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1320
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1321
lemma holding_unique: "\<lbrakk>holding (s::state) th1 cs; holding s th2 cs\<rbrakk> \<Longrightarrow> th1 = th2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1322
  by (unfold s_holding_def cs_holding_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1323
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1324
lemma unique_depend: "\<lbrakk>vt step s; (n, n1) \<in> depend s; (n, n2) \<in> depend s\<rbrakk> \<Longrightarrow> n1 = n2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1325
  apply(unfold s_depend_def, auto, fold waiting_eq holding_eq)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1326
  by(auto elim:waiting_unique holding_unique)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1327
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1328
lemma trancl_split: "(a, b) \<in> r^+ \<Longrightarrow> \<exists> c. (a, c) \<in> r"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1329
by (induct rule:trancl_induct, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1330
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1331
lemma dchain_unique:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1332
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1333
  and th1_d: "(n, Th th1) \<in> (depend s)^+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1334
  and th1_r: "th1 \<in> readys s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1335
  and th2_d: "(n, Th th2) \<in> (depend s)^+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1336
  and th2_r: "th2 \<in> readys s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1337
  shows "th1 = th2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1338
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1339
  { assume neq: "th1 \<noteq> th2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1340
    hence "Th th1 \<noteq> Th th2" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1341
    from unique_chain [OF _ th1_d th2_d this] and unique_depend [OF vt]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1342
    have "(Th th1, Th th2) \<in> (depend s)\<^sup>+ \<or> (Th th2, Th th1) \<in> (depend s)\<^sup>+" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1343
    hence "False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1344
    proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1345
      assume "(Th th1, Th th2) \<in> (depend s)\<^sup>+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1346
      from trancl_split [OF this]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1347
      obtain n where dd: "(Th th1, n) \<in> depend s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1348
      then obtain cs where eq_n: "n = Cs cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1349
        by (auto simp:s_depend_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1350
      from dd eq_n have "th1 \<notin> readys s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1351
        by (auto simp:readys_def s_depend_def s_waiting_def cs_waiting_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1352
      with th1_r show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1353
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1354
      assume "(Th th2, Th th1) \<in> (depend s)\<^sup>+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1355
      from trancl_split [OF this]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1356
      obtain n where dd: "(Th th2, n) \<in> depend s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1357
      then obtain cs where eq_n: "n = Cs cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1358
        by (auto simp:s_depend_def s_holding_def cs_holding_def cs_waiting_def wq_def dest:hd_np_in)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1359
      from dd eq_n have "th2 \<notin> readys s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1360
        by (auto simp:readys_def s_depend_def s_waiting_def cs_waiting_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1361
      with th2_r show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1362
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1363
  } thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1364
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1365
             
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1366
definition count :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> nat"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1367
where "count Q l = length (filter Q l)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1368
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1369
definition cntP :: "state \<Rightarrow> thread \<Rightarrow> nat"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1370
where "cntP s th = count (\<lambda> e. \<exists> cs. e = P th cs) s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1371
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1372
definition cntV :: "state \<Rightarrow> thread \<Rightarrow> nat"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1373
where "cntV s th = count (\<lambda> e. \<exists> cs. e = V th cs) s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1374
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1375
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1376
lemma step_holdents_p_add:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1377
  fixes th cs s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1378
  assumes vt: "vt step (P th cs#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1379
  and "wq s cs = []"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1380
  shows "holdents (P th cs#s) th = holdents s th \<union> {cs}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1381
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1382
  from prems show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1383
  unfolding  holdents_def step_depend_p[OF vt] by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1384
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1385
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1386
lemma step_holdents_p_eq:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1387
  fixes th cs s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1388
  assumes vt: "vt step (P th cs#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1389
  and "wq s cs \<noteq> []"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1390
  shows "holdents (P th cs#s) th = holdents s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1391
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1392
  from prems show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1393
  unfolding  holdents_def step_depend_p[OF vt] by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1394
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1395
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1396
lemma step_holdents_v_minus:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1397
  fixes th cs s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1398
  assumes vt: "vt step (V th cs#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1399
  shows "holdents (V th cs#s) th = holdents s th - {cs}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1400
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1401
  { fix rest l r
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1402
    assume eq_wq: "wq s cs = th # rest" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1403
      and eq_lsp: "lsp (cp s) rest = (l, [th], r)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1404
    have "False" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1405
    proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1406
      from lsp_set_eq [OF eq_lsp] have " rest = l @ [th] @ r" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1407
      with eq_wq have "wq s cs = th#\<dots>" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1408
      with wq_distinct [OF step_back_vt[OF vt], of cs]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1409
      show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1410
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1411
  } thus ?thesis unfolding holdents_def step_depend_v[OF vt] by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1412
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1413
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1414
lemma step_holdents_v_add:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1415
  fixes th th' cs s rest l r
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1416
  assumes vt: "vt step (V th' cs#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1417
  and neq_th: "th \<noteq> th'" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1418
  and eq_wq: "wq s cs = th' # rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1419
  and eq_lsp: "lsp (cp s) rest = (l, [th], r)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1420
  shows "holdents (V th' cs#s) th = holdents s th \<union> {cs}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1421
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1422
  from prems show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1423
  unfolding  holdents_def step_depend_v[OF vt] by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1424
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1425
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1426
lemma step_holdents_v_eq:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1427
  fixes th th' cs s rest l r th''
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1428
  assumes vt: "vt step (V th' cs#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1429
  and neq_th: "th \<noteq> th'" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1430
  and eq_wq: "wq s cs = th' # rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1431
  and eq_lsp: "lsp (cp s) rest = (l, [th''], r)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1432
  and neq_th': "th \<noteq> th''"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1433
  shows "holdents (V th' cs#s) th = holdents s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1434
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1435
  from prems show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1436
  unfolding  holdents_def step_depend_v[OF vt] by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1437
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1438
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1439
definition cntCS :: "state \<Rightarrow> thread \<Rightarrow> nat"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1440
where "cntCS s th = card (holdents s th)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1441
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1442
lemma cntCS_v_eq:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1443
  fixes th thread cs rest
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1444
  assumes neq_th: "th \<noteq> thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1445
  and eq_wq: "wq s cs = thread#rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1446
  and not_in: "th \<notin>  set rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1447
  and vtv: "vt step (V thread cs#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1448
  shows "cntCS (V thread cs#s) th = cntCS s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1449
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1450
  from prems show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1451
    apply (unfold cntCS_def holdents_def step_depend_v)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1452
    apply auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1453
    apply (subgoal_tac "\<not>  (\<exists>l r. lsp (cp s) rest = (l, [th], r))", auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1454
    by (drule_tac lsp_set_eq, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1455
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1456
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1457
lemma cntCS_v_eq_1:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1458
  fixes th thread cs rest
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1459
  assumes neq_th: "th \<noteq> thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1460
  and eq_wq: "wq s cs = thread#rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1461
  and eq_lsp: "lsp (cp s) rest = (l, [th'], r)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1462
  and neq_th': "th \<noteq> th'"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1463
  and vtv: "vt step (V thread cs#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1464
  shows "cntCS (V thread cs#s) th = cntCS s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1465
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1466
  from prems show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1467
    apply (unfold cntCS_def holdents_def step_depend_v)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1468
    by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1469
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1470
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1471
fun the_cs :: "node \<Rightarrow> cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1472
where "the_cs (Cs cs) = cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1473
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1474
lemma cntCS_v_eq_2:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1475
  fixes th thread cs rest
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1476
  assumes neq_th: "th \<noteq> thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1477
  and eq_wq: "wq s cs = thread#rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1478
  and eq_lsp: "lsp (cp s) rest = (l, [th'], r)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1479
  and neq_th': "th = th'"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1480
  and vtv: "vt step (V thread cs#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1481
  shows "cntCS (V thread cs#s) th = 1 + cntCS s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1482
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1483
  have "card {csa. csa = cs \<or> (Cs csa, Th th') \<in> depend s} = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1484
                     Suc (card {cs. (Cs cs, Th th') \<in> depend s})" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1485
    (is "card ?A = Suc (card ?B)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1486
  proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1487
    have h: "?A = insert cs ?B" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1488
    moreover have h1: "?B = ?B - {cs}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1489
    proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1490
      { assume "(Cs cs, Th th') \<in> depend s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1491
        moreover have "(Th th', Cs cs) \<in> depend s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1492
        proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1493
          from wq_distinct [OF step_back_vt[OF vtv], of cs]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1494
          eq_wq lsp_set_eq [OF eq_lsp] show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1495
            apply (auto simp:s_depend_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1496
            by (unfold cs_waiting_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1497
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1498
        moreover note acyclic_depend [OF step_back_vt[OF vtv]]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1499
        ultimately have "False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1500
          apply (auto simp:acyclic_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1501
          apply (erule_tac x="Cs cs" in allE)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1502
          apply (subgoal_tac "(Cs cs, Cs cs) \<in> (depend s)\<^sup>+", simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1503
          by (rule_tac trancl_into_trancl [where b = "Th th'"], auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1504
      } thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1505
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1506
    moreover have "card (insert cs ?B) = Suc (card (?B - {cs}))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1507
    proof(rule card_insert)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1508
      from finite_depend [OF step_back_vt [OF vtv]]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1509
      have fnt: "finite (depend s)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1510
      show " finite {cs. (Cs cs, Th th') \<in> depend s}" (is "finite ?B")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1511
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1512
        have "?B \<subseteq>  (\<lambda> (a, b). the_cs a) ` (depend s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1513
          apply (auto simp:image_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1514
          by (rule_tac x = "(Cs x, Th th')" in bexI, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1515
        with fnt show ?thesis by (auto intro:finite_subset)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1516
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1517
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1518
    ultimately show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1519
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1520
  with prems show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1521
    apply (unfold cntCS_def holdents_def step_depend_v[OF vtv])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1522
    by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1523
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1524
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1525
lemma finite_holding:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1526
  fixes s th cs
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1527
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1528
  shows "finite (holdents s th)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1529
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1530
  let ?F = "\<lambda> (x, y). the_cs x"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1531
  from finite_depend [OF vt]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1532
  have "finite (depend s)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1533
  hence "finite (?F `(depend s))" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1534
  moreover have "{cs . (Cs cs, Th th) \<in> depend s} \<subseteq> \<dots>" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1535
  proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1536
    { have h: "\<And> a A f. a \<in> A \<Longrightarrow> f a \<in> f ` A" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1537
      fix x assume "(Cs x, Th th) \<in> depend s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1538
      hence "?F (Cs x, Th th) \<in> ?F `(depend s)" by (rule h)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1539
      moreover have "?F (Cs x, Th th) = x" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1540
      ultimately have "x \<in> (\<lambda>(x, y). the_cs x) ` depend s" by simp 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1541
    } thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1542
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1543
  ultimately show ?thesis by (unfold holdents_def, auto intro:finite_subset)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1544
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1545
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1546
inductive_cases case_step_v: "step s (V thread cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1547
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1548
lemma cntCS_v_dec: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1549
  fixes s thread cs
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1550
  assumes vtv: "vt step (V thread cs#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1551
  shows "(cntCS (V thread cs#s) thread + 1) = cntCS s thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1552
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1553
  have cs_in: "cs \<in> holdents s thread" using step_back_step[OF vtv]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1554
    apply (erule_tac case_step_v)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1555
    apply (unfold holdents_def s_depend_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1556
    by (unfold cs_holding_def s_holding_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1557
  moreover have cs_not_in: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1558
    "(holdents (V thread cs#s) thread) = holdents s thread - {cs}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1559
    apply (insert wq_distinct[OF step_back_vt[OF vtv], of cs])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1560
    by (unfold holdents_def, unfold step_depend_v[OF vtv], 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1561
            auto dest:lsp_set_eq)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1562
  ultimately 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1563
  have "holdents s thread = insert cs (holdents (V thread cs#s) thread)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1564
    by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1565
  moreover have "card \<dots> = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1566
                    Suc (card ((holdents (V thread cs#s) thread) - {cs}))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1567
  proof(rule card_insert)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1568
    from finite_holding [OF vtv]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1569
    show " finite (holdents (V thread cs # s) thread)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1570
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1571
  moreover from cs_not_in 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1572
  have "cs \<notin> (holdents (V thread cs#s) thread)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1573
  ultimately show ?thesis by (simp add:cntCS_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1574
qed 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1575
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1576
lemma cnp_cnv_cncs:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1577
  fixes s th
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1578
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1579
  shows "cntP s th = cntV s th + (if (th \<in> readys s \<or> th \<notin> threads s) 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1580
                                       then cntCS s th else cntCS s th + 1)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1581
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1582
  from vt show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1583
  proof(induct arbitrary:th)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1584
    case (vt_cons s e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1585
    assume vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1586
    and ih: "\<And>th. cntP s th  = cntV s th +
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1587
               (if (th \<in> readys s \<or> th \<notin> threads s) then cntCS s th else cntCS s th + 1)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1588
    and stp: "step s e"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1589
    from stp show ?case
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1590
    proof(cases)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1591
      case (thread_create prio max_prio thread)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1592
      assume eq_e: "e = Create thread prio"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1593
        and not_in: "thread \<notin> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1594
      show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1595
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1596
        { fix cs 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1597
          assume "thread \<in> set (wq s cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1598
          from wq_threads [OF vt this] have "thread \<in> threads s" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1599
          with not_in have "False" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1600
        } with eq_e have eq_readys: "readys (e#s) = readys s \<union> {thread}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1601
          by (auto simp:readys_def threads.simps s_waiting_def 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1602
            wq_def cs_waiting_def Let_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1603
        from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1604
        from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1605
        have eq_cncs: "cntCS (e#s) th = cntCS s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1606
          unfolding cntCS_def holdents_def
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1607
          by (simp add:depend_create_unchanged eq_e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1608
        { assume "th \<noteq> thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1609
          with eq_readys eq_e
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1610
          have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1611
                      (th \<in> readys (s) \<or> th \<notin> threads (s))" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1612
            by (simp add:threads.simps)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1613
          with eq_cnp eq_cnv eq_cncs ih not_in
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1614
          have ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1615
        } moreover {
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1616
          assume eq_th: "th = thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1617
          with not_in ih have " cntP s th  = cntV s th + cntCS s th" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1618
          moreover from eq_th and eq_readys have "th \<in> readys (e#s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1619
          moreover note eq_cnp eq_cnv eq_cncs
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1620
          ultimately have ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1621
        } ultimately show ?thesis by blast
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1622
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1623
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1624
      case (thread_exit thread)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1625
      assume eq_e: "e = Exit thread" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1626
      and is_runing: "thread \<in> runing s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1627
      and no_hold: "holdents s thread = {}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1628
      from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1629
      from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1630
      have eq_cncs: "cntCS (e#s) th = cntCS s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1631
        unfolding cntCS_def holdents_def
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1632
        by (simp add:depend_exit_unchanged eq_e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1633
      { assume "th \<noteq> thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1634
        with eq_e
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1635
        have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1636
          (th \<in> readys (s) \<or> th \<notin> threads (s))" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1637
          apply (simp add:threads.simps readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1638
          apply (subst s_waiting_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1639
          apply (subst (1 2) wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1640
          apply (simp add:Let_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1641
          apply (subst s_waiting_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1642
          by (fold wq_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1643
        with eq_cnp eq_cnv eq_cncs ih
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1644
        have ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1645
      } moreover {
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1646
        assume eq_th: "th = thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1647
        with ih is_runing have " cntP s th = cntV s th + cntCS s th" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1648
          by (simp add:runing_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1649
        moreover from eq_th eq_e have "th \<notin> threads (e#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1650
          by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1651
        moreover note eq_cnp eq_cnv eq_cncs
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1652
        ultimately have ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1653
      } ultimately show ?thesis by blast
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1654
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1655
      case (thread_P thread cs)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1656
      assume eq_e: "e = P thread cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1657
        and is_runing: "thread \<in> runing s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1658
        and no_dep: "(Cs cs, Th thread) \<notin> (depend s)\<^sup>+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1659
      from prems have vtp: "vt step (P thread cs#s)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1660
      show ?thesis 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1661
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1662
        { have hh: "\<And> A B C. (B = C) \<Longrightarrow> (A \<and> B) = (A \<and> C)" by blast
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1663
          assume neq_th: "th \<noteq> thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1664
          with eq_e
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1665
          have eq_readys: "(th \<in> readys (e#s)) = (th \<in> readys (s))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1666
            apply (simp add:readys_def s_waiting_def wq_def Let_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1667
            apply (rule_tac hh, clarify)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1668
            apply (intro iffI allI, clarify)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1669
            apply (erule_tac x = csa in allE, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1670
            apply (subgoal_tac "waiting_queue (schs s) cs \<noteq> []", auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1671
            apply (erule_tac x = cs in allE, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1672
            by (case_tac "(waiting_queue (schs s) cs)", auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1673
          moreover from neq_th eq_e have "cntCS (e # s) th = cntCS s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1674
            apply (simp add:cntCS_def holdents_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1675
            by (unfold  step_depend_p [OF vtp], auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1676
          moreover from eq_e neq_th have "cntP (e # s) th = cntP s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1677
            by (simp add:cntP_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1678
          moreover from eq_e neq_th have "cntV (e#s) th = cntV s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1679
            by (simp add:cntV_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1680
          moreover from eq_e neq_th have "threads (e#s) = threads s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1681
          moreover note ih [of th] 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1682
          ultimately have ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1683
        } moreover {
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1684
          assume eq_th: "th = thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1685
          have ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1686
          proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1687
            from eq_e eq_th have eq_cnp: "cntP (e # s) th  = 1 + (cntP s th)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1688
              by (simp add:cntP_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1689
            from eq_e eq_th have eq_cnv: "cntV (e#s) th = cntV s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1690
              by (simp add:cntV_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1691
            show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1692
            proof (cases "wq s cs = []")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1693
              case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1694
              with is_runing
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1695
              have "th \<in> readys (e#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1696
                apply (unfold eq_e wq_def, unfold readys_def s_depend_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1697
                apply (simp add: wq_def[symmetric] runing_def eq_th s_waiting_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1698
                by (auto simp:readys_def wq_def Let_def s_waiting_def wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1699
              moreover have "cntCS (e # s) th = 1 + cntCS s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1700
              proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1701
                have "card {csa. csa = cs \<or> (Cs csa, Th thread) \<in> depend s} =
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1702
                  Suc (card {cs. (Cs cs, Th thread) \<in> depend s})" (is "card ?L = Suc (card ?R)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1703
                proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1704
                  have "?L = insert cs ?R" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1705
                  moreover have "card \<dots> = Suc (card (?R - {cs}))" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1706
                  proof(rule card_insert)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1707
                    from finite_holding [OF vt, of thread]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1708
                    show " finite {cs. (Cs cs, Th thread) \<in> depend s}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1709
                      by (unfold holdents_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1710
                  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1711
                  moreover have "?R - {cs} = ?R"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1712
                  proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1713
                    have "cs \<notin> ?R"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1714
                    proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1715
                      assume "cs \<in> {cs. (Cs cs, Th thread) \<in> depend s}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1716
                      with no_dep show False by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1717
                    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1718
                    thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1719
                  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1720
                  ultimately show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1721
                qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1722
                thus ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1723
                  apply (unfold eq_e eq_th cntCS_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1724
                  apply (simp add: holdents_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1725
                  by (unfold step_depend_p [OF vtp], auto simp:True)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1726
              qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1727
              moreover from is_runing have "th \<in> readys s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1728
                by (simp add:runing_def eq_th)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1729
              moreover note eq_cnp eq_cnv ih [of th]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1730
              ultimately show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1731
            next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1732
              case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1733
              have eq_wq: "wq (e#s) cs = wq s cs @ [th]"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1734
                    by (unfold eq_th eq_e wq_def, auto simp:Let_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1735
              have "th \<notin> readys (e#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1736
              proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1737
                assume "th \<in> readys (e#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1738
                hence "\<forall>cs. \<not> waiting (e # s) th cs" by (simp add:readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1739
                from this[rule_format, of cs] have " \<not> waiting (e # s) th cs" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1740
                hence "th \<in> set (wq (e#s) cs) \<Longrightarrow> th = hd (wq (e#s) cs)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1741
                  by (simp add:s_waiting_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1742
                moreover from eq_wq have "th \<in> set (wq (e#s) cs)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1743
                ultimately have "th = hd (wq (e#s) cs)" by blast
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1744
                with eq_wq have "th = hd (wq s cs @ [th])" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1745
                hence "th = hd (wq s cs)" using False by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1746
                with False eq_wq wq_distinct [OF vtp, of cs]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1747
                show False by (fold eq_e, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1748
              qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1749
              moreover from is_runing have "th \<in> threads (e#s)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1750
                by (unfold eq_e, auto simp:runing_def readys_def eq_th)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1751
              moreover have "cntCS (e # s) th = cntCS s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1752
                apply (unfold cntCS_def holdents_def eq_e step_depend_p[OF vtp])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1753
                by (auto simp:False)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1754
              moreover note eq_cnp eq_cnv ih[of th]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1755
              moreover from is_runing have "th \<in> readys s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1756
                by (simp add:runing_def eq_th)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1757
              ultimately show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1758
            qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1759
          qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1760
        } ultimately show ?thesis by blast
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1761
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1762
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1763
      case (thread_V thread cs)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1764
      from prems have vtv: "vt step (V thread cs # s)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1765
      assume eq_e: "e = V thread cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1766
        and is_runing: "thread \<in> runing s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1767
        and hold: "holding s thread cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1768
      from hold obtain rest 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1769
        where eq_wq: "wq s cs = thread # rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1770
        by (case_tac "wq s cs", auto simp:s_holding_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1771
      have eq_threads: "threads (e#s) = threads s" by (simp add: eq_e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1772
      show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1773
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1774
        { assume eq_th: "th = thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1775
          from eq_th have eq_cnp: "cntP (e # s) th = cntP s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1776
            by (unfold eq_e, simp add:cntP_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1777
          moreover from eq_th have eq_cnv: "cntV (e#s) th = 1 + cntV s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1778
            by (unfold eq_e, simp add:cntV_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1779
          moreover from cntCS_v_dec [OF vtv] 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1780
          have "cntCS (e # s) thread + 1 = cntCS s thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1781
            by (simp add:eq_e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1782
          moreover from is_runing have rd_before: "thread \<in> readys s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1783
            by (unfold runing_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1784
          moreover have "thread \<in> readys (e # s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1785
          proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1786
            from is_runing
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1787
            have "thread \<in> threads (e#s)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1788
              by (unfold eq_e, auto simp:runing_def readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1789
            moreover have "\<forall> cs1. \<not> waiting (e#s) thread cs1"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1790
            proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1791
              fix cs1
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1792
              { assume eq_cs: "cs1 = cs" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1793
                have "\<not> waiting (e # s) thread cs1"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1794
                proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1795
                  have "thread \<notin> set (wq (e#s) cs1)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1796
                  proof(cases "lsp (cp s) rest")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1797
                    fix l m r 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1798
                    assume h: "lsp (cp s) rest = (l, m, r)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1799
                    show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1800
                    proof(cases "m")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1801
                      case Nil
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1802
                      from wq_v_eq_nil [OF eq_wq] h Nil eq_e
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1803
                      have " wq (e # s) cs = []" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1804
                      thus ?thesis using eq_cs by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1805
                    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1806
                      case (Cons th' l')
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1807
                      from lsp_mid_length [OF h] and Cons h
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1808
                      have eqh: "lsp (cp s) rest = (l, [th'], r)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1809
                      from wq_v_eq [OF eq_wq this]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1810
                      have "wq (V thread cs # s) cs = th' # l @ r" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1811
                      moreover from lsp_set_eq [OF eqh]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1812
                      have "set rest = set \<dots>" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1813
                      moreover have "thread \<notin> set rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1814
                      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1815
                        from wq_distinct [OF step_back_vt[OF vtv], of cs]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1816
                        and eq_wq show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1817
                      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1818
                      moreover note eq_e eq_cs
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1819
                      ultimately show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1820
                    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1821
                  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1822
                  thus ?thesis by (simp add:s_waiting_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1823
                qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1824
              } moreover {
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1825
                assume neq_cs: "cs1 \<noteq> cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1826
                  have "\<not> waiting (e # s) thread cs1" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1827
                  proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1828
                    from wq_v_neq [OF neq_cs[symmetric]]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1829
                    have "wq (V thread cs # s) cs1 = wq s cs1" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1830
                    moreover have "\<not> waiting s thread cs1" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1831
                    proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1832
                      from runing_ready and is_runing
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1833
                      have "thread \<in> readys s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1834
                      thus ?thesis by (simp add:readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1835
                    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1836
                    ultimately show ?thesis 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1837
                      by (auto simp:s_waiting_def eq_e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1838
                  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1839
              } ultimately show "\<not> waiting (e # s) thread cs1" by blast
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1840
            qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1841
            ultimately show ?thesis by (simp add:readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1842
          qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1843
          moreover note eq_th ih
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1844
          ultimately have ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1845
        } moreover {
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1846
          assume neq_th: "th \<noteq> thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1847
          from neq_th eq_e have eq_cnp: "cntP (e # s) th = cntP s th" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1848
            by (simp add:cntP_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1849
          from neq_th eq_e have eq_cnv: "cntV (e # s) th = cntV s th" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1850
            by (simp add:cntV_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1851
          have ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1852
          proof(cases "th \<in> set rest")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1853
            case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1854
            have "(th \<in> readys (e # s)) = (th \<in> readys s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1855
              by(unfold eq_e, rule readys_v_eq [OF neq_th eq_wq False])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1856
            moreover have "cntCS (e#s) th = cntCS s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1857
              by(unfold eq_e, rule cntCS_v_eq [OF neq_th eq_wq False vtv]) 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1858
            moreover note ih eq_cnp eq_cnv eq_threads
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1859
            ultimately show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1860
          next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1861
            case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1862
            obtain l m r where eq_lsp: "lsp (cp s) rest = (l, m, r)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1863
              by (cases "lsp (cp s) rest", auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1864
            with True have "m \<noteq> []" by  (auto dest:lsp_mid_nil)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1865
            with eq_lsp obtain th' where eq_lsp: "lsp (cp s) rest = (l, [th'], r)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1866
              by (case_tac m, auto dest:lsp_mid_length)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1867
            show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1868
            proof(cases "th = th'")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1869
              case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1870
              have "(th \<in> readys (e # s)) = (th \<in> readys s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1871
                by (unfold eq_e, rule readys_v_eq_1 [OF neq_th eq_wq eq_lsp False])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1872
              moreover have "cntCS (e#s) th = cntCS s th" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1873
                by (unfold eq_e, rule cntCS_v_eq_1[OF neq_th eq_wq eq_lsp False vtv])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1874
              moreover note ih eq_cnp eq_cnv eq_threads
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1875
              ultimately show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1876
            next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1877
              case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1878
              have "th \<in> readys (e # s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1879
                by (unfold eq_e, rule readys_v_eq_2 [OF neq_th eq_wq eq_lsp True vt])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1880
              moreover have "cntP s th = cntV s th + cntCS s th + 1"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1881
              proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1882
                have "th \<notin> readys s" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1883
                proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1884
                  from True eq_wq lsp_set_eq [OF eq_lsp] neq_th
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1885
                  show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1886
                    apply (unfold readys_def s_waiting_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1887
                    by (rule_tac x = cs in exI, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1888
                qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1889
                moreover have "th \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1890
                proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1891
                  from True eq_wq lsp_set_eq [OF eq_lsp] neq_th
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1892
                  have "th \<in> set (wq s cs)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1893
                  from wq_threads [OF step_back_vt[OF vtv] this] 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1894
                  show ?thesis .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1895
                qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1896
                ultimately show ?thesis using ih by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1897
              qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1898
              moreover have "cntCS (e # s) th = 1 + cntCS s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1899
                by (unfold eq_e, rule cntCS_v_eq_2 [OF neq_th eq_wq eq_lsp True vtv])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1900
              moreover note eq_cnp eq_cnv
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1901
              ultimately show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1902
            qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1903
          qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1904
        } ultimately show ?thesis by blast
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1905
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1906
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1907
      case (thread_set thread prio)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1908
      assume eq_e: "e = Set thread prio"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1909
        and is_runing: "thread \<in> runing s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1910
      show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1911
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1912
        from eq_e have eq_cnp: "cntP (e#s) th = cntP s th" by (simp add:cntP_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1913
        from eq_e have eq_cnv: "cntV (e#s) th = cntV s th" by (simp add:cntV_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1914
        have eq_cncs: "cntCS (e#s) th = cntCS s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1915
          unfolding cntCS_def holdents_def
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1916
          by (simp add:depend_set_unchanged eq_e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1917
        from eq_e have eq_readys: "readys (e#s) = readys s" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1918
          by (simp add:readys_def cs_waiting_def s_waiting_def wq_def,
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1919
                  auto simp:Let_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1920
        { assume "th \<noteq> thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1921
          with eq_readys eq_e
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1922
          have "(th \<in> readys (e # s) \<or> th \<notin> threads (e # s)) = 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1923
                      (th \<in> readys (s) \<or> th \<notin> threads (s))" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1924
            by (simp add:threads.simps)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1925
          with eq_cnp eq_cnv eq_cncs ih is_runing
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1926
          have ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1927
        } moreover {
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1928
          assume eq_th: "th = thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1929
          with is_runing ih have " cntP s th  = cntV s th + cntCS s th" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1930
            by (unfold runing_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1931
          moreover from eq_th and eq_readys is_runing have "th \<in> readys (e#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1932
            by (simp add:runing_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1933
          moreover note eq_cnp eq_cnv eq_cncs
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1934
          ultimately have ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1935
        } ultimately show ?thesis by blast
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1936
      qed   
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1937
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1938
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1939
    case vt_nil
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1940
    show ?case 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1941
      by (unfold cntP_def cntV_def cntCS_def, 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1942
        auto simp:count_def holdents_def s_depend_def wq_def cs_holding_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1943
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1944
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1945
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1946
lemma not_thread_cncs:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1947
  fixes th s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1948
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1949
  and not_in: "th \<notin> threads s" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1950
  shows "cntCS s th = 0"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1951
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1952
  from vt not_in show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1953
  proof(induct arbitrary:th)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1954
    case (vt_cons s e th)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1955
    assume vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1956
      and ih: "\<And>th. th \<notin> threads s \<Longrightarrow> cntCS s th = 0"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1957
      and stp: "step s e"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1958
      and not_in: "th \<notin> threads (e # s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1959
    from stp show ?case
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1960
    proof(cases)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1961
      case (thread_create prio max_prio thread)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1962
      assume eq_e: "e = Create thread prio"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1963
        and not_in': "thread \<notin> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1964
      have "cntCS (e # s) th = cntCS s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1965
        apply (unfold eq_e cntCS_def holdents_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1966
        by (simp add:depend_create_unchanged)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1967
      moreover have "th \<notin> threads s" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1968
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1969
        from not_in eq_e show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1970
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1971
      moreover note ih ultimately show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1972
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1973
      case (thread_exit thread)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1974
      assume eq_e: "e = Exit thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1975
      and nh: "holdents s thread = {}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1976
      have eq_cns: "cntCS (e # s) th = cntCS s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1977
        apply (unfold eq_e cntCS_def holdents_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1978
        by (simp add:depend_exit_unchanged)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1979
      show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1980
      proof(cases "th = thread")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1981
        case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1982
        have "cntCS s th = 0" by (unfold cntCS_def, auto simp:nh True)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1983
        with eq_cns show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1984
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1985
        case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1986
        with not_in and eq_e
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1987
        have "th \<notin> threads s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1988
        from ih[OF this] and eq_cns show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1989
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1990
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1991
      case (thread_P thread cs)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1992
      assume eq_e: "e = P thread cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1993
      and is_runing: "thread \<in> runing s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1994
      from prems have vtp: "vt step (P thread cs#s)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1995
      have neq_th: "th \<noteq> thread" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1996
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1997
        from not_in eq_e have "th \<notin> threads s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1998
        moreover from is_runing have "thread \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  1999
          by (simp add:runing_def readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2000
        ultimately show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2001
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2002
      hence "cntCS (e # s) th  = cntCS s th "
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2003
        apply (unfold cntCS_def holdents_def eq_e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2004
        by (unfold step_depend_p[OF vtp], auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2005
      moreover have "cntCS s th = 0"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2006
      proof(rule ih)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2007
        from not_in eq_e show "th \<notin> threads s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2008
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2009
      ultimately show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2010
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2011
      case (thread_V thread cs)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2012
      assume eq_e: "e = V thread cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2013
        and is_runing: "thread \<in> runing s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2014
        and hold: "holding s thread cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2015
      have neq_th: "th \<noteq> thread" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2016
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2017
        from not_in eq_e have "th \<notin> threads s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2018
        moreover from is_runing have "thread \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2019
          by (simp add:runing_def readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2020
        ultimately show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2021
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2022
      from prems have vtv: "vt step (V thread cs#s)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2023
      from hold obtain rest 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2024
        where eq_wq: "wq s cs = thread # rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2025
        by (case_tac "wq s cs", auto simp:s_holding_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2026
      have "cntCS (e # s) th  = cntCS s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2027
      proof(unfold eq_e, rule cntCS_v_eq [OF neq_th eq_wq _ vtv])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2028
        show "th \<notin> set rest" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2029
        proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2030
          assume "th \<in> set rest"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2031
          with eq_wq have "th \<in> set (wq s cs)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2032
          from wq_threads [OF vt this] eq_e not_in 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2033
          show False by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2034
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2035
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2036
      moreover have "cntCS s th = 0"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2037
      proof(rule ih)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2038
        from not_in eq_e show "th \<notin> threads s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2039
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2040
      ultimately show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2041
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2042
      case (thread_set thread prio)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2043
      print_facts
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2044
      assume eq_e: "e = Set thread prio"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2045
        and is_runing: "thread \<in> runing s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2046
      from not_in and eq_e have "th \<notin> threads s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2047
      from ih [OF this] and eq_e
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2048
      show ?thesis 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2049
        apply (unfold eq_e cntCS_def holdents_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2050
        by (simp add:depend_set_unchanged)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2051
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2052
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2053
      case vt_nil
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2054
      show ?case
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2055
      by (unfold cntCS_def, 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2056
        auto simp:count_def holdents_def s_depend_def wq_def cs_holding_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2057
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2058
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2059
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2060
lemma eq_waiting: "waiting (wq (s::state)) th cs = waiting s th cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2061
  by (auto simp:s_waiting_def cs_waiting_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2062
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2063
lemma dm_depend_threads:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2064
  fixes th s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2065
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2066
  and in_dom: "(Th th) \<in> Domain (depend s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2067
  shows "th \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2068
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2069
  from in_dom obtain n where "(Th th, n) \<in> depend s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2070
  moreover from depend_target_th[OF this] obtain cs where "n = Cs cs" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2071
  ultimately have "(Th th, Cs cs) \<in> depend s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2072
  hence "th \<in> set (wq s cs)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2073
    by (unfold s_depend_def, auto simp:cs_waiting_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2074
  from wq_threads [OF vt this] show ?thesis .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2075
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2076
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2077
lemma cp_eq_cpreced: "cp s th = cpreced s (wq s) th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2078
proof(unfold cp_def wq_def, induct s)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2079
  case (Cons e s')
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2080
  show ?case
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2081
    by (auto simp:Let_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2082
next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2083
  case Nil
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2084
  show ?case by (auto simp:Let_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2085
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2086
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2087
fun the_th :: "node \<Rightarrow> thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2088
  where "the_th (Th th) = th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2089
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2090
lemma runing_unique:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2091
  fixes th1 th2 s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2092
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2093
  and runing_1: "th1 \<in> runing s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2094
  and runing_2: "th2 \<in> runing s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2095
  shows "th1 = th2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2096
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2097
  from runing_1 and runing_2 have "cp s th1 = cp s th2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2098
    by (unfold runing_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2099
  hence eq_max: "Max ((\<lambda>th. preced th s) ` ({th1} \<union> dependents (wq s) th1)) =
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2100
                 Max ((\<lambda>th. preced th s) ` ({th2} \<union> dependents (wq s) th2))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2101
    (is "Max (?f ` ?A) = Max (?f ` ?B)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2102
    by (unfold cp_eq_cpreced cpreced_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2103
  obtain th1' where th1_in: "th1' \<in> ?A" and eq_f_th1: "?f th1' = Max (?f ` ?A)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2104
  proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2105
    have h1: "finite (?f ` ?A)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2106
    proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2107
      have "finite ?A" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2108
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2109
        have "finite (dependents (wq s) th1)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2110
        proof-
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2111
          have "finite {th'. (Th th', Th th1) \<in> (depend (wq s))\<^sup>+}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2112
          proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2113
            let ?F = "\<lambda> (x, y). the_th x"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2114
            have "{th'. (Th th', Th th1) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (wq s))\<^sup>+)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2115
              apply (auto simp:image_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2116
              by (rule_tac x = "(Th x, Th th1)" in bexI, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2117
            moreover have "finite \<dots>"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2118
            proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2119
              from finite_depend[OF vt] have "finite (depend s)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2120
              hence "finite ((depend (wq s))\<^sup>+)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2121
                apply (unfold finite_trancl)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2122
                by (auto simp: s_depend_def cs_depend_def wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2123
              thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2124
            qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2125
            ultimately show ?thesis by (auto intro:finite_subset)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2126
          qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2127
          thus ?thesis by (simp add:cs_dependents_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2128
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2129
        thus ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2130
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2131
      thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2132
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2133
    moreover have h2: "(?f ` ?A) \<noteq> {}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2134
    proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2135
      have "?A \<noteq> {}" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2136
      thus ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2137
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2138
    from Max_in [OF h1 h2]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2139
    have "Max (?f ` ?A) \<in> (?f ` ?A)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2140
    thus ?thesis by (auto intro:that)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2141
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2142
  obtain th2' where th2_in: "th2' \<in> ?B" and eq_f_th2: "?f th2' = Max (?f ` ?B)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2143
  proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2144
    have h1: "finite (?f ` ?B)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2145
    proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2146
      have "finite ?B" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2147
      proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2148
        have "finite (dependents (wq s) th2)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2149
        proof-
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2150
          have "finite {th'. (Th th', Th th2) \<in> (depend (wq s))\<^sup>+}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2151
          proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2152
            let ?F = "\<lambda> (x, y). the_th x"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2153
            have "{th'. (Th th', Th th2) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (wq s))\<^sup>+)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2154
              apply (auto simp:image_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2155
              by (rule_tac x = "(Th x, Th th2)" in bexI, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2156
            moreover have "finite \<dots>"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2157
            proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2158
              from finite_depend[OF vt] have "finite (depend s)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2159
              hence "finite ((depend (wq s))\<^sup>+)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2160
                apply (unfold finite_trancl)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2161
                by (auto simp: s_depend_def cs_depend_def wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2162
              thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2163
            qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2164
            ultimately show ?thesis by (auto intro:finite_subset)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2165
          qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2166
          thus ?thesis by (simp add:cs_dependents_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2167
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2168
        thus ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2169
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2170
      thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2171
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2172
    moreover have h2: "(?f ` ?B) \<noteq> {}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2173
    proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2174
      have "?B \<noteq> {}" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2175
      thus ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2176
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2177
    from Max_in [OF h1 h2]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2178
    have "Max (?f ` ?B) \<in> (?f ` ?B)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2179
    thus ?thesis by (auto intro:that)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2180
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2181
  from eq_f_th1 eq_f_th2 eq_max 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2182
  have eq_preced: "preced th1' s = preced th2' s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2183
  hence eq_th12: "th1' = th2'"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2184
  proof (rule preced_unique)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2185
    from th1_in have "th1' = th1 \<or> (th1' \<in> dependents (wq s) th1)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2186
    thus "th1' \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2187
    proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2188
      assume "th1' \<in> dependents (wq s) th1"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2189
      hence "(Th th1') \<in> Domain ((depend s)^+)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2190
        apply (unfold cs_dependents_def cs_depend_def s_depend_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2191
        by (auto simp:Domain_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2192
      hence "(Th th1') \<in> Domain (depend s)" by (simp add:trancl_domain)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2193
      from dm_depend_threads[OF vt this] show ?thesis .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2194
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2195
      assume "th1' = th1"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2196
      with runing_1 show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2197
        by (unfold runing_def readys_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2198
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2199
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2200
    from th2_in have "th2' = th2 \<or> (th2' \<in> dependents (wq s) th2)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2201
    thus "th2' \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2202
    proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2203
      assume "th2' \<in> dependents (wq s) th2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2204
      hence "(Th th2') \<in> Domain ((depend s)^+)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2205
        apply (unfold cs_dependents_def cs_depend_def s_depend_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2206
        by (auto simp:Domain_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2207
      hence "(Th th2') \<in> Domain (depend s)" by (simp add:trancl_domain)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2208
      from dm_depend_threads[OF vt this] show ?thesis .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2209
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2210
      assume "th2' = th2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2211
      with runing_2 show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2212
        by (unfold runing_def readys_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2213
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2214
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2215
  from th1_in have "th1' = th1 \<or> th1' \<in> dependents (wq s) th1" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2216
  thus ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2217
  proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2218
    assume eq_th': "th1' = th1"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2219
    from th2_in have "th2' = th2 \<or> th2' \<in> dependents (wq s) th2" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2220
    thus ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2221
    proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2222
      assume "th2' = th2" thus ?thesis using eq_th' eq_th12 by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2223
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2224
      assume "th2' \<in> dependents (wq s) th2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2225
      with eq_th12 eq_th' have "th1 \<in> dependents (wq s) th2" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2226
      hence "(Th th1, Th th2) \<in> (depend s)^+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2227
        by (unfold cs_dependents_def s_depend_def cs_depend_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2228
      hence "Th th1 \<in> Domain ((depend s)^+)" using Domain_def [of "(depend s)^+"]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2229
        by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2230
      hence "Th th1 \<in> Domain (depend s)" by (simp add:trancl_domain)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2231
      then obtain n where d: "(Th th1, n) \<in> depend s" by (auto simp:Domain_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2232
      from depend_target_th [OF this]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2233
      obtain cs' where "n = Cs cs'" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2234
      with d have "(Th th1, Cs cs') \<in> depend s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2235
      with runing_1 have "False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2236
        apply (unfold runing_def readys_def s_depend_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2237
        by (auto simp:eq_waiting)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2238
      thus ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2239
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2240
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2241
    assume th1'_in: "th1' \<in> dependents (wq s) th1"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2242
    from th2_in have "th2' = th2 \<or> th2' \<in> dependents (wq s) th2" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2243
    thus ?thesis 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2244
    proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2245
      assume "th2' = th2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2246
      with th1'_in eq_th12 have "th2 \<in> dependents (wq s) th1" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2247
      hence "(Th th2, Th th1) \<in> (depend s)^+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2248
        by (unfold cs_dependents_def s_depend_def cs_depend_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2249
      hence "Th th2 \<in> Domain ((depend s)^+)" using Domain_def [of "(depend s)^+"]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2250
        by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2251
      hence "Th th2 \<in> Domain (depend s)" by (simp add:trancl_domain)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2252
      then obtain n where d: "(Th th2, n) \<in> depend s" by (auto simp:Domain_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2253
      from depend_target_th [OF this]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2254
      obtain cs' where "n = Cs cs'" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2255
      with d have "(Th th2, Cs cs') \<in> depend s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2256
      with runing_2 have "False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2257
        apply (unfold runing_def readys_def s_depend_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2258
        by (auto simp:eq_waiting)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2259
      thus ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2260
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2261
      assume "th2' \<in> dependents (wq s) th2"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2262
      with eq_th12 have "th1' \<in> dependents (wq s) th2" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2263
      hence h1: "(Th th1', Th th2) \<in> (depend s)^+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2264
        by (unfold cs_dependents_def s_depend_def cs_depend_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2265
      from th1'_in have h2: "(Th th1', Th th1) \<in> (depend s)^+"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2266
        by (unfold cs_dependents_def s_depend_def cs_depend_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2267
      show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2268
      proof(rule dchain_unique[OF vt h1 _ h2, symmetric])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2269
        from runing_1 show "th1 \<in> readys s" by (simp add:runing_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2270
        from runing_2 show "th2 \<in> readys s" by (simp add:runing_def) 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2271
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2272
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2273
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2274
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2275
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2276
lemma create_pre:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2277
  assumes stp: "step s e"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2278
  and not_in: "th \<notin> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2279
  and is_in: "th \<in> threads (e#s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2280
  obtains prio where "e = Create th prio"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2281
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2282
  from assms  
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2283
  show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2284
  proof(cases)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2285
    case (thread_create prio max_prio thread)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2286
    with is_in not_in have "e = Create th prio" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2287
    from that[OF this] show ?thesis .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2288
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2289
    case (thread_exit thread)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2290
    with assms show ?thesis by (auto intro!:that)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2291
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2292
    case (thread_P thread)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2293
    with assms show ?thesis by (auto intro!:that)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2294
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2295
    case (thread_V thread)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2296
    with assms show ?thesis by (auto intro!:that)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2297
  next 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2298
    case (thread_set thread)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2299
    with assms show ?thesis by (auto intro!:that)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2300
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2301
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2302
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2303
lemma length_down_to_in: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2304
  assumes le_ij: "i \<le> j"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2305
    and le_js: "j \<le> length s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2306
  shows "length (down_to j i s) = j - i"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2307
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2308
  have "length (down_to j i s) = length (from_to i j (rev s))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2309
    by (unfold down_to_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2310
  also have "\<dots> = j - i"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2311
  proof(rule length_from_to_in[OF le_ij])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2312
    from le_js show "j \<le> length (rev s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2313
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2314
  finally show ?thesis .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2315
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2316
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2317
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2318
lemma moment_head: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2319
  assumes le_it: "Suc i \<le> length t"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2320
  obtains e where "moment (Suc i) t = e#moment i t"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2321
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2322
  have "i \<le> Suc i" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2323
  from length_down_to_in [OF this le_it]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2324
  have "length (down_to (Suc i) i t) = 1" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2325
  then obtain e where "down_to (Suc i) i t = [e]"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2326
    apply (cases "(down_to (Suc i) i t)") by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2327
  moreover have "down_to (Suc i) 0 t = down_to (Suc i) i t @ down_to i 0 t"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2328
    by (rule down_to_conc[symmetric], auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2329
  ultimately have eq_me: "moment (Suc i) t = e#(moment i t)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2330
    by (auto simp:down_to_moment)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2331
  from that [OF this] show ?thesis .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2332
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2333
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2334
lemma cnp_cnv_eq:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2335
  fixes th s
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2336
  assumes "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2337
  and "th \<notin> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2338
  shows "cntP s th = cntV s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2339
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2340
  from assms show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2341
  proof(induct)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2342
    case (vt_cons s e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2343
    have ih: "th \<notin> threads s \<Longrightarrow> cntP s th = cntV s th" by fact
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2344
    have not_in: "th \<notin> threads (e # s)" by fact
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2345
    have "step s e" by fact
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2346
    thus ?case proof(cases)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2347
      case (thread_create prio max_prio thread)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2348
      assume eq_e: "e = Create thread prio"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2349
      hence "thread \<in> threads (e#s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2350
      with not_in and eq_e have "th \<notin> threads s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2351
      from ih [OF this] show ?thesis using eq_e
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2352
        by (auto simp:cntP_def cntV_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2353
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2354
      case (thread_exit thread)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2355
      assume eq_e: "e = Exit thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2356
        and not_holding: "holdents s thread = {}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2357
      have vt_s: "vt step s" by fact
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2358
      from finite_holding[OF vt_s] have "finite (holdents s thread)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2359
      with not_holding have "cntCS s thread = 0" by (unfold cntCS_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2360
      moreover have "thread \<in> readys s" using thread_exit by (auto simp:runing_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2361
      moreover note cnp_cnv_cncs[OF vt_s, of thread]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2362
      ultimately have eq_thread: "cntP s thread = cntV s thread" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2363
      show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2364
      proof(cases "th = thread")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2365
        case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2366
        with eq_thread eq_e show ?thesis 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2367
          by (auto simp:cntP_def cntV_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2368
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2369
        case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2370
        with not_in and eq_e have "th \<notin> threads s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2371
        from ih[OF this] and eq_e show ?thesis 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2372
           by (auto simp:cntP_def cntV_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2373
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2374
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2375
      case (thread_P thread cs)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2376
      assume eq_e: "e = P thread cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2377
      have "thread \<in> runing s" by fact
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2378
      with not_in eq_e have neq_th: "thread \<noteq> th" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2379
        by (auto simp:runing_def readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2380
      from not_in eq_e have "th \<notin> threads s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2381
      from ih[OF this] and neq_th and eq_e show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2382
        by (auto simp:cntP_def cntV_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2383
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2384
      case (thread_V thread cs)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2385
      assume eq_e: "e = V thread cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2386
      have "thread \<in> runing s" by fact
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2387
      with not_in eq_e have neq_th: "thread \<noteq> th" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2388
        by (auto simp:runing_def readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2389
      from not_in eq_e have "th \<notin> threads s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2390
      from ih[OF this] and neq_th and eq_e show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2391
        by (auto simp:cntP_def cntV_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2392
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2393
      case (thread_set thread prio)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2394
      assume eq_e: "e = Set thread prio"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2395
        and "thread \<in> runing s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2396
      hence "thread \<in> threads (e#s)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2397
        by (simp add:runing_def readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2398
      with not_in and eq_e have "th \<notin> threads s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2399
      from ih [OF this] show ?thesis using eq_e
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2400
        by (auto simp:cntP_def cntV_def count_def)  
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2401
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2402
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2403
    case vt_nil
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2404
    show ?case by (auto simp:cntP_def cntV_def count_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2405
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2406
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2407
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2408
lemma eq_depend: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2409
  "depend (wq s) = depend s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2410
by (unfold cs_depend_def s_depend_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2411
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2412
lemma count_eq_dependents:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2413
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2414
  and eq_pv: "cntP s th = cntV s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2415
  shows "dependents (wq s) th = {}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2416
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2417
  from cnp_cnv_cncs[OF vt] and eq_pv
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2418
  have "cntCS s th = 0" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2419
    by (auto split:if_splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2420
  moreover have "finite {cs. (Cs cs, Th th) \<in> depend s}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2421
  proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2422
    from finite_holding[OF vt, of th] show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2423
      by (simp add:holdents_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2424
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2425
  ultimately have h: "{cs. (Cs cs, Th th) \<in> depend s} = {}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2426
    by (unfold cntCS_def holdents_def cs_dependents_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2427
  show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2428
  proof(unfold cs_dependents_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2429
    { assume "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<noteq> {}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2430
      then obtain th' where "(Th th', Th th) \<in> (depend (wq s))\<^sup>+" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2431
      hence "False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2432
      proof(cases)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2433
        assume "(Th th', Th th) \<in> depend (wq s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2434
        thus "False" by (auto simp:cs_depend_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2435
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2436
        fix c
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2437
        assume "(c, Th th) \<in> depend (wq s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2438
        with h and eq_depend show "False"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2439
          by (cases c, auto simp:cs_depend_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2440
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2441
    } thus "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} = {}" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2442
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2443
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2444
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2445
lemma dependents_threads:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2446
  fixes s th
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2447
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2448
  shows "dependents (wq s) th \<subseteq> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2449
proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2450
  { fix th th'
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2451
    assume h: "th \<in> {th'a. (Th th'a, Th th') \<in> (depend (wq s))\<^sup>+}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2452
    have "Th th \<in> Domain (depend s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2453
    proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2454
      from h obtain th' where "(Th th, Th th') \<in> (depend (wq s))\<^sup>+" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2455
      hence "(Th th) \<in> Domain ( (depend (wq s))\<^sup>+)" by (auto simp:Domain_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2456
      with trancl_domain have "(Th th) \<in> Domain (depend (wq s))" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2457
      thus ?thesis using eq_depend by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2458
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2459
    from dm_depend_threads[OF vt this]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2460
    have "th \<in> threads s" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2461
  } note hh = this
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2462
  fix th1 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2463
  assume "th1 \<in> dependents (wq s) th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2464
  hence "th1 \<in> {th'a. (Th th'a, Th th) \<in> (depend (wq s))\<^sup>+}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2465
    by (unfold cs_dependents_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2466
  from hh [OF this] show "th1 \<in> threads s" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2467
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2468
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2469
lemma finite_threads:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2470
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2471
  shows "finite (threads s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2472
proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2473
  from vt show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2474
  proof(induct)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2475
    case (vt_cons s e)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2476
    assume vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2477
    and step: "step s e"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2478
    and ih: "finite (threads s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2479
    from step
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2480
    show ?case
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2481
    proof(cases)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2482
      case (thread_create prio max_prio thread)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2483
      assume eq_e: "e = Create thread prio"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2484
      with ih
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2485
      show ?thesis by (unfold eq_e, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2486
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2487
      case (thread_exit thread)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2488
      assume eq_e: "e = Exit thread"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2489
      with ih show ?thesis 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2490
        by (unfold eq_e, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2491
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2492
      case (thread_P thread cs)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2493
      assume eq_e: "e = P thread cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2494
      with ih show ?thesis by (unfold eq_e, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2495
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2496
      case (thread_V thread cs)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2497
      assume eq_e: "e = V thread cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2498
      with ih show ?thesis by (unfold eq_e, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2499
    next 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2500
      case (thread_set thread prio)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2501
      from vt_cons thread_set show ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2502
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2503
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2504
    case vt_nil
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2505
    show ?case by (auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2506
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2507
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2508
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2509
lemma Max_f_mono:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2510
  assumes seq: "A \<subseteq> B"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2511
  and np: "A \<noteq> {}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2512
  and fnt: "finite B"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2513
  shows "Max (f ` A) \<le> Max (f ` B)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2514
proof(rule Max_mono)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2515
  from seq show "f ` A \<subseteq> f ` B" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2516
next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2517
  from np show "f ` A \<noteq> {}" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2518
next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2519
  from fnt and seq show "finite (f ` B)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2520
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2521
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2522
lemma cp_le:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2523
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2524
  and th_in: "th \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2525
  shows "cp s th \<le> Max ((\<lambda> th. (preced th s)) ` threads s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2526
proof(unfold cp_eq_cpreced cpreced_def cs_dependents_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2527
  show "Max ((\<lambda>th. preced th s) ` ({th} \<union> {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+}))
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2528
         \<le> Max ((\<lambda>th. preced th s) ` threads s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2529
    (is "Max (?f ` ?A) \<le> Max (?f ` ?B)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2530
  proof(rule Max_f_mono)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2531
    show "{th} \<union> {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<noteq> {}" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2532
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2533
    from finite_threads [OF vt]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2534
    show "finite (threads s)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2535
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2536
    from th_in
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2537
    show "{th} \<union> {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<subseteq> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2538
      apply (auto simp:Domain_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2539
      apply (rule_tac dm_depend_threads[OF vt])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2540
      apply (unfold trancl_domain [of "depend s", symmetric])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2541
      by (unfold cs_depend_def s_depend_def, auto simp:Domain_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2542
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2543
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2544
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2545
lemma le_cp:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2546
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2547
  shows "preced th s \<le> cp s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2548
proof(unfold cp_eq_cpreced preced_def cpreced_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2549
  show "Prc (original_priority th s) (birthtime th s)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2550
    \<le> Max (insert (Prc (original_priority th s) (birthtime th s))
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2551
            ((\<lambda>th. Prc (original_priority th s) (birthtime th s)) ` dependents (wq s) th))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2552
    (is "?l \<le> Max (insert ?l ?A)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2553
  proof(cases "?A = {}")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2554
    case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2555
    have "finite ?A" (is "finite (?f ` ?B)")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2556
    proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2557
      have "finite ?B" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2558
      proof-
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2559
        have "finite {th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2560
        proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2561
          let ?F = "\<lambda> (x, y). the_th x"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2562
          have "{th'. (Th th', Th th) \<in> (depend (wq s))\<^sup>+} \<subseteq> ?F ` ((depend (wq s))\<^sup>+)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2563
            apply (auto simp:image_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2564
            by (rule_tac x = "(Th x, Th th)" in bexI, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2565
          moreover have "finite \<dots>"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2566
          proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2567
            from finite_depend[OF vt] have "finite (depend s)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2568
            hence "finite ((depend (wq s))\<^sup>+)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2569
              apply (unfold finite_trancl)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2570
              by (auto simp: s_depend_def cs_depend_def wq_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2571
            thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2572
          qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2573
          ultimately show ?thesis by (auto intro:finite_subset)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2574
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2575
        thus ?thesis by (simp add:cs_dependents_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2576
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2577
      thus ?thesis by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2578
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2579
    from Max_insert [OF this False, of ?l] show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2580
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2581
    case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2582
    thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2583
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2584
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2585
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2586
lemma max_cp_eq: 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2587
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2588
  shows "Max ((cp s) ` threads s) = Max ((\<lambda> th. (preced th s)) ` threads s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2589
  (is "?l = ?r")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2590
proof(cases "threads s = {}")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2591
  case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2592
  thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2593
next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2594
  case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2595
  have "?l \<in> ((cp s) ` threads s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2596
  proof(rule Max_in)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2597
    from finite_threads[OF vt] 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2598
    show "finite (cp s ` threads s)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2599
  next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2600
    from False show "cp s ` threads s \<noteq> {}" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2601
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2602
  then obtain th 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2603
    where th_in: "th \<in> threads s" and eq_l: "?l = cp s th" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2604
  have "\<dots> \<le> ?r" by (rule cp_le[OF vt th_in])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2605
  moreover have "?r \<le> cp s th" (is "Max (?f ` ?A) \<le> cp s th")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2606
  proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2607
    have "?r \<in> (?f ` ?A)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2608
    proof(rule Max_in)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2609
      from finite_threads[OF vt]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2610
      show " finite ((\<lambda>th. preced th s) ` threads s)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2611
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2612
      from False show " (\<lambda>th. preced th s) ` threads s \<noteq> {}" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2613
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2614
    then obtain th' where 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2615
      th_in': "th' \<in> ?A " and eq_r: "?r = ?f th'" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2616
    from le_cp [OF vt, of th']  eq_r
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2617
    have "?r \<le> cp s th'" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2618
    moreover have "\<dots> \<le> cp s th"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2619
    proof(fold eq_l)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2620
      show " cp s th' \<le> Max (cp s ` threads s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2621
      proof(rule Max_ge)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2622
        from th_in' show "cp s th' \<in> cp s ` threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2623
          by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2624
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2625
        from finite_threads[OF vt]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2626
        show "finite (cp s ` threads s)" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2627
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2628
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2629
    ultimately show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2630
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2631
  ultimately show ?thesis using eq_l by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2632
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2633
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2634
lemma max_cp_readys_threads_pre:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2635
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2636
  and np: "threads s \<noteq> {}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2637
  shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2638
proof(unfold max_cp_eq[OF vt])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2639
  show "Max (cp s ` readys s) = Max ((\<lambda>th. preced th s) ` threads s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2640
  proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2641
    let ?p = "Max ((\<lambda>th. preced th s) ` threads s)" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2642
    let ?f = "(\<lambda>th. preced th s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2643
    have "?p \<in> ((\<lambda>th. preced th s) ` threads s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2644
    proof(rule Max_in)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2645
      from finite_threads[OF vt] show "finite (?f ` threads s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2646
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2647
      from np show "?f ` threads s \<noteq> {}" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2648
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2649
    then obtain tm where tm_max: "?f tm = ?p" and tm_in: "tm \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2650
      by (auto simp:Image_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2651
    from th_chain_to_ready [OF vt tm_in]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2652
    have "tm \<in> readys s \<or> (\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (depend s)\<^sup>+)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2653
    thus ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2654
    proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2655
      assume "\<exists>th'. th' \<in> readys s \<and> (Th tm, Th th') \<in> (depend s)\<^sup>+ "
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2656
      then obtain th' where th'_in: "th' \<in> readys s" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2657
        and tm_chain:"(Th tm, Th th') \<in> (depend s)\<^sup>+" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2658
      have "cp s th' = ?f tm"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2659
      proof(subst cp_eq_cpreced, subst cpreced_def, rule Max_eqI)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2660
        from dependents_threads[OF vt] finite_threads[OF vt]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2661
        show "finite ((\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th'))" 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2662
          by (auto intro:finite_subset)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2663
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2664
        fix p assume p_in: "p \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2665
        from tm_max have " preced tm s = Max ((\<lambda>th. preced th s) ` threads s)" .
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2666
        moreover have "p \<le> \<dots>"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2667
        proof(rule Max_ge)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2668
          from finite_threads[OF vt]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2669
          show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2670
        next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2671
          from p_in and th'_in and dependents_threads[OF vt, of th']
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2672
          show "p \<in> (\<lambda>th. preced th s) ` threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2673
            by (auto simp:readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2674
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2675
        ultimately show "p \<le> preced tm s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2676
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2677
        show "preced tm s \<in> (\<lambda>th. preced th s) ` ({th'} \<union> dependents (wq s) th')"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2678
        proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2679
          from tm_chain
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2680
          have "tm \<in> dependents (wq s) th'"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2681
            by (unfold cs_dependents_def s_depend_def cs_depend_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2682
          thus ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2683
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2684
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2685
      with tm_max
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2686
      have h: "cp s th' = Max ((\<lambda>th. preced th s) ` threads s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2687
      show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2688
      proof (fold h, rule Max_eqI)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2689
        fix q 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2690
        assume "q \<in> cp s ` readys s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2691
        then obtain th1 where th1_in: "th1 \<in> readys s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2692
          and eq_q: "q = cp s th1" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2693
        show "q \<le> cp s th'"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2694
          apply (unfold h eq_q)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2695
          apply (unfold cp_eq_cpreced cpreced_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2696
          apply (rule Max_mono)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2697
        proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2698
          from dependents_threads [OF vt, of th1] th1_in
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2699
          show "(\<lambda>th. preced th s) ` ({th1} \<union> dependents (wq s) th1) \<subseteq> 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2700
                 (\<lambda>th. preced th s) ` threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2701
            by (auto simp:readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2702
        next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2703
          show "(\<lambda>th. preced th s) ` ({th1} \<union> dependents (wq s) th1) \<noteq> {}" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2704
        next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2705
          from finite_threads[OF vt] 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2706
          show " finite ((\<lambda>th. preced th s) ` threads s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2707
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2708
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2709
        from finite_threads[OF vt]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2710
        show "finite (cp s ` readys s)" by (auto simp:readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2711
      next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2712
        from th'_in
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2713
        show "cp s th' \<in> cp s ` readys s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2714
      qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2715
    next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2716
      assume tm_ready: "tm \<in> readys s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2717
      show ?thesis
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2718
      proof(fold tm_max)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2719
        have cp_eq_p: "cp s tm = preced tm s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2720
        proof(unfold cp_eq_cpreced cpreced_def, rule Max_eqI)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2721
          fix y 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2722
          assume hy: "y \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependents (wq s) tm)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2723
          show "y \<le> preced tm s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2724
          proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2725
            { fix y'
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2726
              assume hy' : "y' \<in> ((\<lambda>th. preced th s) ` dependents (wq s) tm)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2727
              have "y' \<le> preced tm s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2728
              proof(unfold tm_max, rule Max_ge)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2729
                from hy' dependents_threads[OF vt, of tm]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2730
                show "y' \<in> (\<lambda>th. preced th s) ` threads s" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2731
              next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2732
                from finite_threads[OF vt] 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2733
                show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2734
              qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2735
            } with hy show ?thesis by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2736
          qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2737
        next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2738
          from dependents_threads[OF vt, of tm] finite_threads[OF vt]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2739
          show "finite ((\<lambda>th. preced th s) ` ({tm} \<union> dependents (wq s) tm))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2740
            by (auto intro:finite_subset)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2741
        next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2742
          show "preced tm s \<in> (\<lambda>th. preced th s) ` ({tm} \<union> dependents (wq s) tm)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2743
            by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2744
        qed 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2745
        moreover have "Max (cp s ` readys s) = cp s tm"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2746
        proof(rule Max_eqI)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2747
          from tm_ready show "cp s tm \<in> cp s ` readys s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2748
        next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2749
          from finite_threads[OF vt]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2750
          show "finite (cp s ` readys s)" by (auto simp:readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2751
        next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2752
          fix y assume "y \<in> cp s ` readys s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2753
          then obtain th1 where th1_readys: "th1 \<in> readys s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2754
            and h: "y = cp s th1" by auto
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2755
          show "y \<le> cp s tm"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2756
            apply(unfold cp_eq_p h)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2757
            apply(unfold cp_eq_cpreced cpreced_def tm_max, rule Max_mono)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2758
          proof -
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2759
            from finite_threads[OF vt]
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2760
            show "finite ((\<lambda>th. preced th s) ` threads s)" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2761
          next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2762
            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependents (wq s) th1) \<noteq> {}"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2763
              by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2764
          next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2765
            from dependents_threads[OF vt, of th1] th1_readys
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2766
            show "(\<lambda>th. preced th s) ` ({th1} \<union> dependents (wq s) th1) 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2767
                    \<subseteq> (\<lambda>th. preced th s) ` threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2768
              by (auto simp:readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2769
          qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2770
        qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2771
        ultimately show " Max (cp s ` readys s) = preced tm s" by simp
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2772
      qed 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2773
    qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2774
  qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2775
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2776
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2777
lemma max_cp_readys_threads:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2778
  assumes vt: "vt step s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2779
  shows "Max (cp s ` readys s) = Max (cp s ` threads s)"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2780
proof(cases "threads s = {}")
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2781
  case True
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2782
  thus ?thesis 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2783
    by (auto simp:readys_def)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2784
next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2785
  case False
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2786
  show ?thesis by (rule max_cp_readys_threads_pre[OF vt False])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2787
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2788
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2789
lemma readys_threads:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2790
  shows "readys s \<subseteq> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2791
proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2792
  fix th
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2793
  assume "th \<in> readys s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2794
  thus "th \<in> threads s"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2795
    by (unfold readys_def, auto)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2796
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2797
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2798
lemma eq_holding: "holding (wq s) th cs = holding s th cs"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2799
  apply (unfold s_holding_def cs_holding_def, simp)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2800
  done
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2801
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2802
lemma f_image_eq:
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2803
  assumes h: "\<And> a. a \<in> A \<Longrightarrow> f a = g a"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2804
  shows "f ` A = g ` A"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2805
proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2806
  show "f ` A \<subseteq> g ` A"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2807
    by(rule image_subsetI, auto intro:h)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2808
next
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2809
  show "g ` A \<subseteq> f ` A"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2810
   by(rule image_subsetI, auto intro:h[symmetric])
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2811
qed
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2812
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
  2813
end