131 The following lemma shows that @{term "th"} is not in the |
103 The following lemma shows that @{term "th"} is not in the |
132 sub-tree of any other thread. |
104 sub-tree of any other thread. |
133 *} |
105 *} |
134 lemma th_in_no_subtree: |
106 lemma th_in_no_subtree: |
135 assumes "th' \<noteq> th" |
107 assumes "th' \<noteq> th" |
136 shows "Th th \<notin> subtree (RAG s') (Th th')" |
108 shows "Th th \<notin> subtree (RAG s) (Th th')" |
137 proof - |
109 proof - |
138 have "th \<in> readys s'" |
110 from readys_in_no_subtree[OF th_ready_s assms(1)] |
139 proof - |
|
140 from step_back_step [OF vt_s[unfolded s_def]] |
|
141 have "step s' (Set th prio)" . |
|
142 hence "th \<in> runing s'" by (cases, simp) |
|
143 thus ?thesis by (simp add:readys_def runing_def) |
|
144 qed |
|
145 from vat_s'.readys_in_no_subtree[OF this assms(1)] |
|
146 show ?thesis by blast |
111 show ?thesis by blast |
147 qed |
112 qed |
148 |
113 |
149 text {* |
114 text {* |
150 By combining @{thm "eq_cp_pre"} and @{thm "th_in_no_subtree"}, |
115 By combining @{thm "eq_cp_pre"} and @{thm "th_in_no_subtree"}, |
151 it is obvious that the change of priority only affects the @{text "cp"}-value |
116 it is obvious that the change of priority only affects the @{text "cp"}-value |
152 of the initiating thread @{text "th"}. |
117 of the initiating thread @{text "th"}. |
153 *} |
118 *} |
154 lemma eq_cp: |
119 lemma eq_cp: |
155 assumes "th' \<noteq> th" |
120 assumes "th' \<noteq> th" |
156 shows "cp s th' = cp s' th'" |
121 shows "cp (e#s) th' = cp s th'" |
157 by (rule eq_cp_pre[OF th_in_no_subtree[OF assms]]) |
122 by (rule eq_cp_pre[OF th_in_no_subtree[OF assms]]) |
158 |
123 |
159 end |
124 end |
160 |
125 |
161 section {* The @{term V} operation *} |
126 section {* The @{term V} operation *} |
162 |
127 |
163 text {* |
128 text {* |
164 The following @{text "step_v_cps"} is the locale for @{text "V"}-operation. |
129 The following @{text "step_v_cps"} is the locale for @{text "V"}-operation. |
165 *} |
130 *} |
166 |
131 |
167 locale step_v_cps = |
132 |
168 -- {* @{text "th"} is the initiating thread *} |
133 context valid_trace_v |
169 -- {* @{text "cs"} is the critical resource release by the @{text "V"}-operation *} |
134 begin |
170 fixes s' th cs s -- {* @{text "s'"} is the state before operation*} |
135 |
171 defines s_def : "s \<equiv> (V th cs#s')" -- {* @{text "s"} is the state after operation*} |
136 lemma ancestors_th: "ancestors (RAG s) (Th th) = {}" |
172 -- {* @{text "s"} is assumed to be valid, which implies the validity of @{text "s'"} *} |
137 proof - |
173 assumes vt_s: "vt s" |
138 from readys_root[OF th_ready_s] |
174 |
|
175 sublocale step_v_cps < vat_s : valid_trace "s" |
|
176 proof |
|
177 from vt_s show "vt s" . |
|
178 qed |
|
179 |
|
180 sublocale step_v_cps < vat_s' : valid_trace "s'" |
|
181 proof |
|
182 from step_back_vt[OF vt_s[unfolded s_def]] show "vt s'" . |
|
183 qed |
|
184 |
|
185 context step_v_cps |
|
186 begin |
|
187 |
|
188 lemma ready_th_s': "th \<in> readys s'" |
|
189 using step_back_step[OF vt_s[unfolded s_def]] |
|
190 by (cases, simp add:runing_def) |
|
191 |
|
192 lemma ancestors_th: "ancestors (RAG s') (Th th) = {}" |
|
193 proof - |
|
194 from vat_s'.readys_root[OF ready_th_s'] |
|
195 show ?thesis |
139 show ?thesis |
196 by (unfold root_def, simp) |
140 by (unfold root_def, simp) |
197 qed |
141 qed |
198 |
142 |
199 lemma holding_th: "holding s' th cs" |
|
200 proof - |
|
201 from vt_s[unfolded s_def] |
|
202 have " PIP s' (V th cs)" by (cases, simp) |
|
203 thus ?thesis by (cases, auto) |
|
204 qed |
|
205 |
|
206 lemma edge_of_th: |
143 lemma edge_of_th: |
207 "(Cs cs, Th th) \<in> RAG s'" |
144 "(Cs cs, Th th) \<in> RAG s" |
208 proof - |
145 proof - |
209 from holding_th |
146 from holding_th_cs_s |
210 show ?thesis |
147 show ?thesis |
211 by (unfold s_RAG_def holding_eq, auto) |
148 by (unfold s_RAG_def holding_eq, auto) |
212 qed |
149 qed |
213 |
150 |
214 lemma ancestors_cs: |
151 lemma ancestors_cs: |
215 "ancestors (RAG s') (Cs cs) = {Th th}" |
152 "ancestors (RAG s) (Cs cs) = {Th th}" |
216 proof - |
153 proof - |
217 have "ancestors (RAG s') (Cs cs) = ancestors (RAG s') (Th th) \<union> {Th th}" |
154 have "ancestors (RAG s) (Cs cs) = ancestors (RAG s) (Th th) \<union> {Th th}" |
218 proof(rule vat_s'.rtree_RAG.ancestors_accum) |
155 by (rule rtree_RAG.ancestors_accum[OF edge_of_th]) |
219 from vt_s[unfolded s_def] |
|
220 have " PIP s' (V th cs)" by (cases, simp) |
|
221 thus "(Cs cs, Th th) \<in> RAG s'" |
|
222 proof(cases) |
|
223 assume "holding s' th cs" |
|
224 from this[unfolded holding_eq] |
|
225 show ?thesis by (unfold s_RAG_def, auto) |
|
226 qed |
|
227 qed |
|
228 from this[unfolded ancestors_th] show ?thesis by simp |
156 from this[unfolded ancestors_th] show ?thesis by simp |
229 qed |
157 qed |
230 |
|
231 lemma preced_kept: "the_preced s = the_preced s'" |
|
232 by (auto simp: s_def the_preced_def preced_def) |
|
233 |
158 |
234 end |
159 end |
235 |
160 |
236 text {* |
161 text {* |
237 The following @{text "step_v_cps_nt"} is the sub-locale for @{text "V"}-operation, |
162 The following @{text "step_v_cps_nt"} is the sub-locale for @{text "V"}-operation, |
238 which represents the case when there is another thread @{text "th'"} |
163 which represents the case when there is another thread @{text "th'"} |
239 to take over the critical resource released by the initiating thread @{text "th"}. |
164 to take over the critical resource released by the initiating thread @{text "th"}. |
240 *} |
165 *} |
241 locale step_v_cps_nt = step_v_cps + |
166 |
242 fixes th' |
167 context valid_trace_v_n |
243 -- {* @{text "th'"} is assumed to take over @{text "cs"} *} |
168 begin |
244 assumes nt: "next_th s' th cs th'" |
169 |
245 |
170 lemma sub_RAGs': |
246 context step_v_cps_nt |
171 "{(Cs cs, Th th), (Th taker, Cs cs)} \<subseteq> RAG s" |
247 begin |
172 using next_th_RAG[OF next_th_taker] . |
248 |
|
249 text {* |
|
250 Lemma @{text "RAG_s"} confirms the change of RAG: |
|
251 two edges removed and one added, as shown by the following diagram. |
|
252 *} |
|
253 |
|
254 (* |
|
255 RAG before the V-operation |
|
256 th1 ----| |
|
257 | |
|
258 th' ----| |
|
259 |----> cs -----| |
|
260 th2 ----| | |
|
261 | | |
|
262 th3 ----| | |
|
263 |------> th |
|
264 th4 ----| | |
|
265 | | |
|
266 th5 ----| | |
|
267 |----> cs'-----| |
|
268 th6 ----| |
|
269 | |
|
270 th7 ----| |
|
271 |
|
272 RAG after the V-operation |
|
273 th1 ----| |
|
274 | |
|
275 |----> cs ----> th' |
|
276 th2 ----| |
|
277 | |
|
278 th3 ----| |
|
279 |
|
280 th4 ----| |
|
281 | |
|
282 th5 ----| |
|
283 |----> cs'----> th |
|
284 th6 ----| |
|
285 | |
|
286 th7 ----| |
|
287 *) |
|
288 |
|
289 lemma sub_RAGs': "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s'" |
|
290 using next_th_RAG[OF nt] . |
|
291 |
173 |
292 lemma ancestors_th': |
174 lemma ancestors_th': |
293 "ancestors (RAG s') (Th th') = {Th th, Cs cs}" |
175 "ancestors (RAG s) (Th taker) = {Th th, Cs cs}" |
294 proof - |
176 proof - |
295 have "ancestors (RAG s') (Th th') = ancestors (RAG s') (Cs cs) \<union> {Cs cs}" |
177 have "ancestors (RAG s) (Th taker) = ancestors (RAG s) (Cs cs) \<union> {Cs cs}" |
296 proof(rule vat_s'.rtree_RAG.ancestors_accum) |
178 proof(rule rtree_RAG.ancestors_accum) |
297 from sub_RAGs' show "(Th th', Cs cs) \<in> RAG s'" by auto |
179 from sub_RAGs' show "(Th taker, Cs cs) \<in> RAG s" by auto |
298 qed |
180 qed |
299 thus ?thesis using ancestors_th ancestors_cs by auto |
181 thus ?thesis using ancestors_th ancestors_cs by auto |
300 qed |
182 qed |
301 |
183 |
302 lemma RAG_s: |
184 lemma RAG_s: |
303 "RAG s = (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) \<union> |
185 "RAG (e#s) = (RAG s - {(Cs cs, Th th), (Th taker, Cs cs)}) \<union> |
304 {(Cs cs, Th th')}" |
186 {(Cs cs, Th taker)}" |
305 proof - |
187 by (unfold RAG_es waiting_set_eq holding_set_eq, auto) |
306 from step_RAG_v[OF vt_s[unfolded s_def], folded s_def] |
|
307 and nt show ?thesis by (auto intro:next_th_unique) |
|
308 qed |
|
309 |
188 |
310 lemma subtree_kept: (* ddd *) |
189 lemma subtree_kept: (* ddd *) |
311 assumes "th1 \<notin> {th, th'}" |
190 assumes "th1 \<notin> {th, taker}" |
312 shows "subtree (RAG s) (Th th1) = subtree (RAG s') (Th th1)" (is "_ = ?R") |
191 shows "subtree (RAG (e#s)) (Th th1) = |
313 proof - |
192 subtree (RAG s) (Th th1)" (is "_ = ?R") |
314 let ?RAG' = "(RAG s' - {(Cs cs, Th th), (Th th', Cs cs)})" |
193 proof - |
315 let ?RAG'' = "?RAG' \<union> {(Cs cs, Th th')}" |
194 let ?RAG' = "(RAG s - {(Cs cs, Th th), (Th taker, Cs cs)})" |
|
195 let ?RAG'' = "?RAG' \<union> {(Cs cs, Th taker)}" |
316 have "subtree ?RAG' (Th th1) = ?R" |
196 have "subtree ?RAG' (Th th1) = ?R" |
317 proof(rule subset_del_subtree_outside) |
197 proof(rule subset_del_subtree_outside) |
318 show "Range {(Cs cs, Th th), (Th th', Cs cs)} \<inter> subtree (RAG s') (Th th1) = {}" |
198 show "Range {(Cs cs, Th th), (Th taker, Cs cs)} \<inter> subtree (RAG s) (Th th1) = {}" |
319 proof - |
199 proof - |
320 have "(Th th) \<notin> subtree (RAG s') (Th th1)" |
200 have "(Th th) \<notin> subtree (RAG s) (Th th1)" |
321 proof(rule subtree_refute) |
201 proof(rule subtree_refute) |
322 show "Th th1 \<notin> ancestors (RAG s') (Th th)" |
202 show "Th th1 \<notin> ancestors (RAG s) (Th th)" |
323 by (unfold ancestors_th, simp) |
203 by (unfold ancestors_th, simp) |
324 next |
204 next |
325 from assms show "Th th1 \<noteq> Th th" by simp |
205 from assms show "Th th1 \<noteq> Th th" by simp |
326 qed |
206 qed |
327 moreover have "(Cs cs) \<notin> subtree (RAG s') (Th th1)" |
207 moreover have "(Cs cs) \<notin> subtree (RAG s) (Th th1)" |
328 proof(rule subtree_refute) |
208 proof(rule subtree_refute) |
329 show "Th th1 \<notin> ancestors (RAG s') (Cs cs)" |
209 show "Th th1 \<notin> ancestors (RAG s) (Cs cs)" |
330 by (unfold ancestors_cs, insert assms, auto) |
210 by (unfold ancestors_cs, insert assms, auto) |
331 qed simp |
211 qed simp |
332 ultimately have "{Th th, Cs cs} \<inter> subtree (RAG s') (Th th1) = {}" by auto |
212 ultimately have "{Th th, Cs cs} \<inter> subtree (RAG s) (Th th1) = {}" by auto |
333 thus ?thesis by simp |
213 thus ?thesis by simp |
334 qed |
214 qed |
335 qed |
215 qed |
336 moreover have "subtree ?RAG'' (Th th1) = subtree ?RAG' (Th th1)" |
216 moreover have "subtree ?RAG'' (Th th1) = subtree ?RAG' (Th th1)" |
337 proof(rule subtree_insert_next) |
217 proof(rule subtree_insert_next) |
338 show "Th th' \<notin> subtree (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) (Th th1)" |
218 show "Th taker \<notin> subtree (RAG s - {(Cs cs, Th th), (Th taker, Cs cs)}) (Th th1)" |
339 proof(rule subtree_refute) |
219 proof(rule subtree_refute) |
340 show "Th th1 \<notin> ancestors (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) (Th th')" |
220 show "Th th1 \<notin> ancestors (RAG s - {(Cs cs, Th th), (Th taker, Cs cs)}) (Th taker)" |
341 (is "_ \<notin> ?R") |
221 (is "_ \<notin> ?R") |
342 proof - |
222 proof - |
343 have "?R \<subseteq> ancestors (RAG s') (Th th')" by (rule ancestors_mono, auto) |
223 have "?R \<subseteq> ancestors (RAG s) (Th taker)" by (rule ancestors_mono, auto) |
344 moreover have "Th th1 \<notin> ..." using ancestors_th' assms by simp |
224 moreover have "Th th1 \<notin> ..." using ancestors_th' assms by simp |
345 ultimately show ?thesis by auto |
225 ultimately show ?thesis by auto |
346 qed |
226 qed |
347 next |
227 next |
348 from assms show "Th th1 \<noteq> Th th'" by simp |
228 from assms show "Th th1 \<noteq> Th taker" by simp |
349 qed |
229 qed |
350 qed |
230 qed |
351 ultimately show ?thesis by (unfold RAG_s, simp) |
231 ultimately show ?thesis by (unfold RAG_s, simp) |
352 qed |
232 qed |
353 |
233 |
354 lemma cp_kept: |
234 lemma cp_kept: |
355 assumes "th1 \<notin> {th, th'}" |
235 assumes "th1 \<notin> {th, taker}" |
356 shows "cp s th1 = cp s' th1" |
236 shows "cp (e#s) th1 = cp s th1" |
357 by (unfold cp_alt_def preced_kept subtree_kept[OF assms], simp) |
237 by (unfold cp_alt_def the_preced_es subtree_kept[OF assms], simp) |
358 |
238 |
359 end |
239 end |
360 |
240 |
361 locale step_v_cps_nnt = step_v_cps + |
241 |
362 assumes nnt: "\<And> th'. (\<not> next_th s' th cs th')" |
242 context valid_trace_v_e |
363 |
243 begin |
364 context step_v_cps_nnt |
244 |
365 begin |
245 find_theorems RAG s e |
366 |
246 |
367 lemma RAG_s: "RAG s = RAG s' - {(Cs cs, Th th)}" |
247 lemma RAG_s: "RAG (e#s) = RAG s - {(Cs cs, Th th)}" |
368 proof - |
248 by (unfold RAG_es waiting_set_eq holding_set_eq, simp) |
369 from nnt and step_RAG_v[OF vt_s[unfolded s_def], folded s_def] |
|
370 show ?thesis by auto |
|
371 qed |
|
372 |
249 |
373 lemma subtree_kept: |
250 lemma subtree_kept: |
374 assumes "th1 \<noteq> th" |
251 assumes "th1 \<noteq> th" |
375 shows "subtree (RAG s) (Th th1) = subtree (RAG s') (Th th1)" |
252 shows "subtree (RAG (e#s)) (Th th1) = subtree (RAG s) (Th th1)" |
376 proof(unfold RAG_s, rule subset_del_subtree_outside) |
253 proof(unfold RAG_s, rule subset_del_subtree_outside) |
377 show "Range {(Cs cs, Th th)} \<inter> subtree (RAG s') (Th th1) = {}" |
254 show "Range {(Cs cs, Th th)} \<inter> subtree (RAG s) (Th th1) = {}" |
378 proof - |
255 proof - |
379 have "(Th th) \<notin> subtree (RAG s') (Th th1)" |
256 have "(Th th) \<notin> subtree (RAG s) (Th th1)" |
380 proof(rule subtree_refute) |
257 proof(rule subtree_refute) |
381 show "Th th1 \<notin> ancestors (RAG s') (Th th)" |
258 show "Th th1 \<notin> ancestors (RAG s) (Th th)" |
382 by (unfold ancestors_th, simp) |
259 by (unfold ancestors_th, simp) |
383 next |
260 next |
384 from assms show "Th th1 \<noteq> Th th" by simp |
261 from assms show "Th th1 \<noteq> Th th" by simp |
385 qed |
262 qed |
386 thus ?thesis by auto |
263 thus ?thesis by auto |
387 qed |
264 qed |
388 qed |
265 qed |
389 |
266 |
390 lemma cp_kept_1: |
267 lemma cp_kept_1: |
391 assumes "th1 \<noteq> th" |
268 assumes "th1 \<noteq> th" |
392 shows "cp s th1 = cp s' th1" |
269 shows "cp (e#s) th1 = cp s th1" |
393 by (unfold cp_alt_def preced_kept subtree_kept[OF assms], simp) |
270 by (unfold cp_alt_def the_preced_es subtree_kept[OF assms], simp) |
394 |
271 |
395 lemma subtree_cs: "subtree (RAG s') (Cs cs) = {Cs cs}" |
272 lemma subtree_cs: "subtree (RAG s) (Cs cs) = {Cs cs}" |
396 proof - |
273 proof - |
397 { fix n |
274 { fix n |
398 have "(Cs cs) \<notin> ancestors (RAG s') n" |
275 have "(Cs cs) \<notin> ancestors (RAG s) n" |
399 proof |
276 proof |
400 assume "Cs cs \<in> ancestors (RAG s') n" |
277 assume "Cs cs \<in> ancestors (RAG s) n" |
401 hence "(n, Cs cs) \<in> (RAG s')^+" by (auto simp:ancestors_def) |
278 hence "(n, Cs cs) \<in> (RAG s)^+" by (auto simp:ancestors_def) |
402 from tranclE[OF this] obtain nn where h: "(nn, Cs cs) \<in> RAG s'" by auto |
279 from tranclE[OF this] obtain nn where h: "(nn, Cs cs) \<in> RAG s" by auto |
403 then obtain th' where "nn = Th th'" |
280 then obtain th' where "nn = Th th'" |
404 by (unfold s_RAG_def, auto) |
281 by (unfold s_RAG_def, auto) |
405 from h[unfolded this] have "(Th th', Cs cs) \<in> RAG s'" . |
282 from h[unfolded this] have "(Th th', Cs cs) \<in> RAG s" . |
406 from this[unfolded s_RAG_def] |
283 from this[unfolded s_RAG_def] |
407 have "waiting (wq s') th' cs" by auto |
284 have "waiting (wq s) th' cs" by auto |
408 from this[unfolded cs_waiting_def] |
285 from this[unfolded cs_waiting_def] |
409 have "1 < length (wq s' cs)" |
286 have "1 < length (wq s cs)" |
410 by (cases "wq s' cs", auto) |
287 by (cases "wq s cs", auto) |
411 from holding_next_thI[OF holding_th this] |
288 from holding_next_thI[OF holding_th_cs_s this] |
412 obtain th' where "next_th s' th cs th'" by auto |
289 obtain th' where "next_th s th cs th'" by auto |
413 with nnt show False by auto |
290 thus False using no_taker by blast |
414 qed |
291 qed |
415 } note h = this |
292 } note h = this |
416 { fix n |
293 { fix n |
417 assume "n \<in> subtree (RAG s') (Cs cs)" |
294 assume "n \<in> subtree (RAG s) (Cs cs)" |
418 hence "n = (Cs cs)" |
295 hence "n = (Cs cs)" |
419 by (elim subtreeE, insert h, auto) |
296 by (elim subtreeE, insert h, auto) |
420 } moreover have "(Cs cs) \<in> subtree (RAG s') (Cs cs)" |
297 } moreover have "(Cs cs) \<in> subtree (RAG s) (Cs cs)" |
421 by (auto simp:subtree_def) |
298 by (auto simp:subtree_def) |
422 ultimately show ?thesis by auto |
299 ultimately show ?thesis by auto |
423 qed |
300 qed |
424 |
301 |
425 lemma subtree_th: |
302 lemma subtree_th: |
426 "subtree (RAG s) (Th th) = subtree (RAG s') (Th th) - {Cs cs}" |
303 "subtree (RAG (e#s)) (Th th) = subtree (RAG s) (Th th) - {Cs cs}" |
427 proof(unfold RAG_s, fold subtree_cs, rule vat_s'.rtree_RAG.subtree_del_inside) |
304 proof(unfold RAG_s, fold subtree_cs, rule rtree_RAG.subtree_del_inside) |
428 from edge_of_th |
305 from edge_of_th |
429 show "(Cs cs, Th th) \<in> edges_in (RAG s') (Th th)" |
306 show "(Cs cs, Th th) \<in> edges_in (RAG s) (Th th)" |
430 by (unfold edges_in_def, auto simp:subtree_def) |
307 by (unfold edges_in_def, auto simp:subtree_def) |
431 qed |
308 qed |
432 |
309 |
433 lemma cp_kept_2: |
310 lemma cp_kept_2: |
434 shows "cp s th = cp s' th" |
311 shows "cp (e#s) th = cp s th" |
435 by (unfold cp_alt_def subtree_th preced_kept, auto) |
312 by (unfold cp_alt_def subtree_th the_preced_es, auto) |
436 |
313 |
437 lemma eq_cp: |
314 lemma eq_cp: |
438 shows "cp s th' = cp s' th'" |
315 shows "cp (e#s) th' = cp s th'" |
439 using cp_kept_1 cp_kept_2 |
316 using cp_kept_1 cp_kept_2 |
440 by (cases "th' = th", auto) |
317 by (cases "th' = th", auto) |
441 end |
318 |
442 |
319 end |
443 |
320 |
444 locale step_P_cps = |
321 |
445 fixes s' th cs s |
322 section {* The @{term P} operation *} |
446 defines s_def : "s \<equiv> (P th cs#s')" |
323 |
447 assumes vt_s: "vt s" |
324 context valid_trace_p |
448 |
325 begin |
449 sublocale step_P_cps < vat_s : valid_trace "s" |
326 |
|
327 lemma root_th: "root (RAG s) (Th th)" |
|
328 by (simp add: ready_th_s readys_root) |
|
329 |
|
330 lemma in_no_others_subtree: |
|
331 assumes "th' \<noteq> th" |
|
332 shows "Th th \<notin> subtree (RAG s) (Th th')" |
450 proof |
333 proof |
451 from vt_s show "vt s" . |
334 assume "Th th \<in> subtree (RAG s) (Th th')" |
452 qed |
|
453 |
|
454 section {* The @{term P} operation *} |
|
455 |
|
456 sublocale step_P_cps < vat_s' : valid_trace "s'" |
|
457 proof |
|
458 from step_back_vt[OF vt_s[unfolded s_def]] show "vt s'" . |
|
459 qed |
|
460 |
|
461 context step_P_cps |
|
462 begin |
|
463 |
|
464 lemma readys_th: "th \<in> readys s'" |
|
465 proof - |
|
466 from step_back_step [OF vt_s[unfolded s_def]] |
|
467 have "PIP s' (P th cs)" . |
|
468 hence "th \<in> runing s'" by (cases, simp) |
|
469 thus ?thesis by (simp add:readys_def runing_def) |
|
470 qed |
|
471 |
|
472 lemma root_th: "root (RAG s') (Th th)" |
|
473 using readys_root[OF readys_th] . |
|
474 |
|
475 lemma in_no_others_subtree: |
|
476 assumes "th' \<noteq> th" |
|
477 shows "Th th \<notin> subtree (RAG s') (Th th')" |
|
478 proof |
|
479 assume "Th th \<in> subtree (RAG s') (Th th')" |
|
480 thus False |
335 thus False |
481 proof(cases rule:subtreeE) |
336 proof(cases rule:subtreeE) |
482 case 1 |
337 case 1 |
483 with assms show ?thesis by auto |
338 with assms show ?thesis by auto |
484 next |
339 next |
485 case 2 |
340 case 2 |
486 with root_th show ?thesis by (auto simp:root_def) |
341 with root_th show ?thesis by (auto simp:root_def) |
487 qed |
342 qed |
488 qed |
343 qed |
489 |
344 |
490 lemma preced_kept: "the_preced s = the_preced s'" |
345 lemma preced_kept: "the_preced (e#s) = the_preced s" |
491 by (auto simp: s_def the_preced_def preced_def) |
346 proof |
492 |
347 fix th' |
493 end |
348 show "the_preced (e # s) th' = the_preced s th'" |
494 |
349 by (unfold the_preced_def is_p preced_def, simp) |
495 locale step_P_cps_ne =step_P_cps + |
350 qed |
496 fixes th' |
351 |
497 assumes ne: "wq s' cs \<noteq> []" |
352 end |
498 defines th'_def: "th' \<equiv> hd (wq s' cs)" |
353 |
499 |
354 |
500 locale step_P_cps_e =step_P_cps + |
355 context valid_trace_p_h |
501 assumes ee: "wq s' cs = []" |
356 begin |
502 |
|
503 context step_P_cps_e |
|
504 begin |
|
505 |
|
506 lemma RAG_s: "RAG s = RAG s' \<union> {(Cs cs, Th th)}" |
|
507 proof - |
|
508 from ee and step_RAG_p[OF vt_s[unfolded s_def], folded s_def] |
|
509 show ?thesis by auto |
|
510 qed |
|
511 |
357 |
512 lemma subtree_kept: |
358 lemma subtree_kept: |
513 assumes "th' \<noteq> th" |
359 assumes "th' \<noteq> th" |
514 shows "subtree (RAG s) (Th th') = subtree (RAG s') (Th th')" |
360 shows "subtree (RAG (e#s)) (Th th') = subtree (RAG s) (Th th')" |
515 proof(unfold RAG_s, rule subtree_insert_next) |
361 proof(unfold RAG_es, rule subtree_insert_next) |
516 from in_no_others_subtree[OF assms] |
362 from in_no_others_subtree[OF assms] |
517 show "Th th \<notin> subtree (RAG s') (Th th')" . |
363 show "Th th \<notin> subtree (RAG s) (Th th')" . |
518 qed |
364 qed |
519 |
365 |
520 lemma cp_kept: |
366 lemma cp_kept: |
521 assumes "th' \<noteq> th" |
367 assumes "th' \<noteq> th" |
522 shows "cp s th' = cp s' th'" |
368 shows "cp (e#s) th' = cp s th'" |
523 proof - |
369 proof - |
524 have "(the_preced s ` {th'a. Th th'a \<in> subtree (RAG s) (Th th')}) = |
370 have "(the_preced (e#s) ` {th'a. Th th'a \<in> subtree (RAG (e#s)) (Th th')}) = |
525 (the_preced s' ` {th'a. Th th'a \<in> subtree (RAG s') (Th th')})" |
371 (the_preced s ` {th'a. Th th'a \<in> subtree (RAG s) (Th th')})" |
526 by (unfold preced_kept subtree_kept[OF assms], simp) |
372 by (unfold preced_kept subtree_kept[OF assms], simp) |
527 thus ?thesis by (unfold cp_alt_def, simp) |
373 thus ?thesis by (unfold cp_alt_def, simp) |
528 qed |
374 qed |
529 |
375 |
530 end |
376 end |
531 |
377 |
532 context step_P_cps_ne |
378 context valid_trace_p_w |
533 begin |
379 begin |
534 |
380 |
535 lemma RAG_s: "RAG s = RAG s' \<union> {(Th th, Cs cs)}" |
381 interpretation vat_e: valid_trace "e#s" |
536 proof - |
382 by (unfold_locales, insert vt_e, simp) |
537 from step_RAG_p[OF vt_s[unfolded s_def]] and ne |
383 |
538 show ?thesis by (simp add:s_def) |
384 lemma cs_held: "(Cs cs, Th holder) \<in> RAG s" |
539 qed |
385 using holding_s_holder |
540 |
386 by (unfold s_RAG_def, fold holding_eq, auto) |
541 lemma cs_held: "(Cs cs, Th th') \<in> RAG s'" |
387 |
542 proof - |
388 lemma tRAG_s: |
543 have "(Cs cs, Th th') \<in> hRAG s'" |
389 "tRAG (e#s) = tRAG s \<union> {(Th th, Th holder)}" |
|
390 using local.RAG_tRAG_transfer[OF RAG_es cs_held] . |
|
391 |
|
392 lemma cp_kept: |
|
393 assumes "Th th'' \<notin> ancestors (tRAG (e#s)) (Th th)" |
|
394 shows "cp (e#s) th'' = cp s th''" |
|
395 proof - |
|
396 have h: "subtree (tRAG (e#s)) (Th th'') = subtree (tRAG s) (Th th'')" |
544 proof - |
397 proof - |
545 from ne |
398 have "Th holder \<notin> subtree (tRAG s) (Th th'')" |
546 have " holding s' th' cs" |
|
547 by (unfold th'_def holding_eq cs_holding_def, auto) |
|
548 thus ?thesis |
|
549 by (unfold hRAG_def, auto) |
|
550 qed |
|
551 thus ?thesis by (unfold RAG_split, auto) |
|
552 qed |
|
553 |
|
554 lemma tRAG_s: |
|
555 "tRAG s = tRAG s' \<union> {(Th th, Th th')}" |
|
556 using RAG_tRAG_transfer[OF RAG_s cs_held] . |
|
557 |
|
558 lemma cp_kept: |
|
559 assumes "Th th'' \<notin> ancestors (tRAG s) (Th th)" |
|
560 shows "cp s th'' = cp s' th''" |
|
561 proof - |
|
562 have h: "subtree (tRAG s) (Th th'') = subtree (tRAG s') (Th th'')" |
|
563 proof - |
|
564 have "Th th' \<notin> subtree (tRAG s') (Th th'')" |
|
565 proof |
399 proof |
566 assume "Th th' \<in> subtree (tRAG s') (Th th'')" |
400 assume "Th holder \<in> subtree (tRAG s) (Th th'')" |
567 thus False |
401 thus False |
568 proof(rule subtreeE) |
402 proof(rule subtreeE) |
569 assume "Th th' = Th th''" |
403 assume "Th holder = Th th''" |
570 from assms[unfolded tRAG_s ancestors_def, folded this] |
404 from assms[unfolded tRAG_s ancestors_def, folded this] |
571 show ?thesis by auto |
405 show ?thesis by auto |
572 next |
406 next |
573 assume "Th th'' \<in> ancestors (tRAG s') (Th th')" |
407 assume "Th th'' \<in> ancestors (tRAG s) (Th holder)" |
574 moreover have "... \<subseteq> ancestors (tRAG s) (Th th')" |
408 moreover have "... \<subseteq> ancestors (tRAG (e#s)) (Th holder)" |
575 proof(rule ancestors_mono) |
409 proof(rule ancestors_mono) |
576 show "tRAG s' \<subseteq> tRAG s" by (unfold tRAG_s, auto) |
410 show "tRAG s \<subseteq> tRAG (e#s)" by (unfold tRAG_s, auto) |
577 qed |
411 qed |
578 ultimately have "Th th'' \<in> ancestors (tRAG s) (Th th')" by auto |
412 ultimately have "Th th'' \<in> ancestors (tRAG (e#s)) (Th holder)" by auto |
579 moreover have "Th th' \<in> ancestors (tRAG s) (Th th)" |
413 moreover have "Th holder \<in> ancestors (tRAG (e#s)) (Th th)" |
580 by (unfold tRAG_s, auto simp:ancestors_def) |
414 by (unfold tRAG_s, auto simp:ancestors_def) |
581 ultimately have "Th th'' \<in> ancestors (tRAG s) (Th th)" |
415 ultimately have "Th th'' \<in> ancestors (tRAG (e#s)) (Th th)" |
582 by (auto simp:ancestors_def) |
416 by (auto simp:ancestors_def) |
583 with assms show ?thesis by auto |
417 with assms show ?thesis by auto |
584 qed |
418 qed |
585 qed |
419 qed |
586 from subtree_insert_next[OF this] |
420 from subtree_insert_next[OF this] |
587 have "subtree (tRAG s' \<union> {(Th th, Th th')}) (Th th'') = subtree (tRAG s') (Th th'')" . |
421 have "subtree (tRAG s \<union> {(Th th, Th holder)}) (Th th'') = subtree (tRAG s) (Th th'')" . |
588 from this[folded tRAG_s] show ?thesis . |
422 from this[folded tRAG_s] show ?thesis . |
589 qed |
423 qed |
590 show ?thesis by (unfold cp_alt_def1 h preced_kept, simp) |
424 show ?thesis by (unfold cp_alt_def1 h preced_kept, simp) |
591 qed |
425 qed |
592 |
426 |
593 lemma cp_gen_update_stop: (* ddd *) |
427 lemma cp_gen_update_stop: (* ddd *) |
594 assumes "u \<in> ancestors (tRAG s) (Th th)" |
428 assumes "u \<in> ancestors (tRAG (e#s)) (Th th)" |
595 and "cp_gen s u = cp_gen s' u" |
429 and "cp_gen (e#s) u = cp_gen s u" |
596 and "y \<in> ancestors (tRAG s) u" |
430 and "y \<in> ancestors (tRAG (e#s)) u" |
597 shows "cp_gen s y = cp_gen s' y" |
431 shows "cp_gen (e#s) y = cp_gen s y" |
598 using assms(3) |
432 using assms(3) |
599 proof(induct rule:wf_induct[OF vat_s.fsbttRAGs.wf]) |
433 proof(induct rule:wf_induct[OF vat_e.fsbttRAGs.wf]) |
600 case (1 x) |
434 case (1 x) |
601 show ?case (is "?L = ?R") |
435 show ?case (is "?L = ?R") |
602 proof - |
436 proof - |
603 from tRAG_ancestorsE[OF 1(2)] |
437 from tRAG_ancestorsE[OF 1(2)] |
604 obtain th2 where eq_x: "x = Th th2" by blast |
438 obtain th2 where eq_x: "x = Th th2" by blast |
605 from vat_s.cp_gen_rec[OF this] |
439 from vat_e.cp_gen_rec[OF this] |
606 have "?L = |
440 have "?L = |
607 Max ({the_preced s th2} \<union> cp_gen s ` RTree.children (tRAG s) x)" . |
441 Max ({the_preced (e#s) th2} \<union> cp_gen (e#s) ` RTree.children (tRAG (e#s)) x)" . |
608 also have "... = |
442 also have "... = |
609 Max ({the_preced s' th2} \<union> cp_gen s' ` RTree.children (tRAG s') x)" |
443 Max ({the_preced s th2} \<union> cp_gen s ` RTree.children (tRAG s) x)" |
610 |
|
611 proof - |
444 proof - |
612 from preced_kept have "the_preced s th2 = the_preced s' th2" by simp |
445 from preced_kept have "the_preced (e#s) th2 = the_preced s th2" by simp |
613 moreover have "cp_gen s ` RTree.children (tRAG s) x = |
446 moreover have "cp_gen (e#s) ` RTree.children (tRAG (e#s)) x = |
614 cp_gen s' ` RTree.children (tRAG s') x" |
447 cp_gen s ` RTree.children (tRAG s) x" |
615 proof - |
448 proof - |
616 have "RTree.children (tRAG s) x = RTree.children (tRAG s') x" |
449 have "RTree.children (tRAG (e#s)) x = RTree.children (tRAG s) x" |
617 proof(unfold tRAG_s, rule children_union_kept) |
450 proof(unfold tRAG_s, rule children_union_kept) |
618 have start: "(Th th, Th th') \<in> tRAG s" |
451 have start: "(Th th, Th holder) \<in> tRAG (e#s)" |
619 by (unfold tRAG_s, auto) |
452 by (unfold tRAG_s, auto) |
620 note x_u = 1(2) |
453 note x_u = 1(2) |
621 show "x \<notin> Range {(Th th, Th th')}" |
454 show "x \<notin> Range {(Th th, Th holder)}" |
622 proof |
455 proof |
623 assume "x \<in> Range {(Th th, Th th')}" |
456 assume "x \<in> Range {(Th th, Th holder)}" |
624 hence eq_x: "x = Th th'" using RangeE by auto |
457 hence eq_x: "x = Th holder" using RangeE by auto |
625 show False |
458 show False |
626 proof(cases rule:vat_s.rtree_s.ancestors_headE[OF assms(1) start]) |
459 proof(cases rule:vat_e.ancestors_headE[OF assms(1) start]) |
627 case 1 |
460 case 1 |
628 from x_u[folded this, unfolded eq_x] vat_s.acyclic_tRAG |
461 from x_u[folded this, unfolded eq_x] vat_e.acyclic_tRAG |
629 show ?thesis by (auto simp:ancestors_def acyclic_def) |
462 show ?thesis by (auto simp:ancestors_def acyclic_def) |
630 next |
463 next |
631 case 2 |
464 case 2 |
632 with x_u[unfolded eq_x] |
465 with x_u[unfolded eq_x] |
633 have "(Th th', Th th') \<in> (tRAG s)^+" by (auto simp:ancestors_def) |
466 have "(Th holder, Th holder) \<in> (tRAG (e#s))^+" by (auto simp:ancestors_def) |
634 with vat_s.acyclic_tRAG show ?thesis by (auto simp:acyclic_def) |
467 with vat_e.acyclic_tRAG show ?thesis by (auto simp:acyclic_def) |
635 qed |
468 qed |
636 qed |
469 qed |
637 qed |
470 qed |
638 moreover have "cp_gen s ` RTree.children (tRAG s) x = |
471 moreover have "cp_gen (e#s) ` RTree.children (tRAG (e#s)) x = |
639 cp_gen s' ` RTree.children (tRAG s) x" (is "?f ` ?A = ?g ` ?A") |
472 cp_gen s ` RTree.children (tRAG (e#s)) x" (is "?f ` ?A = ?g ` ?A") |
640 proof(rule f_image_eq) |
473 proof(rule f_image_eq) |
641 fix a |
474 fix a |
642 assume a_in: "a \<in> ?A" |
475 assume a_in: "a \<in> ?A" |
643 from 1(2) |
476 from 1(2) |
644 show "?f a = ?g a" |
477 show "?f a = ?g a" |
645 proof(cases rule:vat_s.rtree_s.ancestors_childrenE[case_names in_ch out_ch]) |
478 proof(cases rule:vat_e.rtree_s.ancestors_childrenE[case_names in_ch out_ch]) |
646 case in_ch |
479 case in_ch |
647 show ?thesis |
480 show ?thesis |
648 proof(cases "a = u") |
481 proof(cases "a = u") |
649 case True |
482 case True |
650 from assms(2)[folded this] show ?thesis . |
483 from assms(2)[folded this] show ?thesis . |
651 next |
484 next |
652 case False |
485 case False |
653 have a_not_in: "a \<notin> ancestors (tRAG s) (Th th)" |
486 have a_not_in: "a \<notin> ancestors (tRAG (e#s)) (Th th)" |
654 proof |
487 proof |
655 assume a_in': "a \<in> ancestors (tRAG s) (Th th)" |
488 assume a_in': "a \<in> ancestors (tRAG (e#s)) (Th th)" |
656 have "a = u" |
489 have "a = u" |
657 proof(rule vat_s.rtree_s.ancestors_children_unique) |
490 proof(rule vat_e.rtree_s.ancestors_children_unique) |
658 from a_in' a_in show "a \<in> ancestors (tRAG s) (Th th) \<inter> |
491 from a_in' a_in show "a \<in> ancestors (tRAG (e#s)) (Th th) \<inter> |
659 RTree.children (tRAG s) x" by auto |
492 RTree.children (tRAG (e#s)) x" by auto |
660 next |
493 next |
661 from assms(1) in_ch show "u \<in> ancestors (tRAG s) (Th th) \<inter> |
494 from assms(1) in_ch show "u \<in> ancestors (tRAG (e#s)) (Th th) \<inter> |
662 RTree.children (tRAG s) x" by auto |
495 RTree.children (tRAG (e#s)) x" by auto |
663 qed |
496 qed |
664 with False show False by simp |
497 with False show False by simp |
665 qed |
498 qed |
666 from a_in obtain th_a where eq_a: "a = Th th_a" |
499 from a_in obtain th_a where eq_a: "a = Th th_a" |
667 by (unfold RTree.children_def tRAG_alt_def, auto) |
500 by (unfold RTree.children_def tRAG_alt_def, auto) |
668 from cp_kept[OF a_not_in[unfolded eq_a]] |
501 from cp_kept[OF a_not_in[unfolded eq_a]] |
669 have "cp s th_a = cp s' th_a" . |
502 have "cp (e#s) th_a = cp s th_a" . |
670 from this [unfolded cp_gen_def_cond[OF eq_a], folded eq_a] |
503 from this [unfolded cp_gen_def_cond[OF eq_a], folded eq_a] |
671 show ?thesis . |
504 show ?thesis . |
672 qed |
505 qed |
673 next |
506 next |
674 case (out_ch z) |
507 case (out_ch z) |
675 hence h: "z \<in> ancestors (tRAG s) u" "z \<in> RTree.children (tRAG s) x" by auto |
508 hence h: "z \<in> ancestors (tRAG (e#s)) u" "z \<in> RTree.children (tRAG (e#s)) x" by auto |
676 show ?thesis |
509 show ?thesis |
677 proof(cases "a = z") |
510 proof(cases "a = z") |
678 case True |
511 case True |
679 from h(2) have zx_in: "(z, x) \<in> (tRAG s)" by (auto simp:RTree.children_def) |
512 from h(2) have zx_in: "(z, x) \<in> (tRAG (e#s))" by (auto simp:RTree.children_def) |
680 from 1(1)[rule_format, OF this h(1)] |
513 from 1(1)[rule_format, OF this h(1)] |
681 have eq_cp_gen: "cp_gen s z = cp_gen s' z" . |
514 have eq_cp_gen: "cp_gen (e#s) z = cp_gen s z" . |
682 with True show ?thesis by metis |
515 with True show ?thesis by metis |
683 next |
516 next |
684 case False |
517 case False |
685 from a_in obtain th_a where eq_a: "a = Th th_a" |
518 from a_in obtain th_a where eq_a: "a = Th th_a" |
686 by (auto simp:RTree.children_def tRAG_alt_def) |
519 by (auto simp:RTree.children_def tRAG_alt_def) |
687 have "a \<notin> ancestors (tRAG s) (Th th)" |
520 have "a \<notin> ancestors (tRAG (e#s)) (Th th)" |
688 proof |
521 proof |
689 assume a_in': "a \<in> ancestors (tRAG s) (Th th)" |
522 assume a_in': "a \<in> ancestors (tRAG (e#s)) (Th th)" |
690 have "a = z" |
523 have "a = z" |
691 proof(rule vat_s.rtree_s.ancestors_children_unique) |
524 proof(rule vat_e.rtree_s.ancestors_children_unique) |
692 from assms(1) h(1) have "z \<in> ancestors (tRAG s) (Th th)" |
525 from assms(1) h(1) have "z \<in> ancestors (tRAG (e#s)) (Th th)" |
693 by (auto simp:ancestors_def) |
526 by (auto simp:ancestors_def) |
694 with h(2) show " z \<in> ancestors (tRAG s) (Th th) \<inter> |
527 with h(2) show " z \<in> ancestors (tRAG (e#s)) (Th th) \<inter> |
695 RTree.children (tRAG s) x" by auto |
528 RTree.children (tRAG (e#s)) x" by auto |
696 next |
529 next |
697 from a_in a_in' |
530 from a_in a_in' |
698 show "a \<in> ancestors (tRAG s) (Th th) \<inter> RTree.children (tRAG s) x" |
531 show "a \<in> ancestors (tRAG (e#s)) (Th th) \<inter> RTree.children (tRAG (e#s)) x" |
699 by auto |
532 by auto |
700 qed |
533 qed |
701 with False show False by auto |
534 with False show False by auto |
702 qed |
535 qed |
703 from cp_kept[OF this[unfolded eq_a]] |
536 from cp_kept[OF this[unfolded eq_a]] |
704 have "cp s th_a = cp s' th_a" . |
537 have "cp (e#s) th_a = cp s th_a" . |
705 from this[unfolded cp_gen_def_cond[OF eq_a], folded eq_a] |
538 from this[unfolded cp_gen_def_cond[OF eq_a], folded eq_a] |
706 show ?thesis . |
539 show ?thesis . |
707 qed |
540 qed |
708 qed |
541 qed |
709 qed |
542 qed |
710 ultimately show ?thesis by metis |
543 ultimately show ?thesis by metis |
711 qed |
544 qed |
712 ultimately show ?thesis by simp |
545 ultimately show ?thesis by simp |
713 qed |
546 qed |
714 also have "... = ?R" |
547 also have "... = ?R" |
715 by (fold vat_s'.cp_gen_rec[OF eq_x], simp) |
548 by (fold cp_gen_rec[OF eq_x], simp) |
716 finally show ?thesis . |
549 finally show ?thesis . |
717 qed |
550 qed |
718 qed |
551 qed |
719 |
552 |
720 lemma cp_up: |
553 lemma cp_up: |
721 assumes "(Th th') \<in> ancestors (tRAG s) (Th th)" |
554 assumes "(Th th') \<in> ancestors (tRAG (e#s)) (Th th)" |
722 and "cp s th' = cp s' th'" |
555 and "cp (e#s) th' = cp s th'" |
723 and "(Th th'') \<in> ancestors (tRAG s) (Th th')" |
556 and "(Th th'') \<in> ancestors (tRAG (e#s)) (Th th')" |
724 shows "cp s th'' = cp s' th''" |
557 shows "cp (e#s) th'' = cp s th''" |
725 proof - |
558 proof - |
726 have "cp_gen s (Th th'') = cp_gen s' (Th th'')" |
559 have "cp_gen (e#s) (Th th'') = cp_gen s (Th th'')" |
727 proof(rule cp_gen_update_stop[OF assms(1) _ assms(3)]) |
560 proof(rule cp_gen_update_stop[OF assms(1) _ assms(3)]) |
728 from assms(2) cp_gen_def_cond[OF refl[of "Th th'"]] |
561 from assms(2) cp_gen_def_cond[OF refl[of "Th th'"]] |
729 show "cp_gen s (Th th') = cp_gen s' (Th th')" by metis |
562 show "cp_gen (e#s) (Th th') = cp_gen s (Th th')" by metis |
730 qed |
563 qed |
731 with cp_gen_def_cond[OF refl[of "Th th''"]] |
564 with cp_gen_def_cond[OF refl[of "Th th''"]] |
732 show ?thesis by metis |
565 show ?thesis by metis |
733 qed |
566 qed |
734 |
567 |
735 end |
568 end |
736 |
569 |
737 section {* The @{term Create} operation *} |
570 section {* The @{term Create} operation *} |
738 |
571 |
739 locale step_create_cps = |
572 context valid_trace_create |
740 fixes s' th prio s |
573 begin |
741 defines s_def : "s \<equiv> (Create th prio#s')" |
574 |
742 assumes vt_s: "vt s" |
575 interpretation vat_e: valid_trace "e#s" |
743 |
576 by (unfold_locales, insert vt_e, simp) |
744 sublocale step_create_cps < vat_s: valid_trace "s" |
577 |
745 by (unfold_locales, insert vt_s, simp) |
578 lemma tRAG_kept: "tRAG (e#s) = tRAG s" |
746 |
579 by (unfold tRAG_alt_def RAG_unchanged, auto) |
747 sublocale step_create_cps < vat_s': valid_trace "s'" |
|
748 by (unfold_locales, insert step_back_vt[OF vt_s[unfolded s_def]], simp) |
|
749 |
|
750 context step_create_cps |
|
751 begin |
|
752 |
|
753 lemma RAG_kept: "RAG s = RAG s'" |
|
754 by (unfold s_def RAG_create_unchanged, auto) |
|
755 |
|
756 lemma tRAG_kept: "tRAG s = tRAG s'" |
|
757 by (unfold tRAG_alt_def RAG_kept, auto) |
|
758 |
580 |
759 lemma preced_kept: |
581 lemma preced_kept: |
760 assumes "th' \<noteq> th" |
582 assumes "th' \<noteq> th" |
761 shows "the_preced s th' = the_preced s' th'" |
583 shows "the_preced (e#s) th' = the_preced s th'" |
762 by (unfold s_def the_preced_def preced_def, insert assms, auto) |
584 by (unfold the_preced_def preced_def is_create, insert assms, auto) |
763 |
585 |
764 lemma th_not_in: "Th th \<notin> Field (tRAG s')" |
586 lemma th_not_in: "Th th \<notin> Field (tRAG s)" |
765 proof - |
587 by (meson not_in_thread_isolated subsetCE tRAG_Field th_not_live_s) |
766 from vt_s[unfolded s_def] |
|
767 have "PIP s' (Create th prio)" by (cases, simp) |
|
768 hence "th \<notin> threads s'" by(cases, simp) |
|
769 from vat_s'.not_in_thread_isolated[OF this] |
|
770 have "Th th \<notin> Field (RAG s')" . |
|
771 with tRAG_Field show ?thesis by auto |
|
772 qed |
|
773 |
588 |
774 lemma eq_cp: |
589 lemma eq_cp: |
775 assumes neq_th: "th' \<noteq> th" |
590 assumes neq_th: "th' \<noteq> th" |
776 shows "cp s th' = cp s' th'" |
591 shows "cp (e#s) th' = cp s th'" |
777 proof - |
592 proof - |
778 have "(the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th') = |
593 have "(the_preced (e#s) \<circ> the_thread) ` subtree (tRAG (e#s)) (Th th') = |
779 (the_preced s' \<circ> the_thread) ` subtree (tRAG s') (Th th')" |
594 (the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th')" |
780 proof(unfold tRAG_kept, rule f_image_eq) |
595 proof(unfold tRAG_kept, rule f_image_eq) |
781 fix a |
596 fix a |
782 assume a_in: "a \<in> subtree (tRAG s') (Th th')" |
597 assume a_in: "a \<in> subtree (tRAG s) (Th th')" |
783 then obtain th_a where eq_a: "a = Th th_a" |
598 then obtain th_a where eq_a: "a = Th th_a" |
784 proof(cases rule:subtreeE) |
599 proof(cases rule:subtreeE) |
785 case 2 |
600 case 2 |
786 from ancestors_Field[OF 2(2)] |
601 from ancestors_Field[OF 2(2)] |
787 and that show ?thesis by (unfold tRAG_alt_def, auto) |
602 and that show ?thesis by (unfold tRAG_alt_def, auto) |
788 qed auto |
603 qed auto |
789 have neq_th_a: "th_a \<noteq> th" |
604 have neq_th_a: "th_a \<noteq> th" |
790 proof - |
605 proof - |
791 have "(Th th) \<notin> subtree (tRAG s') (Th th')" |
606 have "(Th th) \<notin> subtree (tRAG s) (Th th')" |
792 proof |
607 proof |
793 assume "Th th \<in> subtree (tRAG s') (Th th')" |
608 assume "Th th \<in> subtree (tRAG s) (Th th')" |
794 thus False |
609 thus False |
795 proof(cases rule:subtreeE) |
610 proof(cases rule:subtreeE) |
796 case 2 |
611 case 2 |
797 from ancestors_Field[OF this(2)] |
612 from ancestors_Field[OF this(2)] |
798 and th_not_in[unfolded Field_def] |
613 and th_not_in[unfolded Field_def] |
800 qed (insert assms, auto) |
615 qed (insert assms, auto) |
801 qed |
616 qed |
802 with a_in[unfolded eq_a] show ?thesis by auto |
617 with a_in[unfolded eq_a] show ?thesis by auto |
803 qed |
618 qed |
804 from preced_kept[OF this] |
619 from preced_kept[OF this] |
805 show "(the_preced s \<circ> the_thread) a = (the_preced s' \<circ> the_thread) a" |
620 show "(the_preced (e#s) \<circ> the_thread) a = (the_preced s \<circ> the_thread) a" |
806 by (unfold eq_a, simp) |
621 by (unfold eq_a, simp) |
807 qed |
622 qed |
808 thus ?thesis by (unfold cp_alt_def1, simp) |
623 thus ?thesis by (unfold cp_alt_def1, simp) |
809 qed |
624 qed |
810 |
625 |
811 lemma children_of_th: "RTree.children (tRAG s) (Th th) = {}" |
626 lemma children_of_th: "RTree.children (tRAG (e#s)) (Th th) = {}" |
812 proof - |
627 proof - |
813 { fix a |
628 { fix a |
814 assume "a \<in> RTree.children (tRAG s) (Th th)" |
629 assume "a \<in> RTree.children (tRAG (e#s)) (Th th)" |
815 hence "(a, Th th) \<in> tRAG s" by (auto simp:RTree.children_def) |
630 hence "(a, Th th) \<in> tRAG (e#s)" by (auto simp:RTree.children_def) |
816 with th_not_in have False |
631 with th_not_in have False |
817 by (unfold Field_def tRAG_kept, auto) |
632 by (unfold Field_def tRAG_kept, auto) |
818 } thus ?thesis by auto |
633 } thus ?thesis by auto |
819 qed |
634 qed |
820 |
635 |
821 lemma eq_cp_th: "cp s th = preced th s" |
636 lemma eq_cp_th: "cp (e#s) th = preced th (e#s)" |
822 by (unfold vat_s.cp_rec children_of_th, simp add:the_preced_def) |
637 by (unfold vat_e.cp_rec children_of_th, simp add:the_preced_def) |
823 |
638 |
824 end |
639 end |
825 |
640 |
826 locale step_exit_cps = |
641 |
827 fixes s' th prio s |
642 context valid_trace_exit |
828 defines s_def : "s \<equiv> Exit th # s'" |
|
829 assumes vt_s: "vt s" |
|
830 |
|
831 sublocale step_exit_cps < vat_s: valid_trace "s" |
|
832 by (unfold_locales, insert vt_s, simp) |
|
833 |
|
834 sublocale step_exit_cps < vat_s': valid_trace "s'" |
|
835 by (unfold_locales, insert step_back_vt[OF vt_s[unfolded s_def]], simp) |
|
836 |
|
837 context step_exit_cps |
|
838 begin |
643 begin |
839 |
644 |
840 lemma preced_kept: |
645 lemma preced_kept: |
841 assumes "th' \<noteq> th" |
646 assumes "th' \<noteq> th" |
842 shows "the_preced s th' = the_preced s' th'" |
647 shows "the_preced (e#s) th' = the_preced s th'" |
843 by (unfold s_def the_preced_def preced_def, insert assms, auto) |
648 using assms |
844 |
649 by (unfold the_preced_def is_exit preced_def, simp) |
845 lemma RAG_kept: "RAG s = RAG s'" |
650 |
846 by (unfold s_def RAG_exit_unchanged, auto) |
651 lemma tRAG_kept: "tRAG (e#s) = tRAG s" |
847 |
652 by (unfold tRAG_alt_def RAG_unchanged, auto) |
848 lemma tRAG_kept: "tRAG s = tRAG s'" |
653 |
849 by (unfold tRAG_alt_def RAG_kept, auto) |
654 lemma th_RAG: "Th th \<notin> Field (RAG s)" |
850 |
655 proof - |
851 lemma th_ready: "th \<in> readys s'" |
656 have "Th th \<notin> Range (RAG s)" |
852 proof - |
|
853 from vt_s[unfolded s_def] |
|
854 have "PIP s' (Exit th)" by (cases, simp) |
|
855 hence h: "th \<in> runing s' \<and> holdents s' th = {}" by (cases, metis) |
|
856 thus ?thesis by (unfold runing_def, auto) |
|
857 qed |
|
858 |
|
859 lemma th_holdents: "holdents s' th = {}" |
|
860 proof - |
|
861 from vt_s[unfolded s_def] |
|
862 have "PIP s' (Exit th)" by (cases, simp) |
|
863 thus ?thesis by (cases, metis) |
|
864 qed |
|
865 |
|
866 lemma th_RAG: "Th th \<notin> Field (RAG s')" |
|
867 proof - |
|
868 have "Th th \<notin> Range (RAG s')" |
|
869 proof |
657 proof |
870 assume "Th th \<in> Range (RAG s')" |
658 assume "Th th \<in> Range (RAG s)" |
871 then obtain cs where "holding (wq s') th cs" |
659 then obtain cs where "holding (wq s) th cs" |
872 by (unfold Range_iff s_RAG_def, auto) |
660 by (unfold Range_iff s_RAG_def, auto) |
873 with th_holdents[unfolded holdents_def] |
661 with holdents_th_s[unfolded holdents_def] |
874 show False by (unfold eq_holding, auto) |
662 show False by (unfold holding_eq, auto) |
875 qed |
663 qed |
876 moreover have "Th th \<notin> Domain (RAG s')" |
664 moreover have "Th th \<notin> Domain (RAG s)" |
877 proof |
665 proof |
878 assume "Th th \<in> Domain (RAG s')" |
666 assume "Th th \<in> Domain (RAG s)" |
879 then obtain cs where "waiting (wq s') th cs" |
667 then obtain cs where "waiting (wq s) th cs" |
880 by (unfold Domain_iff s_RAG_def, auto) |
668 by (unfold Domain_iff s_RAG_def, auto) |
881 with th_ready show False by (unfold readys_def eq_waiting, auto) |
669 with th_ready_s show False by (unfold readys_def waiting_eq, auto) |
882 qed |
670 qed |
883 ultimately show ?thesis by (auto simp:Field_def) |
671 ultimately show ?thesis by (auto simp:Field_def) |
884 qed |
672 qed |
885 |
673 |
886 lemma th_tRAG: "(Th th) \<notin> Field (tRAG s')" |
674 lemma th_tRAG: "(Th th) \<notin> Field (tRAG s)" |
887 using th_RAG tRAG_Field[of s'] by auto |
675 using th_RAG tRAG_Field by auto |
888 |
676 |
889 lemma eq_cp: |
677 lemma eq_cp: |
890 assumes neq_th: "th' \<noteq> th" |
678 assumes neq_th: "th' \<noteq> th" |
891 shows "cp s th' = cp s' th'" |
679 shows "cp (e#s) th' = cp s th'" |
892 proof - |
680 proof - |
893 have "(the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th') = |
681 have "(the_preced (e#s) \<circ> the_thread) ` subtree (tRAG (e#s)) (Th th') = |
894 (the_preced s' \<circ> the_thread) ` subtree (tRAG s') (Th th')" |
682 (the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th')" |
895 proof(unfold tRAG_kept, rule f_image_eq) |
683 proof(unfold tRAG_kept, rule f_image_eq) |
896 fix a |
684 fix a |
897 assume a_in: "a \<in> subtree (tRAG s') (Th th')" |
685 assume a_in: "a \<in> subtree (tRAG s) (Th th')" |
898 then obtain th_a where eq_a: "a = Th th_a" |
686 then obtain th_a where eq_a: "a = Th th_a" |
899 proof(cases rule:subtreeE) |
687 proof(cases rule:subtreeE) |
900 case 2 |
688 case 2 |
901 from ancestors_Field[OF 2(2)] |
689 from ancestors_Field[OF 2(2)] |
902 and that show ?thesis by (unfold tRAG_alt_def, auto) |
690 and that show ?thesis by (unfold tRAG_alt_def, auto) |
903 qed auto |
691 qed auto |
904 have neq_th_a: "th_a \<noteq> th" |
692 have neq_th_a: "th_a \<noteq> th" |
905 proof - |
693 proof - |
906 from vat_s'.readys_in_no_subtree[OF th_ready assms] |
694 from readys_in_no_subtree[OF th_ready_s assms] |
907 have "(Th th) \<notin> subtree (RAG s') (Th th')" . |
695 have "(Th th) \<notin> subtree (RAG s) (Th th')" . |
908 with tRAG_subtree_RAG[of s' "Th th'"] |
696 with tRAG_subtree_RAG[of s "Th th'"] |
909 have "(Th th) \<notin> subtree (tRAG s') (Th th')" by auto |
697 have "(Th th) \<notin> subtree (tRAG s) (Th th')" by auto |
910 with a_in[unfolded eq_a] show ?thesis by auto |
698 with a_in[unfolded eq_a] show ?thesis by auto |
911 qed |
699 qed |
912 from preced_kept[OF this] |
700 from preced_kept[OF this] |
913 show "(the_preced s \<circ> the_thread) a = (the_preced s' \<circ> the_thread) a" |
701 show "(the_preced (e#s) \<circ> the_thread) a = (the_preced s \<circ> the_thread) a" |
914 by (unfold eq_a, simp) |
702 by (unfold eq_a, simp) |
915 qed |
703 qed |
916 thus ?thesis by (unfold cp_alt_def1, simp) |
704 thus ?thesis by (unfold cp_alt_def1, simp) |
917 qed |
705 qed |
918 |
706 |
919 end |
707 end |
920 |
708 |
921 end |
709 end |
922 |
710 |
|
711 |