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