| author | zhangx | 
| Fri, 29 Jan 2016 11:01:13 +0800 | |
| changeset 93 | 524bd3caa6b6 | 
| parent 92 | 4763aa246dbd | 
| child 97 | c7ba70dc49bd | 
| child 102 | 3a801bbd2687 | 
| permissions | -rw-r--r-- | 
| 53 
8142e80f5d58
Finished comments on PrioGDef.thy
 xingyuan zhang <xingyuanzhang@126.com> parents: 
45diff
changeset | 1 | section {*
 | 
| 
8142e80f5d58
Finished comments on PrioGDef.thy
 xingyuan zhang <xingyuanzhang@126.com> parents: 
45diff
changeset | 2 | This file contains lemmas used to guide the recalculation of current precedence | 
| 
8142e80f5d58
Finished comments on PrioGDef.thy
 xingyuan zhang <xingyuanzhang@126.com> parents: 
45diff
changeset | 3 | after every system call (or system operation) | 
| 
8142e80f5d58
Finished comments on PrioGDef.thy
 xingyuan zhang <xingyuanzhang@126.com> parents: 
45diff
changeset | 4 | *} | 
| 93 
524bd3caa6b6
The overwriten original .thy files are working now. The ones in last revision aren't.
 zhangx parents: 
92diff
changeset | 5 | theory Implementation | 
| 
524bd3caa6b6
The overwriten original .thy files are working now. The ones in last revision aren't.
 zhangx parents: 
92diff
changeset | 6 | imports PIPBasics | 
| 63 | 7 | begin | 
| 8 | ||
| 55 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 9 | text {* (* ddd *)
 | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 10 | One beauty of our modelling is that we follow the definitional extension tradition of HOL. | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 11 | The benefit of such a concise and miniature model is that large number of intuitively | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 12 | obvious facts are derived as lemmas, rather than asserted as axioms. | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 13 | *} | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 14 | |
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 15 | text {*
 | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 16 | However, the lemmas in the forthcoming several locales are no longer | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 17 | obvious. These lemmas show how the current precedences should be recalculated | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 18 | after every execution step (in our model, every step is represented by an event, | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 19 | which in turn, represents a system call, or operation). Each operation is | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 20 | treated in a separate locale. | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 21 | |
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 22 | The complication of current precedence recalculation comes | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 23 | because the changing of RAG needs to be taken into account, | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 24 | in addition to the changing of precedence. | 
| 68 | 25 | |
| 55 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 26 | The reason RAG changing affects current precedence is that, | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 27 | according to the definition, current precedence | 
| 68 | 28 | of a thread is the maximum of the precedences of every threads in its subtree, | 
| 29 | where the notion of sub-tree in RAG is defined in RTree.thy. | |
| 55 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 30 | |
| 68 | 31 | Therefore, for each operation, lemmas about the change of precedences | 
| 32 | and RAG are derived first, on which lemmas about current precedence | |
| 33 | recalculation are based on. | |
| 55 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 34 | *} | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 35 | |
| 68 | 36 | section {* The @{term Set} operation *}
 | 
