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