| author | zhangx | 
| Wed, 27 Jan 2016 23:34:23 +0800 | |
| changeset 84 | cfd644dfc3b4 | 
| parent 65 | 633b1fc8631b | 
| permissions | -rw-r--r-- | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 1 | section {*
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 2 | This file contains lemmas used to guide the recalculation of current precedence | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 3 | after every system call (or system operation) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 4 | *} | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 5 | theory Implementation | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 6 | imports PIPBasics | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 7 | begin | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 8 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 9 | text {* (* ddd *)
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 10 | One beauty of our modelling is that we follow the definitional extension tradition of HOL. | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 11 | The benefit of such a concise and miniature model is that large number of intuitively | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 12 | obvious facts are derived as lemmas, rather than asserted as axioms. | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 13 | *} | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 14 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 15 | text {*
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 16 | However, the lemmas in the forthcoming several locales are no longer | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 17 | obvious. These lemmas show how the current precedences should be recalculated | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 18 | after every execution step (in our model, every step is represented by an event, | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 19 | which in turn, represents a system call, or operation). Each operation is | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 20 | treated in a separate locale. | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 21 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 22 | The complication of current precedence recalculation comes | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 23 | because the changing of RAG needs to be taken into account, | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 24 | in addition to the changing of precedence. | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 25 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 26 | The reason RAG changing affects current precedence is that, | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 27 | according to the definition, current precedence | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 28 | of a thread is the maximum of the precedences of every threads in its subtree, | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 29 | where the notion of sub-tree in RAG is defined in RTree.thy. | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 30 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 31 | Therefore, for each operation, lemmas about the change of precedences | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 32 | and RAG are derived first, on which lemmas about current precedence | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 33 | recalculation are based on. | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 34 | *} | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 35 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 36 | section {* The @{term Set} operation *}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 37 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 38 | text {* (* ddd *)
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 39 |   The following locale @{text "step_set_cps"} investigates the recalculation 
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 40 |   after the @{text "Set"} operation.
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 41 | *} | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 42 | locale step_set_cps = | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 43 | fixes s' th prio s | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 44 |   -- {* @{text "s'"} is the system state before the operation *}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 45 |   -- {* @{text "s"} is the system state after the operation *}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 46 | defines s_def : "s \<equiv> (Set th prio#s')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 47 |   -- {* @{text "s"} is assumed to be a legitimate state, from which
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 48 |          the legitimacy of @{text "s"} can be derived. *}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 49 | assumes vt_s: "vt s" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 50 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 51 | sublocale step_set_cps < vat_s : valid_trace "s" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 52 | proof | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 53 | from vt_s show "vt s" . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 54 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 55 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 56 | sublocale step_set_cps < vat_s' : valid_trace "s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 57 | proof | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 58 | from step_back_vt[OF vt_s[unfolded s_def]] show "vt s'" . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 59 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 60 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 61 | context step_set_cps | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 62 | begin | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 63 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 64 | text {* (* ddd *)
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 65 |   The following two lemmas confirm that @{text "Set"}-operation
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 66 | only changes the precedence of the initiating thread (or actor) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 67 | of the operation (or event). | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 68 | *} | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 69 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 70 | lemma eq_preced: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 71 | assumes "th' \<noteq> th" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 72 | shows "preced th' s = preced th' s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 73 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 74 | from assms show ?thesis | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 75 | by (unfold s_def, auto simp:preced_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 76 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 77 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 78 | lemma eq_the_preced: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 79 | assumes "th' \<noteq> th" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 80 | shows "the_preced s th' = the_preced s' th'" | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 81 | using assms | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 82 | by (unfold the_preced_def, intro eq_preced, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 83 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 84 | text {*
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 85 | The following lemma assures that the resetting of priority does not change the RAG. | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 86 | *} | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 87 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 88 | lemma eq_dep: "RAG s = RAG s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 89 | by (unfold s_def RAG_set_unchanged, auto) | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 90 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 91 | text {* (* ddd *)
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 92 |   Th following lemma @{text "eq_cp_pre"} says that the priority change of @{text "th"}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 93 |   only affects those threads, which as @{text "Th th"} in their sub-trees.
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 94 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 95 | The proof of this lemma is simplified by using the alternative definition | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 96 |   of @{text "cp"}. 
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 97 | *} | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 98 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 99 | lemma eq_cp_pre: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 100 | assumes nd: "Th th \<notin> subtree (RAG s') (Th th')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 101 | shows "cp s th' = cp s' th'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 102 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 103 |   -- {* After unfolding using the alternative definition, elements 
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 104 |         affecting the @{term "cp"}-value of threads become explicit. 
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 105 | We only need to prove the following: *} | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 106 |   have "Max (the_preced s ` {th'a. Th th'a \<in> subtree (RAG s) (Th th')}) =
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 107 |         Max (the_preced s' ` {th'a. Th th'a \<in> subtree (RAG s') (Th th')})"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 108 | (is "Max (?f ` ?S1) = Max (?g ` ?S2)") | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 109 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 110 |     -- {* The base sets are equal. *}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 111 | have "?S1 = ?S2" using eq_dep by simp | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 112 |     -- {* The function values on the base set are equal as well. *}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 113 | moreover have "\<forall> e \<in> ?S2. ?f e = ?g e" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 114 | proof | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 115 | fix th1 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 116 | assume "th1 \<in> ?S2" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 117 | with nd have "th1 \<noteq> th" by (auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 118 | from eq_the_preced[OF this] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 119 | show "the_preced s th1 = the_preced s' th1" . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 120 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 121 |     -- {* Therefore, the image of the functions are equal. *}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 122 | ultimately have "(?f ` ?S1) = (?g ` ?S2)" by (auto intro!:f_image_eq) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 123 | thus ?thesis by simp | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 124 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 125 | thus ?thesis by (simp add:cp_alt_def) | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 126 | qed | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 127 | |
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 128 | text {*
 | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 129 |   The following lemma shows that @{term "th"} is not in the 
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 130 | sub-tree of any other thread. | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 131 | *} | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 132 | lemma th_in_no_subtree: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 133 | assumes "th' \<noteq> th" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 134 | shows "Th th \<notin> subtree (RAG s') (Th th')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 135 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 136 | have "th \<in> readys s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 137 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 138 | from step_back_step [OF vt_s[unfolded s_def]] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 139 | have "step s' (Set th prio)" . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 140 | hence "th \<in> runing s'" by (cases, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 141 | thus ?thesis by (simp add:readys_def runing_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 142 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 143 | from vat_s'.readys_in_no_subtree[OF this assms(1)] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 144 | show ?thesis by blast | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 145 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 146 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 147 | text {* 
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 148 |   By combining @{thm "eq_cp_pre"} and @{thm "th_in_no_subtree"}, 
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 149 |   it is obvious that the change of priority only affects the @{text "cp"}-value 
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 150 |   of the initiating thread @{text "th"}.
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 151 | *} | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 152 | lemma eq_cp: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 153 | assumes "th' \<noteq> th" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 154 | shows "cp s th' = cp s' th'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 155 | by (rule eq_cp_pre[OF th_in_no_subtree[OF assms]]) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 156 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 157 | end | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 158 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 159 | section {* The @{term V} operation *}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 160 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 161 | text {*
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 162 |   The following @{text "step_v_cps"} is the locale for @{text "V"}-operation.
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 163 | *} | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 164 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 165 | locale step_v_cps = | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 166 |   -- {* @{text "th"} is the initiating thread *}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 167 |   -- {* @{text "cs"} is the critical resource release by the @{text "V"}-operation *}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 168 |   fixes s' th cs s    -- {* @{text "s'"} is the state before operation*}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 169 |   defines s_def : "s \<equiv> (V th cs#s')" -- {* @{text "s"} is the state after operation*}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 170 |   -- {* @{text "s"} is assumed to be valid, which implies the validity of @{text "s'"} *}
 | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 171 | assumes vt_s: "vt s" | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 172 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 173 | sublocale step_v_cps < vat_s : valid_trace "s" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 174 | proof | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 175 | from vt_s show "vt s" . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 176 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 177 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 178 | sublocale step_v_cps < vat_s' : valid_trace "s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 179 | proof | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 180 | from step_back_vt[OF vt_s[unfolded s_def]] show "vt s'" . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 181 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 182 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 183 | context step_v_cps | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 184 | begin | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 185 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 186 | lemma ready_th_s': "th \<in> readys s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 187 | using step_back_step[OF vt_s[unfolded s_def]] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 188 | by (cases, simp add:runing_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 189 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 190 | lemma ancestors_th: "ancestors (RAG s') (Th th) = {}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 191 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 192 | from vat_s'.readys_root[OF ready_th_s'] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 193 | show ?thesis | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 194 | by (unfold root_def, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 195 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 196 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 197 | lemma holding_th: "holding s' th cs" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 198 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 199 | from vt_s[unfolded s_def] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 200 | have " PIP s' (V th cs)" by (cases, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 201 | thus ?thesis by (cases, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 202 | qed | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 203 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 204 | lemma edge_of_th: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 205 | "(Cs cs, Th th) \<in> RAG s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 206 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 207 | from holding_th | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 208 | show ?thesis | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 209 | by (unfold s_RAG_def holding_eq, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 210 | qed | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 211 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 212 | lemma ancestors_cs: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 213 |   "ancestors (RAG s') (Cs cs) = {Th th}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 214 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 215 |   have "ancestors (RAG s') (Cs cs) = ancestors (RAG s') (Th th)  \<union>  {Th th}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 216 | proof(rule vat_s'.rtree_RAG.ancestors_accum) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 217 | from vt_s[unfolded s_def] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 218 | have " PIP s' (V th cs)" by (cases, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 219 | thus "(Cs cs, Th th) \<in> RAG s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 220 | proof(cases) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 221 | assume "holding s' th cs" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 222 | from this[unfolded holding_eq] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 223 | show ?thesis by (unfold s_RAG_def, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 224 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 225 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 226 | from this[unfolded ancestors_th] show ?thesis by simp | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 227 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 228 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 229 | lemma preced_kept: "the_preced s = the_preced s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 230 | by (auto simp: s_def the_preced_def preced_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 231 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 232 | end | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 233 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 234 | text {*
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 235 |   The following @{text "step_v_cps_nt"} is the sub-locale for @{text "V"}-operation, 
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 236 |   which represents the case when there is another thread @{text "th'"}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 237 |   to take over the critical resource released by the initiating thread @{text "th"}.
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 238 | *} | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 239 | locale step_v_cps_nt = step_v_cps + | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 240 | fixes th' | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 241 |   -- {* @{text "th'"} is assumed to take over @{text "cs"} *}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 242 | assumes nt: "next_th s' th cs th'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 243 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 244 | context step_v_cps_nt | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 245 | begin | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 246 | |
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 247 | text {*
 | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 248 |   Lemma @{text "RAG_s"} confirms the change of RAG:
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 249 | two edges removed and one added, as shown by the following diagram. | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 250 | *} | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 251 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 252 | (* | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 253 | RAG before the V-operation | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 254 | th1 ----| | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 255 | | | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 256 | th' ----| | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 257 | |----> cs -----| | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 258 | th2 ----| | | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 259 | | | | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 260 | th3 ----| | | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 261 | |------> th | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 262 | th4 ----| | | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 263 | | | | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 264 | th5 ----| | | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 265 | |----> cs'-----| | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 266 | th6 ----| | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 267 | | | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 268 | th7 ----| | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 269 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 270 | RAG after the V-operation | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 271 | th1 ----| | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 272 | | | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 273 | |----> cs ----> th' | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 274 | th2 ----| | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 275 | | | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 276 | th3 ----| | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 277 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 278 | th4 ----| | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 279 | | | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 280 | th5 ----| | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 281 | |----> cs'----> th | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 282 | th6 ----| | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 283 | | | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 284 | th7 ----| | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 285 | *) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 286 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 287 | lemma sub_RAGs': "{(Cs cs, Th th), (Th th', Cs cs)} \<subseteq> RAG s'"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 288 | using next_th_RAG[OF nt] . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 289 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 290 | lemma ancestors_th': | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 291 |   "ancestors (RAG s') (Th th') = {Th th, Cs cs}" 
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 292 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 293 |   have "ancestors (RAG s') (Th th') = ancestors (RAG s') (Cs cs) \<union> {Cs cs}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 294 | proof(rule vat_s'.rtree_RAG.ancestors_accum) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 295 | from sub_RAGs' show "(Th th', Cs cs) \<in> RAG s'" by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 296 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 297 | thus ?thesis using ancestors_th ancestors_cs by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 298 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 299 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 300 | lemma RAG_s: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 301 |   "RAG s = (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) \<union>
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 302 |                                          {(Cs cs, Th th')}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 303 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 304 | from step_RAG_v[OF vt_s[unfolded s_def], folded s_def] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 305 | and nt show ?thesis by (auto intro:next_th_unique) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 306 | qed | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 307 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 308 | lemma subtree_kept: (* ddd *) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 309 |   assumes "th1 \<notin> {th, th'}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 310 | shows "subtree (RAG s) (Th th1) = subtree (RAG s') (Th th1)" (is "_ = ?R") | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 311 | proof - | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 312 |   let ?RAG' = "(RAG s' - {(Cs cs, Th th), (Th th', Cs cs)})"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 313 |   let ?RAG'' = "?RAG' \<union> {(Cs cs, Th th')}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 314 | have "subtree ?RAG' (Th th1) = ?R" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 315 | proof(rule subset_del_subtree_outside) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 316 |     show "Range {(Cs cs, Th th), (Th th', Cs cs)} \<inter> subtree (RAG s') (Th th1) = {}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 317 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 318 | have "(Th th) \<notin> subtree (RAG s') (Th th1)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 319 | proof(rule subtree_refute) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 320 | show "Th th1 \<notin> ancestors (RAG s') (Th th)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 321 | by (unfold ancestors_th, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 322 | next | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 323 | from assms show "Th th1 \<noteq> Th th" by simp | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 324 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 325 | moreover have "(Cs cs) \<notin> subtree (RAG s') (Th th1)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 326 | proof(rule subtree_refute) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 327 | show "Th th1 \<notin> ancestors (RAG s') (Cs cs)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 328 | by (unfold ancestors_cs, insert assms, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 329 | qed simp | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 330 |       ultimately have "{Th th, Cs cs} \<inter> subtree (RAG s') (Th th1) = {}" by auto
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 331 | thus ?thesis by simp | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 332 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 333 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 334 | moreover have "subtree ?RAG'' (Th th1) = subtree ?RAG' (Th th1)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 335 | proof(rule subtree_insert_next) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 336 |     show "Th th' \<notin> subtree (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) (Th th1)"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 337 | proof(rule subtree_refute) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 338 |       show "Th th1 \<notin> ancestors (RAG s' - {(Cs cs, Th th), (Th th', Cs cs)}) (Th th')"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 339 | (is "_ \<notin> ?R") | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 340 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 341 | have "?R \<subseteq> ancestors (RAG s') (Th th')" by (rule ancestors_mono, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 342 | moreover have "Th th1 \<notin> ..." using ancestors_th' assms by simp | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 343 | ultimately show ?thesis by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 344 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 345 | next | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 346 | from assms show "Th th1 \<noteq> Th th'" by simp | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 347 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 348 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 349 | ultimately show ?thesis by (unfold RAG_s, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 350 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 351 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 352 | lemma cp_kept: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 353 |   assumes "th1 \<notin> {th, th'}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 354 | shows "cp s th1 = cp s' th1" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 355 | by (unfold cp_alt_def preced_kept subtree_kept[OF assms], simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 356 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 357 | end | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 358 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 359 | locale step_v_cps_nnt = step_v_cps + | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 360 | assumes nnt: "\<And> th'. (\<not> next_th s' th cs th')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 361 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 362 | context step_v_cps_nnt | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 363 | begin | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 364 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 365 | lemma RAG_s: "RAG s = RAG s' - {(Cs cs, Th th)}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 366 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 367 | from nnt and step_RAG_v[OF vt_s[unfolded s_def], folded s_def] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 368 | show ?thesis by auto | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 369 | qed | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 370 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 371 | lemma subtree_kept: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 372 | assumes "th1 \<noteq> th" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 373 | shows "subtree (RAG s) (Th th1) = subtree (RAG s') (Th th1)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 374 | proof(unfold RAG_s, rule subset_del_subtree_outside) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 375 |   show "Range {(Cs cs, Th th)} \<inter> subtree (RAG s') (Th th1) = {}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 376 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 377 | have "(Th th) \<notin> subtree (RAG s') (Th th1)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 378 | proof(rule subtree_refute) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 379 | show "Th th1 \<notin> ancestors (RAG s') (Th th)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 380 | by (unfold ancestors_th, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 381 | next | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 382 | from assms show "Th th1 \<noteq> Th th" by simp | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 383 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 384 | thus ?thesis by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 385 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 386 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 387 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 388 | lemma cp_kept_1: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 389 | assumes "th1 \<noteq> th" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 390 | shows "cp s th1 = cp s' th1" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 391 | by (unfold cp_alt_def preced_kept subtree_kept[OF assms], simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 392 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 393 | lemma subtree_cs: "subtree (RAG s') (Cs cs) = {Cs cs}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 394 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 395 |   { fix n
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 396 | have "(Cs cs) \<notin> ancestors (RAG s') n" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 397 | proof | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 398 | assume "Cs cs \<in> ancestors (RAG s') n" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 399 | hence "(n, Cs cs) \<in> (RAG s')^+" by (auto simp:ancestors_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 400 | from tranclE[OF this] obtain nn where h: "(nn, Cs cs) \<in> RAG s'" by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 401 | then obtain th' where "nn = Th th'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 402 | by (unfold s_RAG_def, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 403 | from h[unfolded this] have "(Th th', Cs cs) \<in> RAG s'" . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 404 | from this[unfolded s_RAG_def] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 405 | have "waiting (wq s') th' cs" by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 406 | from this[unfolded cs_waiting_def] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 407 | have "1 < length (wq s' cs)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 408 | by (cases "wq s' cs", auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 409 | from holding_next_thI[OF holding_th this] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 410 | obtain th' where "next_th s' th cs th'" by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 411 | with nnt show False by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 412 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 413 | } note h = this | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 414 |   {  fix n
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 415 | assume "n \<in> subtree (RAG s') (Cs cs)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 416 | hence "n = (Cs cs)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 417 | by (elim subtreeE, insert h, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 418 | } moreover have "(Cs cs) \<in> subtree (RAG s') (Cs cs)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 419 | by (auto simp:subtree_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 420 | ultimately show ?thesis by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 421 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 422 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 423 | lemma subtree_th: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 424 |   "subtree (RAG s) (Th th) = subtree (RAG s') (Th th) - {Cs cs}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 425 | proof(unfold RAG_s, fold subtree_cs, rule vat_s'.rtree_RAG.subtree_del_inside) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 426 | from edge_of_th | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 427 | show "(Cs cs, Th th) \<in> edges_in (RAG s') (Th th)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 428 | by (unfold edges_in_def, auto simp:subtree_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 429 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 430 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 431 | lemma cp_kept_2: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 432 | shows "cp s th = cp s' th" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 433 | by (unfold cp_alt_def subtree_th preced_kept, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 434 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 435 | lemma eq_cp: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 436 | shows "cp s th' = cp s' th'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 437 | using cp_kept_1 cp_kept_2 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 438 | by (cases "th' = th", auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 439 | end | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 440 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 441 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 442 | locale step_P_cps = | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 443 | fixes s' th cs s | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 444 | defines s_def : "s \<equiv> (P th cs#s')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 445 | assumes vt_s: "vt s" | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 446 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 447 | sublocale step_P_cps < vat_s : valid_trace "s" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 448 | proof | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 449 | from vt_s show "vt s" . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 450 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 451 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 452 | section {* The @{term P} operation *}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 453 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 454 | sublocale step_P_cps < vat_s' : valid_trace "s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 455 | proof | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 456 | from step_back_vt[OF vt_s[unfolded s_def]] show "vt s'" . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 457 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 458 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 459 | context step_P_cps | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 460 | begin | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 461 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 462 | lemma readys_th: "th \<in> readys s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 463 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 464 | from step_back_step [OF vt_s[unfolded s_def]] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 465 | have "PIP s' (P th cs)" . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 466 | hence "th \<in> runing s'" by (cases, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 467 | thus ?thesis by (simp add:readys_def runing_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 468 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 469 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 470 | lemma root_th: "root (RAG s') (Th th)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 471 | using readys_root[OF readys_th] . | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 472 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 473 | lemma in_no_others_subtree: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 474 | assumes "th' \<noteq> th" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 475 | shows "Th th \<notin> subtree (RAG s') (Th th')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 476 | proof | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 477 | assume "Th th \<in> subtree (RAG s') (Th th')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 478 | thus False | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 479 | proof(cases rule:subtreeE) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 480 | case 1 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 481 | with assms show ?thesis by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 482 | next | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 483 | case 2 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 484 | with root_th show ?thesis by (auto simp:root_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 485 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 486 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 487 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 488 | lemma preced_kept: "the_preced s = the_preced s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 489 | by (auto simp: s_def the_preced_def preced_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 490 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 491 | end | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 492 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 493 | locale step_P_cps_ne =step_P_cps + | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 494 | fixes th' | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 495 | assumes ne: "wq s' cs \<noteq> []" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 496 | defines th'_def: "th' \<equiv> hd (wq s' cs)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 497 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 498 | locale step_P_cps_e =step_P_cps + | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 499 | assumes ee: "wq s' cs = []" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 500 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 501 | context step_P_cps_e | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 502 | begin | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 503 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 504 | lemma RAG_s: "RAG s = RAG s' \<union> {(Cs cs, Th th)}"
 | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 505 | proof - | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 506 | from ee and step_RAG_p[OF vt_s[unfolded s_def], folded s_def] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 507 | show ?thesis by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 508 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 509 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 510 | lemma subtree_kept: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 511 | assumes "th' \<noteq> th" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 512 | shows "subtree (RAG s) (Th th') = subtree (RAG s') (Th th')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 513 | proof(unfold RAG_s, rule subtree_insert_next) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 514 | from in_no_others_subtree[OF assms] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 515 | show "Th th \<notin> subtree (RAG s') (Th th')" . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 516 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 517 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 518 | lemma cp_kept: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 519 | assumes "th' \<noteq> th" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 520 | shows "cp s th' = cp s' th'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 521 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 522 |   have "(the_preced s ` {th'a. Th th'a \<in> subtree (RAG s) (Th th')}) =
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 523 |         (the_preced s' ` {th'a. Th th'a \<in> subtree (RAG s') (Th th')})"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 524 | by (unfold preced_kept subtree_kept[OF assms], simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 525 | thus ?thesis by (unfold cp_alt_def, simp) | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 526 | qed | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 527 | |
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 528 | end | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 529 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 530 | context step_P_cps_ne | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 531 | begin | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 532 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 533 | lemma RAG_s: "RAG s = RAG s' \<union> {(Th th, Cs cs)}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 534 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 535 | from step_RAG_p[OF vt_s[unfolded s_def]] and ne | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 536 | show ?thesis by (simp add:s_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 537 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 538 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 539 | lemma cs_held: "(Cs cs, Th th') \<in> RAG s'" | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 540 | proof - | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 541 | have "(Cs cs, Th th') \<in> hRAG s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 542 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 543 | from ne | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 544 | have " holding s' th' cs" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 545 | by (unfold th'_def holding_eq cs_holding_def, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 546 | thus ?thesis | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 547 | by (unfold hRAG_def, auto) | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 548 | qed | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 549 | thus ?thesis by (unfold RAG_split, auto) | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 550 | qed | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 551 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 552 | lemma tRAG_s: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 553 |   "tRAG s = tRAG s' \<union> {(Th th, Th th')}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 554 | using RAG_tRAG_transfer[OF RAG_s cs_held] . | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 555 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 556 | lemma cp_kept: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 557 | assumes "Th th'' \<notin> ancestors (tRAG s) (Th th)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 558 | shows "cp s th'' = cp s' th''" | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 559 | proof - | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 560 | have h: "subtree (tRAG s) (Th th'') = subtree (tRAG s') (Th th'')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 561 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 562 | have "Th th' \<notin> subtree (tRAG s') (Th th'')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 563 | proof | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 564 | assume "Th th' \<in> subtree (tRAG s') (Th th'')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 565 | thus False | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 566 | proof(rule subtreeE) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 567 | assume "Th th' = Th th''" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 568 | from assms[unfolded tRAG_s ancestors_def, folded this] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 569 | show ?thesis by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 570 | next | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 571 | assume "Th th'' \<in> ancestors (tRAG s') (Th th')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 572 | moreover have "... \<subseteq> ancestors (tRAG s) (Th th')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 573 | proof(rule ancestors_mono) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 574 | show "tRAG s' \<subseteq> tRAG s" by (unfold tRAG_s, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 575 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 576 | ultimately have "Th th'' \<in> ancestors (tRAG s) (Th th')" by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 577 | moreover have "Th th' \<in> ancestors (tRAG s) (Th th)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 578 | by (unfold tRAG_s, auto simp:ancestors_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 579 | ultimately have "Th th'' \<in> ancestors (tRAG s) (Th th)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 580 | by (auto simp:ancestors_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 581 | with assms show ?thesis by auto | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 582 | qed | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 583 | qed | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 584 | from subtree_insert_next[OF this] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 585 |     have "subtree (tRAG s' \<union> {(Th th, Th th')}) (Th th'') = subtree (tRAG s') (Th th'')" .
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 586 | from this[folded tRAG_s] show ?thesis . | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 587 | qed | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 588 | show ?thesis by (unfold cp_alt_def1 h preced_kept, simp) | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 589 | qed | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 590 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 591 | lemma cp_gen_update_stop: (* ddd *) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 592 | assumes "u \<in> ancestors (tRAG s) (Th th)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 593 | and "cp_gen s u = cp_gen s' u" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 594 | and "y \<in> ancestors (tRAG s) u" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 595 | shows "cp_gen s y = cp_gen s' y" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 596 | using assms(3) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 597 | proof(induct rule:wf_induct[OF vat_s.fsbttRAGs.wf]) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 598 | case (1 x) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 599 | show ?case (is "?L = ?R") | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 600 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 601 | from tRAG_ancestorsE[OF 1(2)] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 602 | obtain th2 where eq_x: "x = Th th2" by blast | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 603 | from vat_s.cp_gen_rec[OF this] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 604 | have "?L = | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 605 |           Max ({the_preced s th2} \<union> cp_gen s ` RTree.children (tRAG s) x)" .
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 606 | also have "... = | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 607 |           Max ({the_preced s' th2} \<union> cp_gen s' ` RTree.children (tRAG s') x)"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 608 | |
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 609 | proof - | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 610 | from preced_kept have "the_preced s th2 = the_preced s' th2" by simp | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 611 | moreover have "cp_gen s ` RTree.children (tRAG s) x = | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 612 | cp_gen s' ` RTree.children (tRAG s') x" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 613 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 614 | have "RTree.children (tRAG s) x = RTree.children (tRAG s') x" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 615 | proof(unfold tRAG_s, rule children_union_kept) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 616 | have start: "(Th th, Th th') \<in> tRAG s" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 617 | by (unfold tRAG_s, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 618 | note x_u = 1(2) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 619 |           show "x \<notin> Range {(Th th, Th th')}"
 | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 620 | proof | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 621 |             assume "x \<in> Range {(Th th, Th th')}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 622 | hence eq_x: "x = Th th'" using RangeE by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 623 | show False | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 624 | proof(cases rule:vat_s.rtree_s.ancestors_headE[OF assms(1) start]) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 625 | case 1 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 626 | from x_u[folded this, unfolded eq_x] vat_s.acyclic_tRAG | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 627 | show ?thesis by (auto simp:ancestors_def acyclic_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 628 | next | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 629 | case 2 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 630 | with x_u[unfolded eq_x] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 631 | have "(Th th', Th th') \<in> (tRAG s)^+" by (auto simp:ancestors_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 632 | with vat_s.acyclic_tRAG show ?thesis by (auto simp:acyclic_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 633 | qed | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 634 | qed | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 635 | qed | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 636 | moreover have "cp_gen s ` RTree.children (tRAG s) x = | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 637 | cp_gen s' ` RTree.children (tRAG s) x" (is "?f ` ?A = ?g ` ?A") | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 638 | proof(rule f_image_eq) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 639 | fix a | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 640 | assume a_in: "a \<in> ?A" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 641 | from 1(2) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 642 | show "?f a = ?g a" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 643 | proof(cases rule:vat_s.rtree_s.ancestors_childrenE[case_names in_ch out_ch]) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 644 | case in_ch | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 645 | show ?thesis | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 646 | proof(cases "a = u") | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 647 | case True | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 648 | from assms(2)[folded this] show ?thesis . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 649 | next | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 650 | case False | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 651 | have a_not_in: "a \<notin> ancestors (tRAG s) (Th th)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 652 | proof | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 653 | assume a_in': "a \<in> ancestors (tRAG s) (Th th)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 654 | have "a = u" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 655 | proof(rule vat_s.rtree_s.ancestors_children_unique) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 656 | from a_in' a_in show "a \<in> ancestors (tRAG s) (Th th) \<inter> | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 657 | RTree.children (tRAG s) x" by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 658 | next | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 659 | from assms(1) in_ch show "u \<in> ancestors (tRAG s) (Th th) \<inter> | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 660 | RTree.children (tRAG s) x" by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 661 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 662 | with False show False by simp | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 663 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 664 | from a_in obtain th_a where eq_a: "a = Th th_a" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 665 | by (unfold RTree.children_def tRAG_alt_def, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 666 | from cp_kept[OF a_not_in[unfolded eq_a]] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 667 | have "cp s th_a = cp s' th_a" . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 668 | from this [unfolded cp_gen_def_cond[OF eq_a], folded eq_a] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 669 | show ?thesis . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 670 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 671 | next | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 672 | case (out_ch z) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 673 | hence h: "z \<in> ancestors (tRAG s) u" "z \<in> RTree.children (tRAG s) x" by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 674 | show ?thesis | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 675 | proof(cases "a = z") | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 676 | case True | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 677 | from h(2) have zx_in: "(z, x) \<in> (tRAG s)" by (auto simp:RTree.children_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 678 | from 1(1)[rule_format, OF this h(1)] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 679 | have eq_cp_gen: "cp_gen s z = cp_gen s' z" . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 680 | with True show ?thesis by metis | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 681 | next | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 682 | case False | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 683 | from a_in obtain th_a where eq_a: "a = Th th_a" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 684 | by (auto simp:RTree.children_def tRAG_alt_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 685 | have "a \<notin> ancestors (tRAG s) (Th th)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 686 | proof | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 687 | assume a_in': "a \<in> ancestors (tRAG s) (Th th)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 688 | have "a = z" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 689 | proof(rule vat_s.rtree_s.ancestors_children_unique) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 690 | from assms(1) h(1) have "z \<in> ancestors (tRAG s) (Th th)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 691 | by (auto simp:ancestors_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 692 | with h(2) show " z \<in> ancestors (tRAG s) (Th th) \<inter> | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 693 | RTree.children (tRAG s) x" by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 694 | next | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 695 | from a_in a_in' | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 696 | show "a \<in> ancestors (tRAG s) (Th th) \<inter> RTree.children (tRAG s) x" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 697 | by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 698 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 699 | with False show False by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 700 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 701 | from cp_kept[OF this[unfolded eq_a]] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 702 | have "cp s th_a = cp s' th_a" . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 703 | from this[unfolded cp_gen_def_cond[OF eq_a], folded eq_a] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 704 | show ?thesis . | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 705 | qed | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 706 | qed | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 707 | qed | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 708 | ultimately show ?thesis by metis | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 709 | qed | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 710 | ultimately show ?thesis by simp | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 711 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 712 | also have "... = ?R" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 713 | by (fold vat_s'.cp_gen_rec[OF eq_x], simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 714 | finally show ?thesis . | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 715 | qed | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 716 | qed | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 717 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 718 | lemma cp_up: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 719 | assumes "(Th th') \<in> ancestors (tRAG s) (Th th)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 720 | and "cp s th' = cp s' th'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 721 | and "(Th th'') \<in> ancestors (tRAG s) (Th th')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 722 | shows "cp s th'' = cp s' th''" | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 723 | proof - | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 724 | have "cp_gen s (Th th'') = cp_gen s' (Th th'')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 725 | proof(rule cp_gen_update_stop[OF assms(1) _ assms(3)]) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 726 | from assms(2) cp_gen_def_cond[OF refl[of "Th th'"]] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 727 | show "cp_gen s (Th th') = cp_gen s' (Th th')" by metis | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 728 | qed | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 729 | with cp_gen_def_cond[OF refl[of "Th th''"]] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 730 | show ?thesis by metis | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 731 | qed | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 732 | |
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 733 | end | 
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 734 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 735 | section {* The @{term Create} operation *}
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 736 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 737 | locale step_create_cps = | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 738 | fixes s' th prio s | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 739 | defines s_def : "s \<equiv> (Create th prio#s')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 740 | assumes vt_s: "vt s" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 741 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 742 | sublocale step_create_cps < vat_s: valid_trace "s" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 743 | by (unfold_locales, insert vt_s, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 744 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 745 | sublocale step_create_cps < vat_s': valid_trace "s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 746 | by (unfold_locales, insert step_back_vt[OF vt_s[unfolded s_def]], simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 747 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 748 | context step_create_cps | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 749 | begin | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 750 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 751 | lemma RAG_kept: "RAG s = RAG s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 752 | by (unfold s_def RAG_create_unchanged, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 753 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 754 | lemma tRAG_kept: "tRAG s = tRAG s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 755 | by (unfold tRAG_alt_def RAG_kept, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 756 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 757 | lemma preced_kept: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 758 | assumes "th' \<noteq> th" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 759 | shows "the_preced s th' = the_preced s' th'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 760 | by (unfold s_def the_preced_def preced_def, insert assms, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 761 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 762 | lemma th_not_in: "Th th \<notin> Field (tRAG s')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 763 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 764 | from vt_s[unfolded s_def] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 765 | have "PIP s' (Create th prio)" by (cases, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 766 | hence "th \<notin> threads s'" by(cases, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 767 | from vat_s'.not_in_thread_isolated[OF this] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 768 | have "Th th \<notin> Field (RAG s')" . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 769 | with tRAG_Field show ?thesis by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 770 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 771 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 772 | lemma eq_cp: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 773 | assumes neq_th: "th' \<noteq> th" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 774 | shows "cp s th' = cp s' th'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 775 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 776 | have "(the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th') = | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 777 | (the_preced s' \<circ> the_thread) ` subtree (tRAG s') (Th th')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 778 | proof(unfold tRAG_kept, rule f_image_eq) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 779 | fix a | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 780 | assume a_in: "a \<in> subtree (tRAG s') (Th th')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 781 | then obtain th_a where eq_a: "a = Th th_a" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 782 | proof(cases rule:subtreeE) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 783 | case 2 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 784 | from ancestors_Field[OF 2(2)] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 785 | and that show ?thesis by (unfold tRAG_alt_def, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 786 | qed auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 787 | have neq_th_a: "th_a \<noteq> th" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 788 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 789 | have "(Th th) \<notin> subtree (tRAG s') (Th th')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 790 | proof | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 791 | assume "Th th \<in> subtree (tRAG s') (Th th')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 792 | thus False | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 793 | proof(cases rule:subtreeE) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 794 | case 2 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 795 | from ancestors_Field[OF this(2)] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 796 | and th_not_in[unfolded Field_def] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 797 | show ?thesis by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 798 | qed (insert assms, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 799 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 800 | with a_in[unfolded eq_a] show ?thesis by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 801 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 802 | from preced_kept[OF this] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 803 | show "(the_preced s \<circ> the_thread) a = (the_preced s' \<circ> the_thread) a" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 804 | by (unfold eq_a, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 805 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 806 | thus ?thesis by (unfold cp_alt_def1, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 807 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 808 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 809 | lemma children_of_th: "RTree.children (tRAG s) (Th th) = {}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 810 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 811 |   { fix a
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 812 | assume "a \<in> RTree.children (tRAG s) (Th th)" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 813 | hence "(a, Th th) \<in> tRAG s" by (auto simp:RTree.children_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 814 | with th_not_in have False | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 815 | by (unfold Field_def tRAG_kept, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 816 | } thus ?thesis by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 817 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 818 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 819 | lemma eq_cp_th: "cp s th = preced th s" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 820 | by (unfold vat_s.cp_rec children_of_th, simp add:the_preced_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 821 | |
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 822 | end | 
| 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 823 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 824 | locale step_exit_cps = | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 825 | fixes s' th prio s | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 826 | defines s_def : "s \<equiv> Exit th # s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 827 | assumes vt_s: "vt s" | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 828 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 829 | sublocale step_exit_cps < vat_s: valid_trace "s" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 830 | by (unfold_locales, insert vt_s, simp) | 
| 65 
633b1fc8631b
Reorganization completed, added "scripts_structure.pdf" and "scirpts_structure.pptx".
 zhangx parents: diff
changeset | 831 | |
| 84 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 832 | sublocale step_exit_cps < vat_s': valid_trace "s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 833 | by (unfold_locales, insert step_back_vt[OF vt_s[unfolded s_def]], simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 834 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 835 | context step_exit_cps | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 836 | begin | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 837 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 838 | lemma preced_kept: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 839 | assumes "th' \<noteq> th" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 840 | shows "the_preced s th' = the_preced s' th'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 841 | by (unfold s_def the_preced_def preced_def, insert assms, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 842 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 843 | lemma RAG_kept: "RAG s = RAG s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 844 | by (unfold s_def RAG_exit_unchanged, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 845 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 846 | lemma tRAG_kept: "tRAG s = tRAG s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 847 | by (unfold tRAG_alt_def RAG_kept, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 848 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 849 | lemma th_ready: "th \<in> readys s'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 850 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 851 | from vt_s[unfolded s_def] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 852 | have "PIP s' (Exit th)" by (cases, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 853 |   hence h: "th \<in> runing s' \<and> holdents s' th = {}" by (cases, metis)
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 854 | thus ?thesis by (unfold runing_def, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 855 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 856 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 857 | lemma th_holdents: "holdents s' th = {}"
 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 858 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 859 | from vt_s[unfolded s_def] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 860 | have "PIP s' (Exit th)" by (cases, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 861 | thus ?thesis by (cases, metis) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 862 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 863 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 864 | lemma th_RAG: "Th th \<notin> Field (RAG s')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 865 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 866 | have "Th th \<notin> Range (RAG s')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 867 | proof | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 868 | assume "Th th \<in> Range (RAG s')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 869 | then obtain cs where "holding (wq s') th cs" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 870 | by (unfold Range_iff s_RAG_def, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 871 | with th_holdents[unfolded holdents_def] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 872 | show False by (unfold eq_holding, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 873 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 874 | moreover have "Th th \<notin> Domain (RAG s')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 875 | proof | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 876 | assume "Th th \<in> Domain (RAG s')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 877 | then obtain cs where "waiting (wq s') th cs" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 878 | by (unfold Domain_iff s_RAG_def, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 879 | with th_ready show False by (unfold readys_def eq_waiting, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 880 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 881 | ultimately show ?thesis by (auto simp:Field_def) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 882 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 883 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 884 | lemma th_tRAG: "(Th th) \<notin> Field (tRAG s')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 885 | using th_RAG tRAG_Field[of s'] by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 886 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 887 | lemma eq_cp: | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 888 | assumes neq_th: "th' \<noteq> th" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 889 | shows "cp s th' = cp s' th'" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 890 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 891 | have "(the_preced s \<circ> the_thread) ` subtree (tRAG s) (Th th') = | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 892 | (the_preced s' \<circ> the_thread) ` subtree (tRAG s') (Th th')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 893 | proof(unfold tRAG_kept, rule f_image_eq) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 894 | fix a | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 895 | assume a_in: "a \<in> subtree (tRAG s') (Th th')" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 896 | then obtain th_a where eq_a: "a = Th th_a" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 897 | proof(cases rule:subtreeE) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 898 | case 2 | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 899 | from ancestors_Field[OF 2(2)] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 900 | and that show ?thesis by (unfold tRAG_alt_def, auto) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 901 | qed auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 902 | have neq_th_a: "th_a \<noteq> th" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 903 | proof - | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 904 | from vat_s'.readys_in_no_subtree[OF th_ready assms] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 905 | have "(Th th) \<notin> subtree (RAG s') (Th th')" . | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 906 | with tRAG_subtree_RAG[of s' "Th th'"] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 907 | have "(Th th) \<notin> subtree (tRAG s') (Th th')" by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 908 | with a_in[unfolded eq_a] show ?thesis by auto | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 909 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 910 | from preced_kept[OF this] | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 911 | show "(the_preced s \<circ> the_thread) a = (the_preced s' \<circ> the_thread) a" | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 912 | by (unfold eq_a, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 913 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 914 | thus ?thesis by (unfold cp_alt_def1, simp) | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 915 | qed | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 916 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 917 | end | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 918 | |
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 919 | end | 
| 
cfd644dfc3b4
The parallel of Implementation.thy, i.e. ExtGG.thy has been updated. And some errors in
 zhangx parents: 
65diff
changeset | 920 |