| 37 | ||
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 38 | context valid_trace_set | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 39 | begin | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 40 | |
| 55 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 41 | text {* (* ddd *)
 | 
| 68 | 42 |   The following two lemmas confirm that @{text "Set"}-operation
 | 
| 43 | only changes the precedence of the initiating thread (or actor) | |
| 44 | of the operation (or event). | |
| 55 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 45 | *} | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 46 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 47 | |
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 48 | lemma eq_preced: | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 49 | assumes "th' \<noteq> th" | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 50 | shows "preced th' (e#s) = preced th' s" | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 51 | proof - | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 52 | from assms show ?thesis | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 53 | by (unfold is_set, auto simp:preced_def) | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 54 | qed | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 55 | |
| 58 | 56 | lemma eq_the_preced: | 
| 57 | assumes "th' \<noteq> th" | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 58 | shows "the_preced (e#s) th' = the_preced s th'" | 
| 58 | 59 | using assms | 
| 60 | by (unfold the_preced_def, intro eq_preced, simp) | |
| 61 | ||
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 62 | |
| 58 | 63 | text {* (* ddd *)
 | 
| 68 | 64 |   Th following lemma @{text "eq_cp_pre"} says that the priority change of @{text "th"}
 | 
| 58 | 65 |   only affects those threads, which as @{text "Th th"} in their sub-trees.
 | 
| 66 | ||
| 68 | 67 | The proof of this lemma is simplified by using the alternative definition | 
| 68 |   of @{text "cp"}. 
 | |
| 55 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 69 | *} | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 70 | |
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 71 | lemma eq_cp_pre: | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 72 | assumes nd: "Th th \<notin> subtree (RAG s) (Th th')" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 73 | shows "cp (e#s) th' = cp s th'" | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 74 | proof - | 
| 58 | 75 |   -- {* After unfolding using the alternative definition, elements 
 | 
| 76 |         affecting the @{term "cp"}-value of threads become explicit. 
 | |
| 77 | We only need to prove the following: *} | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 78 |   have "Max (the_preced (e#s) ` {th'a. Th th'a \<in> subtree (RAG (e#s)) (Th th')}) =
 | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 79 |         Max (the_preced s ` {th'a. Th th'a \<in> subtree (RAG s) (Th th')})"
 | 
| 58 | 80 | (is "Max (?f ` ?S1) = Max (?g ` ?S2)") | 
| 55 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 81 | proof - | 
| 58 | 82 |     -- {* The base sets are equal. *}
 | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 83 | have "?S1 = ?S2" using RAG_unchanged by simp | 
| 58 | 84 |     -- {* The function values on the base set are equal as well. *}
 | 
| 85 | moreover have "\<forall> e \<in> ?S2. ?f e = ?g e" | |
| 86 | proof | |
| 87 | fix th1 | |
| 88 | assume "th1 \<in> ?S2" | |
| 89 | with nd have "th1 \<noteq> th" by (auto) | |
| 90 | from eq_the_preced[OF this] | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 91 | show "the_preced (e#s) th1 = the_preced s th1" . | 
| 58 | 92 | qed | 
| 93 |     -- {* Therefore, the image of the functions are equal. *}
 | |
| 94 | ultimately have "(?f ` ?S1) = (?g ` ?S2)" by (auto intro!:f_image_eq) | |
| 95 | thus ?thesis by simp | |
| 96 | qed | |
| 97 | thus ?thesis by (simp add:cp_alt_def) | |
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 98 | qed | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 99 | |
| 55 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 100 | text {*
 | 
| 58 | 101 |   The following lemma shows that @{term "th"} is not in the 
 | 
| 102 | sub-tree of any other thread. | |
| 55 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 103 | *} | 
| 58 | 104 | lemma th_in_no_subtree: | 
| 105 | assumes "th' \<noteq> th" | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 106 | shows "Th th \<notin> subtree (RAG s) (Th th')" | 
| 58 | 107 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 108 | from readys_in_no_subtree[OF th_ready_s assms(1)] | 
| 58 | 109 | show ?thesis by blast | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 110 | qed | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 111 | |
| 55 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 112 | text {* 
 | 
| 58 | 113 |   By combining @{thm "eq_cp_pre"} and @{thm "th_in_no_subtree"}, 
 | 
| 114 |   it is obvious that the change of priority only affects the @{text "cp"}-value 
 | |
| 115 |   of the initiating thread @{text "th"}.
 | |
| 55 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 116 | *} | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 117 | lemma eq_cp: | 
| 55 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 118 | assumes "th' \<noteq> th" | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 119 | shows "cp (e#s) th' = cp s th'" | 
| 58 | 120 | by (rule eq_cp_pre[OF th_in_no_subtree[OF assms]]) | 
| 55 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 121 | |
| 56 
0fd478e14e87
Before switching to generic theory of relational trees.
 xingyuan zhang <xingyuanzhang@126.com> parents: 
55diff
changeset | 122 | end | 
| 55 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 123 | |
| 68 | 124 | section {* The @{term V} operation *}
 | 
| 125 | ||
| 56 
0fd478e14e87
Before switching to generic theory of relational trees.
 xingyuan zhang <xingyuanzhang@126.com> parents: 
55diff
changeset | 126 | text {*
 | 
| 
0fd478e14e87
Before switching to generic theory of relational trees.
 xingyuan zhang <xingyuanzhang@126.com> parents: 
55diff
changeset | 127 |   The following @{text "step_v_cps"} is the locale for @{text "V"}-operation.
 | 
| 55 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 128 | *} | 
| 
b85cfbd58f59
Comments for Set-operation finished
 xingyuan zhang <xingyuanzhang@126.com> parents: 
53diff
changeset | 129 | |
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 130 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 131 | context valid_trace_v | 
| 61 | 132 | begin | 
| 58 | 133 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 134 | lemma ancestors_th: "ancestors (RAG s) (Th th) = {}"
 | 
| 58 | 135 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 136 | from readys_root[OF th_ready_s] | 
| 58 | 137 | show ?thesis | 
| 138 | by (unfold root_def, simp) | |
| 139 | qed | |
| 140 | ||
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 141 | lemma edge_of_th: | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 142 | "(Cs cs, Th th) \<in> RAG s" | 
| 58 | 143 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 144 | from holding_th_cs_s | 
| 58 | 145 | show ?thesis | 
| 146 | by (unfold s_RAG_def holding_eq, auto) | |
| 147 | qed | |
| 148 | ||
| 149 | lemma ancestors_cs: | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 150 |   "ancestors (RAG s) (Cs cs) = {Th th}"
 | 
| 58 | 151 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 152 |   have "ancestors (RAG s) (Cs cs) = ancestors (RAG s) (Th th)  \<union>  {Th th}"
 | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 153 | by (rule rtree_RAG.ancestors_accum[OF edge_of_th]) | 
| 58 | 154 | from this[unfolded ancestors_th] show ?thesis by simp | 
| 155 | qed | |
| 156 | ||
| 157 | end | |
| 158 | ||
| 56 
0fd478e14e87
Before switching to generic theory of relational trees.
 xingyuan zhang <xingyuanzhang@126.com> parents: 
55diff
changeset | 159 | text {*
 | 
| 
0fd478e14e87
Before switching to generic theory of relational trees.
 xingyuan zhang <xingyuanzhang@126.com> parents: 
55diff
changeset | 160 |   The following @{text "step_v_cps_nt"} is the sub-locale for @{text "V"}-operation, 
 | 
| 
0fd478e14e87
Before switching to generic theory of relational trees.
 xingyuan zhang <xingyuanzhang@126.com> parents: 
55diff
changeset | 161 |   which represents the case when there is another thread @{text "th'"}
 | 
| 
0fd478e14e87
Before switching to generic theory of relational trees.
 xingyuan zhang <xingyuanzhang@126.com> parents: 
55diff
changeset | 162 |   to take over the critical resource released by the initiating thread @{text "th"}.
 | 
| 
0fd478e14e87
Before switching to generic theory of relational trees.
 xingyuan zhang <xingyuanzhang@126.com> parents: 
55diff
changeset | 163 | *} | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 164 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 165 | context valid_trace_v_n | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 166 | begin | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 167 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 168 | lemma sub_RAGs': | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 169 |   "{(Cs cs, Th th), (Th taker, Cs cs)} \<subseteq> RAG s"
 | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 170 | using next_th_RAG[OF next_th_taker] . | 
| 58 | 171 | |
| 172 | lemma ancestors_th': | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 173 |   "ancestors (RAG s) (Th taker) = {Th th, Cs cs}" 
 | 
| 58 | 174 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 175 |   have "ancestors (RAG s) (Th taker) = ancestors (RAG s) (Cs cs) \<union> {Cs cs}"
 | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 176 | proof(rule rtree_RAG.ancestors_accum) | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 177 | from sub_RAGs' show "(Th taker, Cs cs) \<in> RAG s" by auto | 
| 58 | 178 | qed | 
| 179 | thus ?thesis using ancestors_th ancestors_cs by auto | |
| 180 | qed | |
| 181 | ||
| 35 
92f61f6a0fe7
added a bit more text to the paper and separated a theory about Max
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: 
33diff
changeset | 182 | lemma RAG_s: | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 183 |   "RAG (e#s) = (RAG s - {(Cs cs, Th th), (Th taker, Cs cs)}) \<union>
 | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 184 |                                          {(Cs cs, Th taker)}"
 | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 185 | by (unfold RAG_es waiting_set_eq holding_set_eq, auto) | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 186 | |
| 68 | 187 | lemma subtree_kept: (* ddd *) | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 188 |   assumes "th1 \<notin> {th, taker}"
 | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 189 | shows "subtree (RAG (e#s)) (Th th1) = | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 190 | subtree (RAG s) (Th th1)" (is "_ = ?R") | 
| 58 | 191 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 192 |   let ?RAG' = "(RAG s - {(Cs cs, Th th), (Th taker, Cs cs)})"
 | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 193 |   let ?RAG'' = "?RAG' \<union> {(Cs cs, Th taker)}"
 | 
| 58 | 194 | have "subtree ?RAG' (Th th1) = ?R" | 
| 195 | proof(rule subset_del_subtree_outside) | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 196 |     show "Range {(Cs cs, Th th), (Th taker, Cs cs)} \<inter> subtree (RAG s) (Th th1) = {}"
 | 
| 58 | 197 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 198 | have "(Th th) \<notin> subtree (RAG s) (Th th1)" | 
| 58 | 199 | proof(rule subtree_refute) | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 200 | show "Th th1 \<notin> ancestors (RAG s) (Th th)" | 
| 58 | 201 | by (unfold ancestors_th, simp) | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 202 | next | 
| 58 | 203 | from assms show "Th th1 \<noteq> Th th" by simp | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 204 | qed | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 205 | moreover have "(Cs cs) \<notin> subtree (RAG s) (Th th1)" | 
| 58 | 206 | proof(rule subtree_refute) | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 207 | show "Th th1 \<notin> ancestors (RAG s) (Cs cs)" | 
| 58 | 208 | by (unfold ancestors_cs, insert assms, auto) | 
| 209 | qed simp | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 210 |       ultimately have "{Th th, Cs cs} \<inter> subtree (RAG s) (Th th1) = {}" by auto
 | 
| 58 | 211 | thus ?thesis by simp | 
| 212 | qed | |
| 213 | qed | |
| 214 | moreover have "subtree ?RAG'' (Th th1) = subtree ?RAG' (Th th1)" | |
| 215 | proof(rule subtree_insert_next) | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 216 |     show "Th taker \<notin> subtree (RAG s - {(Cs cs, Th th), (Th taker, Cs cs)}) (Th th1)"
 | 
| 58 | 217 | proof(rule subtree_refute) | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 218 |       show "Th th1 \<notin> ancestors (RAG s - {(Cs cs, Th th), (Th taker, Cs cs)}) (Th taker)"
 | 
| 58 | 219 | (is "_ \<notin> ?R") | 
| 220 | proof - | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 221 | have "?R \<subseteq> ancestors (RAG s) (Th taker)" by (rule ancestors_mono, auto) | 
| 58 | 222 | moreover have "Th th1 \<notin> ..." using ancestors_th' assms by simp | 
| 223 | ultimately show ?thesis by auto | |
| 224 | qed | |
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 225 | next | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 226 | from assms show "Th th1 \<noteq> Th taker" by simp | 
| 58 | 227 | qed | 
| 228 | qed | |
| 229 | ultimately show ?thesis by (unfold RAG_s, simp) | |
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 230 | qed | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 231 | |
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 232 | lemma cp_kept: | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 233 |   assumes "th1 \<notin> {th, taker}"
 | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 234 | shows "cp (e#s) th1 = cp s th1" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 235 | by (unfold cp_alt_def the_preced_es subtree_kept[OF assms], simp) | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 236 | |
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 237 | end | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 238 | |
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 239 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 240 | context valid_trace_v_e | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 241 | begin | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 242 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 243 | find_theorems RAG s e | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 244 | |
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 245 | lemma RAG_s: "RAG (e#s) = RAG s - {(Cs cs, Th th)}"
 | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 246 | by (unfold RAG_es waiting_set_eq holding_set_eq, simp) | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 247 | |
| 58 | 248 | lemma subtree_kept: | 
| 249 | assumes "th1 \<noteq> th" | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 250 | shows "subtree (RAG (e#s)) (Th th1) = subtree (RAG s) (Th th1)" | 
| 58 | 251 | proof(unfold RAG_s, rule subset_del_subtree_outside) | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 252 |   show "Range {(Cs cs, Th th)} \<inter> subtree (RAG s) (Th th1) = {}"
 | 
| 58 | 253 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 254 | have "(Th th) \<notin> subtree (RAG s) (Th th1)" | 
| 58 | 255 | proof(rule subtree_refute) | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 256 | show "Th th1 \<notin> ancestors (RAG s) (Th th)" | 
| 58 | 257 | by (unfold ancestors_th, simp) | 
| 258 | next | |
| 259 | from assms show "Th th1 \<noteq> Th th" by simp | |
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 260 | qed | 
| 58 | 261 | thus ?thesis by auto | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 262 | qed | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 263 | qed | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 264 | |
| 58 | 265 | lemma cp_kept_1: | 
| 266 | assumes "th1 \<noteq> th" | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 267 | shows "cp (e#s) th1 = cp s th1" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 268 | by (unfold cp_alt_def the_preced_es subtree_kept[OF assms], simp) | 
| 58 | 269 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 270 | lemma subtree_cs: "subtree (RAG s) (Cs cs) = {Cs cs}"
 | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 271 | proof - | 
| 58 | 272 |   { fix n
 | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 273 | have "(Cs cs) \<notin> ancestors (RAG s) n" | 
| 58 | 274 | proof | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 275 | assume "Cs cs \<in> ancestors (RAG s) n" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 276 | hence "(n, Cs cs) \<in> (RAG s)^+" by (auto simp:ancestors_def) | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 277 | from tranclE[OF this] obtain nn where h: "(nn, Cs cs) \<in> RAG s" by auto | 
| 58 | 278 | then obtain th' where "nn = Th th'" | 
| 279 | by (unfold s_RAG_def, auto) | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 280 | from h[unfolded this] have "(Th th', Cs cs) \<in> RAG s" . | 
| 58 | 281 | from this[unfolded s_RAG_def] | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 282 | have "waiting (wq s) th' cs" by auto | 
| 58 | 283 | from this[unfolded cs_waiting_def] | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 284 | have "1 < length (wq s cs)" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 285 | by (cases "wq s cs", auto) | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 286 | from holding_next_thI[OF holding_th_cs_s this] | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 287 | obtain th' where "next_th s th cs th'" by auto | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 288 | thus False using no_taker by blast | 
| 58 | 289 | qed | 
| 290 | } note h = this | |
| 291 |   {  fix n
 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 292 | assume "n \<in> subtree (RAG s) (Cs cs)" | 
| 58 | 293 | hence "n = (Cs cs)" | 
| 294 | by (elim subtreeE, insert h, auto) | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 295 | } moreover have "(Cs cs) \<in> subtree (RAG s) (Cs cs)" | 
| 58 | 296 | by (auto simp:subtree_def) | 
| 297 | ultimately show ?thesis by auto | |
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 298 | qed | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 299 | |
| 58 | 300 | lemma subtree_th: | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 301 |   "subtree (RAG (e#s)) (Th th) = subtree (RAG s) (Th th) - {Cs cs}"
 | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 302 | proof(unfold RAG_s, fold subtree_cs, rule rtree_RAG.subtree_del_inside) | 
| 58 | 303 | from edge_of_th | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 304 | show "(Cs cs, Th th) \<in> edges_in (RAG s) (Th th)" | 
| 58 | 305 | by (unfold edges_in_def, auto simp:subtree_def) | 
| 306 | qed | |
| 307 | ||
| 308 | lemma cp_kept_2: | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 309 | shows "cp (e#s) th = cp s th" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 310 | by (unfold cp_alt_def subtree_th the_preced_es, auto) | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 311 | |
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 312 | lemma eq_cp: | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 313 | shows "cp (e#s) th' = cp s th'" | 
| 58 | 314 | using cp_kept_1 cp_kept_2 | 
| 315 | by (cases "th' = th", auto) | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 316 | |
| 58 | 317 | end | 
| 318 | ||
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 319 | |
| 68 | 320 | section {* The @{term P} operation *}
 | 
| 321 | ||
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 322 | context valid_trace_p | 
| 58 | 323 | begin | 
| 324 | ||
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 325 | lemma root_th: "root (RAG s) (Th th)" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 326 | by (simp add: ready_th_s readys_root) | 
| 60 | 327 | |
| 328 | lemma in_no_others_subtree: | |
| 329 | assumes "th' \<noteq> th" | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 330 | shows "Th th \<notin> subtree (RAG s) (Th th')" | 
| 58 | 331 | proof | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 332 | assume "Th th \<in> subtree (RAG s) (Th th')" | 
| 60 | 333 | thus False | 
| 334 | proof(cases rule:subtreeE) | |
| 335 | case 1 | |
| 336 | with assms show ?thesis by auto | |
| 337 | next | |
| 338 | case 2 | |
| 339 | with root_th show ?thesis by (auto simp:root_def) | |
| 340 | qed | |
| 58 | 341 | qed | 
| 342 | ||
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 343 | lemma preced_kept: "the_preced (e#s) = the_preced s" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 344 | proof | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 345 | fix th' | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 346 | show "the_preced (e # s) th' = the_preced s th'" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 347 | by (unfold the_preced_def is_p preced_def, simp) | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 348 | qed | 
| 58 | 349 | |
| 350 | end | |
| 351 | ||
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 352 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 353 | context valid_trace_p_h | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 354 | begin | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 355 | |
| 60 | 356 | lemma subtree_kept: | 
| 357 | assumes "th' \<noteq> th" | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 358 | shows "subtree (RAG (e#s)) (Th th') = subtree (RAG s) (Th th')" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 359 | proof(unfold RAG_es, rule subtree_insert_next) | 
| 60 | 360 | from in_no_others_subtree[OF assms] | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 361 | show "Th th \<notin> subtree (RAG s) (Th th')" . | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 362 | qed | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 363 | |
| 60 | 364 | lemma cp_kept: | 
| 365 | assumes "th' \<noteq> th" | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 366 | shows "cp (e#s) th' = cp s th'" | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 367 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 368 |   have "(the_preced (e#s) ` {th'a. Th th'a \<in> subtree (RAG (e#s)) (Th th')}) =
 | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 369 |         (the_preced s ` {th'a. Th th'a \<in> subtree (RAG s) (Th th')})"
 | 
| 60 | 370 | by (unfold preced_kept subtree_kept[OF assms], simp) | 
| 371 | thus ?thesis by (unfold cp_alt_def, simp) | |
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 372 | qed | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 373 | |
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 374 | end | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 375 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 376 | context valid_trace_p_w | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 377 | begin | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 378 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 379 | interpretation vat_e: valid_trace "e#s" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 380 | by (unfold_locales, insert vt_e, simp) | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 381 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 382 | lemma cs_held: "(Cs cs, Th holder) \<in> RAG s" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 383 | using holding_s_holder | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 384 | by (unfold s_RAG_def, fold holding_eq, auto) | 
| 58 | 385 | |
| 386 | lemma tRAG_s: | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 387 |   "tRAG (e#s) = tRAG s \<union> {(Th th, Th holder)}"
 | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 388 | using local.RAG_tRAG_transfer[OF RAG_es cs_held] . | 
| 58 | 389 | |
| 390 | lemma cp_kept: | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 391 | assumes "Th th'' \<notin> ancestors (tRAG (e#s)) (Th th)" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 392 | shows "cp (e#s) th'' = cp s th''" | 
| 58 | 393 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 394 | have h: "subtree (tRAG (e#s)) (Th th'') = subtree (tRAG s) (Th th'')" | 
| 58 | 395 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 396 | have "Th holder \<notin> subtree (tRAG s) (Th th'')" | 
| 58 | 397 | proof | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 398 | assume "Th holder \<in> subtree (tRAG s) (Th th'')" | 
| 58 | 399 | thus False | 
| 400 | proof(rule subtreeE) | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 401 | assume "Th holder = Th th''" | 
| 58 | 402 | from assms[unfolded tRAG_s ancestors_def, folded this] | 
| 403 | show ?thesis by auto | |
| 404 | next | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 405 | assume "Th th'' \<in> ancestors (tRAG s) (Th holder)" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 406 | moreover have "... \<subseteq> ancestors (tRAG (e#s)) (Th holder)" | 
| 58 | 407 | proof(rule ancestors_mono) | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 408 | show "tRAG s \<subseteq> tRAG (e#s)" by (unfold tRAG_s, auto) | 
| 58 | 409 | qed | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 410 | ultimately have "Th th'' \<in> ancestors (tRAG (e#s)) (Th holder)" by auto | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 411 | moreover have "Th holder \<in> ancestors (tRAG (e#s)) (Th th)" | 
| 58 | 412 | by (unfold tRAG_s, auto simp:ancestors_def) | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 413 | ultimately have "Th th'' \<in> ancestors (tRAG (e#s)) (Th th)" | 
| 58 | 414 | by (auto simp:ancestors_def) | 
| 415 | with assms show ?thesis by auto | |
| 416 | qed | |
| 417 | qed | |
| 418 | from subtree_insert_next[OF this] | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 419 |     have "subtree (tRAG s \<union> {(Th th, Th holder)}) (Th th'') = subtree (tRAG s) (Th th'')" .
 | 
| 58 | 420 | from this[folded tRAG_s] show ?thesis . | 
| 421 | qed | |
| 422 | show ?thesis by (unfold cp_alt_def1 h preced_kept, simp) | |
| 423 | qed | |
| 424 | ||
| 60 | 425 | lemma cp_gen_update_stop: (* ddd *) | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 426 | assumes "u \<in> ancestors (tRAG (e#s)) (Th th)" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 427 | and "cp_gen (e#s) u = cp_gen s u" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 428 | and "y \<in> ancestors (tRAG (e#s)) u" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 429 | shows "cp_gen (e#s) y = cp_gen s y" | 
| 58 | 430 | using assms(3) | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 431 | proof(induct rule:wf_induct[OF vat_e.fsbttRAGs.wf]) | 
| 58 | 432 | case (1 x) | 
| 433 | show ?case (is "?L = ?R") | |
| 434 | proof - | |
| 435 | from tRAG_ancestorsE[OF 1(2)] | |
| 436 | obtain th2 where eq_x: "x = Th th2" by blast | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 437 | from vat_e.cp_gen_rec[OF this] | 
| 58 | 438 | have "?L = | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 439 |           Max ({the_preced (e#s) th2} \<union> cp_gen (e#s) ` RTree.children (tRAG (e#s)) x)" .
 | 
| 58 | 440 | also have "... = | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 441 |           Max ({the_preced s th2} \<union> cp_gen s ` RTree.children (tRAG s) x)"
 | 
| 58 | 442 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 443 | from preced_kept have "the_preced (e#s) th2 = the_preced s th2" by simp | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 444 | moreover have "cp_gen (e#s) ` RTree.children (tRAG (e#s)) x = | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 445 | cp_gen s ` RTree.children (tRAG s) x" | 
| 58 | 446 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 447 | have "RTree.children (tRAG (e#s)) x = RTree.children (tRAG s) x" | 
| 58 | 448 | proof(unfold tRAG_s, rule children_union_kept) | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 449 | have start: "(Th th, Th holder) \<in> tRAG (e#s)" | 
| 58 | 450 | by (unfold tRAG_s, auto) | 
| 451 | note x_u = 1(2) | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 452 |           show "x \<notin> Range {(Th th, Th holder)}"
 | 
| 58 | 453 | proof | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 454 |             assume "x \<in> Range {(Th th, Th holder)}"
 | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 455 | hence eq_x: "x = Th holder" using RangeE by auto | 
| 58 | 456 | show False | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 457 | proof(cases rule:vat_e.ancestors_headE[OF assms(1) start]) | 
| 58 | 458 | case 1 | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 459 | from x_u[folded this, unfolded eq_x] vat_e.acyclic_tRAG | 
| 58 | 460 | show ?thesis by (auto simp:ancestors_def acyclic_def) | 
| 461 | next | |
| 462 | case 2 | |
| 463 | with x_u[unfolded eq_x] | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 464 | have "(Th holder, Th holder) \<in> (tRAG (e#s))^+" by (auto simp:ancestors_def) | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 465 | with vat_e.acyclic_tRAG show ?thesis by (auto simp:acyclic_def) | 
| 58 | 466 | qed | 
| 467 | qed | |
| 468 | qed | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 469 | moreover have "cp_gen (e#s) ` RTree.children (tRAG (e#s)) x = | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 470 | cp_gen s ` RTree.children (tRAG (e#s)) x" (is "?f ` ?A = ?g ` ?A") | 
| 58 | 471 | proof(rule f_image_eq) | 
| 472 | fix a | |
| 473 | assume a_in: "a \<in> ?A" | |
| 474 | from 1(2) | |
| 475 | show "?f a = ?g a" | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 476 | proof(cases rule:vat_e.rtree_s.ancestors_childrenE[case_names in_ch out_ch]) | 
| 58 | 477 | case in_ch | 
| 478 | show ?thesis | |
| 479 | proof(cases "a = u") | |
| 480 | case True | |
| 481 | from assms(2)[folded this] show ?thesis . | |
| 482 | next | |
| 483 | case False | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 484 | have a_not_in: "a \<notin> ancestors (tRAG (e#s)) (Th th)" | 
| 58 | 485 | proof | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 486 | assume a_in': "a \<in> ancestors (tRAG (e#s)) (Th th)" | 
| 58 | 487 | have "a = u" | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 488 | proof(rule vat_e.rtree_s.ancestors_children_unique) | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 489 | from a_in' a_in show "a \<in> ancestors (tRAG (e#s)) (Th th) \<inter> | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 490 | RTree.children (tRAG (e#s)) x" by auto | 
| 58 | 491 | next | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 492 | from assms(1) in_ch show "u \<in> ancestors (tRAG (e#s)) (Th th) \<inter> | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 493 | RTree.children (tRAG (e#s)) x" by auto | 
| 58 | 494 | qed | 
| 495 | with False show False by simp | |
| 496 | qed | |
| 497 | from a_in obtain th_a where eq_a: "a = Th th_a" | |
| 498 | by (unfold RTree.children_def tRAG_alt_def, auto) | |
| 499 | from cp_kept[OF a_not_in[unfolded eq_a]] | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 500 | have "cp (e#s) th_a = cp s th_a" . | 
| 58 | 501 | from this [unfolded cp_gen_def_cond[OF eq_a], folded eq_a] | 
| 502 | show ?thesis . | |
| 503 | qed | |
| 504 | next | |
| 505 | case (out_ch z) | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 506 | hence h: "z \<in> ancestors (tRAG (e#s)) u" "z \<in> RTree.children (tRAG (e#s)) x" by auto | 
| 58 | 507 | show ?thesis | 
| 508 | proof(cases "a = z") | |
| 509 | case True | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 510 | from h(2) have zx_in: "(z, x) \<in> (tRAG (e#s))" by (auto simp:RTree.children_def) | 
| 58 | 511 | from 1(1)[rule_format, OF this h(1)] | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 512 | have eq_cp_gen: "cp_gen (e#s) z = cp_gen s z" . | 
| 58 | 513 | with True show ?thesis by metis | 
| 514 | next | |
| 515 | case False | |
| 516 | from a_in obtain th_a where eq_a: "a = Th th_a" | |
| 517 | by (auto simp:RTree.children_def tRAG_alt_def) | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 518 | have "a \<notin> ancestors (tRAG (e#s)) (Th th)" | 
| 60 | 519 | proof | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 520 | assume a_in': "a \<in> ancestors (tRAG (e#s)) (Th th)" | 
| 60 | 521 | have "a = z" | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 522 | proof(rule vat_e.rtree_s.ancestors_children_unique) | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 523 | from assms(1) h(1) have "z \<in> ancestors (tRAG (e#s)) (Th th)" | 
| 60 | 524 | by (auto simp:ancestors_def) | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 525 | with h(2) show " z \<in> ancestors (tRAG (e#s)) (Th th) \<inter> | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 526 | RTree.children (tRAG (e#s)) x" by auto | 
| 60 | 527 | next | 
| 528 | from a_in a_in' | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 529 | show "a \<in> ancestors (tRAG (e#s)) (Th th) \<inter> RTree.children (tRAG (e#s)) x" | 
| 60 | 530 | by auto | 
| 531 | qed | |
| 532 | with False show False by auto | |
| 533 | qed | |
| 58 | 534 | from cp_kept[OF this[unfolded eq_a]] | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 535 | have "cp (e#s) th_a = cp s th_a" . | 
| 58 | 536 | from this[unfolded cp_gen_def_cond[OF eq_a], folded eq_a] | 
| 537 | show ?thesis . | |
| 538 | qed | |
| 539 | qed | |
| 540 | qed | |
| 541 | ultimately show ?thesis by metis | |
| 542 | qed | |
| 543 | ultimately show ?thesis by simp | |
| 544 | qed | |
| 545 | also have "... = ?R" | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 546 | by (fold cp_gen_rec[OF eq_x], simp) | 
| 58 | 547 | finally show ?thesis . | 
| 548 | qed | |
| 549 | qed | |
| 550 | ||
| 60 | 551 | lemma cp_up: | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 552 | assumes "(Th th') \<in> ancestors (tRAG (e#s)) (Th th)" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 553 | and "cp (e#s) th' = cp s th'" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 554 | and "(Th th'') \<in> ancestors (tRAG (e#s)) (Th th')" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 555 | shows "cp (e#s) th'' = cp s th''" | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 556 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 557 | have "cp_gen (e#s) (Th th'') = cp_gen s (Th th'')" | 
| 60 | 558 | proof(rule cp_gen_update_stop[OF assms(1) _ assms(3)]) | 
| 559 | from assms(2) cp_gen_def_cond[OF refl[of "Th th'"]] | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 560 | show "cp_gen (e#s) (Th th') = cp_gen s (Th th')" by metis | 
| 60 | 561 | qed | 
| 562 | with cp_gen_def_cond[OF refl[of "Th th''"]] | |
| 563 | show ?thesis by metis | |
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 564 | qed | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 565 | |
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 566 | end | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 567 | |
| 68 | 568 | section {* The @{term Create} operation *}
 | 
| 569 | ||
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 570 | context valid_trace_create | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 571 | begin | 
| 60 | 572 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 573 | interpretation vat_e: valid_trace "e#s" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 574 | by (unfold_locales, insert vt_e, simp) | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 575 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 576 | lemma tRAG_kept: "tRAG (e#s) = tRAG s" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 577 | by (unfold tRAG_alt_def RAG_unchanged, auto) | 
| 60 | 578 | |
| 579 | lemma preced_kept: | |
| 580 | assumes "th' \<noteq> th" | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 581 | shows "the_preced (e#s) th' = the_preced s th'" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 582 | by (unfold the_preced_def preced_def is_create, insert assms, auto) | 
| 60 | 583 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 584 | lemma th_not_in: "Th th \<notin> Field (tRAG s)" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 585 | by (meson not_in_thread_isolated subsetCE tRAG_Field th_not_live_s) | 
| 60 | 586 | |
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 587 | lemma eq_cp: | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 588 | assumes neq_th: "th' \<noteq> th" | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 589 | shows "cp (e#s) th' = cp s th'" | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 590 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 591 | have "(the_preced (e#s) \<circ> the_thread) ` subtree (tRAG (e#s)) (Th th') = | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 592 | (the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th')" | 
| 60 | 593 | proof(unfold tRAG_kept, rule f_image_eq) | 
| 594 | fix a | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 595 | assume a_in: "a \<in> subtree (tRAG s) (Th th')" | 
| 60 | 596 | then obtain th_a where eq_a: "a = Th th_a" | 
| 597 | proof(cases rule:subtreeE) | |
| 598 | case 2 | |
| 599 | from ancestors_Field[OF 2(2)] | |
| 600 | and that show ?thesis by (unfold tRAG_alt_def, auto) | |
| 601 | qed auto | |
| 602 | have neq_th_a: "th_a \<noteq> th" | |
| 603 | proof - | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 604 | have "(Th th) \<notin> subtree (tRAG s) (Th th')" | 
| 60 | 605 | proof | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 606 | assume "Th th \<in> subtree (tRAG s) (Th th')" | 
| 60 | 607 | thus False | 
| 608 | proof(cases rule:subtreeE) | |
| 609 | case 2 | |
| 610 | from ancestors_Field[OF this(2)] | |
| 611 | and th_not_in[unfolded Field_def] | |
| 612 | show ?thesis by auto | |
| 613 | qed (insert assms, auto) | |
| 614 | qed | |
| 615 | with a_in[unfolded eq_a] show ?thesis by auto | |
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 616 | qed | 
| 60 | 617 | from preced_kept[OF this] | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 618 | show "(the_preced (e#s) \<circ> the_thread) a = (the_preced s \<circ> the_thread) a" | 
| 60 | 619 | by (unfold eq_a, simp) | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 620 | qed | 
| 60 | 621 | thus ?thesis by (unfold cp_alt_def1, simp) | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 622 | qed | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 623 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 624 | lemma children_of_th: "RTree.children (tRAG (e#s)) (Th th) = {}"
 | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 625 | proof - | 
| 60 | 626 |   { fix a
 | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 627 | assume "a \<in> RTree.children (tRAG (e#s)) (Th th)" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 628 | hence "(a, Th th) \<in> tRAG (e#s)" by (auto simp:RTree.children_def) | 
| 60 | 629 | with th_not_in have False | 
| 630 | by (unfold Field_def tRAG_kept, auto) | |
| 631 | } thus ?thesis by auto | |
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 632 | qed | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 633 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 634 | lemma eq_cp_th: "cp (e#s) th = preced th (e#s)" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 635 | by (unfold vat_e.cp_rec children_of_th, simp add:the_preced_def) | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 636 | |
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 637 | end | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 638 | |
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 639 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 640 | context valid_trace_exit | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 641 | begin | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 642 | |
| 60 | 643 | lemma preced_kept: | 
| 644 | assumes "th' \<noteq> th" | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 645 | shows "the_preced (e#s) th' = the_preced s th'" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 646 | using assms | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 647 | by (unfold the_preced_def is_exit preced_def, simp) | 
| 60 | 648 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 649 | lemma tRAG_kept: "tRAG (e#s) = tRAG s" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 650 | by (unfold tRAG_alt_def RAG_unchanged, auto) | 
| 60 | 651 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 652 | lemma th_RAG: "Th th \<notin> Field (RAG s)" | 
| 60 | 653 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 654 | have "Th th \<notin> Range (RAG s)" | 
| 60 | 655 | proof | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 656 | assume "Th th \<in> Range (RAG s)" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 657 | then obtain cs where "holding (wq s) th cs" | 
| 60 | 658 | by (unfold Range_iff s_RAG_def, auto) | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 659 | with holdents_th_s[unfolded holdents_def] | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 660 | show False by (unfold holding_eq, auto) | 
| 60 | 661 | qed | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 662 | moreover have "Th th \<notin> Domain (RAG s)" | 
| 60 | 663 | proof | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 664 | assume "Th th \<in> Domain (RAG s)" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 665 | then obtain cs where "waiting (wq s) th cs" | 
| 60 | 666 | by (unfold Domain_iff s_RAG_def, auto) | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 667 | with th_ready_s show False by (unfold readys_def waiting_eq, auto) | 
| 60 | 668 | qed | 
| 669 | ultimately show ?thesis by (auto simp:Field_def) | |
| 670 | qed | |
| 671 | ||
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 672 | lemma th_tRAG: "(Th th) \<notin> Field (tRAG s)" | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 673 | using th_RAG tRAG_Field by auto | 
| 60 | 674 | |
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 675 | lemma eq_cp: | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 676 | assumes neq_th: "th' \<noteq> th" | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 677 | shows "cp (e#s) th' = cp s th'" | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 678 | proof - | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 679 | have "(the_preced (e#s) \<circ> the_thread) ` subtree (tRAG (e#s)) (Th th') = | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 680 | (the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th')" | 
| 60 | 681 | proof(unfold tRAG_kept, rule f_image_eq) | 
| 682 | fix a | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 683 | assume a_in: "a \<in> subtree (tRAG s) (Th th')" | 
| 60 | 684 | then obtain th_a where eq_a: "a = Th th_a" | 
| 685 | proof(cases rule:subtreeE) | |
| 686 | case 2 | |
| 687 | from ancestors_Field[OF 2(2)] | |
| 688 | and that show ?thesis by (unfold tRAG_alt_def, auto) | |
| 689 | qed auto | |
| 690 | have neq_th_a: "th_a \<noteq> th" | |
| 691 | proof - | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 692 | from readys_in_no_subtree[OF th_ready_s assms] | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 693 | have "(Th th) \<notin> subtree (RAG s) (Th th')" . | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 694 | with tRAG_subtree_RAG[of s "Th th'"] | 
| 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 695 | have "(Th th) \<notin> subtree (tRAG s) (Th th')" by auto | 
| 60 | 696 | with a_in[unfolded eq_a] show ?thesis by auto | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 697 | qed | 
| 60 | 698 | from preced_kept[OF this] | 
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 699 | show "(the_preced (e#s) \<circ> the_thread) a = (the_preced s \<circ> the_thread) a" | 
| 60 | 700 | by (unfold eq_a, simp) | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 701 | qed | 
| 60 | 702 | thus ?thesis by (unfold cp_alt_def1, simp) | 
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 703 | qed | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 704 | |
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 705 | end | 
| 60 | 706 | |
| 0 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 707 | end | 
| 
110247f9d47e
added
 Christian Urban <christian dot urban at kcl dot ac dot uk> parents: diff
changeset | 708 | |
| 92 
4763aa246dbd
Original files overwrite by their parallels (Correctness.thy v.s. PrioG.thy, PIPBasics.thy v.s. CpsG.thy, Implementation v.s. ExtGG.thy).
 zhangx parents: 
68diff
changeset | 709 |