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